From 50ea35fceb52c5f66ccbc4f709df3a3471b12647 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 19 Mar 2019 19:28:03 +0100 Subject: added helper functions but strange idiv.c: error: __compcert_i32_sdiv: missing or incorrect declaration --- backend/SelectDiv.vp | 11 +++----- backend/Selection.v | 7 ++++- backend/SplitLong.vp | 30 --------------------- backend/SplitLongproof.v | 51 ----------------------------------- cfrontend/C2C.ml | 18 ++++++++++++- mppa_k1c/SelectOp.v | 49 ++++++++++++++++++++++++++++++++- mppa_k1c/SelectOp.vp | 48 ++++++++++++++++++++++++++++++++- mppa_k1c/SelectOpproof.v | 70 +++++++++++++++++++++++++++++++++++++++++++++--- 8 files changed, 189 insertions(+), 95 deletions(-) diff --git a/backend/SelectDiv.vp b/backend/SelectDiv.vp index 357fab5e..6ddcd6ac 100644 --- a/backend/SelectDiv.vp +++ b/backend/SelectDiv.vp @@ -19,6 +19,9 @@ Require Import Op CminorSel SelectOp SplitLong SelectLong. Local Open Scope cminorsel_scope. +Section SELECT. +Context {hf: helper_functions}. + Definition is_intconst (e: expr) : option int := match e with | Eop (Ointconst n) _ => Some n @@ -221,10 +224,6 @@ Definition mods (e1: expr) (e2: expr) := (** 64-bit integer divisions *) -Section SELECT. - -Context {hf: helper_functions}. - Definition modl_from_divl (equo: expr) (n: int64) := subl (Eletvar O) (mullimm n equo). @@ -311,8 +310,6 @@ Definition modls (e1 e2: expr) := | _, _ => modls_base e1 e2 end. -End SELECT. - (** Floating-point division by a constant can also be turned into a FP multiplication by the inverse constant, but only for powers of 2. *) @@ -340,4 +337,4 @@ Nondetfunction divfs (e1: expr) (e2: expr) := | _ => Eop Odivfs (e1 ::: e2 ::: Enil) end. - +End SELECT. \ No newline at end of file diff --git a/backend/Selection.v b/backend/Selection.v index 4520cb0c..fc9315dc 100644 --- a/backend/Selection.v +++ b/backend/Selection.v @@ -391,11 +391,16 @@ Definition get_helpers (defmap: PTree.t globdef) : res helper_functions := do i64_sar <- lookup_helper globs "__compcert_i64_sar" sig_li_l ; do i64_umulh <- lookup_helper globs "__compcert_i64_umulh" sig_ll_l ; do i64_smulh <- lookup_helper globs "__compcert_i64_smulh" sig_ll_l ; + do i32_sdiv <- lookup_helper globs "__compcert_i32_sdiv" sig_ii_i ; + do i32_udiv <- lookup_helper globs "__compcert_i32_udiv" sig_ii_i ; + do i32_smod <- lookup_helper globs "__compcert_i32_smod" sig_ii_i ; + do i32_umod <- lookup_helper globs "__compcert_i32_umod" sig_ii_i ; OK (mk_helper_functions i64_dtos i64_dtou i64_stod i64_utod i64_stof i64_utof i64_sdiv i64_udiv i64_smod i64_umod i64_shl i64_shr i64_sar - i64_umulh i64_smulh). + i64_umulh i64_smulh + i32_sdiv i32_udiv i32_smod i32_umod). (** Conversion of programs. *) diff --git a/backend/SplitLong.vp b/backend/SplitLong.vp index de954482..6a7d4f5c 100644 --- a/backend/SplitLong.vp +++ b/backend/SplitLong.vp @@ -21,36 +21,6 @@ Require Import SelectOp. Local Open Scope cminorsel_scope. Local Open Scope string_scope. -(** Some operations on 64-bit integers are transformed into calls to - runtime library functions. The following type class collects - the names of these functions. *) - -Class helper_functions := mk_helper_functions { - i64_dtos: ident; (**r float64 -> signed long *) - i64_dtou: ident; (**r float64 -> unsigned long *) - i64_stod: ident; (**r signed long -> float64 *) - i64_utod: ident; (**r unsigned long -> float64 *) - i64_stof: ident; (**r signed long -> float32 *) - i64_utof: ident; (**r unsigned long -> float32 *) - i64_sdiv: ident; (**r signed division *) - i64_udiv: ident; (**r unsigned division *) - i64_smod: ident; (**r signed remainder *) - i64_umod: ident; (**r unsigned remainder *) - i64_shl: ident; (**r shift left *) - i64_shr: ident; (**r shift right unsigned *) - i64_sar: ident; (**r shift right signed *) - i64_umulh: ident; (**r unsigned multiply high *) - i64_smulh: ident; (**r signed multiply high *) -}. - -Definition sig_l_l := mksignature (Tlong :: nil) (Some Tlong) cc_default. -Definition sig_l_f := mksignature (Tlong :: nil) (Some Tfloat) cc_default. -Definition sig_l_s := mksignature (Tlong :: nil) (Some Tsingle) cc_default. -Definition sig_f_l := mksignature (Tfloat :: nil) (Some Tlong) cc_default. -Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_default. -Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default. -Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default. - Section SELECT. Context {hf: helper_functions}. diff --git a/backend/SplitLongproof.v b/backend/SplitLongproof.v index f1e8b590..3b74b216 100644 --- a/backend/SplitLongproof.v +++ b/backend/SplitLongproof.v @@ -21,57 +21,6 @@ Require Import SelectOp SelectOpproof SplitLong. Local Open Scope cminorsel_scope. Local Open Scope string_scope. -(** * Axiomatization of the helper functions *) - -Definition external_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop := - forall F V (ge: Genv.t F V) m, - external_call (EF_runtime name sg) ge vargs m E0 vres m. - -Definition builtin_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop := - forall F V (ge: Genv.t F V) m, - external_call (EF_builtin name sg) ge vargs m E0 vres m. - -Axiom i64_helpers_correct : - (forall x z, Val.longoffloat x = Some z -> external_implements "__compcert_i64_dtos" sig_f_l (x::nil) z) - /\ (forall x z, Val.longuoffloat x = Some z -> external_implements "__compcert_i64_dtou" sig_f_l (x::nil) z) - /\ (forall x z, Val.floatoflong x = Some z -> external_implements "__compcert_i64_stod" sig_l_f (x::nil) z) - /\ (forall x z, Val.floatoflongu x = Some z -> external_implements "__compcert_i64_utod" sig_l_f (x::nil) z) - /\ (forall x z, Val.singleoflong x = Some z -> external_implements "__compcert_i64_stof" sig_l_s (x::nil) z) - /\ (forall x z, Val.singleoflongu x = Some z -> external_implements "__compcert_i64_utof" sig_l_s (x::nil) z) - /\ (forall x, builtin_implements "__builtin_negl" sig_l_l (x::nil) (Val.negl x)) - /\ (forall x y, builtin_implements "__builtin_addl" sig_ll_l (x::y::nil) (Val.addl x y)) - /\ (forall x y, builtin_implements "__builtin_subl" sig_ll_l (x::y::nil) (Val.subl x y)) - /\ (forall x y, builtin_implements "__builtin_mull" sig_ii_l (x::y::nil) (Val.mull' x y)) - /\ (forall x y z, Val.divls x y = Some z -> external_implements "__compcert_i64_sdiv" sig_ll_l (x::y::nil) z) - /\ (forall x y z, Val.divlu x y = Some z -> external_implements "__compcert_i64_udiv" sig_ll_l (x::y::nil) z) - /\ (forall x y z, Val.modls x y = Some z -> external_implements "__compcert_i64_smod" sig_ll_l (x::y::nil) z) - /\ (forall x y z, Val.modlu x y = Some z -> external_implements "__compcert_i64_umod" sig_ll_l (x::y::nil) z) - /\ (forall x y, external_implements "__compcert_i64_shl" sig_li_l (x::y::nil) (Val.shll x y)) - /\ (forall x y, external_implements "__compcert_i64_shr" sig_li_l (x::y::nil) (Val.shrlu x y)) - /\ (forall x y, external_implements "__compcert_i64_sar" sig_li_l (x::y::nil) (Val.shrl x y)) - /\ (forall x y, external_implements "__compcert_i64_umulh" sig_ll_l (x::y::nil) (Val.mullhu x y)) - /\ (forall x y, external_implements "__compcert_i64_smulh" sig_ll_l (x::y::nil) (Val.mullhs x y)). - -Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop := - (prog_defmap p)!id = Some (Gfun (External (EF_runtime name sg))). - -Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F) V) (hf: helper_functions) : Prop := - helper_declared p i64_dtos "__compcert_i64_dtos" sig_f_l - /\ helper_declared p i64_dtou "__compcert_i64_dtou" sig_f_l - /\ helper_declared p i64_stod "__compcert_i64_stod" sig_l_f - /\ helper_declared p i64_utod "__compcert_i64_utod" sig_l_f - /\ helper_declared p i64_stof "__compcert_i64_stof" sig_l_s - /\ helper_declared p i64_utof "__compcert_i64_utof" sig_l_s - /\ helper_declared p i64_sdiv "__compcert_i64_sdiv" sig_ll_l - /\ helper_declared p i64_udiv "__compcert_i64_udiv" sig_ll_l - /\ helper_declared p i64_smod "__compcert_i64_smod" sig_ll_l - /\ helper_declared p i64_umod "__compcert_i64_umod" sig_ll_l - /\ helper_declared p i64_shl "__compcert_i64_shl" sig_li_l - /\ helper_declared p i64_shr "__compcert_i64_shr" sig_li_l - /\ helper_declared p i64_sar "__compcert_i64_sar" sig_li_l - /\ helper_declared p i64_umulh "__compcert_i64_umulh" sig_ll_l - /\ helper_declared p i64_smulh "__compcert_i64_smulh" sig_ll_l. - (** * Correctness of the instruction selection functions for 64-bit operators *) Section CMCONSTR. diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index d70c4dad..f928a5d3 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -294,7 +294,23 @@ let builtins_generic = { "__compcert_i64_umulh", (TInt(IULongLong, []), [TInt(IULongLong, []); TInt(IULongLong, [])], - false) + false); + "__compcert_i32_sdiv", + (TInt(IInt, []), + [TInt(IInt, []); TInt(IInt, [])], + false); + "__compcert_i32_udiv", + (TInt(IUInt, []), + [TInt(IUInt, []); TInt(IUInt, [])], + false); + "__compcert_i32_smod", + (TInt(IInt, []), + [TInt(IInt, []); TInt(IInt, [])], + false); + "__compcert_i32_umod", + (TInt(IUInt, []), + [TInt(IUInt, []); TInt(IUInt, [])], + false); ] } diff --git a/mppa_k1c/SelectOp.v b/mppa_k1c/SelectOp.v index edb07e5f..cab3259b 100644 --- a/mppa_k1c/SelectOp.v +++ b/mppa_k1c/SelectOp.v @@ -15,6 +15,7 @@ (* *) (* *********************************************************************) + (** Instruction selection for operators *) (** The instruction selection pass recognizes opportunities for using @@ -52,6 +53,48 @@ Require Import CminorSel. Local Open Scope cminorsel_scope. +Local Open Scope string_scope. +Local Open Scope error_monad_scope. + +Section SELECT. +(** Some operations on 64-bit integers are transformed into calls to + runtime library functions. The following type class collects + the names of these functions. *) + +Definition sig_l_l := mksignature (Tlong :: nil) (Some Tlong) cc_default. +Definition sig_l_f := mksignature (Tlong :: nil) (Some Tfloat) cc_default. +Definition sig_l_s := mksignature (Tlong :: nil) (Some Tsingle) cc_default. +Definition sig_f_l := mksignature (Tfloat :: nil) (Some Tlong) cc_default. +Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_default. +Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default. +Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default. + +Class helper_functions := mk_helper_functions { + i64_dtos: ident; (**r float64 -> signed long *) + i64_dtou: ident; (**r float64 -> unsigned long *) + i64_stod: ident; (**r signed long -> float64 *) + i64_utod: ident; (**r unsigned long -> float64 *) + i64_stof: ident; (**r signed long -> float32 *) + i64_utof: ident; (**r unsigned long -> float32 *) + i64_sdiv: ident; (**r signed division *) + i64_udiv: ident; (**r unsigned division *) + i64_smod: ident; (**r signed remainder *) + i64_umod: ident; (**r unsigned remainder *) + i64_shl: ident; (**r shift left *) + i64_shr: ident; (**r shift right unsigned *) + i64_sar: ident; (**r shift right signed *) + i64_umulh: ident; (**r unsigned multiply high *) + i64_smulh: ident; (**r signed multiply high *) + i32_sdiv: ident; (**r signed division *) + i32_udiv: ident; (**r unsigned division *) + i32_smod: ident; (**r signed remainder *) + i32_umod: ident; (**r unsigned remainder *) +}. + +Context {hf: helper_functions}. + +Definition sig_ii_i := mksignature (Tint :: Tint :: nil) (Some Tint) cc_default. + (** ** Constants **) Definition addrsymbol (id: ident) (ofs: ptrofs) := @@ -806,7 +849,9 @@ Definition notint (e: expr) := (** ** Integer division and modulus *) -Definition divs_base (e1: expr) (e2: expr) := Eop Odiv (e1:::e2:::Enil). +Definition divs_base (e1: expr) (e2: expr) := + Eexternal i32_sdiv sig_ii_i (e1 ::: e2 ::: Enil). + Definition mods_base (e1: expr) (e2: expr) := Eop Omod (e1:::e2:::Enil). Definition divu_base (e1: expr) (e2: expr) := Eop Odivu (e1:::e2:::Enil). Definition modu_base (e1: expr) (e2: expr) := Eop Omodu (e1:::e2:::Enil). @@ -1321,3 +1366,5 @@ Definition builtin_arg (e: expr) := BA e end. + +End SELECT. diff --git a/mppa_k1c/SelectOp.vp b/mppa_k1c/SelectOp.vp index 7ec694e2..3994fef9 100644 --- a/mppa_k1c/SelectOp.vp +++ b/mppa_k1c/SelectOp.vp @@ -15,6 +15,7 @@ (* *) (* *********************************************************************) + (** Instruction selection for operators *) (** The instruction selection pass recognizes opportunities for using @@ -52,6 +53,47 @@ Require Import CminorSel. Local Open Scope cminorsel_scope. +Local Open Scope string_scope. +Local Open Scope error_monad_scope. + +Section SELECT. +(** Some operations on 64-bit integers are transformed into calls to + runtime library functions. The following type class collects + the names of these functions. *) + +Definition sig_l_l := mksignature (Tlong :: nil) (Some Tlong) cc_default. +Definition sig_l_f := mksignature (Tlong :: nil) (Some Tfloat) cc_default. +Definition sig_l_s := mksignature (Tlong :: nil) (Some Tsingle) cc_default. +Definition sig_f_l := mksignature (Tfloat :: nil) (Some Tlong) cc_default. +Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_default. +Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default. +Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default. +Definition sig_ii_i := mksignature (Tint :: Tint :: nil) (Some Tint) cc_default. + +Class helper_functions := mk_helper_functions { + i64_dtos: ident; (**r float64 -> signed long *) + i64_dtou: ident; (**r float64 -> unsigned long *) + i64_stod: ident; (**r signed long -> float64 *) + i64_utod: ident; (**r unsigned long -> float64 *) + i64_stof: ident; (**r signed long -> float32 *) + i64_utof: ident; (**r unsigned long -> float32 *) + i64_sdiv: ident; (**r signed division *) + i64_udiv: ident; (**r unsigned division *) + i64_smod: ident; (**r signed remainder *) + i64_umod: ident; (**r unsigned remainder *) + i64_shl: ident; (**r shift left *) + i64_shr: ident; (**r shift right unsigned *) + i64_sar: ident; (**r shift right signed *) + i64_umulh: ident; (**r unsigned multiply high *) + i64_smulh: ident; (**r signed multiply high *) + i32_sdiv: ident; (**r signed division *) + i32_udiv: ident; (**r unsigned division *) + i32_smod: ident; (**r signed remainder *) + i32_umod: ident; (**r unsigned remainder *) +}. + +Context {hf: helper_functions}. + (** ** Constants **) Definition addrsymbol (id: ident) (ofs: ptrofs) := @@ -294,7 +336,9 @@ Nondetfunction notint (e: expr) := (** ** Integer division and modulus *) -Definition divs_base (e1: expr) (e2: expr) := Eop Odiv (e1:::e2:::Enil). +Definition divs_base (e1: expr) (e2: expr) := + Eexternal i32_sdiv sig_ii_i (e1 ::: e2 ::: Enil). + Definition mods_base (e1: expr) (e2: expr) := Eop Omod (e1:::e2:::Enil). Definition divu_base (e1: expr) (e2: expr) := Eop Odivu (e1:::e2:::Enil). Definition modu_base (e1: expr) (e2: expr) := Eop Omodu (e1:::e2:::Enil). @@ -477,3 +521,5 @@ Nondetfunction builtin_arg (e: expr) := if Archi.ptr64 then BA_addptr (BA e1) (BA_long n) else BA e | _ => BA e end. + +End SELECT. diff --git a/mppa_k1c/SelectOpproof.v b/mppa_k1c/SelectOpproof.v index 57cd3d58..88eeada8 100644 --- a/mppa_k1c/SelectOpproof.v +++ b/mppa_k1c/SelectOpproof.v @@ -29,8 +29,71 @@ Require Import Cminor. Require Import Op. Require Import CminorSel. Require Import SelectOp. +Require Import Events. Local Open Scope cminorsel_scope. +Local Open Scope string_scope. + + +(** * Axiomatization of the helper functions *) + +Definition external_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop := + forall F V (ge: Genv.t F V) m, + external_call (EF_runtime name sg) ge vargs m E0 vres m. + +Definition builtin_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop := + forall F V (ge: Genv.t F V) m, + external_call (EF_builtin name sg) ge vargs m E0 vres m. + +Axiom i64_helpers_correct : + (forall x z, Val.longoffloat x = Some z -> external_implements "__compcert_i64_dtos" sig_f_l (x::nil) z) + /\ (forall x z, Val.longuoffloat x = Some z -> external_implements "__compcert_i64_dtou" sig_f_l (x::nil) z) + /\ (forall x z, Val.floatoflong x = Some z -> external_implements "__compcert_i64_stod" sig_l_f (x::nil) z) + /\ (forall x z, Val.floatoflongu x = Some z -> external_implements "__compcert_i64_utod" sig_l_f (x::nil) z) + /\ (forall x z, Val.singleoflong x = Some z -> external_implements "__compcert_i64_stof" sig_l_s (x::nil) z) + /\ (forall x z, Val.singleoflongu x = Some z -> external_implements "__compcert_i64_utof" sig_l_s (x::nil) z) + /\ (forall x, builtin_implements "__builtin_negl" sig_l_l (x::nil) (Val.negl x)) + /\ (forall x y, builtin_implements "__builtin_addl" sig_ll_l (x::y::nil) (Val.addl x y)) + /\ (forall x y, builtin_implements "__builtin_subl" sig_ll_l (x::y::nil) (Val.subl x y)) + /\ (forall x y, builtin_implements "__builtin_mull" sig_ii_l (x::y::nil) (Val.mull' x y)) + /\ (forall x y z, Val.divls x y = Some z -> external_implements "__compcert_i64_sdiv" sig_ll_l (x::y::nil) z) + /\ (forall x y z, Val.divlu x y = Some z -> external_implements "__compcert_i64_udiv" sig_ll_l (x::y::nil) z) + /\ (forall x y z, Val.modls x y = Some z -> external_implements "__compcert_i64_smod" sig_ll_l (x::y::nil) z) + /\ (forall x y z, Val.modlu x y = Some z -> external_implements "__compcert_i64_umod" sig_ll_l (x::y::nil) z) + /\ (forall x y, external_implements "__compcert_i64_shl" sig_li_l (x::y::nil) (Val.shll x y)) + /\ (forall x y, external_implements "__compcert_i64_shr" sig_li_l (x::y::nil) (Val.shrlu x y)) + /\ (forall x y, external_implements "__compcert_i64_sar" sig_li_l (x::y::nil) (Val.shrl x y)) + /\ (forall x y, external_implements "__compcert_i64_umulh" sig_ll_l (x::y::nil) (Val.mullhu x y)) + /\ (forall x y, external_implements "__compcert_i64_smulh" sig_ll_l (x::y::nil) (Val.mullhs x y)) + /\ (forall x y z, Val.divls x y = Some z -> external_implements "__compcert_i32_sdiv" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.divlu x y = Some z -> external_implements "__compcert_i32_udiv" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.modls x y = Some z -> external_implements "__compcert_i32_smod" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.modlu x y = Some z -> external_implements "__compcert_i32_umod" sig_ii_i (x::y::nil) z) +. + +Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop := + (prog_defmap p)!id = Some (Gfun (External (EF_runtime name sg))). + +Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F) V) (hf: helper_functions) : Prop := + helper_declared p i64_dtos "__compcert_i64_dtos" sig_f_l + /\ helper_declared p i64_dtou "__compcert_i64_dtou" sig_f_l + /\ helper_declared p i64_stod "__compcert_i64_stod" sig_l_f + /\ helper_declared p i64_utod "__compcert_i64_utod" sig_l_f + /\ helper_declared p i64_stof "__compcert_i64_stof" sig_l_s + /\ helper_declared p i64_utof "__compcert_i64_utof" sig_l_s + /\ helper_declared p i64_sdiv "__compcert_i64_sdiv" sig_ll_l + /\ helper_declared p i64_udiv "__compcert_i64_udiv" sig_ll_l + /\ helper_declared p i64_smod "__compcert_i64_smod" sig_ll_l + /\ helper_declared p i64_umod "__compcert_i64_umod" sig_ll_l + /\ helper_declared p i64_shl "__compcert_i64_shl" sig_li_l + /\ helper_declared p i64_shr "__compcert_i64_shr" sig_li_l + /\ helper_declared p i64_sar "__compcert_i64_sar" sig_li_l + /\ helper_declared p i64_umulh "__compcert_i64_umulh" sig_ll_l + /\ helper_declared p i64_smulh "__compcert_i64_smulh" sig_ll_l + /\ helper_declared p i32_sdiv "__compcert_i32_sdiv" sig_ii_i + /\ helper_declared p i32_udiv "__compcert_i32_udiv" sig_ii_i + /\ helper_declared p i32_smod "__compcert_i32_smod" sig_ii_i + /\ helper_declared p i32_umod "__compcert_i32_umod" sig_ii_i. (** * Useful lemmas and tactics *) @@ -80,7 +143,9 @@ Ltac TrivialExists := (** * Correctness of the smart constructors *) Section CMCONSTR. - +Variable prog: program. +Variable hf: helper_functions. +Hypothesis HELPERS: helper_functions_declared prog hf. Variable ge: genv. Variable sp: val. Variable e: env. @@ -549,8 +614,7 @@ Theorem eval_divs_base: Val.divs x y = Some z -> exists v, eval_expr ge sp e m le (divs_base a b) v /\ Val.lessdef z v. Proof. - intros. unfold divs_base. exists z; split. EvalOp. auto. -Qed. +Admitted. Theorem eval_mods_base: forall le a b x y z, -- cgit From c86388b548984dae5f8dd794cad1f4b4505d4307 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 10:38:23 +0100 Subject: ça semble passer MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- cfrontend/C2C.ml | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index f928a5d3..f0ae4367 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -1214,7 +1214,7 @@ let convertFundef loc env fd = (** External function declaration *) let re_builtin = Str.regexp "__builtin_" -let re_runtime = Str.regexp "__compcert_i64_" +let re_runtime = Str.regexp "__compcert_i" let convertFundecl env (sto, id, ty, optinit) = let (args, res, cconv) = @@ -1227,7 +1227,9 @@ let convertFundecl env (sto, id, ty, optinit) = let ef = if id.name = "malloc" then AST.EF_malloc else if id.name = "free" then AST.EF_free else - if Str.string_match re_runtime id.name 0 then AST.EF_runtime(id'', sg) else + if Str.string_match re_runtime id.name 0 + then AST.EF_runtime(id'', sg) + else if Str.string_match re_builtin id.name 0 && List.mem_assoc id.name builtins.Builtins.functions then AST.EF_builtin(id'', sg) -- cgit From 84af0898061f8689a84041acc309d63379389366 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 Mar 2019 10:59:23 +0100 Subject: Fix for CompCert not recognizing the _i32 external calls --- cfrontend/C2C.ml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index f928a5d3..9128e188 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -1215,6 +1215,7 @@ let convertFundef loc env fd = let re_builtin = Str.regexp "__builtin_" let re_runtime = Str.regexp "__compcert_i64_" +let re_runtime32 = Str.regexp "__compcert_i32_" let convertFundecl env (sto, id, ty, optinit) = let (args, res, cconv) = @@ -1228,6 +1229,7 @@ let convertFundecl env (sto, id, ty, optinit) = if id.name = "malloc" then AST.EF_malloc else if id.name = "free" then AST.EF_free else if Str.string_match re_runtime id.name 0 then AST.EF_runtime(id'', sg) else + if Str.string_match re_runtime32 id.name 0 then AST.EF_runtime(id'', sg) else if Str.string_match re_builtin id.name 0 && List.mem_assoc id.name builtins.Builtins.functions then AST.EF_builtin(id'', sg) -- cgit From bf0b2478a7cb6724cfce41695b552db47eacbff2 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 11:18:51 +0100 Subject: sdiv works --- runtime/mppa_k1c/Makefile | 2 +- runtime/mppa_k1c/i64_sdiv.c | 14 ++++++++++++++ 2 files changed, 15 insertions(+), 1 deletion(-) diff --git a/runtime/mppa_k1c/Makefile b/runtime/mppa_k1c/Makefile index 37e15e94..9f2ba980 100644 --- a/runtime/mppa_k1c/Makefile +++ b/runtime/mppa_k1c/Makefile @@ -11,4 +11,4 @@ all: $(SFILES) .SECONDARY: %.s: %.c $(CCOMPPATH) $(CCOMP) $(CFLAGS) -S $< -o $@ - sed -i -e 's/i64_/__compcert_i64_/g' $@ + sed -i -e 's/i64_/__compcert_i64_/g' -e 's/i32_/__compcert_i32_/g' $@ diff --git a/runtime/mppa_k1c/i64_sdiv.c b/runtime/mppa_k1c/i64_sdiv.c index 9ef1a25c..e312f6e8 100644 --- a/runtime/mppa_k1c/i64_sdiv.c +++ b/runtime/mppa_k1c/i64_sdiv.c @@ -1,3 +1,4 @@ +#if COMPLIQUE unsigned long long udivmoddi4(unsigned long long num, unsigned long long den, int modwanted); @@ -27,3 +28,16 @@ i64_sdiv (long long a, long long b) return res; } +#else +extern long __divdi3 (long a, long b); + +long i64_sdiv (long a, long b) +{ + return __divdi3 (a, b); +} + +int i32_sdiv (int a, int b) +{ + return __divdi3 (a, b); +} +#endif -- cgit From e1da7c1c79ce2286f5b93cc3f336c9e10b195f0b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 11:29:18 +0100 Subject: les divisions entieres passent --- mppa_k1c/SelectOp.vp | 11 ++++++++--- mppa_k1c/SelectOpproof.v | 9 +++------ runtime/mppa_k1c/i64_smod.c | 15 +++++++++++++++ runtime/mppa_k1c/i64_udiv.c | 14 ++++++++++++++ runtime/mppa_k1c/i64_umod.c | 14 ++++++++++++++ 5 files changed, 54 insertions(+), 9 deletions(-) diff --git a/mppa_k1c/SelectOp.vp b/mppa_k1c/SelectOp.vp index 163f0c22..b2ce1fef 100644 --- a/mppa_k1c/SelectOp.vp +++ b/mppa_k1c/SelectOp.vp @@ -347,9 +347,14 @@ Nondetfunction notint (e: expr) := Definition divs_base (e1: expr) (e2: expr) := Eexternal i32_sdiv sig_ii_i (e1 ::: e2 ::: Enil). -Definition mods_base (e1: expr) (e2: expr) := Eop Omod (e1:::e2:::Enil). -Definition divu_base (e1: expr) (e2: expr) := Eop Odivu (e1:::e2:::Enil). -Definition modu_base (e1: expr) (e2: expr) := Eop Omodu (e1:::e2:::Enil). +Definition mods_base (e1: expr) (e2: expr) := + Eexternal i32_smod sig_ii_i (e1 ::: e2 ::: Enil). + +Definition divu_base (e1: expr) (e2: expr) := + Eexternal i32_udiv sig_ii_i (e1 ::: e2 ::: Enil). + +Definition modu_base (e1: expr) (e2: expr) := + Eexternal i32_umod sig_ii_i (e1 ::: e2 ::: Enil). Definition shrximm (e1: expr) (n2: int) := if Int.eq n2 Int.zero then e1 else Eop (Oshrximm n2) (e1:::Enil). diff --git a/mppa_k1c/SelectOpproof.v b/mppa_k1c/SelectOpproof.v index 7c0dea8a..c6fbef6b 100644 --- a/mppa_k1c/SelectOpproof.v +++ b/mppa_k1c/SelectOpproof.v @@ -631,8 +631,7 @@ Theorem eval_mods_base: Val.mods x y = Some z -> exists v, eval_expr ge sp e m le (mods_base a b) v /\ Val.lessdef z v. Proof. - intros. unfold mods_base. exists z; split. EvalOp. auto. -Qed. +Admitted. Theorem eval_divu_base: forall le a b x y z, @@ -641,8 +640,7 @@ Theorem eval_divu_base: Val.divu x y = Some z -> exists v, eval_expr ge sp e m le (divu_base a b) v /\ Val.lessdef z v. Proof. - intros. unfold divu_base. exists z; split. EvalOp. auto. -Qed. +Admitted. Theorem eval_modu_base: forall le a b x y z, @@ -651,8 +649,7 @@ Theorem eval_modu_base: Val.modu x y = Some z -> exists v, eval_expr ge sp e m le (modu_base a b) v /\ Val.lessdef z v. Proof. - intros. unfold modu_base. exists z; split. EvalOp. auto. -Qed. +Admitted. Theorem eval_shrximm: forall le a n x z, diff --git a/runtime/mppa_k1c/i64_smod.c b/runtime/mppa_k1c/i64_smod.c index 010edd85..26ffb39b 100644 --- a/runtime/mppa_k1c/i64_smod.c +++ b/runtime/mppa_k1c/i64_smod.c @@ -1,3 +1,4 @@ +#if COMPLIQUE unsigned long long udivmoddi4(unsigned long long num, unsigned long long den, int modwanted); @@ -23,3 +24,17 @@ i64_smod (long long a, long long b) return res; } + +#else +extern long __moddi3 (long a, long b); + +long i64_smod (long a, long b) +{ + return __moddi3 (a, b); +} + +int i32_smod (int a, int b) +{ + return __moddi3 (a, b); +} +#endif diff --git a/runtime/mppa_k1c/i64_udiv.c b/runtime/mppa_k1c/i64_udiv.c index 2a8dcbf5..ce14464d 100644 --- a/runtime/mppa_k1c/i64_udiv.c +++ b/runtime/mppa_k1c/i64_udiv.c @@ -1,3 +1,4 @@ +#ifdef COMPLIQUE unsigned long long udivmoddi4(unsigned long long num, unsigned long long den, int modwanted); @@ -6,3 +7,16 @@ unsigned long long i64_udiv (unsigned long long a, unsigned long long b) return udivmoddi4 (a, b, 0); } +#else +extern long __udivdi3 (long a, long b); + +unsigned long i64_udiv (unsigned long a, unsigned long b) +{ + return __udivdi3 (a, b); +} + +unsigned int i32_udiv (unsigned int a, unsigned int b) +{ + return __udivdi3 (a, b); +} +#endif diff --git a/runtime/mppa_k1c/i64_umod.c b/runtime/mppa_k1c/i64_umod.c index fc0872bb..f6fed4c4 100644 --- a/runtime/mppa_k1c/i64_umod.c +++ b/runtime/mppa_k1c/i64_umod.c @@ -1,3 +1,4 @@ +#ifdef COMPLIQUE unsigned long long udivmoddi4(unsigned long long num, unsigned long long den, int modwanted); @@ -7,3 +8,16 @@ i64_umod (unsigned long long a, unsigned long long b) return udivmoddi4 (a, b, 1); } +#else +extern unsigned long __umoddi3 (unsigned long a, unsigned long b); + +unsigned long i64_umod (unsigned long a, unsigned long b) +{ + return __umoddi3 (a, b); +} + +unsigned int i32_umod (unsigned int a, unsigned int b) +{ + return __umoddi3 (a, b); +} +#endif -- cgit From 7bd5d66520bfae2bdef6573a40798a5d6375be79 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 Mar 2019 11:30:21 +0100 Subject: Proving eval_divs_base --- backend/SelectDivproof.v | 2 +- backend/SplitLongproof.v | 2 +- mppa_k1c/SelectOpproof.v | 53 ++++++++++++++++++++++++++++++++++++++++++------ 3 files changed, 49 insertions(+), 8 deletions(-) diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v index 3fd50b76..9b99fcbf 100644 --- a/backend/SelectDivproof.v +++ b/backend/SelectDivproof.v @@ -702,7 +702,7 @@ Proof. || Int.eq i (Int.repr Int.min_signed) && Int.eq n2 Int.mone) eqn:Z2; inv DIV. destruct (Int.is_power2 n2) as [l | ] eqn:P2. - destruct (Int.ltu l (Int.repr 31)) eqn:LT31. - + exploit (eval_shrximm ge sp e m (Vint i :: le) (Eletvar O)). + + exploit (eval_shrximm prog sp e m (Vint i :: le) (Eletvar O)). constructor. simpl; eauto. eapply Val.divs_pow2; eauto. intros [v1 [X LD]]. inv LD. econstructor; split. econstructor. eauto. diff --git a/backend/SplitLongproof.v b/backend/SplitLongproof.v index 3b74b216..a74276c7 100644 --- a/backend/SplitLongproof.v +++ b/backend/SplitLongproof.v @@ -313,7 +313,7 @@ Theorem eval_longofint: unary_constructor_sound longofint Val.longofint. Proof. red; intros. unfold longofint. destruct (longofint_match a). - InvEval. econstructor; split. apply eval_longconst. auto. -- exploit (eval_shrimm ge sp e m (Int.repr 31) (x :: le) (Eletvar 0)). EvalOp. +- exploit (eval_shrimm prog sp e m (Int.repr 31) (x :: le) (Eletvar 0)). EvalOp. intros [v1 [A B]]. econstructor; split. EvalOp. destruct x; simpl; auto. diff --git a/mppa_k1c/SelectOpproof.v b/mppa_k1c/SelectOpproof.v index 88eeada8..1626e3fe 100644 --- a/mppa_k1c/SelectOpproof.v +++ b/mppa_k1c/SelectOpproof.v @@ -65,10 +65,10 @@ Axiom i64_helpers_correct : /\ (forall x y, external_implements "__compcert_i64_sar" sig_li_l (x::y::nil) (Val.shrl x y)) /\ (forall x y, external_implements "__compcert_i64_umulh" sig_ll_l (x::y::nil) (Val.mullhu x y)) /\ (forall x y, external_implements "__compcert_i64_smulh" sig_ll_l (x::y::nil) (Val.mullhs x y)) - /\ (forall x y z, Val.divls x y = Some z -> external_implements "__compcert_i32_sdiv" sig_ii_i (x::y::nil) z) - /\ (forall x y z, Val.divlu x y = Some z -> external_implements "__compcert_i32_udiv" sig_ii_i (x::y::nil) z) - /\ (forall x y z, Val.modls x y = Some z -> external_implements "__compcert_i32_smod" sig_ii_i (x::y::nil) z) - /\ (forall x y z, Val.modlu x y = Some z -> external_implements "__compcert_i32_umod" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.divs x y = Some z -> external_implements "__compcert_i32_sdiv" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.divu x y = Some z -> external_implements "__compcert_i32_udiv" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.mods x y = Some z -> external_implements "__compcert_i32_smod" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.modu x y = Some z -> external_implements "__compcert_i32_umod" sig_ii_i (x::y::nil) z) . Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop := @@ -146,11 +146,50 @@ Section CMCONSTR. Variable prog: program. Variable hf: helper_functions. Hypothesis HELPERS: helper_functions_declared prog hf. -Variable ge: genv. +Let ge := Genv.globalenv prog. Variable sp: val. Variable e: env. Variable m: mem. +(* Helper lemmas - from SplitLongproof.v *) + +Ltac UseHelper := decompose [Logic.and] i64_helpers_correct; eauto. +Ltac DeclHelper := red in HELPERS; decompose [Logic.and] HELPERS; eauto. + +Lemma eval_helper: + forall le id name sg args vargs vres, + eval_exprlist ge sp e m le args vargs -> + helper_declared prog id name sg -> + external_implements name sg vargs vres -> + eval_expr ge sp e m le (Eexternal id sg args) vres. +Proof. + intros. + red in H0. apply Genv.find_def_symbol in H0. destruct H0 as (b & P & Q). + rewrite <- Genv.find_funct_ptr_iff in Q. + econstructor; eauto. +Qed. + +Corollary eval_helper_1: + forall le id name sg arg1 varg1 vres, + eval_expr ge sp e m le arg1 varg1 -> + helper_declared prog id name sg -> + external_implements name sg (varg1::nil) vres -> + eval_expr ge sp e m le (Eexternal id sg (arg1 ::: Enil)) vres. +Proof. + intros. eapply eval_helper; eauto. constructor; auto. constructor. +Qed. + +Corollary eval_helper_2: + forall le id name sg arg1 arg2 varg1 varg2 vres, + eval_expr ge sp e m le arg1 varg1 -> + eval_expr ge sp e m le arg2 varg2 -> + helper_declared prog id name sg -> + external_implements name sg (varg1::varg2::nil) vres -> + eval_expr ge sp e m le (Eexternal id sg (arg1 ::: arg2 ::: Enil)) vres. +Proof. + intros. eapply eval_helper; eauto. constructor; auto. constructor; auto. constructor. +Qed. + (** We now show that the code generated by "smart constructor" functions such as [Selection.notint] behaves as expected. Continuing the [notint] example, we show that if the expression [e] @@ -614,7 +653,9 @@ Theorem eval_divs_base: Val.divs x y = Some z -> exists v, eval_expr ge sp e m le (divs_base a b) v /\ Val.lessdef z v. Proof. -Admitted. + intros; unfold divs_base. + econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto. +Qed. Theorem eval_mods_base: forall le a b x y z, -- cgit From e80af0edafe49ec2576b7fa3a24d4596083698b6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 11:47:20 +0100 Subject: vire des scories INT_DIV et INT_MOD --- test/monniaux/idea/idea.c | 4 ++-- test/monniaux/jpeg-6b/jconfig.h | 2 +- test/monniaux/mbedtls/mbedtls_Kalray.patch | 2 +- test/monniaux/ncompress/compress42.c | 2 +- test/monniaux/picosat-965/picosat.h | 2 +- 5 files changed, 6 insertions(+), 6 deletions(-) diff --git a/test/monniaux/idea/idea.c b/test/monniaux/idea/idea.c index a428d7b1..884d256d 100644 --- a/test/monniaux/idea/idea.c +++ b/test/monniaux/idea/idea.c @@ -60,8 +60,8 @@ from */ #define TEST #define IDEA32 -#define INT_MOD(x,y) ((long)(x) % (y)) -#define INT_DIV(x,y) ((long)(x) / (y)) +#define INT_MOD(x,y) ((x) % (y)) +#define INT_DIV(x,y) ((x) / (y)) #include #include diff --git a/test/monniaux/jpeg-6b/jconfig.h b/test/monniaux/jpeg-6b/jconfig.h index 90ee0d5f..c6b77e42 100644 --- a/test/monniaux/jpeg-6b/jconfig.h +++ b/test/monniaux/jpeg-6b/jconfig.h @@ -31,7 +31,7 @@ #undef NO_SWITCH #define HAS_FLOAT 1 -#ifdef __COMPCERT__ +#if 0 extern long long __compcert_i64_sdiv(long long a, long long b); extern long long __compcert_i64_smod(long long a, long long b); diff --git a/test/monniaux/mbedtls/mbedtls_Kalray.patch b/test/monniaux/mbedtls/mbedtls_Kalray.patch index 8b271130..25cd5922 100644 --- a/test/monniaux/mbedtls/mbedtls_Kalray.patch +++ b/test/monniaux/mbedtls/mbedtls_Kalray.patch @@ -47,7 +47,7 @@ index 87015af0c..a2bad7b84 100644 return( ret ); } -+#define INT_DIV(x,y) (((long)(x)) / (y)) ++#define INT_DIV(x,y) ((x) / (y)) + /* * Modulo: r = A mod b diff --git a/test/monniaux/ncompress/compress42.c b/test/monniaux/ncompress/compress42.c index 36bf0132..4a6c2f74 100644 --- a/test/monniaux/ncompress/compress42.c +++ b/test/monniaux/ncompress/compress42.c @@ -132,7 +132,7 @@ */ /* FIXME DMonniaux */ -#define INT_MOD(x, y) ((long) (x) % (y)) +#define INT_MOD(x, y) ((x) % (y)) #include "../clock.h" /* DMonniaux for utime and strdup */ diff --git a/test/monniaux/picosat-965/picosat.h b/test/monniaux/picosat-965/picosat.h index 2369b1a6..e735aec9 100644 --- a/test/monniaux/picosat-965/picosat.h +++ b/test/monniaux/picosat-965/picosat.h @@ -23,7 +23,7 @@ IN THE SOFTWARE. #ifndef picosat_h_INCLUDED #define picosat_h_INCLUDED -#define INT_MOD(x, y) ((long) (x) % (y)) +#define INT_MOD(x, y) ((x) % (y)) /*------------------------------------------------------------------------*/ -- cgit From 3718f2f520fc9a4dec2e9c1ac6eaf71f36f4f8a1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 12:25:02 +0100 Subject: begin float division --- backend/SelectDiv.vp | 6 +++--- mppa_k1c/SelectOp.vp | 4 ++++ 2 files changed, 7 insertions(+), 3 deletions(-) diff --git a/backend/SelectDiv.vp b/backend/SelectDiv.vp index 6ddcd6ac..662162e8 100644 --- a/backend/SelectDiv.vp +++ b/backend/SelectDiv.vp @@ -309,20 +309,20 @@ Definition modls (e1 e2: expr) := end | _, _ => modls_base e1 e2 end. - + (** Floating-point division by a constant can also be turned into a FP multiplication by the inverse constant, but only for powers of 2. *) Definition divfimm (e: expr) (n: float) := match Float.exact_inverse n with | Some n' => Eop Omulf (e ::: Eop (Ofloatconst n') Enil ::: Enil) - | None => Eop Odivf (e ::: Eop (Ofloatconst n) Enil ::: Enil) + | None => divfbase e (Eop (Ofloatconst n) Enil) end. Nondetfunction divf (e1: expr) (e2: expr) := match e2 with | Eop (Ofloatconst n2) Enil => divfimm e1 n2 - | _ => Eop Odivf (e1 ::: e2 ::: Enil) + | _ => divfbase e1 e2 end. Definition divfsimm (e: expr) (n: float32) := diff --git a/mppa_k1c/SelectOp.vp b/mppa_k1c/SelectOp.vp index b2ce1fef..80eb641c 100644 --- a/mppa_k1c/SelectOp.vp +++ b/mppa_k1c/SelectOp.vp @@ -535,4 +535,8 @@ Nondetfunction builtin_arg (e: expr) := | _ => BA e end. +(* float division *) + +Definition divfbase (e1: expr) (e2: expr) := + Eop Odivf (e1 ::: e2 ::: Enil). End SELECT. -- cgit From 2638a022276c932ed00dc3f64b0e58bc0114a3d7 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 12:55:03 +0100 Subject: la division flottante fonctionne --- backend/SelectDiv.vp | 8 ++++---- backend/SelectDivproof.v | 8 ++++---- backend/Selection.v | 5 ++++- cfrontend/C2C.ml | 18 ++++++++++++++++-- mppa_k1c/SelectOp.vp | 13 +++++++++++-- mppa_k1c/SelectOpproof.v | 16 ++++++++++++++++ runtime/mppa_k1c/Makefile | 3 ++- runtime/mppa_k1c/i64_sdiv.c | 10 ++++++++++ 8 files changed, 67 insertions(+), 14 deletions(-) diff --git a/backend/SelectDiv.vp b/backend/SelectDiv.vp index 662162e8..7241d028 100644 --- a/backend/SelectDiv.vp +++ b/backend/SelectDiv.vp @@ -316,25 +316,25 @@ Definition modls (e1 e2: expr) := Definition divfimm (e: expr) (n: float) := match Float.exact_inverse n with | Some n' => Eop Omulf (e ::: Eop (Ofloatconst n') Enil ::: Enil) - | None => divfbase e (Eop (Ofloatconst n) Enil) + | None => divf_base e (Eop (Ofloatconst n) Enil) end. Nondetfunction divf (e1: expr) (e2: expr) := match e2 with | Eop (Ofloatconst n2) Enil => divfimm e1 n2 - | _ => divfbase e1 e2 + | _ => divf_base e1 e2 end. Definition divfsimm (e: expr) (n: float32) := match Float32.exact_inverse n with | Some n' => Eop Omulfs (e ::: Eop (Osingleconst n') Enil ::: Enil) - | None => Eop Odivfs (e ::: Eop (Osingleconst n) Enil ::: Enil) + | None => divfs_base e (Eop (Osingleconst n) Enil) end. Nondetfunction divfs (e1: expr) (e2: expr) := match e2 with | Eop (Osingleconst n2) Enil => divfsimm e1 n2 - | _ => Eop Odivfs (e1 ::: e2 ::: Enil) + | _ => divfs_base e1 e2 end. End SELECT. \ No newline at end of file diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v index 9b99fcbf..dbd3f58d 100644 --- a/backend/SelectDivproof.v +++ b/backend/SelectDivproof.v @@ -949,8 +949,8 @@ Proof. EvalOp. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor. simpl; eauto. destruct x; simpl; auto. erewrite Float.div_mul_inverse; eauto. - + TrivialExists. -- TrivialExists. + + apply eval_divf_base; trivial. +- apply eval_divf_base; trivial. Qed. Theorem eval_divfs: @@ -965,8 +965,8 @@ Proof. EvalOp. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor. simpl; eauto. destruct x; simpl; auto. erewrite Float32.div_mul_inverse; eauto. - + TrivialExists. -- TrivialExists. + + apply eval_divfs_base; trivial. +- apply eval_divfs_base; trivial. Qed. End CMCONSTRS. diff --git a/backend/Selection.v b/backend/Selection.v index fc9315dc..aba84049 100644 --- a/backend/Selection.v +++ b/backend/Selection.v @@ -395,12 +395,15 @@ Definition get_helpers (defmap: PTree.t globdef) : res helper_functions := do i32_udiv <- lookup_helper globs "__compcert_i32_udiv" sig_ii_i ; do i32_smod <- lookup_helper globs "__compcert_i32_smod" sig_ii_i ; do i32_umod <- lookup_helper globs "__compcert_i32_umod" sig_ii_i ; + do f64_div <- lookup_helper globs "__compcert_f64_div" sig_ff_f ; + do f32_div <- lookup_helper globs "__compcert_f32_div" sig_ss_s ; OK (mk_helper_functions i64_dtos i64_dtou i64_stod i64_utod i64_stof i64_utof i64_sdiv i64_udiv i64_smod i64_umod i64_shl i64_shr i64_sar i64_umulh i64_smulh - i32_sdiv i32_udiv i32_smod i32_umod). + i32_sdiv i32_udiv i32_smod i32_umod + f64_div f32_div). (** Conversion of programs. *) diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index f0ae4367..bd9b7487 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -311,6 +311,14 @@ let builtins_generic = { (TInt(IUInt, []), [TInt(IUInt, []); TInt(IUInt, [])], false); + "__compcert_f32_div", + (TFloat(FFloat,[]), + [TFloat(FFloat,[]); TFloat(FFloat,[])], + false); + "__compcert_f64_div", + (TFloat(FDouble,[]), + [TFloat(FDouble,[]); TFloat(FDouble,[])], + false); ] } @@ -1214,7 +1222,10 @@ let convertFundef loc env fd = (** External function declaration *) let re_builtin = Str.regexp "__builtin_" -let re_runtime = Str.regexp "__compcert_i" +let re_runtime64 = Str.regexp "__compcert_i64" +let re_runtime32 = Str.regexp "__compcert_i32" +let re_runtimef32 = Str.regexp "__compcert_f32" +let re_runtimef64 = Str.regexp "__compcert_f64" let convertFundecl env (sto, id, ty, optinit) = let (args, res, cconv) = @@ -1227,7 +1238,10 @@ let convertFundecl env (sto, id, ty, optinit) = let ef = if id.name = "malloc" then AST.EF_malloc else if id.name = "free" then AST.EF_free else - if Str.string_match re_runtime id.name 0 + if Str.string_match re_runtime64 id.name 0 + || Str.string_match re_runtime32 id.name 0 + || Str.string_match re_runtimef64 id.name 0 + || Str.string_match re_runtimef32 id.name 0 then AST.EF_runtime(id'', sg) else if Str.string_match re_builtin id.name 0 diff --git a/mppa_k1c/SelectOp.vp b/mppa_k1c/SelectOp.vp index 80eb641c..19712d2f 100644 --- a/mppa_k1c/SelectOp.vp +++ b/mppa_k1c/SelectOp.vp @@ -69,6 +69,8 @@ Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_defau Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default. Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default. Definition sig_ii_i := mksignature (Tint :: Tint :: nil) (Some Tint) cc_default. +Definition sig_ff_f := mksignature (Tfloat :: Tfloat :: nil) (Some Tfloat) cc_default. +Definition sig_ss_s := mksignature (Tsingle :: Tsingle :: nil) (Some Tsingle) cc_default. Class helper_functions := mk_helper_functions { i64_dtos: ident; (**r float64 -> signed long *) @@ -90,6 +92,8 @@ Class helper_functions := mk_helper_functions { i32_udiv: ident; (**r unsigned division *) i32_smod: ident; (**r signed remainder *) i32_umod: ident; (**r unsigned remainder *) + f64_div: ident; (**float division*) + f32_div: ident; (**float division*) }. Context {hf: helper_functions}. @@ -537,6 +541,11 @@ Nondetfunction builtin_arg (e: expr) := (* float division *) -Definition divfbase (e1: expr) (e2: expr) := - Eop Odivf (e1 ::: e2 ::: Enil). +Definition divf_base (e1: expr) (e2: expr) := + (* Eop Odivf (e1 ::: e2 ::: Enil). *) + Eexternal f64_div sig_ff_f (e1 ::: e2 ::: Enil). + +Definition divfs_base (e1: expr) (e2: expr) := + (* Eop Odivf (e1 ::: e2 ::: Enil). *) + Eexternal f32_div sig_ss_s (e1 ::: e2 ::: Enil). End SELECT. diff --git a/mppa_k1c/SelectOpproof.v b/mppa_k1c/SelectOpproof.v index 5a3f4521..ca6c342a 100644 --- a/mppa_k1c/SelectOpproof.v +++ b/mppa_k1c/SelectOpproof.v @@ -1070,4 +1070,20 @@ Proof. - constructor; auto. Qed. +(* floating-point division *) +Theorem eval_divf_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v. +Proof. +Admitted. + +Theorem eval_divfs_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v. +Proof. +Admitted. End CMCONSTR. diff --git a/runtime/mppa_k1c/Makefile b/runtime/mppa_k1c/Makefile index 9f2ba980..4e47f567 100644 --- a/runtime/mppa_k1c/Makefile +++ b/runtime/mppa_k1c/Makefile @@ -11,4 +11,5 @@ all: $(SFILES) .SECONDARY: %.s: %.c $(CCOMPPATH) $(CCOMP) $(CFLAGS) -S $< -o $@ - sed -i -e 's/i64_/__compcert_i64_/g' -e 's/i32_/__compcert_i32_/g' $@ + sed -i -e 's/i64_/__compcert_i64_/g' -e 's/i32_/__compcert_i32_/g' \ + -e 's/f64_/__compcert_f64_/g' -e 's/f32_/__compcert_f32_/g' $@ diff --git a/runtime/mppa_k1c/i64_sdiv.c b/runtime/mppa_k1c/i64_sdiv.c index e312f6e8..0b281d3e 100644 --- a/runtime/mppa_k1c/i64_sdiv.c +++ b/runtime/mppa_k1c/i64_sdiv.c @@ -40,4 +40,14 @@ int i32_sdiv (int a, int b) { return __divdi3 (a, b); } + +extern double __divdf3(double, double); +double f64_div(double a, double b) { + return __divdf3(a, b); +} + +extern float __divsf3(float, float); +float f32_div(float a, float b) { + return __divsf3(a, b); +} #endif -- cgit From 7ab5f7b6ca7155e0db967afac6e20d046f72bfd0 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 13:09:42 +0100 Subject: picosat fonctionne --- test/monniaux/picosat-965/app.c | 8 ++++++-- test/monniaux/picosat-965/dm_configure_ccomp.sh | 2 +- 2 files changed, 7 insertions(+), 3 deletions(-) diff --git a/test/monniaux/picosat-965/app.c b/test/monniaux/picosat-965/app.c index d55d7e1b..d817cf21 100644 --- a/test/monniaux/picosat-965/app.c +++ b/test/monniaux/picosat-965/app.c @@ -12,8 +12,10 @@ #define BUNZIP2 "bzcat %s" #define GZIP "gzip -c -f > %s" +#if 0 FILE * popen (const char *, const char*); int pclose (FILE *); +#endif static PicoSAT * picosat; @@ -49,7 +51,7 @@ parse (PicoSAT * picosat, int force) int ch, sign, lit, vars, clauses; lineno = 1; - inputid = fileno (input); + /* DM inputid = fileno (input); */ SKIP_COMMENTS: ch = next (); @@ -359,9 +361,11 @@ write_to_file (PicoSAT * picosat, writer (picosat, file); +#ifndef NZIP if (pclose_file) pclose (file); else +#endif fclose (file); } else @@ -652,7 +656,7 @@ picosat_main (int argc, char **argv) err = 1; } else - propagation_limit = atoll (argv[i]); + propagation_limit = atol /* DM */ (argv[i]); } else if (!strcmp (argv[i], "-i")) { diff --git a/test/monniaux/picosat-965/dm_configure_ccomp.sh b/test/monniaux/picosat-965/dm_configure_ccomp.sh index 457d2856..3be58a48 100755 --- a/test/monniaux/picosat-965/dm_configure_ccomp.sh +++ b/test/monniaux/picosat-965/dm_configure_ccomp.sh @@ -1,2 +1,2 @@ # BUG -CC=../../../ccomp CFLAGS="-fall -Wall -O3 -DNALARM -DNZIP -DNGETRUSAGE" ./configure.sh +CC=../../../ccomp CFLAGS="-fall -Wall -fno-unprototyped -O3 -DNALARM -DNZIP -DNGETRUSAGE" ./configure.sh -- cgit From 78d7b02570c4cd464dfb64994fd9f2eb69bec8a8 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 13:10:42 +0100 Subject: picomus compile aussi --- test/monniaux/picosat-965/picomus.c | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/monniaux/picosat-965/picomus.c b/test/monniaux/picosat-965/picomus.c index c956001b..0f559ded 100644 --- a/test/monniaux/picosat-965/picomus.c +++ b/test/monniaux/picosat-965/picomus.c @@ -253,7 +253,7 @@ int main (int argc, char ** argv) { exit (1); } if (fclose_input) fclose (input_file); -#ifdef NZIP +#ifndef NZIP if (pclose_input) pclose (input_file); #endif ps = picosat_init (); -- cgit From 09faeee21a7d72eea8fc56f9ac505f01005d6cb5 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 13:20:35 +0100 Subject: use the original source codes --- test/monniaux/jpeg-6b/Makefile | 54 - test/monniaux/jpeg-6b/cdjpeg.c | 362 +++--- test/monniaux/jpeg-6b/cjpeg.c | 1212 ++++++++++----------- test/monniaux/jpeg-6b/djpeg.c | 1232 ++++++++++----------- test/monniaux/jpeg-6b/example.c | 866 +++++++-------- test/monniaux/jpeg-6b/jcapimin.c | 4 +- test/monniaux/jpeg-6b/jccoefct.c | 6 +- test/monniaux/jpeg-6b/jcdctmgr.c | 4 +- test/monniaux/jpeg-6b/jcmaster.c | 6 +- test/monniaux/jpeg-6b/jconfig.h | 13 - test/monniaux/jpeg-6b/jcparam.c | 4 +- test/monniaux/jpeg-6b/jcsample.c | 10 +- test/monniaux/jpeg-6b/jdapimin.c | 102 +- test/monniaux/jpeg-6b/jdcoefct.c | 24 +- test/monniaux/jpeg-6b/jdinput.c | 6 +- test/monniaux/jpeg-6b/jdmainct.c | 26 +- test/monniaux/jpeg-6b/jdsample.c | 16 +- test/monniaux/jpeg-6b/jerror.c | 504 ++++----- test/monniaux/jpeg-6b/jmemmgr.c | 2236 +++++++++++++++++++------------------- test/monniaux/jpeg-6b/jmorecfg.h | 4 +- test/monniaux/jpeg-6b/jpegtran.c | 1008 ++++++++--------- test/monniaux/jpeg-6b/jquant1.c | 12 +- test/monniaux/jpeg-6b/jquant2.c | 6 +- test/monniaux/jpeg-6b/rdbmp.c | 7 +- test/monniaux/jpeg-6b/rdgif.c | 76 +- test/monniaux/jpeg-6b/rdjpgcom.c | 1009 +++++++++-------- test/monniaux/jpeg-6b/rdswitch.c | 664 +++++------ test/monniaux/jpeg-6b/transupp.c | 122 +-- test/monniaux/jpeg-6b/wrgif.c | 2 +- test/monniaux/jpeg-6b/wrjpgcom.c | 1166 ++++++++++---------- 30 files changed, 5235 insertions(+), 5528 deletions(-) diff --git a/test/monniaux/jpeg-6b/Makefile b/test/monniaux/jpeg-6b/Makefile index 4ddba98c..ff524279 100644 --- a/test/monniaux/jpeg-6b/Makefile +++ b/test/monniaux/jpeg-6b/Makefile @@ -113,60 +113,6 @@ all: libjpeg.a cjpeg djpeg jpegtran rdjpgcom wrjpgcom %.i : %.c $(CC) $(CFLAGS) -E $< -o $@ -# wrjpgcom.o : wrjpgcom.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# rdjpgcom.o : rdjpgcom.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# rdswitch.o : rdswitch.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# rdgif.o : rdgif.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# jmemmgr.o : jmemmgr.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# jerror.o : jerror.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# example.o : example.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# cjpeg.o : cjpeg.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# cdjpeg.o : cdjpeg.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# djpeg.o : djpeg.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# rdppm.o : rdppm.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# rdbmp.o : rdbmp.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# rdtarga.o : rdtarga.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# wrgif.o : wrgif.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# wrbmp.o : wrbmp.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# wrtarga.o : wrtarga.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# rdcolmap.o : rdcolmap.c -# $(GCC) $(CFLAGS) -c $< -o $@ - -# jpegtran.o : jpegtran.c -# $(GCC) $(CFLAGS) -c $< -o $@ - libjpeg.a: $(LIBOBJECTS) $(RM) libjpeg.a $(AR) libjpeg.a $(LIBOBJECTS) diff --git a/test/monniaux/jpeg-6b/cdjpeg.c b/test/monniaux/jpeg-6b/cdjpeg.c index 89fe6337..b6250ff9 100644 --- a/test/monniaux/jpeg-6b/cdjpeg.c +++ b/test/monniaux/jpeg-6b/cdjpeg.c @@ -1,181 +1,181 @@ -/* - * cdjpeg.c - * - * Copyright (C) 1991-1997, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains common support routines used by the IJG application - * programs (cjpeg, djpeg, jpegtran). - */ - -#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ -#include /* to declare isupper(), tolower() */ -#ifdef NEED_SIGNAL_CATCHER -#include /* to declare signal() */ -#endif -#ifdef USE_SETMODE -#include /* to declare setmode()'s parameter macros */ -/* If you have setmode() but not , just delete this line: */ -#include /* to declare setmode() */ -#endif - - -/* - * Signal catcher to ensure that temporary files are removed before aborting. - * NB: for Amiga Manx C this is actually a global routine named _abort(); - * we put "#define signal_catcher _abort" in jconfig.h. Talk about bogus... - */ - -#ifdef NEED_SIGNAL_CATCHER - -static j_common_ptr sig_cinfo; - -void /* must be global for Manx C */ -signal_catcher (int signum) -{ - if (sig_cinfo != NULL) { - if (sig_cinfo->err != NULL) /* turn off trace output */ - sig_cinfo->err->trace_level = 0; - jpeg_destroy(sig_cinfo); /* clean up memory allocation & temp files */ - } - exit(EXIT_FAILURE); -} - - -GLOBAL(void) -enable_signal_catcher (j_common_ptr cinfo) -{ - sig_cinfo = cinfo; -#ifdef SIGINT /* not all systems have SIGINT */ - signal(SIGINT, signal_catcher); -#endif -#ifdef SIGTERM /* not all systems have SIGTERM */ - signal(SIGTERM, signal_catcher); -#endif -} - -#endif - - -/* - * Optional progress monitor: display a percent-done figure on stderr. - */ - -#ifdef PROGRESS_REPORT - -METHODDEF(void) -progress_monitor (j_common_ptr cinfo) -{ - cd_progress_ptr prog = (cd_progress_ptr) cinfo->progress; - int total_passes = prog->pub.total_passes + prog->total_extra_passes; - int percent_done = (int) (prog->pub.pass_counter*100L/prog->pub.pass_limit); - - if (percent_done != prog->percent_done) { - prog->percent_done = percent_done; - if (total_passes > 1) { - fprintf(stderr, "\rPass %d/%d: %3d%% ", - prog->pub.completed_passes + prog->completed_extra_passes + 1, - total_passes, percent_done); - } else { - fprintf(stderr, "\r %3d%% ", percent_done); - } - fflush(stderr); - } -} - - -GLOBAL(void) -start_progress_monitor (j_common_ptr cinfo, cd_progress_ptr progress) -{ - /* Enable progress display, unless trace output is on */ - if (cinfo->err->trace_level == 0) { - progress->pub.progress_monitor = progress_monitor; - progress->completed_extra_passes = 0; - progress->total_extra_passes = 0; - progress->percent_done = -1; - cinfo->progress = &progress->pub; - } -} - - -GLOBAL(void) -end_progress_monitor (j_common_ptr cinfo) -{ - /* Clear away progress display */ - if (cinfo->err->trace_level == 0) { - fprintf(stderr, "\r \r"); - fflush(stderr); - } -} - -#endif - - -/* - * Case-insensitive matching of possibly-abbreviated keyword switches. - * keyword is the constant keyword (must be lower case already), - * minchars is length of minimum legal abbreviation. - */ - -GLOBAL(boolean) -keymatch (char * arg, const char * keyword, int minchars) -{ - register int ca, ck; - register int nmatched = 0; - - while ((ca = *arg++) != '\0') { - if ((ck = *keyword++) == '\0') - return FALSE; /* arg longer than keyword, no good */ - if (isupper(ca)) /* force arg to lcase (assume ck is already) */ - ca = tolower(ca); - if (ca != ck) - return FALSE; /* no good */ - nmatched++; /* count matched characters */ - } - /* reached end of argument; fail if it's too short for unique abbrev */ - if (nmatched < minchars) - return FALSE; - return TRUE; /* A-OK */ -} - - -/* - * Routines to establish binary I/O mode for stdin and stdout. - * Non-Unix systems often require some hacking to get out of text mode. - */ - -GLOBAL(FILE *) -read_stdin (void) -{ - FILE * input_file = stdin; - -#ifdef USE_SETMODE /* need to hack file mode? */ - setmode(fileno(stdin), O_BINARY); -#endif -#ifdef USE_FDOPEN /* need to re-open in binary mode? */ - if ((input_file = fdopen(fileno(stdin), READ_BINARY)) == NULL) { - fprintf(stderr, "Cannot reopen stdin\n"); - exit(EXIT_FAILURE); - } -#endif - return input_file; -} - - -GLOBAL(FILE *) -write_stdout (void) -{ - FILE * output_file = stdout; - -#ifdef USE_SETMODE /* need to hack file mode? */ - setmode(fileno(stdout), O_BINARY); -#endif -#ifdef USE_FDOPEN /* need to re-open in binary mode? */ - if ((output_file = fdopen(fileno(stdout), WRITE_BINARY)) == NULL) { - fprintf(stderr, "Cannot reopen stdout\n"); - exit(EXIT_FAILURE); - } -#endif - return output_file; -} +/* + * cdjpeg.c + * + * Copyright (C) 1991-1997, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains common support routines used by the IJG application + * programs (cjpeg, djpeg, jpegtran). + */ + +#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ +#include /* to declare isupper(), tolower() */ +#ifdef NEED_SIGNAL_CATCHER +#include /* to declare signal() */ +#endif +#ifdef USE_SETMODE +#include /* to declare setmode()'s parameter macros */ +/* If you have setmode() but not , just delete this line: */ +#include /* to declare setmode() */ +#endif + + +/* + * Signal catcher to ensure that temporary files are removed before aborting. + * NB: for Amiga Manx C this is actually a global routine named _abort(); + * we put "#define signal_catcher _abort" in jconfig.h. Talk about bogus... + */ + +#ifdef NEED_SIGNAL_CATCHER + +static j_common_ptr sig_cinfo; + +void /* must be global for Manx C */ +signal_catcher (int signum) +{ + if (sig_cinfo != NULL) { + if (sig_cinfo->err != NULL) /* turn off trace output */ + sig_cinfo->err->trace_level = 0; + jpeg_destroy(sig_cinfo); /* clean up memory allocation & temp files */ + } + exit(EXIT_FAILURE); +} + + +GLOBAL(void) +enable_signal_catcher (j_common_ptr cinfo) +{ + sig_cinfo = cinfo; +#ifdef SIGINT /* not all systems have SIGINT */ + signal(SIGINT, signal_catcher); +#endif +#ifdef SIGTERM /* not all systems have SIGTERM */ + signal(SIGTERM, signal_catcher); +#endif +} + +#endif + + +/* + * Optional progress monitor: display a percent-done figure on stderr. + */ + +#ifdef PROGRESS_REPORT + +METHODDEF(void) +progress_monitor (j_common_ptr cinfo) +{ + cd_progress_ptr prog = (cd_progress_ptr) cinfo->progress; + int total_passes = prog->pub.total_passes + prog->total_extra_passes; + int percent_done = (int) (prog->pub.pass_counter*100L/prog->pub.pass_limit); + + if (percent_done != prog->percent_done) { + prog->percent_done = percent_done; + if (total_passes > 1) { + fprintf(stderr, "\rPass %d/%d: %3d%% ", + prog->pub.completed_passes + prog->completed_extra_passes + 1, + total_passes, percent_done); + } else { + fprintf(stderr, "\r %3d%% ", percent_done); + } + fflush(stderr); + } +} + + +GLOBAL(void) +start_progress_monitor (j_common_ptr cinfo, cd_progress_ptr progress) +{ + /* Enable progress display, unless trace output is on */ + if (cinfo->err->trace_level == 0) { + progress->pub.progress_monitor = progress_monitor; + progress->completed_extra_passes = 0; + progress->total_extra_passes = 0; + progress->percent_done = -1; + cinfo->progress = &progress->pub; + } +} + + +GLOBAL(void) +end_progress_monitor (j_common_ptr cinfo) +{ + /* Clear away progress display */ + if (cinfo->err->trace_level == 0) { + fprintf(stderr, "\r \r"); + fflush(stderr); + } +} + +#endif + + +/* + * Case-insensitive matching of possibly-abbreviated keyword switches. + * keyword is the constant keyword (must be lower case already), + * minchars is length of minimum legal abbreviation. + */ + +GLOBAL(boolean) +keymatch (char * arg, const char * keyword, int minchars) +{ + register int ca, ck; + register int nmatched = 0; + + while ((ca = *arg++) != '\0') { + if ((ck = *keyword++) == '\0') + return FALSE; /* arg longer than keyword, no good */ + if (isupper(ca)) /* force arg to lcase (assume ck is already) */ + ca = tolower(ca); + if (ca != ck) + return FALSE; /* no good */ + nmatched++; /* count matched characters */ + } + /* reached end of argument; fail if it's too short for unique abbrev */ + if (nmatched < minchars) + return FALSE; + return TRUE; /* A-OK */ +} + + +/* + * Routines to establish binary I/O mode for stdin and stdout. + * Non-Unix systems often require some hacking to get out of text mode. + */ + +GLOBAL(FILE *) +read_stdin (void) +{ + FILE * input_file = stdin; + +#ifdef USE_SETMODE /* need to hack file mode? */ + setmode(fileno(stdin), O_BINARY); +#endif +#ifdef USE_FDOPEN /* need to re-open in binary mode? */ + if ((input_file = fdopen(fileno(stdin), READ_BINARY)) == NULL) { + fprintf(stderr, "Cannot reopen stdin\n"); + exit(EXIT_FAILURE); + } +#endif + return input_file; +} + + +GLOBAL(FILE *) +write_stdout (void) +{ + FILE * output_file = stdout; + +#ifdef USE_SETMODE /* need to hack file mode? */ + setmode(fileno(stdout), O_BINARY); +#endif +#ifdef USE_FDOPEN /* need to re-open in binary mode? */ + if ((output_file = fdopen(fileno(stdout), WRITE_BINARY)) == NULL) { + fprintf(stderr, "Cannot reopen stdout\n"); + exit(EXIT_FAILURE); + } +#endif + return output_file; +} diff --git a/test/monniaux/jpeg-6b/cjpeg.c b/test/monniaux/jpeg-6b/cjpeg.c index b9102ee9..f2a929f0 100644 --- a/test/monniaux/jpeg-6b/cjpeg.c +++ b/test/monniaux/jpeg-6b/cjpeg.c @@ -1,606 +1,606 @@ -/* - * cjpeg.c - * - * Copyright (C) 1991-1998, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains a command-line user interface for the JPEG compressor. - * It should work on any system with Unix- or MS-DOS-style command lines. - * - * Two different command line styles are permitted, depending on the - * compile-time switch TWO_FILE_COMMANDLINE: - * cjpeg [options] inputfile outputfile - * cjpeg [options] [inputfile] - * In the second style, output is always to standard output, which you'd - * normally redirect to a file or pipe to some other program. Input is - * either from a named file or from standard input (typically redirected). - * The second style is convenient on Unix but is unhelpful on systems that - * don't support pipes. Also, you MUST use the first style if your system - * doesn't do binary I/O to stdin/stdout. - * To simplify script writing, the "-outfile" switch is provided. The syntax - * cjpeg [options] -outfile outputfile inputfile - * works regardless of which command line style is used. - */ - -#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ -#include "jversion.h" /* for version message */ - -#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ -#ifdef __MWERKS__ -#include /* Metrowerks needs this */ -#include /* ... and this */ -#endif -#ifdef THINK_C -#include /* Think declares it here */ -#endif -#endif - - -/* Create the add-on message string table. */ - -#define JMESSAGE(code,string) string , - -static const char * const cdjpeg_message_table[] = { -#include "cderror.h" - NULL -}; - - -/* - * This routine determines what format the input file is, - * and selects the appropriate input-reading module. - * - * To determine which family of input formats the file belongs to, - * we may look only at the first byte of the file, since C does not - * guarantee that more than one character can be pushed back with ungetc. - * Looking at additional bytes would require one of these approaches: - * 1) assume we can fseek() the input file (fails for piped input); - * 2) assume we can push back more than one character (works in - * some C implementations, but unportable); - * 3) provide our own buffering (breaks input readers that want to use - * stdio directly, such as the RLE library); - * or 4) don't put back the data, and modify the input_init methods to assume - * they start reading after the start of file (also breaks RLE library). - * #1 is attractive for MS-DOS but is untenable on Unix. - * - * The most portable solution for file types that can't be identified by their - * first byte is to make the user tell us what they are. This is also the - * only approach for "raw" file types that contain only arbitrary values. - * We presently apply this method for Targa files. Most of the time Targa - * files start with 0x00, so we recognize that case. Potentially, however, - * a Targa file could start with any byte value (byte 0 is the length of the - * seldom-used ID field), so we provide a switch to force Targa input mode. - */ - -static boolean is_targa; /* records user -targa switch */ - - -LOCAL(cjpeg_source_ptr) -select_file_type (j_compress_ptr cinfo, FILE * infile) -{ - int c; - - if (is_targa) { -#ifdef TARGA_SUPPORTED - return jinit_read_targa(cinfo); -#else - ERREXIT(cinfo, JERR_TGA_NOTCOMP); -#endif - } - - if ((c = getc(infile)) == EOF) - ERREXIT(cinfo, JERR_INPUT_EMPTY); - if (ungetc(c, infile) == EOF) - ERREXIT(cinfo, JERR_UNGETC_FAILED); - - switch (c) { -#ifdef BMP_SUPPORTED - case 'B': - return jinit_read_bmp(cinfo); -#endif -#ifdef GIF_SUPPORTED - case 'G': - return jinit_read_gif(cinfo); -#endif -#ifdef PPM_SUPPORTED - case 'P': - return jinit_read_ppm(cinfo); -#endif -#ifdef RLE_SUPPORTED - case 'R': - return jinit_read_rle(cinfo); -#endif -#ifdef TARGA_SUPPORTED - case 0x00: - return jinit_read_targa(cinfo); -#endif - default: - ERREXIT(cinfo, JERR_UNKNOWN_FORMAT); - break; - } - - return NULL; /* suppress compiler warnings */ -} - - -/* - * Argument-parsing code. - * The switch parser is designed to be useful with DOS-style command line - * syntax, ie, intermixed switches and file names, where only the switches - * to the left of a given file name affect processing of that file. - * The main program in this file doesn't actually use this capability... - */ - - -static const char * progname; /* program name for error messages */ -static char * outfilename; /* for -outfile switch */ - - -LOCAL(void) -usage (void) -/* complain about bad command line */ -{ - fprintf(stderr, "usage: %s [switches] ", progname); -#ifdef TWO_FILE_COMMANDLINE - fprintf(stderr, "inputfile outputfile\n"); -#else - fprintf(stderr, "[inputfile]\n"); -#endif - - fprintf(stderr, "Switches (names may be abbreviated):\n"); - fprintf(stderr, " -quality N Compression quality (0..100; 5-95 is useful range)\n"); - fprintf(stderr, " -grayscale Create monochrome JPEG file\n"); -#ifdef ENTROPY_OPT_SUPPORTED - fprintf(stderr, " -optimize Optimize Huffman table (smaller file, but slow compression)\n"); -#endif -#ifdef C_PROGRESSIVE_SUPPORTED - fprintf(stderr, " -progressive Create progressive JPEG file\n"); -#endif -#ifdef TARGA_SUPPORTED - fprintf(stderr, " -targa Input file is Targa format (usually not needed)\n"); -#endif - fprintf(stderr, "Switches for advanced users:\n"); -#ifdef DCT_ISLOW_SUPPORTED - fprintf(stderr, " -dct int Use integer DCT method%s\n", - (JDCT_DEFAULT == JDCT_ISLOW ? " (default)" : "")); -#endif -#ifdef DCT_IFAST_SUPPORTED - fprintf(stderr, " -dct fast Use fast integer DCT (less accurate)%s\n", - (JDCT_DEFAULT == JDCT_IFAST ? " (default)" : "")); -#endif -#ifdef DCT_FLOAT_SUPPORTED - fprintf(stderr, " -dct float Use floating-point DCT method%s\n", - (JDCT_DEFAULT == JDCT_FLOAT ? " (default)" : "")); -#endif - fprintf(stderr, " -restart N Set restart interval in rows, or in blocks with B\n"); -#ifdef INPUT_SMOOTHING_SUPPORTED - fprintf(stderr, " -smooth N Smooth dithered input (N=1..100 is strength)\n"); -#endif - fprintf(stderr, " -maxmemory N Maximum memory to use (in kbytes)\n"); - fprintf(stderr, " -outfile name Specify name for output file\n"); - fprintf(stderr, " -verbose or -debug Emit debug output\n"); - fprintf(stderr, "Switches for wizards:\n"); -#ifdef C_ARITH_CODING_SUPPORTED - fprintf(stderr, " -arithmetic Use arithmetic coding\n"); -#endif - fprintf(stderr, " -baseline Force baseline quantization tables\n"); - fprintf(stderr, " -qtables file Use quantization tables given in file\n"); - fprintf(stderr, " -qslots N[,...] Set component quantization tables\n"); - fprintf(stderr, " -sample HxV[,...] Set component sampling factors\n"); -#ifdef C_MULTISCAN_FILES_SUPPORTED - fprintf(stderr, " -scans file Create multi-scan JPEG per script file\n"); -#endif - exit(EXIT_FAILURE); -} - - -LOCAL(int) -parse_switches (j_compress_ptr cinfo, int argc, char **argv, - int last_file_arg_seen, boolean for_real) -/* Parse optional switches. - * Returns argv[] index of first file-name argument (== argc if none). - * Any file names with indexes <= last_file_arg_seen are ignored; - * they have presumably been processed in a previous iteration. - * (Pass 0 for last_file_arg_seen on the first or only iteration.) - * for_real is FALSE on the first (dummy) pass; we may skip any expensive - * processing. - */ -{ - int argn; - char * arg; - int quality; /* -quality parameter */ - int q_scale_factor; /* scaling percentage for -qtables */ - boolean force_baseline; - boolean simple_progressive; - char * qtablefile = NULL; /* saves -qtables filename if any */ - char * qslotsarg = NULL; /* saves -qslots parm if any */ - char * samplearg = NULL; /* saves -sample parm if any */ - char * scansarg = NULL; /* saves -scans parm if any */ - - /* Set up default JPEG parameters. */ - /* Note that default -quality level need not, and does not, - * match the default scaling for an explicit -qtables argument. - */ - quality = 75; /* default -quality value */ - q_scale_factor = 100; /* default to no scaling for -qtables */ - force_baseline = FALSE; /* by default, allow 16-bit quantizers */ - simple_progressive = FALSE; - is_targa = FALSE; - outfilename = NULL; - cinfo->err->trace_level = 0; - - /* Scan command line options, adjust parameters */ - - for (argn = 1; argn < argc; argn++) { - arg = argv[argn]; - if (*arg != '-') { - /* Not a switch, must be a file name argument */ - if (argn <= last_file_arg_seen) { - outfilename = NULL; /* -outfile applies to just one input file */ - continue; /* ignore this name if previously processed */ - } - break; /* else done parsing switches */ - } - arg++; /* advance past switch marker character */ - - if (keymatch(arg, "arithmetic", 1)) { - /* Use arithmetic coding. */ -#ifdef C_ARITH_CODING_SUPPORTED - cinfo->arith_code = TRUE; -#else - fprintf(stderr, "%s: sorry, arithmetic coding not supported\n", - progname); - exit(EXIT_FAILURE); -#endif - - } else if (keymatch(arg, "baseline", 1)) { - /* Force baseline-compatible output (8-bit quantizer values). */ - force_baseline = TRUE; - - } else if (keymatch(arg, "dct", 2)) { - /* Select DCT algorithm. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (keymatch(argv[argn], "int", 1)) { - cinfo->dct_method = JDCT_ISLOW; - } else if (keymatch(argv[argn], "fast", 2)) { - cinfo->dct_method = JDCT_IFAST; - } else if (keymatch(argv[argn], "float", 2)) { - cinfo->dct_method = JDCT_FLOAT; - } else - usage(); - - } else if (keymatch(arg, "debug", 1) || keymatch(arg, "verbose", 1)) { - /* Enable debug printouts. */ - /* On first -d, print version identification */ - static boolean printed_version = FALSE; - - if (! printed_version) { - fprintf(stderr, "Independent JPEG Group's CJPEG, version %s\n%s\n", - JVERSION, JCOPYRIGHT); - printed_version = TRUE; - } - cinfo->err->trace_level++; - - } else if (keymatch(arg, "grayscale", 2) || keymatch(arg, "greyscale",2)) { - /* Force a monochrome JPEG file to be generated. */ - jpeg_set_colorspace(cinfo, JCS_GRAYSCALE); - - } else if (keymatch(arg, "maxmemory", 3)) { - /* Maximum memory in Kb (or Mb with 'm'). */ - long lval; - char ch = 'x'; - - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) - usage(); - if (ch == 'm' || ch == 'M') - lval *= 1000L; - cinfo->mem->max_memory_to_use = lval * 1000L; - - } else if (keymatch(arg, "optimize", 1) || keymatch(arg, "optimise", 1)) { - /* Enable entropy parm optimization. */ -#ifdef ENTROPY_OPT_SUPPORTED - cinfo->optimize_coding = TRUE; -#else - fprintf(stderr, "%s: sorry, entropy optimization was not compiled\n", - progname); - exit(EXIT_FAILURE); -#endif - - } else if (keymatch(arg, "outfile", 4)) { - /* Set output file name. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - outfilename = argv[argn]; /* save it away for later use */ - - } else if (keymatch(arg, "progressive", 1)) { - /* Select simple progressive mode. */ -#ifdef C_PROGRESSIVE_SUPPORTED - simple_progressive = TRUE; - /* We must postpone execution until num_components is known. */ -#else - fprintf(stderr, "%s: sorry, progressive output was not compiled\n", - progname); - exit(EXIT_FAILURE); -#endif - - } else if (keymatch(arg, "quality", 1)) { - /* Quality factor (quantization table scaling factor). */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%d", &quality) != 1) - usage(); - /* Change scale factor in case -qtables is present. */ - q_scale_factor = jpeg_quality_scaling(quality); - - } else if (keymatch(arg, "qslots", 2)) { - /* Quantization table slot numbers. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - qslotsarg = argv[argn]; - /* Must delay setting qslots until after we have processed any - * colorspace-determining switches, since jpeg_set_colorspace sets - * default quant table numbers. - */ - - } else if (keymatch(arg, "qtables", 2)) { - /* Quantization tables fetched from file. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - qtablefile = argv[argn]; - /* We postpone actually reading the file in case -quality comes later. */ - - } else if (keymatch(arg, "restart", 1)) { - /* Restart interval in MCU rows (or in MCUs with 'b'). */ - long lval; - char ch = 'x'; - - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) - usage(); - if (lval < 0 || lval > 65535L) - usage(); - if (ch == 'b' || ch == 'B') { - cinfo->restart_interval = (unsigned int) lval; - cinfo->restart_in_rows = 0; /* else prior '-restart n' overrides me */ - } else { - cinfo->restart_in_rows = (int) lval; - /* restart_interval will be computed during startup */ - } - - } else if (keymatch(arg, "sample", 2)) { - /* Set sampling factors. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - samplearg = argv[argn]; - /* Must delay setting sample factors until after we have processed any - * colorspace-determining switches, since jpeg_set_colorspace sets - * default sampling factors. - */ - - } else if (keymatch(arg, "scans", 2)) { - /* Set scan script. */ -#ifdef C_MULTISCAN_FILES_SUPPORTED - if (++argn >= argc) /* advance to next argument */ - usage(); - scansarg = argv[argn]; - /* We must postpone reading the file in case -progressive appears. */ -#else - fprintf(stderr, "%s: sorry, multi-scan output was not compiled\n", - progname); - exit(EXIT_FAILURE); -#endif - - } else if (keymatch(arg, "smooth", 2)) { - /* Set input smoothing factor. */ - int val; - - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%d", &val) != 1) - usage(); - if (val < 0 || val > 100) - usage(); - cinfo->smoothing_factor = val; - - } else if (keymatch(arg, "targa", 1)) { - /* Input file is Targa format. */ - is_targa = TRUE; - - } else { - usage(); /* bogus switch */ - } - } - - /* Post-switch-scanning cleanup */ - - if (for_real) { - - /* Set quantization tables for selected quality. */ - /* Some or all may be overridden if -qtables is present. */ - jpeg_set_quality(cinfo, quality, force_baseline); - - if (qtablefile != NULL) /* process -qtables if it was present */ - if (! read_quant_tables(cinfo, qtablefile, - q_scale_factor, force_baseline)) - usage(); - - if (qslotsarg != NULL) /* process -qslots if it was present */ - if (! set_quant_slots(cinfo, qslotsarg)) - usage(); - - if (samplearg != NULL) /* process -sample if it was present */ - if (! set_sample_factors(cinfo, samplearg)) - usage(); - -#ifdef C_PROGRESSIVE_SUPPORTED - if (simple_progressive) /* process -progressive; -scans can override */ - jpeg_simple_progression(cinfo); -#endif - -#ifdef C_MULTISCAN_FILES_SUPPORTED - if (scansarg != NULL) /* process -scans if it was present */ - if (! read_scan_script(cinfo, scansarg)) - usage(); -#endif - } - - return argn; /* return index of next arg (file name) */ -} - - -/* - * The main program. - */ - -int -main (int argc, char **argv) -{ - struct jpeg_compress_struct cinfo; - struct jpeg_error_mgr jerr; -#ifdef PROGRESS_REPORT - struct cdjpeg_progress_mgr progress; -#endif - int file_index; - cjpeg_source_ptr src_mgr; - FILE * input_file; - FILE * output_file; - JDIMENSION num_scanlines; - - /* On Mac, fetch a command line. */ -#ifdef USE_CCOMMAND - argc = ccommand(&argv); -#endif - - progname = argv[0]; - if (progname == NULL || progname[0] == 0) - progname = "cjpeg"; /* in case C library doesn't provide it */ - - /* Initialize the JPEG compression object with default error handling. */ - cinfo.err = jpeg_std_error(&jerr); - jpeg_create_compress(&cinfo); - /* Add some application-specific error messages (from cderror.h) */ - jerr.addon_message_table = cdjpeg_message_table; - jerr.first_addon_message = JMSG_FIRSTADDONCODE; - jerr.last_addon_message = JMSG_LASTADDONCODE; - - /* Now safe to enable signal catcher. */ -#ifdef NEED_SIGNAL_CATCHER - enable_signal_catcher((j_common_ptr) &cinfo); -#endif - - /* Initialize JPEG parameters. - * Much of this may be overridden later. - * In particular, we don't yet know the input file's color space, - * but we need to provide some value for jpeg_set_defaults() to work. - */ - - cinfo.in_color_space = JCS_RGB; /* arbitrary guess */ - jpeg_set_defaults(&cinfo); - - /* Scan command line to find file names. - * It is convenient to use just one switch-parsing routine, but the switch - * values read here are ignored; we will rescan the switches after opening - * the input file. - */ - - file_index = parse_switches(&cinfo, argc, argv, 0, FALSE); - -#ifdef TWO_FILE_COMMANDLINE - /* Must have either -outfile switch or explicit output file name */ - if (outfilename == NULL) { - if (file_index != argc-2) { - fprintf(stderr, "%s: must name one input and one output file\n", - progname); - usage(); - } - outfilename = argv[file_index+1]; - } else { - if (file_index != argc-1) { - fprintf(stderr, "%s: must name one input and one output file\n", - progname); - usage(); - } - } -#else - /* Unix style: expect zero or one file name */ - if (file_index < argc-1) { - fprintf(stderr, "%s: only one input file\n", progname); - usage(); - } -#endif /* TWO_FILE_COMMANDLINE */ - - /* Open the input file. */ - if (file_index < argc) { - if ((input_file = fopen(argv[file_index], READ_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, argv[file_index]); - exit(EXIT_FAILURE); - } - } else { - /* default input file is stdin */ - input_file = read_stdin(); - } - - /* Open the output file. */ - if (outfilename != NULL) { - if ((output_file = fopen(outfilename, WRITE_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, outfilename); - exit(EXIT_FAILURE); - } - } else { - /* default output file is stdout */ - output_file = write_stdout(); - } - -#ifdef PROGRESS_REPORT - start_progress_monitor((j_common_ptr) &cinfo, &progress); -#endif - - /* Figure out the input file format, and set up to read it. */ - src_mgr = select_file_type(&cinfo, input_file); - src_mgr->input_file = input_file; - - /* Read the input file header to obtain file size & colorspace. */ - (*src_mgr->start_input) (&cinfo, src_mgr); - - /* Now that we know input colorspace, fix colorspace-dependent defaults */ - jpeg_default_colorspace(&cinfo); - - /* Adjust default compression parameters by re-parsing the options */ - file_index = parse_switches(&cinfo, argc, argv, 0, TRUE); - - /* Specify data destination for compression */ - jpeg_stdio_dest(&cinfo, output_file); - - /* Start compressor */ - jpeg_start_compress(&cinfo, TRUE); - - /* Process data */ - while (cinfo.next_scanline < cinfo.image_height) { - num_scanlines = (*src_mgr->get_pixel_rows) (&cinfo, src_mgr); - (void) jpeg_write_scanlines(&cinfo, src_mgr->buffer, num_scanlines); - } - - /* Finish compression and release memory */ - (*src_mgr->finish_input) (&cinfo, src_mgr); - jpeg_finish_compress(&cinfo); - jpeg_destroy_compress(&cinfo); - - /* Close files, if we opened them */ - if (input_file != stdin) - fclose(input_file); - if (output_file != stdout) - fclose(output_file); - -#ifdef PROGRESS_REPORT - end_progress_monitor((j_common_ptr) &cinfo); -#endif - - /* All done. */ - exit(jerr.num_warnings ? EXIT_WARNING : EXIT_SUCCESS); - return 0; /* suppress no-return-value warnings */ -} +/* + * cjpeg.c + * + * Copyright (C) 1991-1998, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains a command-line user interface for the JPEG compressor. + * It should work on any system with Unix- or MS-DOS-style command lines. + * + * Two different command line styles are permitted, depending on the + * compile-time switch TWO_FILE_COMMANDLINE: + * cjpeg [options] inputfile outputfile + * cjpeg [options] [inputfile] + * In the second style, output is always to standard output, which you'd + * normally redirect to a file or pipe to some other program. Input is + * either from a named file or from standard input (typically redirected). + * The second style is convenient on Unix but is unhelpful on systems that + * don't support pipes. Also, you MUST use the first style if your system + * doesn't do binary I/O to stdin/stdout. + * To simplify script writing, the "-outfile" switch is provided. The syntax + * cjpeg [options] -outfile outputfile inputfile + * works regardless of which command line style is used. + */ + +#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ +#include "jversion.h" /* for version message */ + +#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ +#ifdef __MWERKS__ +#include /* Metrowerks needs this */ +#include /* ... and this */ +#endif +#ifdef THINK_C +#include /* Think declares it here */ +#endif +#endif + + +/* Create the add-on message string table. */ + +#define JMESSAGE(code,string) string , + +static const char * const cdjpeg_message_table[] = { +#include "cderror.h" + NULL +}; + + +/* + * This routine determines what format the input file is, + * and selects the appropriate input-reading module. + * + * To determine which family of input formats the file belongs to, + * we may look only at the first byte of the file, since C does not + * guarantee that more than one character can be pushed back with ungetc. + * Looking at additional bytes would require one of these approaches: + * 1) assume we can fseek() the input file (fails for piped input); + * 2) assume we can push back more than one character (works in + * some C implementations, but unportable); + * 3) provide our own buffering (breaks input readers that want to use + * stdio directly, such as the RLE library); + * or 4) don't put back the data, and modify the input_init methods to assume + * they start reading after the start of file (also breaks RLE library). + * #1 is attractive for MS-DOS but is untenable on Unix. + * + * The most portable solution for file types that can't be identified by their + * first byte is to make the user tell us what they are. This is also the + * only approach for "raw" file types that contain only arbitrary values. + * We presently apply this method for Targa files. Most of the time Targa + * files start with 0x00, so we recognize that case. Potentially, however, + * a Targa file could start with any byte value (byte 0 is the length of the + * seldom-used ID field), so we provide a switch to force Targa input mode. + */ + +static boolean is_targa; /* records user -targa switch */ + + +LOCAL(cjpeg_source_ptr) +select_file_type (j_compress_ptr cinfo, FILE * infile) +{ + int c; + + if (is_targa) { +#ifdef TARGA_SUPPORTED + return jinit_read_targa(cinfo); +#else + ERREXIT(cinfo, JERR_TGA_NOTCOMP); +#endif + } + + if ((c = getc(infile)) == EOF) + ERREXIT(cinfo, JERR_INPUT_EMPTY); + if (ungetc(c, infile) == EOF) + ERREXIT(cinfo, JERR_UNGETC_FAILED); + + switch (c) { +#ifdef BMP_SUPPORTED + case 'B': + return jinit_read_bmp(cinfo); +#endif +#ifdef GIF_SUPPORTED + case 'G': + return jinit_read_gif(cinfo); +#endif +#ifdef PPM_SUPPORTED + case 'P': + return jinit_read_ppm(cinfo); +#endif +#ifdef RLE_SUPPORTED + case 'R': + return jinit_read_rle(cinfo); +#endif +#ifdef TARGA_SUPPORTED + case 0x00: + return jinit_read_targa(cinfo); +#endif + default: + ERREXIT(cinfo, JERR_UNKNOWN_FORMAT); + break; + } + + return NULL; /* suppress compiler warnings */ +} + + +/* + * Argument-parsing code. + * The switch parser is designed to be useful with DOS-style command line + * syntax, ie, intermixed switches and file names, where only the switches + * to the left of a given file name affect processing of that file. + * The main program in this file doesn't actually use this capability... + */ + + +static const char * progname; /* program name for error messages */ +static char * outfilename; /* for -outfile switch */ + + +LOCAL(void) +usage (void) +/* complain about bad command line */ +{ + fprintf(stderr, "usage: %s [switches] ", progname); +#ifdef TWO_FILE_COMMANDLINE + fprintf(stderr, "inputfile outputfile\n"); +#else + fprintf(stderr, "[inputfile]\n"); +#endif + + fprintf(stderr, "Switches (names may be abbreviated):\n"); + fprintf(stderr, " -quality N Compression quality (0..100; 5-95 is useful range)\n"); + fprintf(stderr, " -grayscale Create monochrome JPEG file\n"); +#ifdef ENTROPY_OPT_SUPPORTED + fprintf(stderr, " -optimize Optimize Huffman table (smaller file, but slow compression)\n"); +#endif +#ifdef C_PROGRESSIVE_SUPPORTED + fprintf(stderr, " -progressive Create progressive JPEG file\n"); +#endif +#ifdef TARGA_SUPPORTED + fprintf(stderr, " -targa Input file is Targa format (usually not needed)\n"); +#endif + fprintf(stderr, "Switches for advanced users:\n"); +#ifdef DCT_ISLOW_SUPPORTED + fprintf(stderr, " -dct int Use integer DCT method%s\n", + (JDCT_DEFAULT == JDCT_ISLOW ? " (default)" : "")); +#endif +#ifdef DCT_IFAST_SUPPORTED + fprintf(stderr, " -dct fast Use fast integer DCT (less accurate)%s\n", + (JDCT_DEFAULT == JDCT_IFAST ? " (default)" : "")); +#endif +#ifdef DCT_FLOAT_SUPPORTED + fprintf(stderr, " -dct float Use floating-point DCT method%s\n", + (JDCT_DEFAULT == JDCT_FLOAT ? " (default)" : "")); +#endif + fprintf(stderr, " -restart N Set restart interval in rows, or in blocks with B\n"); +#ifdef INPUT_SMOOTHING_SUPPORTED + fprintf(stderr, " -smooth N Smooth dithered input (N=1..100 is strength)\n"); +#endif + fprintf(stderr, " -maxmemory N Maximum memory to use (in kbytes)\n"); + fprintf(stderr, " -outfile name Specify name for output file\n"); + fprintf(stderr, " -verbose or -debug Emit debug output\n"); + fprintf(stderr, "Switches for wizards:\n"); +#ifdef C_ARITH_CODING_SUPPORTED + fprintf(stderr, " -arithmetic Use arithmetic coding\n"); +#endif + fprintf(stderr, " -baseline Force baseline quantization tables\n"); + fprintf(stderr, " -qtables file Use quantization tables given in file\n"); + fprintf(stderr, " -qslots N[,...] Set component quantization tables\n"); + fprintf(stderr, " -sample HxV[,...] Set component sampling factors\n"); +#ifdef C_MULTISCAN_FILES_SUPPORTED + fprintf(stderr, " -scans file Create multi-scan JPEG per script file\n"); +#endif + exit(EXIT_FAILURE); +} + + +LOCAL(int) +parse_switches (j_compress_ptr cinfo, int argc, char **argv, + int last_file_arg_seen, boolean for_real) +/* Parse optional switches. + * Returns argv[] index of first file-name argument (== argc if none). + * Any file names with indexes <= last_file_arg_seen are ignored; + * they have presumably been processed in a previous iteration. + * (Pass 0 for last_file_arg_seen on the first or only iteration.) + * for_real is FALSE on the first (dummy) pass; we may skip any expensive + * processing. + */ +{ + int argn; + char * arg; + int quality; /* -quality parameter */ + int q_scale_factor; /* scaling percentage for -qtables */ + boolean force_baseline; + boolean simple_progressive; + char * qtablefile = NULL; /* saves -qtables filename if any */ + char * qslotsarg = NULL; /* saves -qslots parm if any */ + char * samplearg = NULL; /* saves -sample parm if any */ + char * scansarg = NULL; /* saves -scans parm if any */ + + /* Set up default JPEG parameters. */ + /* Note that default -quality level need not, and does not, + * match the default scaling for an explicit -qtables argument. + */ + quality = 75; /* default -quality value */ + q_scale_factor = 100; /* default to no scaling for -qtables */ + force_baseline = FALSE; /* by default, allow 16-bit quantizers */ + simple_progressive = FALSE; + is_targa = FALSE; + outfilename = NULL; + cinfo->err->trace_level = 0; + + /* Scan command line options, adjust parameters */ + + for (argn = 1; argn < argc; argn++) { + arg = argv[argn]; + if (*arg != '-') { + /* Not a switch, must be a file name argument */ + if (argn <= last_file_arg_seen) { + outfilename = NULL; /* -outfile applies to just one input file */ + continue; /* ignore this name if previously processed */ + } + break; /* else done parsing switches */ + } + arg++; /* advance past switch marker character */ + + if (keymatch(arg, "arithmetic", 1)) { + /* Use arithmetic coding. */ +#ifdef C_ARITH_CODING_SUPPORTED + cinfo->arith_code = TRUE; +#else + fprintf(stderr, "%s: sorry, arithmetic coding not supported\n", + progname); + exit(EXIT_FAILURE); +#endif + + } else if (keymatch(arg, "baseline", 1)) { + /* Force baseline-compatible output (8-bit quantizer values). */ + force_baseline = TRUE; + + } else if (keymatch(arg, "dct", 2)) { + /* Select DCT algorithm. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (keymatch(argv[argn], "int", 1)) { + cinfo->dct_method = JDCT_ISLOW; + } else if (keymatch(argv[argn], "fast", 2)) { + cinfo->dct_method = JDCT_IFAST; + } else if (keymatch(argv[argn], "float", 2)) { + cinfo->dct_method = JDCT_FLOAT; + } else + usage(); + + } else if (keymatch(arg, "debug", 1) || keymatch(arg, "verbose", 1)) { + /* Enable debug printouts. */ + /* On first -d, print version identification */ + static boolean printed_version = FALSE; + + if (! printed_version) { + fprintf(stderr, "Independent JPEG Group's CJPEG, version %s\n%s\n", + JVERSION, JCOPYRIGHT); + printed_version = TRUE; + } + cinfo->err->trace_level++; + + } else if (keymatch(arg, "grayscale", 2) || keymatch(arg, "greyscale",2)) { + /* Force a monochrome JPEG file to be generated. */ + jpeg_set_colorspace(cinfo, JCS_GRAYSCALE); + + } else if (keymatch(arg, "maxmemory", 3)) { + /* Maximum memory in Kb (or Mb with 'm'). */ + long lval; + char ch = 'x'; + + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) + usage(); + if (ch == 'm' || ch == 'M') + lval *= 1000L; + cinfo->mem->max_memory_to_use = lval * 1000L; + + } else if (keymatch(arg, "optimize", 1) || keymatch(arg, "optimise", 1)) { + /* Enable entropy parm optimization. */ +#ifdef ENTROPY_OPT_SUPPORTED + cinfo->optimize_coding = TRUE; +#else + fprintf(stderr, "%s: sorry, entropy optimization was not compiled\n", + progname); + exit(EXIT_FAILURE); +#endif + + } else if (keymatch(arg, "outfile", 4)) { + /* Set output file name. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + outfilename = argv[argn]; /* save it away for later use */ + + } else if (keymatch(arg, "progressive", 1)) { + /* Select simple progressive mode. */ +#ifdef C_PROGRESSIVE_SUPPORTED + simple_progressive = TRUE; + /* We must postpone execution until num_components is known. */ +#else + fprintf(stderr, "%s: sorry, progressive output was not compiled\n", + progname); + exit(EXIT_FAILURE); +#endif + + } else if (keymatch(arg, "quality", 1)) { + /* Quality factor (quantization table scaling factor). */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%d", &quality) != 1) + usage(); + /* Change scale factor in case -qtables is present. */ + q_scale_factor = jpeg_quality_scaling(quality); + + } else if (keymatch(arg, "qslots", 2)) { + /* Quantization table slot numbers. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + qslotsarg = argv[argn]; + /* Must delay setting qslots until after we have processed any + * colorspace-determining switches, since jpeg_set_colorspace sets + * default quant table numbers. + */ + + } else if (keymatch(arg, "qtables", 2)) { + /* Quantization tables fetched from file. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + qtablefile = argv[argn]; + /* We postpone actually reading the file in case -quality comes later. */ + + } else if (keymatch(arg, "restart", 1)) { + /* Restart interval in MCU rows (or in MCUs with 'b'). */ + long lval; + char ch = 'x'; + + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) + usage(); + if (lval < 0 || lval > 65535L) + usage(); + if (ch == 'b' || ch == 'B') { + cinfo->restart_interval = (unsigned int) lval; + cinfo->restart_in_rows = 0; /* else prior '-restart n' overrides me */ + } else { + cinfo->restart_in_rows = (int) lval; + /* restart_interval will be computed during startup */ + } + + } else if (keymatch(arg, "sample", 2)) { + /* Set sampling factors. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + samplearg = argv[argn]; + /* Must delay setting sample factors until after we have processed any + * colorspace-determining switches, since jpeg_set_colorspace sets + * default sampling factors. + */ + + } else if (keymatch(arg, "scans", 2)) { + /* Set scan script. */ +#ifdef C_MULTISCAN_FILES_SUPPORTED + if (++argn >= argc) /* advance to next argument */ + usage(); + scansarg = argv[argn]; + /* We must postpone reading the file in case -progressive appears. */ +#else + fprintf(stderr, "%s: sorry, multi-scan output was not compiled\n", + progname); + exit(EXIT_FAILURE); +#endif + + } else if (keymatch(arg, "smooth", 2)) { + /* Set input smoothing factor. */ + int val; + + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%d", &val) != 1) + usage(); + if (val < 0 || val > 100) + usage(); + cinfo->smoothing_factor = val; + + } else if (keymatch(arg, "targa", 1)) { + /* Input file is Targa format. */ + is_targa = TRUE; + + } else { + usage(); /* bogus switch */ + } + } + + /* Post-switch-scanning cleanup */ + + if (for_real) { + + /* Set quantization tables for selected quality. */ + /* Some or all may be overridden if -qtables is present. */ + jpeg_set_quality(cinfo, quality, force_baseline); + + if (qtablefile != NULL) /* process -qtables if it was present */ + if (! read_quant_tables(cinfo, qtablefile, + q_scale_factor, force_baseline)) + usage(); + + if (qslotsarg != NULL) /* process -qslots if it was present */ + if (! set_quant_slots(cinfo, qslotsarg)) + usage(); + + if (samplearg != NULL) /* process -sample if it was present */ + if (! set_sample_factors(cinfo, samplearg)) + usage(); + +#ifdef C_PROGRESSIVE_SUPPORTED + if (simple_progressive) /* process -progressive; -scans can override */ + jpeg_simple_progression(cinfo); +#endif + +#ifdef C_MULTISCAN_FILES_SUPPORTED + if (scansarg != NULL) /* process -scans if it was present */ + if (! read_scan_script(cinfo, scansarg)) + usage(); +#endif + } + + return argn; /* return index of next arg (file name) */ +} + + +/* + * The main program. + */ + +int +main (int argc, char **argv) +{ + struct jpeg_compress_struct cinfo; + struct jpeg_error_mgr jerr; +#ifdef PROGRESS_REPORT + struct cdjpeg_progress_mgr progress; +#endif + int file_index; + cjpeg_source_ptr src_mgr; + FILE * input_file; + FILE * output_file; + JDIMENSION num_scanlines; + + /* On Mac, fetch a command line. */ +#ifdef USE_CCOMMAND + argc = ccommand(&argv); +#endif + + progname = argv[0]; + if (progname == NULL || progname[0] == 0) + progname = "cjpeg"; /* in case C library doesn't provide it */ + + /* Initialize the JPEG compression object with default error handling. */ + cinfo.err = jpeg_std_error(&jerr); + jpeg_create_compress(&cinfo); + /* Add some application-specific error messages (from cderror.h) */ + jerr.addon_message_table = cdjpeg_message_table; + jerr.first_addon_message = JMSG_FIRSTADDONCODE; + jerr.last_addon_message = JMSG_LASTADDONCODE; + + /* Now safe to enable signal catcher. */ +#ifdef NEED_SIGNAL_CATCHER + enable_signal_catcher((j_common_ptr) &cinfo); +#endif + + /* Initialize JPEG parameters. + * Much of this may be overridden later. + * In particular, we don't yet know the input file's color space, + * but we need to provide some value for jpeg_set_defaults() to work. + */ + + cinfo.in_color_space = JCS_RGB; /* arbitrary guess */ + jpeg_set_defaults(&cinfo); + + /* Scan command line to find file names. + * It is convenient to use just one switch-parsing routine, but the switch + * values read here are ignored; we will rescan the switches after opening + * the input file. + */ + + file_index = parse_switches(&cinfo, argc, argv, 0, FALSE); + +#ifdef TWO_FILE_COMMANDLINE + /* Must have either -outfile switch or explicit output file name */ + if (outfilename == NULL) { + if (file_index != argc-2) { + fprintf(stderr, "%s: must name one input and one output file\n", + progname); + usage(); + } + outfilename = argv[file_index+1]; + } else { + if (file_index != argc-1) { + fprintf(stderr, "%s: must name one input and one output file\n", + progname); + usage(); + } + } +#else + /* Unix style: expect zero or one file name */ + if (file_index < argc-1) { + fprintf(stderr, "%s: only one input file\n", progname); + usage(); + } +#endif /* TWO_FILE_COMMANDLINE */ + + /* Open the input file. */ + if (file_index < argc) { + if ((input_file = fopen(argv[file_index], READ_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, argv[file_index]); + exit(EXIT_FAILURE); + } + } else { + /* default input file is stdin */ + input_file = read_stdin(); + } + + /* Open the output file. */ + if (outfilename != NULL) { + if ((output_file = fopen(outfilename, WRITE_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, outfilename); + exit(EXIT_FAILURE); + } + } else { + /* default output file is stdout */ + output_file = write_stdout(); + } + +#ifdef PROGRESS_REPORT + start_progress_monitor((j_common_ptr) &cinfo, &progress); +#endif + + /* Figure out the input file format, and set up to read it. */ + src_mgr = select_file_type(&cinfo, input_file); + src_mgr->input_file = input_file; + + /* Read the input file header to obtain file size & colorspace. */ + (*src_mgr->start_input) (&cinfo, src_mgr); + + /* Now that we know input colorspace, fix colorspace-dependent defaults */ + jpeg_default_colorspace(&cinfo); + + /* Adjust default compression parameters by re-parsing the options */ + file_index = parse_switches(&cinfo, argc, argv, 0, TRUE); + + /* Specify data destination for compression */ + jpeg_stdio_dest(&cinfo, output_file); + + /* Start compressor */ + jpeg_start_compress(&cinfo, TRUE); + + /* Process data */ + while (cinfo.next_scanline < cinfo.image_height) { + num_scanlines = (*src_mgr->get_pixel_rows) (&cinfo, src_mgr); + (void) jpeg_write_scanlines(&cinfo, src_mgr->buffer, num_scanlines); + } + + /* Finish compression and release memory */ + (*src_mgr->finish_input) (&cinfo, src_mgr); + jpeg_finish_compress(&cinfo); + jpeg_destroy_compress(&cinfo); + + /* Close files, if we opened them */ + if (input_file != stdin) + fclose(input_file); + if (output_file != stdout) + fclose(output_file); + +#ifdef PROGRESS_REPORT + end_progress_monitor((j_common_ptr) &cinfo); +#endif + + /* All done. */ + exit(jerr.num_warnings ? EXIT_WARNING : EXIT_SUCCESS); + return 0; /* suppress no-return-value warnings */ +} diff --git a/test/monniaux/jpeg-6b/djpeg.c b/test/monniaux/jpeg-6b/djpeg.c index a63523c5..e099e90a 100644 --- a/test/monniaux/jpeg-6b/djpeg.c +++ b/test/monniaux/jpeg-6b/djpeg.c @@ -1,616 +1,616 @@ -/* - * djpeg.c - * - * Copyright (C) 1991-1997, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains a command-line user interface for the JPEG decompressor. - * It should work on any system with Unix- or MS-DOS-style command lines. - * - * Two different command line styles are permitted, depending on the - * compile-time switch TWO_FILE_COMMANDLINE: - * djpeg [options] inputfile outputfile - * djpeg [options] [inputfile] - * In the second style, output is always to standard output, which you'd - * normally redirect to a file or pipe to some other program. Input is - * either from a named file or from standard input (typically redirected). - * The second style is convenient on Unix but is unhelpful on systems that - * don't support pipes. Also, you MUST use the first style if your system - * doesn't do binary I/O to stdin/stdout. - * To simplify script writing, the "-outfile" switch is provided. The syntax - * djpeg [options] -outfile outputfile inputfile - * works regardless of which command line style is used. - */ - -#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ -#include "jversion.h" /* for version message */ - -#include /* to declare isprint() */ - -#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ -#ifdef __MWERKS__ -#include /* Metrowerks needs this */ -#include /* ... and this */ -#endif -#ifdef THINK_C -#include /* Think declares it here */ -#endif -#endif - - -/* Create the add-on message string table. */ - -#define JMESSAGE(code,string) string , - -static const char * const cdjpeg_message_table[] = { -#include "cderror.h" - NULL -}; - - -/* - * This list defines the known output image formats - * (not all of which need be supported by a given version). - * You can change the default output format by defining DEFAULT_FMT; - * indeed, you had better do so if you undefine PPM_SUPPORTED. - */ - -typedef enum { - FMT_BMP, /* BMP format (Windows flavor) */ - FMT_GIF, /* GIF format */ - FMT_OS2, /* BMP format (OS/2 flavor) */ - FMT_PPM, /* PPM/PGM (PBMPLUS formats) */ - FMT_RLE, /* RLE format */ - FMT_TARGA, /* Targa format */ - FMT_TIFF /* TIFF format */ -} IMAGE_FORMATS; - -#ifndef DEFAULT_FMT /* so can override from CFLAGS in Makefile */ -#define DEFAULT_FMT FMT_PPM -#endif - -static IMAGE_FORMATS requested_fmt; - - -/* - * Argument-parsing code. - * The switch parser is designed to be useful with DOS-style command line - * syntax, ie, intermixed switches and file names, where only the switches - * to the left of a given file name affect processing of that file. - * The main program in this file doesn't actually use this capability... - */ - - -static const char * progname; /* program name for error messages */ -static char * outfilename; /* for -outfile switch */ - - -LOCAL(void) -usage (void) -/* complain about bad command line */ -{ - fprintf(stderr, "usage: %s [switches] ", progname); -#ifdef TWO_FILE_COMMANDLINE - fprintf(stderr, "inputfile outputfile\n"); -#else - fprintf(stderr, "[inputfile]\n"); -#endif - - fprintf(stderr, "Switches (names may be abbreviated):\n"); - fprintf(stderr, " -colors N Reduce image to no more than N colors\n"); - fprintf(stderr, " -fast Fast, low-quality processing\n"); - fprintf(stderr, " -grayscale Force grayscale output\n"); -#ifdef IDCT_SCALING_SUPPORTED - fprintf(stderr, " -scale M/N Scale output image by fraction M/N, eg, 1/8\n"); -#endif -#ifdef BMP_SUPPORTED - fprintf(stderr, " -bmp Select BMP output format (Windows style)%s\n", - (DEFAULT_FMT == FMT_BMP ? " (default)" : "")); -#endif -#ifdef GIF_SUPPORTED - fprintf(stderr, " -gif Select GIF output format%s\n", - (DEFAULT_FMT == FMT_GIF ? " (default)" : "")); -#endif -#ifdef BMP_SUPPORTED - fprintf(stderr, " -os2 Select BMP output format (OS/2 style)%s\n", - (DEFAULT_FMT == FMT_OS2 ? " (default)" : "")); -#endif -#ifdef PPM_SUPPORTED - fprintf(stderr, " -pnm Select PBMPLUS (PPM/PGM) output format%s\n", - (DEFAULT_FMT == FMT_PPM ? " (default)" : "")); -#endif -#ifdef RLE_SUPPORTED - fprintf(stderr, " -rle Select Utah RLE output format%s\n", - (DEFAULT_FMT == FMT_RLE ? " (default)" : "")); -#endif -#ifdef TARGA_SUPPORTED - fprintf(stderr, " -targa Select Targa output format%s\n", - (DEFAULT_FMT == FMT_TARGA ? " (default)" : "")); -#endif - fprintf(stderr, "Switches for advanced users:\n"); -#ifdef DCT_ISLOW_SUPPORTED - fprintf(stderr, " -dct int Use integer DCT method%s\n", - (JDCT_DEFAULT == JDCT_ISLOW ? " (default)" : "")); -#endif -#ifdef DCT_IFAST_SUPPORTED - fprintf(stderr, " -dct fast Use fast integer DCT (less accurate)%s\n", - (JDCT_DEFAULT == JDCT_IFAST ? " (default)" : "")); -#endif -#ifdef DCT_FLOAT_SUPPORTED - fprintf(stderr, " -dct float Use floating-point DCT method%s\n", - (JDCT_DEFAULT == JDCT_FLOAT ? " (default)" : "")); -#endif - fprintf(stderr, " -dither fs Use F-S dithering (default)\n"); - fprintf(stderr, " -dither none Don't use dithering in quantization\n"); - fprintf(stderr, " -dither ordered Use ordered dither (medium speed, quality)\n"); -#ifdef QUANT_2PASS_SUPPORTED - fprintf(stderr, " -map FILE Map to colors used in named image file\n"); -#endif - fprintf(stderr, " -nosmooth Don't use high-quality upsampling\n"); -#ifdef QUANT_1PASS_SUPPORTED - fprintf(stderr, " -onepass Use 1-pass quantization (fast, low quality)\n"); -#endif - fprintf(stderr, " -maxmemory N Maximum memory to use (in kbytes)\n"); - fprintf(stderr, " -outfile name Specify name for output file\n"); - fprintf(stderr, " -verbose or -debug Emit debug output\n"); - exit(EXIT_FAILURE); -} - - -LOCAL(int) -parse_switches (j_decompress_ptr cinfo, int argc, char **argv, - int last_file_arg_seen, boolean for_real) -/* Parse optional switches. - * Returns argv[] index of first file-name argument (== argc if none). - * Any file names with indexes <= last_file_arg_seen are ignored; - * they have presumably been processed in a previous iteration. - * (Pass 0 for last_file_arg_seen on the first or only iteration.) - * for_real is FALSE on the first (dummy) pass; we may skip any expensive - * processing. - */ -{ - int argn; - char * arg; - - /* Set up default JPEG parameters. */ - requested_fmt = DEFAULT_FMT; /* set default output file format */ - outfilename = NULL; - cinfo->err->trace_level = 0; - - /* Scan command line options, adjust parameters */ - - for (argn = 1; argn < argc; argn++) { - arg = argv[argn]; - if (*arg != '-') { - /* Not a switch, must be a file name argument */ - if (argn <= last_file_arg_seen) { - outfilename = NULL; /* -outfile applies to just one input file */ - continue; /* ignore this name if previously processed */ - } - break; /* else done parsing switches */ - } - arg++; /* advance past switch marker character */ - - if (keymatch(arg, "bmp", 1)) { - /* BMP output format. */ - requested_fmt = FMT_BMP; - - } else if (keymatch(arg, "colors", 1) || keymatch(arg, "colours", 1) || - keymatch(arg, "quantize", 1) || keymatch(arg, "quantise", 1)) { - /* Do color quantization. */ - int val; - - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%d", &val) != 1) - usage(); - cinfo->desired_number_of_colors = val; - cinfo->quantize_colors = TRUE; - - } else if (keymatch(arg, "dct", 2)) { - /* Select IDCT algorithm. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (keymatch(argv[argn], "int", 1)) { - cinfo->dct_method = JDCT_ISLOW; - } else if (keymatch(argv[argn], "fast", 2)) { - cinfo->dct_method = JDCT_IFAST; - } else if (keymatch(argv[argn], "float", 2)) { - cinfo->dct_method = JDCT_FLOAT; - } else - usage(); - - } else if (keymatch(arg, "dither", 2)) { - /* Select dithering algorithm. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (keymatch(argv[argn], "fs", 2)) { - cinfo->dither_mode = JDITHER_FS; - } else if (keymatch(argv[argn], "none", 2)) { - cinfo->dither_mode = JDITHER_NONE; - } else if (keymatch(argv[argn], "ordered", 2)) { - cinfo->dither_mode = JDITHER_ORDERED; - } else - usage(); - - } else if (keymatch(arg, "debug", 1) || keymatch(arg, "verbose", 1)) { - /* Enable debug printouts. */ - /* On first -d, print version identification */ - static boolean printed_version = FALSE; - - if (! printed_version) { - fprintf(stderr, "Independent JPEG Group's DJPEG, version %s\n%s\n", - JVERSION, JCOPYRIGHT); - printed_version = TRUE; - } - cinfo->err->trace_level++; - - } else if (keymatch(arg, "fast", 1)) { - /* Select recommended processing options for quick-and-dirty output. */ - cinfo->two_pass_quantize = FALSE; - cinfo->dither_mode = JDITHER_ORDERED; - if (! cinfo->quantize_colors) /* don't override an earlier -colors */ - cinfo->desired_number_of_colors = 216; - cinfo->dct_method = JDCT_FASTEST; - cinfo->do_fancy_upsampling = FALSE; - - } else if (keymatch(arg, "gif", 1)) { - /* GIF output format. */ - requested_fmt = FMT_GIF; - - } else if (keymatch(arg, "grayscale", 2) || keymatch(arg, "greyscale",2)) { - /* Force monochrome output. */ - cinfo->out_color_space = JCS_GRAYSCALE; - - } else if (keymatch(arg, "map", 3)) { - /* Quantize to a color map taken from an input file. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (for_real) { /* too expensive to do twice! */ -#ifdef QUANT_2PASS_SUPPORTED /* otherwise can't quantize to supplied map */ - FILE * mapfile; - - if ((mapfile = fopen(argv[argn], READ_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, argv[argn]); - exit(EXIT_FAILURE); - } - read_color_map(cinfo, mapfile); - fclose(mapfile); - cinfo->quantize_colors = TRUE; -#else - ERREXIT(cinfo, JERR_NOT_COMPILED); -#endif - } - - } else if (keymatch(arg, "maxmemory", 3)) { - /* Maximum memory in Kb (or Mb with 'm'). */ - long lval; - char ch = 'x'; - - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) - usage(); - if (ch == 'm' || ch == 'M') - lval *= 1000L; - cinfo->mem->max_memory_to_use = lval * 1000L; - - } else if (keymatch(arg, "nosmooth", 3)) { - /* Suppress fancy upsampling */ - cinfo->do_fancy_upsampling = FALSE; - - } else if (keymatch(arg, "onepass", 3)) { - /* Use fast one-pass quantization. */ - cinfo->two_pass_quantize = FALSE; - - } else if (keymatch(arg, "os2", 3)) { - /* BMP output format (OS/2 flavor). */ - requested_fmt = FMT_OS2; - - } else if (keymatch(arg, "outfile", 4)) { - /* Set output file name. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - outfilename = argv[argn]; /* save it away for later use */ - - } else if (keymatch(arg, "pnm", 1) || keymatch(arg, "ppm", 1)) { - /* PPM/PGM output format. */ - requested_fmt = FMT_PPM; - - } else if (keymatch(arg, "rle", 1)) { - /* RLE output format. */ - requested_fmt = FMT_RLE; - - } else if (keymatch(arg, "scale", 1)) { - /* Scale the output image by a fraction M/N. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%d/%d", - &cinfo->scale_num, &cinfo->scale_denom) != 2) - usage(); - - } else if (keymatch(arg, "targa", 1)) { - /* Targa output format. */ - requested_fmt = FMT_TARGA; - - } else { - usage(); /* bogus switch */ - } - } - - return argn; /* return index of next arg (file name) */ -} - - -/* - * Marker processor for COM and interesting APPn markers. - * This replaces the library's built-in processor, which just skips the marker. - * We want to print out the marker as text, to the extent possible. - * Note this code relies on a non-suspending data source. - */ - -LOCAL(unsigned int) -jpeg_getc (j_decompress_ptr cinfo) -/* Read next byte */ -{ - struct jpeg_source_mgr * datasrc = cinfo->src; - - if (datasrc->bytes_in_buffer == 0) { - if (! (*datasrc->fill_input_buffer) (cinfo)) - ERREXIT(cinfo, JERR_CANT_SUSPEND); - } - datasrc->bytes_in_buffer--; - return GETJOCTET(*datasrc->next_input_byte++); -} - - -METHODDEF(boolean) -print_text_marker (j_decompress_ptr cinfo) -{ - boolean traceit = (cinfo->err->trace_level >= 1); - INT32 length; - unsigned int ch; - unsigned int lastch = 0; - - length = jpeg_getc(cinfo) << 8; - length += jpeg_getc(cinfo); - length -= 2; /* discount the length word itself */ - - if (traceit) { - if (cinfo->unread_marker == JPEG_COM) - fprintf(stderr, "Comment, length %ld:\n", (long) length); - else /* assume it is an APPn otherwise */ - fprintf(stderr, "APP%d, length %ld:\n", - cinfo->unread_marker - JPEG_APP0, (long) length); - } - - while (--length >= 0) { - ch = jpeg_getc(cinfo); - if (traceit) { - /* Emit the character in a readable form. - * Nonprintables are converted to \nnn form, - * while \ is converted to \\. - * Newlines in CR, CR/LF, or LF form will be printed as one newline. - */ - if (ch == '\r') { - fprintf(stderr, "\n"); - } else if (ch == '\n') { - if (lastch != '\r') - fprintf(stderr, "\n"); - } else if (ch == '\\') { - fprintf(stderr, "\\\\"); - } else if (isprint(ch)) { - putc(ch, stderr); - } else { - fprintf(stderr, "\\%03o", ch); - } - lastch = ch; - } - } - - if (traceit) - fprintf(stderr, "\n"); - - return TRUE; -} - - -/* - * The main program. - */ - -int -main (int argc, char **argv) -{ - struct jpeg_decompress_struct cinfo; - struct jpeg_error_mgr jerr; -#ifdef PROGRESS_REPORT - struct cdjpeg_progress_mgr progress; -#endif - int file_index; - djpeg_dest_ptr dest_mgr = NULL; - FILE * input_file; - FILE * output_file; - JDIMENSION num_scanlines; - - /* On Mac, fetch a command line. */ -#ifdef USE_CCOMMAND - argc = ccommand(&argv); -#endif - - progname = argv[0]; - if (progname == NULL || progname[0] == 0) - progname = "djpeg"; /* in case C library doesn't provide it */ - - /* Initialize the JPEG decompression object with default error handling. */ - cinfo.err = jpeg_std_error(&jerr); - jpeg_create_decompress(&cinfo); - /* Add some application-specific error messages (from cderror.h) */ - jerr.addon_message_table = cdjpeg_message_table; - jerr.first_addon_message = JMSG_FIRSTADDONCODE; - jerr.last_addon_message = JMSG_LASTADDONCODE; - - /* Insert custom marker processor for COM and APP12. - * APP12 is used by some digital camera makers for textual info, - * so we provide the ability to display it as text. - * If you like, additional APPn marker types can be selected for display, - * but don't try to override APP0 or APP14 this way (see libjpeg.doc). - */ - jpeg_set_marker_processor(&cinfo, JPEG_COM, print_text_marker); - jpeg_set_marker_processor(&cinfo, JPEG_APP0+12, print_text_marker); - - /* Now safe to enable signal catcher. */ -#ifdef NEED_SIGNAL_CATCHER - enable_signal_catcher((j_common_ptr) &cinfo); -#endif - - /* Scan command line to find file names. */ - /* It is convenient to use just one switch-parsing routine, but the switch - * values read here are ignored; we will rescan the switches after opening - * the input file. - * (Exception: tracing level set here controls verbosity for COM markers - * found during jpeg_read_header...) - */ - - file_index = parse_switches(&cinfo, argc, argv, 0, FALSE); - -#ifdef TWO_FILE_COMMANDLINE - /* Must have either -outfile switch or explicit output file name */ - if (outfilename == NULL) { - if (file_index != argc-2) { - fprintf(stderr, "%s: must name one input and one output file\n", - progname); - usage(); - } - outfilename = argv[file_index+1]; - } else { - if (file_index != argc-1) { - fprintf(stderr, "%s: must name one input and one output file\n", - progname); - usage(); - } - } -#else - /* Unix style: expect zero or one file name */ - if (file_index < argc-1) { - fprintf(stderr, "%s: only one input file\n", progname); - usage(); - } -#endif /* TWO_FILE_COMMANDLINE */ - - /* Open the input file. */ - if (file_index < argc) { - if ((input_file = fopen(argv[file_index], READ_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, argv[file_index]); - exit(EXIT_FAILURE); - } - } else { - /* default input file is stdin */ - input_file = read_stdin(); - } - - /* Open the output file. */ - if (outfilename != NULL) { - if ((output_file = fopen(outfilename, WRITE_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, outfilename); - exit(EXIT_FAILURE); - } - } else { - /* default output file is stdout */ - output_file = write_stdout(); - } - -#ifdef PROGRESS_REPORT - start_progress_monitor((j_common_ptr) &cinfo, &progress); -#endif - - /* Specify data source for decompression */ - jpeg_stdio_src(&cinfo, input_file); - - /* Read file header, set default decompression parameters */ - (void) jpeg_read_header(&cinfo, TRUE); - - /* Adjust default decompression parameters by re-parsing the options */ - file_index = parse_switches(&cinfo, argc, argv, 0, TRUE); - - /* Initialize the output module now to let it override any crucial - * option settings (for instance, GIF wants to force color quantization). - */ - switch (requested_fmt) { -#ifdef BMP_SUPPORTED - case FMT_BMP: - dest_mgr = jinit_write_bmp(&cinfo, FALSE); - break; - case FMT_OS2: - dest_mgr = jinit_write_bmp(&cinfo, TRUE); - break; -#endif -#ifdef GIF_SUPPORTED - case FMT_GIF: - dest_mgr = jinit_write_gif(&cinfo); - break; -#endif -#ifdef PPM_SUPPORTED - case FMT_PPM: - dest_mgr = jinit_write_ppm(&cinfo); - break; -#endif -#ifdef RLE_SUPPORTED - case FMT_RLE: - dest_mgr = jinit_write_rle(&cinfo); - break; -#endif -#ifdef TARGA_SUPPORTED - case FMT_TARGA: - dest_mgr = jinit_write_targa(&cinfo); - break; -#endif - default: - ERREXIT(&cinfo, JERR_UNSUPPORTED_FORMAT); - break; - } - dest_mgr->output_file = output_file; - - /* Start decompressor */ - (void) jpeg_start_decompress(&cinfo); - - /* Write output file header */ - (*dest_mgr->start_output) (&cinfo, dest_mgr); - - /* Process data */ - while (cinfo.output_scanline < cinfo.output_height) { - num_scanlines = jpeg_read_scanlines(&cinfo, dest_mgr->buffer, - dest_mgr->buffer_height); - (*dest_mgr->put_pixel_rows) (&cinfo, dest_mgr, num_scanlines); - } - -#ifdef PROGRESS_REPORT - /* Hack: count final pass as done in case finish_output does an extra pass. - * The library won't have updated completed_passes. - */ - progress.pub.completed_passes = progress.pub.total_passes; -#endif - - /* Finish decompression and release memory. - * I must do it in this order because output module has allocated memory - * of lifespan JPOOL_IMAGE; it needs to finish before releasing memory. - */ - (*dest_mgr->finish_output) (&cinfo, dest_mgr); - (void) jpeg_finish_decompress(&cinfo); - jpeg_destroy_decompress(&cinfo); - - /* Close files, if we opened them */ - if (input_file != stdin) - fclose(input_file); - if (output_file != stdout) - fclose(output_file); - -#ifdef PROGRESS_REPORT - end_progress_monitor((j_common_ptr) &cinfo); -#endif - - /* All done. */ - exit(jerr.num_warnings ? EXIT_WARNING : EXIT_SUCCESS); - return 0; /* suppress no-return-value warnings */ -} +/* + * djpeg.c + * + * Copyright (C) 1991-1997, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains a command-line user interface for the JPEG decompressor. + * It should work on any system with Unix- or MS-DOS-style command lines. + * + * Two different command line styles are permitted, depending on the + * compile-time switch TWO_FILE_COMMANDLINE: + * djpeg [options] inputfile outputfile + * djpeg [options] [inputfile] + * In the second style, output is always to standard output, which you'd + * normally redirect to a file or pipe to some other program. Input is + * either from a named file or from standard input (typically redirected). + * The second style is convenient on Unix but is unhelpful on systems that + * don't support pipes. Also, you MUST use the first style if your system + * doesn't do binary I/O to stdin/stdout. + * To simplify script writing, the "-outfile" switch is provided. The syntax + * djpeg [options] -outfile outputfile inputfile + * works regardless of which command line style is used. + */ + +#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ +#include "jversion.h" /* for version message */ + +#include /* to declare isprint() */ + +#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ +#ifdef __MWERKS__ +#include /* Metrowerks needs this */ +#include /* ... and this */ +#endif +#ifdef THINK_C +#include /* Think declares it here */ +#endif +#endif + + +/* Create the add-on message string table. */ + +#define JMESSAGE(code,string) string , + +static const char * const cdjpeg_message_table[] = { +#include "cderror.h" + NULL +}; + + +/* + * This list defines the known output image formats + * (not all of which need be supported by a given version). + * You can change the default output format by defining DEFAULT_FMT; + * indeed, you had better do so if you undefine PPM_SUPPORTED. + */ + +typedef enum { + FMT_BMP, /* BMP format (Windows flavor) */ + FMT_GIF, /* GIF format */ + FMT_OS2, /* BMP format (OS/2 flavor) */ + FMT_PPM, /* PPM/PGM (PBMPLUS formats) */ + FMT_RLE, /* RLE format */ + FMT_TARGA, /* Targa format */ + FMT_TIFF /* TIFF format */ +} IMAGE_FORMATS; + +#ifndef DEFAULT_FMT /* so can override from CFLAGS in Makefile */ +#define DEFAULT_FMT FMT_PPM +#endif + +static IMAGE_FORMATS requested_fmt; + + +/* + * Argument-parsing code. + * The switch parser is designed to be useful with DOS-style command line + * syntax, ie, intermixed switches and file names, where only the switches + * to the left of a given file name affect processing of that file. + * The main program in this file doesn't actually use this capability... + */ + + +static const char * progname; /* program name for error messages */ +static char * outfilename; /* for -outfile switch */ + + +LOCAL(void) +usage (void) +/* complain about bad command line */ +{ + fprintf(stderr, "usage: %s [switches] ", progname); +#ifdef TWO_FILE_COMMANDLINE + fprintf(stderr, "inputfile outputfile\n"); +#else + fprintf(stderr, "[inputfile]\n"); +#endif + + fprintf(stderr, "Switches (names may be abbreviated):\n"); + fprintf(stderr, " -colors N Reduce image to no more than N colors\n"); + fprintf(stderr, " -fast Fast, low-quality processing\n"); + fprintf(stderr, " -grayscale Force grayscale output\n"); +#ifdef IDCT_SCALING_SUPPORTED + fprintf(stderr, " -scale M/N Scale output image by fraction M/N, eg, 1/8\n"); +#endif +#ifdef BMP_SUPPORTED + fprintf(stderr, " -bmp Select BMP output format (Windows style)%s\n", + (DEFAULT_FMT == FMT_BMP ? " (default)" : "")); +#endif +#ifdef GIF_SUPPORTED + fprintf(stderr, " -gif Select GIF output format%s\n", + (DEFAULT_FMT == FMT_GIF ? " (default)" : "")); +#endif +#ifdef BMP_SUPPORTED + fprintf(stderr, " -os2 Select BMP output format (OS/2 style)%s\n", + (DEFAULT_FMT == FMT_OS2 ? " (default)" : "")); +#endif +#ifdef PPM_SUPPORTED + fprintf(stderr, " -pnm Select PBMPLUS (PPM/PGM) output format%s\n", + (DEFAULT_FMT == FMT_PPM ? " (default)" : "")); +#endif +#ifdef RLE_SUPPORTED + fprintf(stderr, " -rle Select Utah RLE output format%s\n", + (DEFAULT_FMT == FMT_RLE ? " (default)" : "")); +#endif +#ifdef TARGA_SUPPORTED + fprintf(stderr, " -targa Select Targa output format%s\n", + (DEFAULT_FMT == FMT_TARGA ? " (default)" : "")); +#endif + fprintf(stderr, "Switches for advanced users:\n"); +#ifdef DCT_ISLOW_SUPPORTED + fprintf(stderr, " -dct int Use integer DCT method%s\n", + (JDCT_DEFAULT == JDCT_ISLOW ? " (default)" : "")); +#endif +#ifdef DCT_IFAST_SUPPORTED + fprintf(stderr, " -dct fast Use fast integer DCT (less accurate)%s\n", + (JDCT_DEFAULT == JDCT_IFAST ? " (default)" : "")); +#endif +#ifdef DCT_FLOAT_SUPPORTED + fprintf(stderr, " -dct float Use floating-point DCT method%s\n", + (JDCT_DEFAULT == JDCT_FLOAT ? " (default)" : "")); +#endif + fprintf(stderr, " -dither fs Use F-S dithering (default)\n"); + fprintf(stderr, " -dither none Don't use dithering in quantization\n"); + fprintf(stderr, " -dither ordered Use ordered dither (medium speed, quality)\n"); +#ifdef QUANT_2PASS_SUPPORTED + fprintf(stderr, " -map FILE Map to colors used in named image file\n"); +#endif + fprintf(stderr, " -nosmooth Don't use high-quality upsampling\n"); +#ifdef QUANT_1PASS_SUPPORTED + fprintf(stderr, " -onepass Use 1-pass quantization (fast, low quality)\n"); +#endif + fprintf(stderr, " -maxmemory N Maximum memory to use (in kbytes)\n"); + fprintf(stderr, " -outfile name Specify name for output file\n"); + fprintf(stderr, " -verbose or -debug Emit debug output\n"); + exit(EXIT_FAILURE); +} + + +LOCAL(int) +parse_switches (j_decompress_ptr cinfo, int argc, char **argv, + int last_file_arg_seen, boolean for_real) +/* Parse optional switches. + * Returns argv[] index of first file-name argument (== argc if none). + * Any file names with indexes <= last_file_arg_seen are ignored; + * they have presumably been processed in a previous iteration. + * (Pass 0 for last_file_arg_seen on the first or only iteration.) + * for_real is FALSE on the first (dummy) pass; we may skip any expensive + * processing. + */ +{ + int argn; + char * arg; + + /* Set up default JPEG parameters. */ + requested_fmt = DEFAULT_FMT; /* set default output file format */ + outfilename = NULL; + cinfo->err->trace_level = 0; + + /* Scan command line options, adjust parameters */ + + for (argn = 1; argn < argc; argn++) { + arg = argv[argn]; + if (*arg != '-') { + /* Not a switch, must be a file name argument */ + if (argn <= last_file_arg_seen) { + outfilename = NULL; /* -outfile applies to just one input file */ + continue; /* ignore this name if previously processed */ + } + break; /* else done parsing switches */ + } + arg++; /* advance past switch marker character */ + + if (keymatch(arg, "bmp", 1)) { + /* BMP output format. */ + requested_fmt = FMT_BMP; + + } else if (keymatch(arg, "colors", 1) || keymatch(arg, "colours", 1) || + keymatch(arg, "quantize", 1) || keymatch(arg, "quantise", 1)) { + /* Do color quantization. */ + int val; + + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%d", &val) != 1) + usage(); + cinfo->desired_number_of_colors = val; + cinfo->quantize_colors = TRUE; + + } else if (keymatch(arg, "dct", 2)) { + /* Select IDCT algorithm. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (keymatch(argv[argn], "int", 1)) { + cinfo->dct_method = JDCT_ISLOW; + } else if (keymatch(argv[argn], "fast", 2)) { + cinfo->dct_method = JDCT_IFAST; + } else if (keymatch(argv[argn], "float", 2)) { + cinfo->dct_method = JDCT_FLOAT; + } else + usage(); + + } else if (keymatch(arg, "dither", 2)) { + /* Select dithering algorithm. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (keymatch(argv[argn], "fs", 2)) { + cinfo->dither_mode = JDITHER_FS; + } else if (keymatch(argv[argn], "none", 2)) { + cinfo->dither_mode = JDITHER_NONE; + } else if (keymatch(argv[argn], "ordered", 2)) { + cinfo->dither_mode = JDITHER_ORDERED; + } else + usage(); + + } else if (keymatch(arg, "debug", 1) || keymatch(arg, "verbose", 1)) { + /* Enable debug printouts. */ + /* On first -d, print version identification */ + static boolean printed_version = FALSE; + + if (! printed_version) { + fprintf(stderr, "Independent JPEG Group's DJPEG, version %s\n%s\n", + JVERSION, JCOPYRIGHT); + printed_version = TRUE; + } + cinfo->err->trace_level++; + + } else if (keymatch(arg, "fast", 1)) { + /* Select recommended processing options for quick-and-dirty output. */ + cinfo->two_pass_quantize = FALSE; + cinfo->dither_mode = JDITHER_ORDERED; + if (! cinfo->quantize_colors) /* don't override an earlier -colors */ + cinfo->desired_number_of_colors = 216; + cinfo->dct_method = JDCT_FASTEST; + cinfo->do_fancy_upsampling = FALSE; + + } else if (keymatch(arg, "gif", 1)) { + /* GIF output format. */ + requested_fmt = FMT_GIF; + + } else if (keymatch(arg, "grayscale", 2) || keymatch(arg, "greyscale",2)) { + /* Force monochrome output. */ + cinfo->out_color_space = JCS_GRAYSCALE; + + } else if (keymatch(arg, "map", 3)) { + /* Quantize to a color map taken from an input file. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (for_real) { /* too expensive to do twice! */ +#ifdef QUANT_2PASS_SUPPORTED /* otherwise can't quantize to supplied map */ + FILE * mapfile; + + if ((mapfile = fopen(argv[argn], READ_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, argv[argn]); + exit(EXIT_FAILURE); + } + read_color_map(cinfo, mapfile); + fclose(mapfile); + cinfo->quantize_colors = TRUE; +#else + ERREXIT(cinfo, JERR_NOT_COMPILED); +#endif + } + + } else if (keymatch(arg, "maxmemory", 3)) { + /* Maximum memory in Kb (or Mb with 'm'). */ + long lval; + char ch = 'x'; + + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) + usage(); + if (ch == 'm' || ch == 'M') + lval *= 1000L; + cinfo->mem->max_memory_to_use = lval * 1000L; + + } else if (keymatch(arg, "nosmooth", 3)) { + /* Suppress fancy upsampling */ + cinfo->do_fancy_upsampling = FALSE; + + } else if (keymatch(arg, "onepass", 3)) { + /* Use fast one-pass quantization. */ + cinfo->two_pass_quantize = FALSE; + + } else if (keymatch(arg, "os2", 3)) { + /* BMP output format (OS/2 flavor). */ + requested_fmt = FMT_OS2; + + } else if (keymatch(arg, "outfile", 4)) { + /* Set output file name. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + outfilename = argv[argn]; /* save it away for later use */ + + } else if (keymatch(arg, "pnm", 1) || keymatch(arg, "ppm", 1)) { + /* PPM/PGM output format. */ + requested_fmt = FMT_PPM; + + } else if (keymatch(arg, "rle", 1)) { + /* RLE output format. */ + requested_fmt = FMT_RLE; + + } else if (keymatch(arg, "scale", 1)) { + /* Scale the output image by a fraction M/N. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%d/%d", + &cinfo->scale_num, &cinfo->scale_denom) != 2) + usage(); + + } else if (keymatch(arg, "targa", 1)) { + /* Targa output format. */ + requested_fmt = FMT_TARGA; + + } else { + usage(); /* bogus switch */ + } + } + + return argn; /* return index of next arg (file name) */ +} + + +/* + * Marker processor for COM and interesting APPn markers. + * This replaces the library's built-in processor, which just skips the marker. + * We want to print out the marker as text, to the extent possible. + * Note this code relies on a non-suspending data source. + */ + +LOCAL(unsigned int) +jpeg_getc (j_decompress_ptr cinfo) +/* Read next byte */ +{ + struct jpeg_source_mgr * datasrc = cinfo->src; + + if (datasrc->bytes_in_buffer == 0) { + if (! (*datasrc->fill_input_buffer) (cinfo)) + ERREXIT(cinfo, JERR_CANT_SUSPEND); + } + datasrc->bytes_in_buffer--; + return GETJOCTET(*datasrc->next_input_byte++); +} + + +METHODDEF(boolean) +print_text_marker (j_decompress_ptr cinfo) +{ + boolean traceit = (cinfo->err->trace_level >= 1); + INT32 length; + unsigned int ch; + unsigned int lastch = 0; + + length = jpeg_getc(cinfo) << 8; + length += jpeg_getc(cinfo); + length -= 2; /* discount the length word itself */ + + if (traceit) { + if (cinfo->unread_marker == JPEG_COM) + fprintf(stderr, "Comment, length %ld:\n", (long) length); + else /* assume it is an APPn otherwise */ + fprintf(stderr, "APP%d, length %ld:\n", + cinfo->unread_marker - JPEG_APP0, (long) length); + } + + while (--length >= 0) { + ch = jpeg_getc(cinfo); + if (traceit) { + /* Emit the character in a readable form. + * Nonprintables are converted to \nnn form, + * while \ is converted to \\. + * Newlines in CR, CR/LF, or LF form will be printed as one newline. + */ + if (ch == '\r') { + fprintf(stderr, "\n"); + } else if (ch == '\n') { + if (lastch != '\r') + fprintf(stderr, "\n"); + } else if (ch == '\\') { + fprintf(stderr, "\\\\"); + } else if (isprint(ch)) { + putc(ch, stderr); + } else { + fprintf(stderr, "\\%03o", ch); + } + lastch = ch; + } + } + + if (traceit) + fprintf(stderr, "\n"); + + return TRUE; +} + + +/* + * The main program. + */ + +int +main (int argc, char **argv) +{ + struct jpeg_decompress_struct cinfo; + struct jpeg_error_mgr jerr; +#ifdef PROGRESS_REPORT + struct cdjpeg_progress_mgr progress; +#endif + int file_index; + djpeg_dest_ptr dest_mgr = NULL; + FILE * input_file; + FILE * output_file; + JDIMENSION num_scanlines; + + /* On Mac, fetch a command line. */ +#ifdef USE_CCOMMAND + argc = ccommand(&argv); +#endif + + progname = argv[0]; + if (progname == NULL || progname[0] == 0) + progname = "djpeg"; /* in case C library doesn't provide it */ + + /* Initialize the JPEG decompression object with default error handling. */ + cinfo.err = jpeg_std_error(&jerr); + jpeg_create_decompress(&cinfo); + /* Add some application-specific error messages (from cderror.h) */ + jerr.addon_message_table = cdjpeg_message_table; + jerr.first_addon_message = JMSG_FIRSTADDONCODE; + jerr.last_addon_message = JMSG_LASTADDONCODE; + + /* Insert custom marker processor for COM and APP12. + * APP12 is used by some digital camera makers for textual info, + * so we provide the ability to display it as text. + * If you like, additional APPn marker types can be selected for display, + * but don't try to override APP0 or APP14 this way (see libjpeg.doc). + */ + jpeg_set_marker_processor(&cinfo, JPEG_COM, print_text_marker); + jpeg_set_marker_processor(&cinfo, JPEG_APP0+12, print_text_marker); + + /* Now safe to enable signal catcher. */ +#ifdef NEED_SIGNAL_CATCHER + enable_signal_catcher((j_common_ptr) &cinfo); +#endif + + /* Scan command line to find file names. */ + /* It is convenient to use just one switch-parsing routine, but the switch + * values read here are ignored; we will rescan the switches after opening + * the input file. + * (Exception: tracing level set here controls verbosity for COM markers + * found during jpeg_read_header...) + */ + + file_index = parse_switches(&cinfo, argc, argv, 0, FALSE); + +#ifdef TWO_FILE_COMMANDLINE + /* Must have either -outfile switch or explicit output file name */ + if (outfilename == NULL) { + if (file_index != argc-2) { + fprintf(stderr, "%s: must name one input and one output file\n", + progname); + usage(); + } + outfilename = argv[file_index+1]; + } else { + if (file_index != argc-1) { + fprintf(stderr, "%s: must name one input and one output file\n", + progname); + usage(); + } + } +#else + /* Unix style: expect zero or one file name */ + if (file_index < argc-1) { + fprintf(stderr, "%s: only one input file\n", progname); + usage(); + } +#endif /* TWO_FILE_COMMANDLINE */ + + /* Open the input file. */ + if (file_index < argc) { + if ((input_file = fopen(argv[file_index], READ_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, argv[file_index]); + exit(EXIT_FAILURE); + } + } else { + /* default input file is stdin */ + input_file = read_stdin(); + } + + /* Open the output file. */ + if (outfilename != NULL) { + if ((output_file = fopen(outfilename, WRITE_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, outfilename); + exit(EXIT_FAILURE); + } + } else { + /* default output file is stdout */ + output_file = write_stdout(); + } + +#ifdef PROGRESS_REPORT + start_progress_monitor((j_common_ptr) &cinfo, &progress); +#endif + + /* Specify data source for decompression */ + jpeg_stdio_src(&cinfo, input_file); + + /* Read file header, set default decompression parameters */ + (void) jpeg_read_header(&cinfo, TRUE); + + /* Adjust default decompression parameters by re-parsing the options */ + file_index = parse_switches(&cinfo, argc, argv, 0, TRUE); + + /* Initialize the output module now to let it override any crucial + * option settings (for instance, GIF wants to force color quantization). + */ + switch (requested_fmt) { +#ifdef BMP_SUPPORTED + case FMT_BMP: + dest_mgr = jinit_write_bmp(&cinfo, FALSE); + break; + case FMT_OS2: + dest_mgr = jinit_write_bmp(&cinfo, TRUE); + break; +#endif +#ifdef GIF_SUPPORTED + case FMT_GIF: + dest_mgr = jinit_write_gif(&cinfo); + break; +#endif +#ifdef PPM_SUPPORTED + case FMT_PPM: + dest_mgr = jinit_write_ppm(&cinfo); + break; +#endif +#ifdef RLE_SUPPORTED + case FMT_RLE: + dest_mgr = jinit_write_rle(&cinfo); + break; +#endif +#ifdef TARGA_SUPPORTED + case FMT_TARGA: + dest_mgr = jinit_write_targa(&cinfo); + break; +#endif + default: + ERREXIT(&cinfo, JERR_UNSUPPORTED_FORMAT); + break; + } + dest_mgr->output_file = output_file; + + /* Start decompressor */ + (void) jpeg_start_decompress(&cinfo); + + /* Write output file header */ + (*dest_mgr->start_output) (&cinfo, dest_mgr); + + /* Process data */ + while (cinfo.output_scanline < cinfo.output_height) { + num_scanlines = jpeg_read_scanlines(&cinfo, dest_mgr->buffer, + dest_mgr->buffer_height); + (*dest_mgr->put_pixel_rows) (&cinfo, dest_mgr, num_scanlines); + } + +#ifdef PROGRESS_REPORT + /* Hack: count final pass as done in case finish_output does an extra pass. + * The library won't have updated completed_passes. + */ + progress.pub.completed_passes = progress.pub.total_passes; +#endif + + /* Finish decompression and release memory. + * I must do it in this order because output module has allocated memory + * of lifespan JPOOL_IMAGE; it needs to finish before releasing memory. + */ + (*dest_mgr->finish_output) (&cinfo, dest_mgr); + (void) jpeg_finish_decompress(&cinfo); + jpeg_destroy_decompress(&cinfo); + + /* Close files, if we opened them */ + if (input_file != stdin) + fclose(input_file); + if (output_file != stdout) + fclose(output_file); + +#ifdef PROGRESS_REPORT + end_progress_monitor((j_common_ptr) &cinfo); +#endif + + /* All done. */ + exit(jerr.num_warnings ? EXIT_WARNING : EXIT_SUCCESS); + return 0; /* suppress no-return-value warnings */ +} diff --git a/test/monniaux/jpeg-6b/example.c b/test/monniaux/jpeg-6b/example.c index 9dcbd7d9..7fc354f0 100644 --- a/test/monniaux/jpeg-6b/example.c +++ b/test/monniaux/jpeg-6b/example.c @@ -1,433 +1,433 @@ -/* - * example.c - * - * This file illustrates how to use the IJG code as a subroutine library - * to read or write JPEG image files. You should look at this code in - * conjunction with the documentation file libjpeg.doc. - * - * This code will not do anything useful as-is, but it may be helpful as a - * skeleton for constructing routines that call the JPEG library. - * - * We present these routines in the same coding style used in the JPEG code - * (ANSI function definitions, etc); but you are of course free to code your - * routines in a different style if you prefer. - */ - -#include - -/* - * Include file for users of JPEG library. - * You will need to have included system headers that define at least - * the typedefs FILE and size_t before you can include jpeglib.h. - * (stdio.h is sufficient on ANSI-conforming systems.) - * You may also wish to include "jerror.h". - */ - -#include "jpeglib.h" - -/* - * is used for the optional error recovery mechanism shown in - * the second part of the example. - */ - -#include - - - -/******************** JPEG COMPRESSION SAMPLE INTERFACE *******************/ - -/* This half of the example shows how to feed data into the JPEG compressor. - * We present a minimal version that does not worry about refinements such - * as error recovery (the JPEG code will just exit() if it gets an error). - */ - - -/* - * IMAGE DATA FORMATS: - * - * The standard input image format is a rectangular array of pixels, with - * each pixel having the same number of "component" values (color channels). - * Each pixel row is an array of JSAMPLEs (which typically are unsigned chars). - * If you are working with color data, then the color values for each pixel - * must be adjacent in the row; for example, R,G,B,R,G,B,R,G,B,... for 24-bit - * RGB color. - * - * For this example, we'll assume that this data structure matches the way - * our application has stored the image in memory, so we can just pass a - * pointer to our image buffer. In particular, let's say that the image is - * RGB color and is described by: - */ - -extern JSAMPLE * image_buffer; /* Points to large array of R,G,B-order data */ -extern int image_height; /* Number of rows in image */ -extern int image_width; /* Number of columns in image */ - - -/* - * Sample routine for JPEG compression. We assume that the target file name - * and a compression quality factor are passed in. - */ - -GLOBAL(void) -write_JPEG_file (char * filename, int quality) -{ - /* This struct contains the JPEG compression parameters and pointers to - * working space (which is allocated as needed by the JPEG library). - * It is possible to have several such structures, representing multiple - * compression/decompression processes, in existence at once. We refer - * to any one struct (and its associated working data) as a "JPEG object". - */ - struct jpeg_compress_struct cinfo; - /* This struct represents a JPEG error handler. It is declared separately - * because applications often want to supply a specialized error handler - * (see the second half of this file for an example). But here we just - * take the easy way out and use the standard error handler, which will - * print a message on stderr and call exit() if compression fails. - * Note that this struct must live as long as the main JPEG parameter - * struct, to avoid dangling-pointer problems. - */ - struct jpeg_error_mgr jerr; - /* More stuff */ - FILE * outfile; /* target file */ - JSAMPROW row_pointer[1]; /* pointer to JSAMPLE row[s] */ - int row_stride; /* physical row width in image buffer */ - - /* Step 1: allocate and initialize JPEG compression object */ - - /* We have to set up the error handler first, in case the initialization - * step fails. (Unlikely, but it could happen if you are out of memory.) - * This routine fills in the contents of struct jerr, and returns jerr's - * address which we place into the link field in cinfo. - */ - cinfo.err = jpeg_std_error(&jerr); - /* Now we can initialize the JPEG compression object. */ - jpeg_create_compress(&cinfo); - - /* Step 2: specify data destination (eg, a file) */ - /* Note: steps 2 and 3 can be done in either order. */ - - /* Here we use the library-supplied code to send compressed data to a - * stdio stream. You can also write your own code to do something else. - * VERY IMPORTANT: use "b" option to fopen() if you are on a machine that - * requires it in order to write binary files. - */ - if ((outfile = fopen(filename, "wb")) == NULL) { - fprintf(stderr, "can't open %s\n", filename); - exit(1); - } - jpeg_stdio_dest(&cinfo, outfile); - - /* Step 3: set parameters for compression */ - - /* First we supply a description of the input image. - * Four fields of the cinfo struct must be filled in: - */ - cinfo.image_width = image_width; /* image width and height, in pixels */ - cinfo.image_height = image_height; - cinfo.input_components = 3; /* # of color components per pixel */ - cinfo.in_color_space = JCS_RGB; /* colorspace of input image */ - /* Now use the library's routine to set default compression parameters. - * (You must set at least cinfo.in_color_space before calling this, - * since the defaults depend on the source color space.) - */ - jpeg_set_defaults(&cinfo); - /* Now you can set any non-default parameters you wish to. - * Here we just illustrate the use of quality (quantization table) scaling: - */ - jpeg_set_quality(&cinfo, quality, TRUE /* limit to baseline-JPEG values */); - - /* Step 4: Start compressor */ - - /* TRUE ensures that we will write a complete interchange-JPEG file. - * Pass TRUE unless you are very sure of what you're doing. - */ - jpeg_start_compress(&cinfo, TRUE); - - /* Step 5: while (scan lines remain to be written) */ - /* jpeg_write_scanlines(...); */ - - /* Here we use the library's state variable cinfo.next_scanline as the - * loop counter, so that we don't have to keep track ourselves. - * To keep things simple, we pass one scanline per call; you can pass - * more if you wish, though. - */ - row_stride = image_width * 3; /* JSAMPLEs per row in image_buffer */ - - while (cinfo.next_scanline < cinfo.image_height) { - /* jpeg_write_scanlines expects an array of pointers to scanlines. - * Here the array is only one element long, but you could pass - * more than one scanline at a time if that's more convenient. - */ - row_pointer[0] = & image_buffer[cinfo.next_scanline * row_stride]; - (void) jpeg_write_scanlines(&cinfo, row_pointer, 1); - } - - /* Step 6: Finish compression */ - - jpeg_finish_compress(&cinfo); - /* After finish_compress, we can close the output file. */ - fclose(outfile); - - /* Step 7: release JPEG compression object */ - - /* This is an important step since it will release a good deal of memory. */ - jpeg_destroy_compress(&cinfo); - - /* And we're done! */ -} - - -/* - * SOME FINE POINTS: - * - * In the above loop, we ignored the return value of jpeg_write_scanlines, - * which is the number of scanlines actually written. We could get away - * with this because we were only relying on the value of cinfo.next_scanline, - * which will be incremented correctly. If you maintain additional loop - * variables then you should be careful to increment them properly. - * Actually, for output to a stdio stream you needn't worry, because - * then jpeg_write_scanlines will write all the lines passed (or else exit - * with a fatal error). Partial writes can only occur if you use a data - * destination module that can demand suspension of the compressor. - * (If you don't know what that's for, you don't need it.) - * - * If the compressor requires full-image buffers (for entropy-coding - * optimization or a multi-scan JPEG file), it will create temporary - * files for anything that doesn't fit within the maximum-memory setting. - * (Note that temp files are NOT needed if you use the default parameters.) - * On some systems you may need to set up a signal handler to ensure that - * temporary files are deleted if the program is interrupted. See libjpeg.doc. - * - * Scanlines MUST be supplied in top-to-bottom order if you want your JPEG - * files to be compatible with everyone else's. If you cannot readily read - * your data in that order, you'll need an intermediate array to hold the - * image. See rdtarga.c or rdbmp.c for examples of handling bottom-to-top - * source data using the JPEG code's internal virtual-array mechanisms. - */ - - - -/******************** JPEG DECOMPRESSION SAMPLE INTERFACE *******************/ - -/* This half of the example shows how to read data from the JPEG decompressor. - * It's a bit more refined than the above, in that we show: - * (a) how to modify the JPEG library's standard error-reporting behavior; - * (b) how to allocate workspace using the library's memory manager. - * - * Just to make this example a little different from the first one, we'll - * assume that we do not intend to put the whole image into an in-memory - * buffer, but to send it line-by-line someplace else. We need a one- - * scanline-high JSAMPLE array as a work buffer, and we will let the JPEG - * memory manager allocate it for us. This approach is actually quite useful - * because we don't need to remember to deallocate the buffer separately: it - * will go away automatically when the JPEG object is cleaned up. - */ - - -/* - * ERROR HANDLING: - * - * The JPEG library's standard error handler (jerror.c) is divided into - * several "methods" which you can override individually. This lets you - * adjust the behavior without duplicating a lot of code, which you might - * have to update with each future release. - * - * Our example here shows how to override the "error_exit" method so that - * control is returned to the library's caller when a fatal error occurs, - * rather than calling exit() as the standard error_exit method does. - * - * We use C's setjmp/longjmp facility to return control. This means that the - * routine which calls the JPEG library must first execute a setjmp() call to - * establish the return point. We want the replacement error_exit to do a - * longjmp(). But we need to make the setjmp buffer accessible to the - * error_exit routine. To do this, we make a private extension of the - * standard JPEG error handler object. (If we were using C++, we'd say we - * were making a subclass of the regular error handler.) - * - * Here's the extended error handler struct: - */ - -struct my_error_mgr { - struct jpeg_error_mgr pub; /* "public" fields */ - - jmp_buf setjmp_buffer; /* for return to caller */ -}; - -typedef struct my_error_mgr * my_error_ptr; - -/* - * Here's the routine that will replace the standard error_exit method: - */ - -METHODDEF(void) -my_error_exit (j_common_ptr cinfo) -{ - /* cinfo->err really points to a my_error_mgr struct, so coerce pointer */ - my_error_ptr myerr = (my_error_ptr) cinfo->err; - - /* Always display the message. */ - /* We could postpone this until after returning, if we chose. */ - (*cinfo->err->output_message) (cinfo); - - /* Return control to the setjmp point */ - longjmp(myerr->setjmp_buffer, 1); -} - - -/* - * Sample routine for JPEG decompression. We assume that the source file name - * is passed in. We want to return 1 on success, 0 on error. - */ - - -GLOBAL(int) -read_JPEG_file (char * filename) -{ - /* This struct contains the JPEG decompression parameters and pointers to - * working space (which is allocated as needed by the JPEG library). - */ - struct jpeg_decompress_struct cinfo; - /* We use our private extension JPEG error handler. - * Note that this struct must live as long as the main JPEG parameter - * struct, to avoid dangling-pointer problems. - */ - struct my_error_mgr jerr; - /* More stuff */ - FILE * infile; /* source file */ - JSAMPARRAY buffer; /* Output row buffer */ - int row_stride; /* physical row width in output buffer */ - - /* In this example we want to open the input file before doing anything else, - * so that the setjmp() error recovery below can assume the file is open. - * VERY IMPORTANT: use "b" option to fopen() if you are on a machine that - * requires it in order to read binary files. - */ - - if ((infile = fopen(filename, "rb")) == NULL) { - fprintf(stderr, "can't open %s\n", filename); - return 0; - } - - /* Step 1: allocate and initialize JPEG decompression object */ - - /* We set up the normal JPEG error routines, then override error_exit. */ - cinfo.err = jpeg_std_error(&jerr.pub); - jerr.pub.error_exit = my_error_exit; - /* Establish the setjmp return context for my_error_exit to use. */ - if (setjmp(jerr.setjmp_buffer)) { - /* If we get here, the JPEG code has signaled an error. - * We need to clean up the JPEG object, close the input file, and return. - */ - jpeg_destroy_decompress(&cinfo); - fclose(infile); - return 0; - } - /* Now we can initialize the JPEG decompression object. */ - jpeg_create_decompress(&cinfo); - - /* Step 2: specify data source (eg, a file) */ - - jpeg_stdio_src(&cinfo, infile); - - /* Step 3: read file parameters with jpeg_read_header() */ - - (void) jpeg_read_header(&cinfo, TRUE); - /* We can ignore the return value from jpeg_read_header since - * (a) suspension is not possible with the stdio data source, and - * (b) we passed TRUE to reject a tables-only JPEG file as an error. - * See libjpeg.doc for more info. - */ - - /* Step 4: set parameters for decompression */ - - /* In this example, we don't need to change any of the defaults set by - * jpeg_read_header(), so we do nothing here. - */ - - /* Step 5: Start decompressor */ - - (void) jpeg_start_decompress(&cinfo); - /* We can ignore the return value since suspension is not possible - * with the stdio data source. - */ - - /* We may need to do some setup of our own at this point before reading - * the data. After jpeg_start_decompress() we have the correct scaled - * output image dimensions available, as well as the output colormap - * if we asked for color quantization. - * In this example, we need to make an output work buffer of the right size. - */ - /* JSAMPLEs per row in output buffer */ - row_stride = cinfo.output_width * cinfo.output_components; - /* Make a one-row-high sample array that will go away when done with image */ - buffer = (*cinfo.mem->alloc_sarray) - ((j_common_ptr) &cinfo, JPOOL_IMAGE, row_stride, 1); - - /* Step 6: while (scan lines remain to be read) */ - /* jpeg_read_scanlines(...); */ - - /* Here we use the library's state variable cinfo.output_scanline as the - * loop counter, so that we don't have to keep track ourselves. - */ - while (cinfo.output_scanline < cinfo.output_height) { - /* jpeg_read_scanlines expects an array of pointers to scanlines. - * Here the array is only one element long, but you could ask for - * more than one scanline at a time if that's more convenient. - */ - (void) jpeg_read_scanlines(&cinfo, buffer, 1); - /* Assume put_scanline_someplace wants a pointer and sample count. */ - put_scanline_someplace(buffer[0], row_stride); - } - - /* Step 7: Finish decompression */ - - (void) jpeg_finish_decompress(&cinfo); - /* We can ignore the return value since suspension is not possible - * with the stdio data source. - */ - - /* Step 8: Release JPEG decompression object */ - - /* This is an important step since it will release a good deal of memory. */ - jpeg_destroy_decompress(&cinfo); - - /* After finish_decompress, we can close the input file. - * Here we postpone it until after no more JPEG errors are possible, - * so as to simplify the setjmp error logic above. (Actually, I don't - * think that jpeg_destroy can do an error exit, but why assume anything...) - */ - fclose(infile); - - /* At this point you may want to check to see whether any corrupt-data - * warnings occurred (test whether jerr.pub.num_warnings is nonzero). - */ - - /* And we're done! */ - return 1; -} - - -/* - * SOME FINE POINTS: - * - * In the above code, we ignored the return value of jpeg_read_scanlines, - * which is the number of scanlines actually read. We could get away with - * this because we asked for only one line at a time and we weren't using - * a suspending data source. See libjpeg.doc for more info. - * - * We cheated a bit by calling alloc_sarray() after jpeg_start_decompress(); - * we should have done it beforehand to ensure that the space would be - * counted against the JPEG max_memory setting. In some systems the above - * code would risk an out-of-memory error. However, in general we don't - * know the output image dimensions before jpeg_start_decompress(), unless we - * call jpeg_calc_output_dimensions(). See libjpeg.doc for more about this. - * - * Scanlines are returned in the same order as they appear in the JPEG file, - * which is standardly top-to-bottom. If you must emit data bottom-to-top, - * you can use one of the virtual arrays provided by the JPEG memory manager - * to invert the data. See wrbmp.c for an example. - * - * As with compression, some operating modes may require temporary files. - * On some systems you may need to set up a signal handler to ensure that - * temporary files are deleted if the program is interrupted. See libjpeg.doc. - */ +/* + * example.c + * + * This file illustrates how to use the IJG code as a subroutine library + * to read or write JPEG image files. You should look at this code in + * conjunction with the documentation file libjpeg.doc. + * + * This code will not do anything useful as-is, but it may be helpful as a + * skeleton for constructing routines that call the JPEG library. + * + * We present these routines in the same coding style used in the JPEG code + * (ANSI function definitions, etc); but you are of course free to code your + * routines in a different style if you prefer. + */ + +#include + +/* + * Include file for users of JPEG library. + * You will need to have included system headers that define at least + * the typedefs FILE and size_t before you can include jpeglib.h. + * (stdio.h is sufficient on ANSI-conforming systems.) + * You may also wish to include "jerror.h". + */ + +#include "jpeglib.h" + +/* + * is used for the optional error recovery mechanism shown in + * the second part of the example. + */ + +#include + + + +/******************** JPEG COMPRESSION SAMPLE INTERFACE *******************/ + +/* This half of the example shows how to feed data into the JPEG compressor. + * We present a minimal version that does not worry about refinements such + * as error recovery (the JPEG code will just exit() if it gets an error). + */ + + +/* + * IMAGE DATA FORMATS: + * + * The standard input image format is a rectangular array of pixels, with + * each pixel having the same number of "component" values (color channels). + * Each pixel row is an array of JSAMPLEs (which typically are unsigned chars). + * If you are working with color data, then the color values for each pixel + * must be adjacent in the row; for example, R,G,B,R,G,B,R,G,B,... for 24-bit + * RGB color. + * + * For this example, we'll assume that this data structure matches the way + * our application has stored the image in memory, so we can just pass a + * pointer to our image buffer. In particular, let's say that the image is + * RGB color and is described by: + */ + +extern JSAMPLE * image_buffer; /* Points to large array of R,G,B-order data */ +extern int image_height; /* Number of rows in image */ +extern int image_width; /* Number of columns in image */ + + +/* + * Sample routine for JPEG compression. We assume that the target file name + * and a compression quality factor are passed in. + */ + +GLOBAL(void) +write_JPEG_file (char * filename, int quality) +{ + /* This struct contains the JPEG compression parameters and pointers to + * working space (which is allocated as needed by the JPEG library). + * It is possible to have several such structures, representing multiple + * compression/decompression processes, in existence at once. We refer + * to any one struct (and its associated working data) as a "JPEG object". + */ + struct jpeg_compress_struct cinfo; + /* This struct represents a JPEG error handler. It is declared separately + * because applications often want to supply a specialized error handler + * (see the second half of this file for an example). But here we just + * take the easy way out and use the standard error handler, which will + * print a message on stderr and call exit() if compression fails. + * Note that this struct must live as long as the main JPEG parameter + * struct, to avoid dangling-pointer problems. + */ + struct jpeg_error_mgr jerr; + /* More stuff */ + FILE * outfile; /* target file */ + JSAMPROW row_pointer[1]; /* pointer to JSAMPLE row[s] */ + int row_stride; /* physical row width in image buffer */ + + /* Step 1: allocate and initialize JPEG compression object */ + + /* We have to set up the error handler first, in case the initialization + * step fails. (Unlikely, but it could happen if you are out of memory.) + * This routine fills in the contents of struct jerr, and returns jerr's + * address which we place into the link field in cinfo. + */ + cinfo.err = jpeg_std_error(&jerr); + /* Now we can initialize the JPEG compression object. */ + jpeg_create_compress(&cinfo); + + /* Step 2: specify data destination (eg, a file) */ + /* Note: steps 2 and 3 can be done in either order. */ + + /* Here we use the library-supplied code to send compressed data to a + * stdio stream. You can also write your own code to do something else. + * VERY IMPORTANT: use "b" option to fopen() if you are on a machine that + * requires it in order to write binary files. + */ + if ((outfile = fopen(filename, "wb")) == NULL) { + fprintf(stderr, "can't open %s\n", filename); + exit(1); + } + jpeg_stdio_dest(&cinfo, outfile); + + /* Step 3: set parameters for compression */ + + /* First we supply a description of the input image. + * Four fields of the cinfo struct must be filled in: + */ + cinfo.image_width = image_width; /* image width and height, in pixels */ + cinfo.image_height = image_height; + cinfo.input_components = 3; /* # of color components per pixel */ + cinfo.in_color_space = JCS_RGB; /* colorspace of input image */ + /* Now use the library's routine to set default compression parameters. + * (You must set at least cinfo.in_color_space before calling this, + * since the defaults depend on the source color space.) + */ + jpeg_set_defaults(&cinfo); + /* Now you can set any non-default parameters you wish to. + * Here we just illustrate the use of quality (quantization table) scaling: + */ + jpeg_set_quality(&cinfo, quality, TRUE /* limit to baseline-JPEG values */); + + /* Step 4: Start compressor */ + + /* TRUE ensures that we will write a complete interchange-JPEG file. + * Pass TRUE unless you are very sure of what you're doing. + */ + jpeg_start_compress(&cinfo, TRUE); + + /* Step 5: while (scan lines remain to be written) */ + /* jpeg_write_scanlines(...); */ + + /* Here we use the library's state variable cinfo.next_scanline as the + * loop counter, so that we don't have to keep track ourselves. + * To keep things simple, we pass one scanline per call; you can pass + * more if you wish, though. + */ + row_stride = image_width * 3; /* JSAMPLEs per row in image_buffer */ + + while (cinfo.next_scanline < cinfo.image_height) { + /* jpeg_write_scanlines expects an array of pointers to scanlines. + * Here the array is only one element long, but you could pass + * more than one scanline at a time if that's more convenient. + */ + row_pointer[0] = & image_buffer[cinfo.next_scanline * row_stride]; + (void) jpeg_write_scanlines(&cinfo, row_pointer, 1); + } + + /* Step 6: Finish compression */ + + jpeg_finish_compress(&cinfo); + /* After finish_compress, we can close the output file. */ + fclose(outfile); + + /* Step 7: release JPEG compression object */ + + /* This is an important step since it will release a good deal of memory. */ + jpeg_destroy_compress(&cinfo); + + /* And we're done! */ +} + + +/* + * SOME FINE POINTS: + * + * In the above loop, we ignored the return value of jpeg_write_scanlines, + * which is the number of scanlines actually written. We could get away + * with this because we were only relying on the value of cinfo.next_scanline, + * which will be incremented correctly. If you maintain additional loop + * variables then you should be careful to increment them properly. + * Actually, for output to a stdio stream you needn't worry, because + * then jpeg_write_scanlines will write all the lines passed (or else exit + * with a fatal error). Partial writes can only occur if you use a data + * destination module that can demand suspension of the compressor. + * (If you don't know what that's for, you don't need it.) + * + * If the compressor requires full-image buffers (for entropy-coding + * optimization or a multi-scan JPEG file), it will create temporary + * files for anything that doesn't fit within the maximum-memory setting. + * (Note that temp files are NOT needed if you use the default parameters.) + * On some systems you may need to set up a signal handler to ensure that + * temporary files are deleted if the program is interrupted. See libjpeg.doc. + * + * Scanlines MUST be supplied in top-to-bottom order if you want your JPEG + * files to be compatible with everyone else's. If you cannot readily read + * your data in that order, you'll need an intermediate array to hold the + * image. See rdtarga.c or rdbmp.c for examples of handling bottom-to-top + * source data using the JPEG code's internal virtual-array mechanisms. + */ + + + +/******************** JPEG DECOMPRESSION SAMPLE INTERFACE *******************/ + +/* This half of the example shows how to read data from the JPEG decompressor. + * It's a bit more refined than the above, in that we show: + * (a) how to modify the JPEG library's standard error-reporting behavior; + * (b) how to allocate workspace using the library's memory manager. + * + * Just to make this example a little different from the first one, we'll + * assume that we do not intend to put the whole image into an in-memory + * buffer, but to send it line-by-line someplace else. We need a one- + * scanline-high JSAMPLE array as a work buffer, and we will let the JPEG + * memory manager allocate it for us. This approach is actually quite useful + * because we don't need to remember to deallocate the buffer separately: it + * will go away automatically when the JPEG object is cleaned up. + */ + + +/* + * ERROR HANDLING: + * + * The JPEG library's standard error handler (jerror.c) is divided into + * several "methods" which you can override individually. This lets you + * adjust the behavior without duplicating a lot of code, which you might + * have to update with each future release. + * + * Our example here shows how to override the "error_exit" method so that + * control is returned to the library's caller when a fatal error occurs, + * rather than calling exit() as the standard error_exit method does. + * + * We use C's setjmp/longjmp facility to return control. This means that the + * routine which calls the JPEG library must first execute a setjmp() call to + * establish the return point. We want the replacement error_exit to do a + * longjmp(). But we need to make the setjmp buffer accessible to the + * error_exit routine. To do this, we make a private extension of the + * standard JPEG error handler object. (If we were using C++, we'd say we + * were making a subclass of the regular error handler.) + * + * Here's the extended error handler struct: + */ + +struct my_error_mgr { + struct jpeg_error_mgr pub; /* "public" fields */ + + jmp_buf setjmp_buffer; /* for return to caller */ +}; + +typedef struct my_error_mgr * my_error_ptr; + +/* + * Here's the routine that will replace the standard error_exit method: + */ + +METHODDEF(void) +my_error_exit (j_common_ptr cinfo) +{ + /* cinfo->err really points to a my_error_mgr struct, so coerce pointer */ + my_error_ptr myerr = (my_error_ptr) cinfo->err; + + /* Always display the message. */ + /* We could postpone this until after returning, if we chose. */ + (*cinfo->err->output_message) (cinfo); + + /* Return control to the setjmp point */ + longjmp(myerr->setjmp_buffer, 1); +} + + +/* + * Sample routine for JPEG decompression. We assume that the source file name + * is passed in. We want to return 1 on success, 0 on error. + */ + + +GLOBAL(int) +read_JPEG_file (char * filename) +{ + /* This struct contains the JPEG decompression parameters and pointers to + * working space (which is allocated as needed by the JPEG library). + */ + struct jpeg_decompress_struct cinfo; + /* We use our private extension JPEG error handler. + * Note that this struct must live as long as the main JPEG parameter + * struct, to avoid dangling-pointer problems. + */ + struct my_error_mgr jerr; + /* More stuff */ + FILE * infile; /* source file */ + JSAMPARRAY buffer; /* Output row buffer */ + int row_stride; /* physical row width in output buffer */ + + /* In this example we want to open the input file before doing anything else, + * so that the setjmp() error recovery below can assume the file is open. + * VERY IMPORTANT: use "b" option to fopen() if you are on a machine that + * requires it in order to read binary files. + */ + + if ((infile = fopen(filename, "rb")) == NULL) { + fprintf(stderr, "can't open %s\n", filename); + return 0; + } + + /* Step 1: allocate and initialize JPEG decompression object */ + + /* We set up the normal JPEG error routines, then override error_exit. */ + cinfo.err = jpeg_std_error(&jerr.pub); + jerr.pub.error_exit = my_error_exit; + /* Establish the setjmp return context for my_error_exit to use. */ + if (setjmp(jerr.setjmp_buffer)) { + /* If we get here, the JPEG code has signaled an error. + * We need to clean up the JPEG object, close the input file, and return. + */ + jpeg_destroy_decompress(&cinfo); + fclose(infile); + return 0; + } + /* Now we can initialize the JPEG decompression object. */ + jpeg_create_decompress(&cinfo); + + /* Step 2: specify data source (eg, a file) */ + + jpeg_stdio_src(&cinfo, infile); + + /* Step 3: read file parameters with jpeg_read_header() */ + + (void) jpeg_read_header(&cinfo, TRUE); + /* We can ignore the return value from jpeg_read_header since + * (a) suspension is not possible with the stdio data source, and + * (b) we passed TRUE to reject a tables-only JPEG file as an error. + * See libjpeg.doc for more info. + */ + + /* Step 4: set parameters for decompression */ + + /* In this example, we don't need to change any of the defaults set by + * jpeg_read_header(), so we do nothing here. + */ + + /* Step 5: Start decompressor */ + + (void) jpeg_start_decompress(&cinfo); + /* We can ignore the return value since suspension is not possible + * with the stdio data source. + */ + + /* We may need to do some setup of our own at this point before reading + * the data. After jpeg_start_decompress() we have the correct scaled + * output image dimensions available, as well as the output colormap + * if we asked for color quantization. + * In this example, we need to make an output work buffer of the right size. + */ + /* JSAMPLEs per row in output buffer */ + row_stride = cinfo.output_width * cinfo.output_components; + /* Make a one-row-high sample array that will go away when done with image */ + buffer = (*cinfo.mem->alloc_sarray) + ((j_common_ptr) &cinfo, JPOOL_IMAGE, row_stride, 1); + + /* Step 6: while (scan lines remain to be read) */ + /* jpeg_read_scanlines(...); */ + + /* Here we use the library's state variable cinfo.output_scanline as the + * loop counter, so that we don't have to keep track ourselves. + */ + while (cinfo.output_scanline < cinfo.output_height) { + /* jpeg_read_scanlines expects an array of pointers to scanlines. + * Here the array is only one element long, but you could ask for + * more than one scanline at a time if that's more convenient. + */ + (void) jpeg_read_scanlines(&cinfo, buffer, 1); + /* Assume put_scanline_someplace wants a pointer and sample count. */ + put_scanline_someplace(buffer[0], row_stride); + } + + /* Step 7: Finish decompression */ + + (void) jpeg_finish_decompress(&cinfo); + /* We can ignore the return value since suspension is not possible + * with the stdio data source. + */ + + /* Step 8: Release JPEG decompression object */ + + /* This is an important step since it will release a good deal of memory. */ + jpeg_destroy_decompress(&cinfo); + + /* After finish_decompress, we can close the input file. + * Here we postpone it until after no more JPEG errors are possible, + * so as to simplify the setjmp error logic above. (Actually, I don't + * think that jpeg_destroy can do an error exit, but why assume anything...) + */ + fclose(infile); + + /* At this point you may want to check to see whether any corrupt-data + * warnings occurred (test whether jerr.pub.num_warnings is nonzero). + */ + + /* And we're done! */ + return 1; +} + + +/* + * SOME FINE POINTS: + * + * In the above code, we ignored the return value of jpeg_read_scanlines, + * which is the number of scanlines actually read. We could get away with + * this because we asked for only one line at a time and we weren't using + * a suspending data source. See libjpeg.doc for more info. + * + * We cheated a bit by calling alloc_sarray() after jpeg_start_decompress(); + * we should have done it beforehand to ensure that the space would be + * counted against the JPEG max_memory setting. In some systems the above + * code would risk an out-of-memory error. However, in general we don't + * know the output image dimensions before jpeg_start_decompress(), unless we + * call jpeg_calc_output_dimensions(). See libjpeg.doc for more about this. + * + * Scanlines are returned in the same order as they appear in the JPEG file, + * which is standardly top-to-bottom. If you must emit data bottom-to-top, + * you can use one of the virtual arrays provided by the JPEG memory manager + * to invert the data. See wrbmp.c for an example. + * + * As with compression, some operating modes may require temporary files. + * On some systems you may need to set up a signal handler to ensure that + * temporary files are deleted if the program is interrupted. See libjpeg.doc. + */ diff --git a/test/monniaux/jpeg-6b/jcapimin.c b/test/monniaux/jpeg-6b/jcapimin.c index 71334e6e..54fb8c58 100644 --- a/test/monniaux/jpeg-6b/jcapimin.c +++ b/test/monniaux/jpeg-6b/jcapimin.c @@ -73,10 +73,8 @@ jpeg_CreateCompress (j_compress_ptr cinfo, int version, size_t structsize) cinfo->script_space = NULL; -#ifdef HAVE_FLOAT cinfo->input_gamma = 1.0; /* in case application forgets */ -#endif - + /* OK, I'm ready */ cinfo->global_state = CSTATE_START; } diff --git a/test/monniaux/jpeg-6b/jccoefct.c b/test/monniaux/jpeg-6b/jccoefct.c index 2a0812fd..1963ddb6 100644 --- a/test/monniaux/jpeg-6b/jccoefct.c +++ b/test/monniaux/jpeg-6b/jccoefct.c @@ -265,13 +265,13 @@ compress_first_pass (j_compress_ptr cinfo, JSAMPIMAGE input_buf) block_rows = compptr->v_samp_factor; else { /* NB: can't use last_row_height here, since may not be set! */ - block_rows = (int) INT_UMOD(compptr->height_in_blocks, compptr->v_samp_factor); + block_rows = (int) (compptr->height_in_blocks % compptr->v_samp_factor); if (block_rows == 0) block_rows = compptr->v_samp_factor; } blocks_across = compptr->width_in_blocks; h_samp_factor = compptr->h_samp_factor; /* Count number of dummy blocks to be added at the right margin. */ - ndummy = (int) INT_UMOD(blocks_across, h_samp_factor); + ndummy = (int) (blocks_across % h_samp_factor); if (ndummy > 0) ndummy = h_samp_factor - ndummy; /* Perform DCT for all non-dummy blocks in this iMCU row. Each call @@ -300,7 +300,7 @@ compress_first_pass (j_compress_ptr cinfo, JSAMPIMAGE input_buf) */ if (coef->iMCU_row_num == last_iMCU_row) { blocks_across += ndummy; /* include lower right corner */ - MCUs_across = INT_DIV(blocks_across, h_samp_factor); + MCUs_across = blocks_across / h_samp_factor; for (block_row = block_rows; block_row < compptr->v_samp_factor; block_row++) { thisblockrow = buffer[block_row]; diff --git a/test/monniaux/jpeg-6b/jcdctmgr.c b/test/monniaux/jpeg-6b/jcdctmgr.c index b85b0832..61fa79b9 100644 --- a/test/monniaux/jpeg-6b/jcdctmgr.c +++ b/test/monniaux/jpeg-6b/jcdctmgr.c @@ -244,9 +244,9 @@ forward_DCT (j_compress_ptr cinfo, jpeg_component_info * compptr, * If your machine's division is fast enough, define FAST_DIVIDE. */ #ifdef FAST_DIVIDE -#define DIVIDE_BY(a,b) a = INT_DIV(a, b) +#define DIVIDE_BY(a,b) a /= b #else -#define DIVIDE_BY(a,b) if (a >= b) a = INT_DIV(a, b); else a = 0 +#define DIVIDE_BY(a,b) if (a >= b) a /= b; else a = 0 #endif if (temp < 0) { temp = -temp; diff --git a/test/monniaux/jpeg-6b/jcmaster.c b/test/monniaux/jpeg-6b/jcmaster.c index 0f206557..aab4020b 100644 --- a/test/monniaux/jpeg-6b/jcmaster.c +++ b/test/monniaux/jpeg-6b/jcmaster.c @@ -330,7 +330,7 @@ per_scan_setup (j_compress_ptr cinfo) /* For noninterleaved scans, it is convenient to define last_row_height * as the number of block rows present in the last iMCU row. */ - tmp = (int) INT_UMOD(compptr->height_in_blocks, compptr->v_samp_factor); + tmp = (int) (compptr->height_in_blocks % compptr->v_samp_factor); if (tmp == 0) tmp = compptr->v_samp_factor; compptr->last_row_height = tmp; @@ -363,10 +363,10 @@ per_scan_setup (j_compress_ptr cinfo) compptr->MCU_blocks = compptr->MCU_width * compptr->MCU_height; compptr->MCU_sample_width = compptr->MCU_width * DCTSIZE; /* Figure number of non-dummy blocks in last MCU column & row */ - tmp = (int) INT_UMOD(compptr->width_in_blocks, compptr->MCU_width); + tmp = (int) (compptr->width_in_blocks % compptr->MCU_width); if (tmp == 0) tmp = compptr->MCU_width; compptr->last_col_width = tmp; - tmp = (int) INT_UMOD(compptr->height_in_blocks, compptr->MCU_height); + tmp = (int) (compptr->height_in_blocks % compptr->MCU_height); if (tmp == 0) tmp = compptr->MCU_height; compptr->last_row_height = tmp; /* Prepare array describing MCU composition */ diff --git a/test/monniaux/jpeg-6b/jconfig.h b/test/monniaux/jpeg-6b/jconfig.h index c6b77e42..e45abb67 100644 --- a/test/monniaux/jpeg-6b/jconfig.h +++ b/test/monniaux/jpeg-6b/jconfig.h @@ -29,25 +29,12 @@ #define HAVE_PROTOTYPES #undef NO_SWITCH -#define HAS_FLOAT 1 -#if 0 -extern long long __compcert_i64_sdiv(long long a, long long b); -extern long long __compcert_i64_smod(long long a, long long b); - -#define INT_UMOD(a, b) __compcert_i64_smod(a, b) -#define INT_MOD(a, b) __compcert_i64_smod(a, b) -#define INT_UDIV(a, b) __compcert_i64_sdiv(a, b) -#define INT_DIV(a, b) __compcert_i64_sdiv(a, b) -#define LONG_DIV(a, b) __compcert_i64_sdiv(a, b) - -#else #define INT_UMOD(a, b) ((a) % (b)) #define INT_MOD(a, b) ((a) % (b)) #define INT_UDIV(a, b) ((a) / (b)) #define INT_DIV(a, b) ((a) / (b)) #define LONG_DIV(a, b) ((a) / (b)) -#endif /* Does your compiler support the declaration "unsigned char" ? * How about "unsigned short" ? diff --git a/test/monniaux/jpeg-6b/jcparam.c b/test/monniaux/jpeg-6b/jcparam.c index d462afed..6fc48f53 100644 --- a/test/monniaux/jpeg-6b/jcparam.c +++ b/test/monniaux/jpeg-6b/jcparam.c @@ -46,7 +46,7 @@ jpeg_add_quant_table (j_compress_ptr cinfo, int which_tbl, *qtblptr = jpeg_alloc_quant_table((j_common_ptr) cinfo); for (i = 0; i < DCTSIZE2; i++) { - temp = LONG_DIV((basic_table[i] * scale_factor + 50L), 100L); + temp = ((long) basic_table[i] * scale_factor + 50L) / 100L; /* limit the values to the valid range */ if (temp <= 0L) temp = 1L; if (temp > 32767L) temp = 32767L; /* max quantizer needed for 12 bits */ @@ -120,7 +120,7 @@ jpeg_quality_scaling (int quality) * Qualities 1..50 are converted to scaling percentage 5000/Q. */ if (quality < 50) - quality = INT_DIV(5000, quality); + quality = 5000 / quality; else quality = 200 - quality*2; diff --git a/test/monniaux/jpeg-6b/jcsample.c b/test/monniaux/jpeg-6b/jcsample.c index d2693198..212ec875 100644 --- a/test/monniaux/jpeg-6b/jcsample.c +++ b/test/monniaux/jpeg-6b/jcsample.c @@ -146,8 +146,8 @@ int_downsample (j_compress_ptr cinfo, jpeg_component_info * compptr, JSAMPROW inptr, outptr; INT32 outvalue; - h_expand = INT_DIV(cinfo->max_h_samp_factor, compptr->h_samp_factor); - v_expand = INT_DIV(cinfo->max_v_samp_factor, compptr->v_samp_factor); + h_expand = cinfo->max_h_samp_factor / compptr->h_samp_factor; + v_expand = cinfo->max_v_samp_factor / compptr->v_samp_factor; numpix = h_expand * v_expand; numpix2 = numpix/2; @@ -170,7 +170,7 @@ int_downsample (j_compress_ptr cinfo, jpeg_component_info * compptr, outvalue += (INT32) GETJSAMPLE(*inptr++); } } - *outptr++ = (JSAMPLE) INT_DIV((outvalue + numpix2), numpix); + *outptr++ = (JSAMPLE) ((outvalue + numpix2) / numpix); } inrow += v_expand; } @@ -504,8 +504,8 @@ jinit_downsampler (j_compress_ptr cinfo) } else #endif downsample->methods[ci] = h2v2_downsample; - } else if (INT_UMOD(cinfo->max_h_samp_factor, compptr->h_samp_factor) == 0 && - INT_UMOD(cinfo->max_v_samp_factor, compptr->v_samp_factor) == 0) { + } else if ((cinfo->max_h_samp_factor % compptr->h_samp_factor) == 0 && + (cinfo->max_v_samp_factor % compptr->v_samp_factor) == 0) { smoothok = FALSE; downsample->methods[ci] = int_downsample; } else diff --git a/test/monniaux/jpeg-6b/jdapimin.c b/test/monniaux/jpeg-6b/jdapimin.c index cf608443..cadb59fc 100644 --- a/test/monniaux/jpeg-6b/jdapimin.c +++ b/test/monniaux/jpeg-6b/jdapimin.c @@ -116,63 +116,6 @@ default_decompress_parms (j_decompress_ptr cinfo) /* Guess the input colorspace, and set output colorspace accordingly. */ /* (Wish JPEG committee had provided a real way to specify this...) */ /* Note application may override our guesses. */ - -#ifdef NO_SWITCH - int choice = cinfo->num_components; - if (choice == 1) { - cinfo->jpeg_color_space = JCS_GRAYSCALE; - cinfo->out_color_space = JCS_GRAYSCALE; - } else if (choice == 3) { - if (cinfo->saw_JFIF_marker) { - cinfo->jpeg_color_space = JCS_YCbCr; /* JFIF implies YCbCr */ - } else if (cinfo->saw_Adobe_marker) { - int choice2 = cinfo->Adobe_transform; - if (choice2 == 0) { - cinfo->jpeg_color_space = JCS_RGB; - } else if (choice2 == 1) { - cinfo->jpeg_color_space = JCS_YCbCr; - } else { - WARNMS1(cinfo, JWRN_ADOBE_XFORM, cinfo->Adobe_transform); - cinfo->jpeg_color_space = JCS_YCbCr; /* assume it's YCbCr */ - } - } else { - /* Saw no special markers, try to guess from the component IDs */ - int cid0 = cinfo->comp_info[0].component_id; - int cid1 = cinfo->comp_info[1].component_id; - int cid2 = cinfo->comp_info[2].component_id; - - if (cid0 == 1 && cid1 == 2 && cid2 == 3) - cinfo->jpeg_color_space = JCS_YCbCr; /* assume JFIF w/out marker */ - else if (cid0 == 82 && cid1 == 71 && cid2 == 66) - cinfo->jpeg_color_space = JCS_RGB; /* ASCII 'R', 'G', 'B' */ - else { - TRACEMS3(cinfo, 1, JTRC_UNKNOWN_IDS, cid0, cid1, cid2); - cinfo->jpeg_color_space = JCS_YCbCr; /* assume it's YCbCr */ - } - } - /* Always guess RGB is proper output colorspace. */ - cinfo->out_color_space = JCS_RGB; - } else if (choice == 4) { - if (cinfo->saw_Adobe_marker) { - int choice2 = cinfo->Adobe_transform; - if (choice2 == 0) { - cinfo->jpeg_color_space = JCS_CMYK; - } else if (choice2 == 2) { - cinfo->jpeg_color_space = JCS_YCCK; - } else { - WARNMS1(cinfo, JWRN_ADOBE_XFORM, cinfo->Adobe_transform); - cinfo->jpeg_color_space = JCS_YCCK; /* assume it's YCCK */ - } - } else { - /* No special markers, assume straight CMYK. */ - cinfo->jpeg_color_space = JCS_CMYK; - } - cinfo->out_color_space = JCS_CMYK; - } else { - cinfo->jpeg_color_space = JCS_UNKNOWN; - cinfo->out_color_space = JCS_UNKNOWN; - } -#else switch (cinfo->num_components) { case 1: cinfo->jpeg_color_space = JCS_GRAYSCALE; @@ -240,14 +183,11 @@ default_decompress_parms (j_decompress_ptr cinfo) cinfo->out_color_space = JCS_UNKNOWN; break; } -#endif - + /* Set defaults for other decompression parameters. */ cinfo->scale_num = 1; /* 1:1 scaling */ cinfo->scale_denom = 1; -#ifdef HAS_FLOAT cinfo->output_gamma = 1.0; -#endif cinfo->buffered_image = FALSE; cinfo->raw_data_out = FALSE; cinfo->dct_method = JDCT_DEFAULT; @@ -349,45 +289,6 @@ jpeg_consume_input (j_decompress_ptr cinfo) int retcode = JPEG_SUSPENDED; /* NB: every possible DSTATE value should be listed in this switch */ -#ifdef NO_SWITCH - int choice = cinfo->global_state; - if (choice == DSTATE_START) { - /* Start-of-datastream actions: reset appropriate modules */ - (*cinfo->inputctl->reset_input_controller) (cinfo); - /* Initialize application's data source module */ - (*cinfo->src->init_source) (cinfo); - cinfo->global_state = DSTATE_INHEADER; - /*FALLTHROUGH*/ - retcode = (*cinfo->inputctl->consume_input) (cinfo); - if (retcode == JPEG_REACHED_SOS) { /* Found SOS, prepare to decompress */ - /* Set up default parameters based on header data */ - default_decompress_parms(cinfo); - /* Set global state: ready for start_decompress */ - cinfo->global_state = DSTATE_READY; - } - } else if (choice == DSTATE_INHEADER) { - retcode = (*cinfo->inputctl->consume_input) (cinfo); - if (retcode == JPEG_REACHED_SOS) { /* Found SOS, prepare to decompress */ - /* Set up default parameters based on header data */ - default_decompress_parms(cinfo); - /* Set global state: ready for start_decompress */ - cinfo->global_state = DSTATE_READY; - } - } else if (choice == DSTATE_READY) { - /* Can't advance past first SOS until start_decompress is called */ - retcode = JPEG_REACHED_SOS; - } else if (choice == DSTATE_PRELOAD || - choice == DSTATE_PRESCAN || - choice == DSTATE_SCANNING || - choice == DSTATE_RAW_OK || - choice == DSTATE_BUFIMAGE || - choice == DSTATE_BUFPOST || - choice == DSTATE_STOPPING) { - retcode = (*cinfo->inputctl->consume_input) (cinfo); - } else { - ERREXIT1(cinfo, JERR_BAD_STATE, cinfo->global_state); - } -#else switch (cinfo->global_state) { case DSTATE_START: /* Start-of-datastream actions: reset appropriate modules */ @@ -421,7 +322,6 @@ jpeg_consume_input (j_decompress_ptr cinfo) default: ERREXIT1(cinfo, JERR_BAD_STATE, cinfo->global_state); } -#endif return retcode; } diff --git a/test/monniaux/jpeg-6b/jdcoefct.c b/test/monniaux/jpeg-6b/jdcoefct.c index 54888720..4938d20f 100644 --- a/test/monniaux/jpeg-6b/jdcoefct.c +++ b/test/monniaux/jpeg-6b/jdcoefct.c @@ -349,7 +349,7 @@ decompress_data (j_decompress_ptr cinfo, JSAMPIMAGE output_buf) block_rows = compptr->v_samp_factor; else { /* NB: can't use last_row_height here; it is input-side-dependent! */ - block_rows = (int) INT_MOD(compptr->height_in_blocks, compptr->v_samp_factor); + block_rows = (int) (compptr->height_in_blocks % compptr->v_samp_factor); if (block_rows == 0) block_rows = compptr->v_samp_factor; } inverse_DCT = cinfo->idct->inverse_DCT[ci]; @@ -508,7 +508,7 @@ decompress_smooth_data (j_decompress_ptr cinfo, JSAMPIMAGE output_buf) last_row = FALSE; } else { /* NB: can't use last_row_height here; it is input-side-dependent! */ - block_rows = (int) INT_MOD(compptr->height_in_blocks, compptr->v_samp_factor); + block_rows = (int) (compptr->height_in_blocks % compptr->v_samp_factor); if (block_rows == 0) block_rows = compptr->v_samp_factor; access_rows = block_rows; /* this iMCU row only */ last_row = TRUE; @@ -575,11 +575,11 @@ decompress_smooth_data (j_decompress_ptr cinfo, JSAMPIMAGE output_buf) if ((Al=coef_bits[1]) != 0 && workspace[1] == 0) { num = 36 * Q00 * (DC4 - DC6); if (num >= 0) { - pred = (int) INT_DIV(((Q01<<7) + num), (Q01<<8)); + pred = (int) (((Q01<<7) + num) / (Q01<<8)); if (Al > 0 && pred >= (1< 0 && pred >= (1<= 0) { - pred = (int) INT_DIV(((Q10<<7) + num), (Q10<<8)); + pred = (int) (((Q10<<7) + num) / (Q10<<8)); if (Al > 0 && pred >= (1< 0 && pred >= (1<= 0) { - pred = (int) INT_DIV(((Q20<<7) + num), (Q20<<8)); + pred = (int) (((Q20<<7) + num) / (Q20<<8)); if (Al > 0 && pred >= (1< 0 && pred >= (1<= 0) { - pred = (int) INT_DIV(((Q11<<7) + num), (Q11<<8)); + pred = (int) (((Q11<<7) + num) / (Q11<<8)); if (Al > 0 && pred >= (1< 0 && pred >= (1<= 0) { - pred = (int) INT_DIV(((Q02<<7) + num), (Q02<<8)); + pred = (int) (((Q02<<7) + num) / (Q02<<8)); if (Al > 0 && pred >= (1< 0 && pred >= (1<height_in_blocks, compptr->v_samp_factor); + tmp = (int) (compptr->height_in_blocks % compptr->v_samp_factor); if (tmp == 0) tmp = compptr->v_samp_factor; compptr->last_row_height = tmp; @@ -176,10 +176,10 @@ per_scan_setup (j_decompress_ptr cinfo) compptr->MCU_blocks = compptr->MCU_width * compptr->MCU_height; compptr->MCU_sample_width = compptr->MCU_width * compptr->DCT_scaled_size; /* Figure number of non-dummy blocks in last MCU column & row */ - tmp = (int) INT_MOD(compptr->width_in_blocks, compptr->MCU_width); + tmp = (int) (compptr->width_in_blocks % compptr->MCU_width); if (tmp == 0) tmp = compptr->MCU_width; compptr->last_col_width = tmp; - tmp = (int) INT_MOD(compptr->height_in_blocks, compptr->MCU_height); + tmp = (int) (compptr->height_in_blocks % compptr->MCU_height); if (tmp == 0) tmp = compptr->MCU_height; compptr->last_row_height = tmp; /* Prepare array describing MCU composition */ diff --git a/test/monniaux/jpeg-6b/jdmainct.c b/test/monniaux/jpeg-6b/jdmainct.c index 0de6a394..13c956f5 100644 --- a/test/monniaux/jpeg-6b/jdmainct.c +++ b/test/monniaux/jpeg-6b/jdmainct.c @@ -175,9 +175,8 @@ alloc_funny_pointers (j_decompress_ptr cinfo) for (ci = 0, compptr = cinfo->comp_info; ci < cinfo->num_components; ci++, compptr++) { - rgroup = INT_DIV((compptr->v_samp_factor * compptr->DCT_scaled_size), - cinfo->min_DCT_scaled_size); - /* height of a row group of component */ + rgroup = (compptr->v_samp_factor * compptr->DCT_scaled_size) / + cinfo->min_DCT_scaled_size; /* height of a row group of component */ /* Get space for pointer lists --- M+4 row groups in each list. * We alloc both pointer lists with one call to save a few cycles. */ @@ -209,9 +208,8 @@ make_funny_pointers (j_decompress_ptr cinfo) for (ci = 0, compptr = cinfo->comp_info; ci < cinfo->num_components; ci++, compptr++) { - rgroup = INT_DIV((compptr->v_samp_factor * compptr->DCT_scaled_size), - cinfo->min_DCT_scaled_size); - /* height of a row group of component */ + rgroup = (compptr->v_samp_factor * compptr->DCT_scaled_size) / + cinfo->min_DCT_scaled_size; /* height of a row group of component */ xbuf0 = main->xbuffer[0][ci]; xbuf1 = main->xbuffer[1][ci]; /* First copy the workspace pointers as-is */ @@ -250,9 +248,8 @@ set_wraparound_pointers (j_decompress_ptr cinfo) for (ci = 0, compptr = cinfo->comp_info; ci < cinfo->num_components; ci++, compptr++) { - rgroup = INT_DIV((compptr->v_samp_factor * compptr->DCT_scaled_size), - cinfo->min_DCT_scaled_size); - /* height of a row group of component */ + rgroup = (compptr->v_samp_factor * compptr->DCT_scaled_size) / + cinfo->min_DCT_scaled_size; /* height of a row group of component */ xbuf0 = main->xbuffer[0][ci]; xbuf1 = main->xbuffer[1][ci]; for (i = 0; i < rgroup; i++) { @@ -281,15 +278,15 @@ set_bottom_pointers (j_decompress_ptr cinfo) ci++, compptr++) { /* Count sample rows in one iMCU row and in one row group */ iMCUheight = compptr->v_samp_factor * compptr->DCT_scaled_size; - rgroup = INT_DIV(iMCUheight, cinfo->min_DCT_scaled_size); + rgroup = iMCUheight / cinfo->min_DCT_scaled_size; /* Count nondummy sample rows remaining for this component */ - rows_left = (int) INT_MOD(compptr->downsampled_height, (JDIMENSION) iMCUheight); + rows_left = (int) (compptr->downsampled_height % (JDIMENSION) iMCUheight); if (rows_left == 0) rows_left = iMCUheight; /* Count nondummy row groups. Should get same answer for each component, * so we need only do it once. */ if (ci == 0) { - main->rowgroups_avail = (JDIMENSION) (INT_DIV((rows_left-1), rgroup) + 1); + main->rowgroups_avail = (JDIMENSION) ((rows_left-1) / rgroup + 1); } /* Duplicate the last real sample row rgroup*2 times; this pads out the * last partial rowgroup and ensures at least one full rowgroup of context. @@ -505,9 +502,8 @@ jinit_d_main_controller (j_decompress_ptr cinfo, boolean need_full_buffer) for (ci = 0, compptr = cinfo->comp_info; ci < cinfo->num_components; ci++, compptr++) { - rgroup = INT_DIV((compptr->v_samp_factor * compptr->DCT_scaled_size), - cinfo->min_DCT_scaled_size); - /* height of a row group of component */ + rgroup = (compptr->v_samp_factor * compptr->DCT_scaled_size) / + cinfo->min_DCT_scaled_size; /* height of a row group of component */ main->buffer[ci] = (*cinfo->mem->alloc_sarray) ((j_common_ptr) cinfo, JPOOL_IMAGE, compptr->width_in_blocks * compptr->DCT_scaled_size, diff --git a/test/monniaux/jpeg-6b/jdsample.c b/test/monniaux/jpeg-6b/jdsample.c index 96d56f0c..80ffefb2 100644 --- a/test/monniaux/jpeg-6b/jdsample.c +++ b/test/monniaux/jpeg-6b/jdsample.c @@ -428,10 +428,10 @@ jinit_upsampler (j_decompress_ptr cinfo) /* Compute size of an "input group" after IDCT scaling. This many samples * are to be converted to max_h_samp_factor * max_v_samp_factor pixels. */ - h_in_group = INT_DIV((compptr->h_samp_factor * compptr->DCT_scaled_size), - cinfo->min_DCT_scaled_size); - v_in_group = INT_DIV((compptr->v_samp_factor * compptr->DCT_scaled_size), - cinfo->min_DCT_scaled_size); + h_in_group = (compptr->h_samp_factor * compptr->DCT_scaled_size) / + cinfo->min_DCT_scaled_size; + v_in_group = (compptr->v_samp_factor * compptr->DCT_scaled_size) / + cinfo->min_DCT_scaled_size; h_out_group = cinfo->max_h_samp_factor; v_out_group = cinfo->max_v_samp_factor; upsample->rowgroup_height[ci] = v_in_group; /* save for use later */ @@ -459,12 +459,12 @@ jinit_upsampler (j_decompress_ptr cinfo) upsample->pub.need_context_rows = TRUE; } else upsample->methods[ci] = h2v2_upsample; - } else if (INT_MOD(h_out_group, h_in_group) == 0 && - INT_MOD(v_out_group, v_in_group) == 0) { + } else if ((h_out_group % h_in_group) == 0 && + (v_out_group % v_in_group) == 0) { /* Generic integral-factors upsampling method */ upsample->methods[ci] = int_upsample; - upsample->h_expand[ci] = (UINT8) INT_DIV(h_out_group, h_in_group); - upsample->v_expand[ci] = (UINT8) INT_DIV(v_out_group, v_in_group); + upsample->h_expand[ci] = (UINT8) (h_out_group / h_in_group); + upsample->v_expand[ci] = (UINT8) (v_out_group / v_in_group); } else ERREXIT(cinfo, JERR_FRACT_SAMPLE_NOTIMPL); if (need_buffer) { diff --git a/test/monniaux/jpeg-6b/jerror.c b/test/monniaux/jpeg-6b/jerror.c index c98aed76..3da7be86 100644 --- a/test/monniaux/jpeg-6b/jerror.c +++ b/test/monniaux/jpeg-6b/jerror.c @@ -1,252 +1,252 @@ -/* - * jerror.c - * - * Copyright (C) 1991-1998, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains simple error-reporting and trace-message routines. - * These are suitable for Unix-like systems and others where writing to - * stderr is the right thing to do. Many applications will want to replace - * some or all of these routines. - * - * If you define USE_WINDOWS_MESSAGEBOX in jconfig.h or in the makefile, - * you get a Windows-specific hack to display error messages in a dialog box. - * It ain't much, but it beats dropping error messages into the bit bucket, - * which is what happens to output to stderr under most Windows C compilers. - * - * These routines are used by both the compression and decompression code. - */ - -/* this is not a core library module, so it doesn't define JPEG_INTERNALS */ -#include "jinclude.h" -#include "jpeglib.h" -#include "jversion.h" -#include "jerror.h" - -#ifdef USE_WINDOWS_MESSAGEBOX -#include -#endif - -#ifndef EXIT_FAILURE /* define exit() codes if not provided */ -#define EXIT_FAILURE 1 -#endif - - -/* - * Create the message string table. - * We do this from the master message list in jerror.h by re-reading - * jerror.h with a suitable definition for macro JMESSAGE. - * The message table is made an external symbol just in case any applications - * want to refer to it directly. - */ - -#ifdef NEED_SHORT_EXTERNAL_NAMES -#define jpeg_std_message_table jMsgTable -#endif - -#define JMESSAGE(code,string) string , - -const char * const jpeg_std_message_table[] = { -#include "jerror.h" - NULL -}; - - -/* - * Error exit handler: must not return to caller. - * - * Applications may override this if they want to get control back after - * an error. Typically one would longjmp somewhere instead of exiting. - * The setjmp buffer can be made a private field within an expanded error - * handler object. Note that the info needed to generate an error message - * is stored in the error object, so you can generate the message now or - * later, at your convenience. - * You should make sure that the JPEG object is cleaned up (with jpeg_abort - * or jpeg_destroy) at some point. - */ - -METHODDEF(void) -error_exit (j_common_ptr cinfo) -{ - /* Always display the message */ - (*cinfo->err->output_message) (cinfo); - - /* Let the memory manager delete any temp files before we die */ - jpeg_destroy(cinfo); - - exit(EXIT_FAILURE); -} - - -/* - * Actual output of an error or trace message. - * Applications may override this method to send JPEG messages somewhere - * other than stderr. - * - * On Windows, printing to stderr is generally completely useless, - * so we provide optional code to produce an error-dialog popup. - * Most Windows applications will still prefer to override this routine, - * but if they don't, it'll do something at least marginally useful. - * - * NOTE: to use the library in an environment that doesn't support the - * C stdio library, you may have to delete the call to fprintf() entirely, - * not just not use this routine. - */ - -METHODDEF(void) -output_message (j_common_ptr cinfo) -{ - char buffer[JMSG_LENGTH_MAX]; - - /* Create the message */ - (*cinfo->err->format_message) (cinfo, buffer); - -#ifdef USE_WINDOWS_MESSAGEBOX - /* Display it in a message dialog box */ - MessageBox(GetActiveWindow(), buffer, "JPEG Library Error", - MB_OK | MB_ICONERROR); -#else - /* Send it to stderr, adding a newline */ - fprintf(stderr, "%s\n", buffer); -#endif -} - - -/* - * Decide whether to emit a trace or warning message. - * msg_level is one of: - * -1: recoverable corrupt-data warning, may want to abort. - * 0: important advisory messages (always display to user). - * 1: first level of tracing detail. - * 2,3,...: successively more detailed tracing messages. - * An application might override this method if it wanted to abort on warnings - * or change the policy about which messages to display. - */ - -METHODDEF(void) -emit_message (j_common_ptr cinfo, int msg_level) -{ - struct jpeg_error_mgr * err = cinfo->err; - - if (msg_level < 0) { - /* It's a warning message. Since corrupt files may generate many warnings, - * the policy implemented here is to show only the first warning, - * unless trace_level >= 3. - */ - if (err->num_warnings == 0 || err->trace_level >= 3) - (*err->output_message) (cinfo); - /* Always count warnings in num_warnings. */ - err->num_warnings++; - } else { - /* It's a trace message. Show it if trace_level >= msg_level. */ - if (err->trace_level >= msg_level) - (*err->output_message) (cinfo); - } -} - - -/* - * Format a message string for the most recent JPEG error or message. - * The message is stored into buffer, which should be at least JMSG_LENGTH_MAX - * characters. Note that no '\n' character is added to the string. - * Few applications should need to override this method. - */ - -METHODDEF(void) -format_message (j_common_ptr cinfo, char * buffer) -{ - struct jpeg_error_mgr * err = cinfo->err; - int msg_code = err->msg_code; - const char * msgtext = NULL; - const char * msgptr; - char ch; - boolean isstring; - - /* Look up message string in proper table */ - if (msg_code > 0 && msg_code <= err->last_jpeg_message) { - msgtext = err->jpeg_message_table[msg_code]; - } else if (err->addon_message_table != NULL && - msg_code >= err->first_addon_message && - msg_code <= err->last_addon_message) { - msgtext = err->addon_message_table[msg_code - err->first_addon_message]; - } - - /* Defend against bogus message number */ - if (msgtext == NULL) { - err->msg_parm.i[0] = msg_code; - msgtext = err->jpeg_message_table[0]; - } - - /* Check for string parameter, as indicated by %s in the message text */ - isstring = FALSE; - msgptr = msgtext; - while ((ch = *msgptr++) != '\0') { - if (ch == '%') { - if (*msgptr == 's') isstring = TRUE; - break; - } - } - - /* Format the message into the passed buffer */ - if (isstring) - sprintf(buffer, msgtext, err->msg_parm.s); - else - sprintf(buffer, msgtext, - err->msg_parm.i[0], err->msg_parm.i[1], - err->msg_parm.i[2], err->msg_parm.i[3], - err->msg_parm.i[4], err->msg_parm.i[5], - err->msg_parm.i[6], err->msg_parm.i[7]); -} - - -/* - * Reset error state variables at start of a new image. - * This is called during compression startup to reset trace/error - * processing to default state, without losing any application-specific - * method pointers. An application might possibly want to override - * this method if it has additional error processing state. - */ - -METHODDEF(void) -reset_error_mgr (j_common_ptr cinfo) -{ - cinfo->err->num_warnings = 0; - /* trace_level is not reset since it is an application-supplied parameter */ - cinfo->err->msg_code = 0; /* may be useful as a flag for "no error" */ -} - - -/* - * Fill in the standard error-handling methods in a jpeg_error_mgr object. - * Typical call is: - * struct jpeg_compress_struct cinfo; - * struct jpeg_error_mgr err; - * - * cinfo.err = jpeg_std_error(&err); - * after which the application may override some of the methods. - */ - -GLOBAL(struct jpeg_error_mgr *) -jpeg_std_error (struct jpeg_error_mgr * err) -{ - err->error_exit = error_exit; - err->emit_message = emit_message; - err->output_message = output_message; - err->format_message = format_message; - err->reset_error_mgr = reset_error_mgr; - - err->trace_level = 0; /* default = no tracing */ - err->num_warnings = 0; /* no warnings emitted yet */ - err->msg_code = 0; /* may be useful as a flag for "no error" */ - - /* Initialize message table pointers */ - err->jpeg_message_table = jpeg_std_message_table; - err->last_jpeg_message = (int) JMSG_LASTMSGCODE - 1; - - err->addon_message_table = NULL; - err->first_addon_message = 0; /* for safety */ - err->last_addon_message = 0; - - return err; -} +/* + * jerror.c + * + * Copyright (C) 1991-1998, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains simple error-reporting and trace-message routines. + * These are suitable for Unix-like systems and others where writing to + * stderr is the right thing to do. Many applications will want to replace + * some or all of these routines. + * + * If you define USE_WINDOWS_MESSAGEBOX in jconfig.h or in the makefile, + * you get a Windows-specific hack to display error messages in a dialog box. + * It ain't much, but it beats dropping error messages into the bit bucket, + * which is what happens to output to stderr under most Windows C compilers. + * + * These routines are used by both the compression and decompression code. + */ + +/* this is not a core library module, so it doesn't define JPEG_INTERNALS */ +#include "jinclude.h" +#include "jpeglib.h" +#include "jversion.h" +#include "jerror.h" + +#ifdef USE_WINDOWS_MESSAGEBOX +#include +#endif + +#ifndef EXIT_FAILURE /* define exit() codes if not provided */ +#define EXIT_FAILURE 1 +#endif + + +/* + * Create the message string table. + * We do this from the master message list in jerror.h by re-reading + * jerror.h with a suitable definition for macro JMESSAGE. + * The message table is made an external symbol just in case any applications + * want to refer to it directly. + */ + +#ifdef NEED_SHORT_EXTERNAL_NAMES +#define jpeg_std_message_table jMsgTable +#endif + +#define JMESSAGE(code,string) string , + +const char * const jpeg_std_message_table[] = { +#include "jerror.h" + NULL +}; + + +/* + * Error exit handler: must not return to caller. + * + * Applications may override this if they want to get control back after + * an error. Typically one would longjmp somewhere instead of exiting. + * The setjmp buffer can be made a private field within an expanded error + * handler object. Note that the info needed to generate an error message + * is stored in the error object, so you can generate the message now or + * later, at your convenience. + * You should make sure that the JPEG object is cleaned up (with jpeg_abort + * or jpeg_destroy) at some point. + */ + +METHODDEF(void) +error_exit (j_common_ptr cinfo) +{ + /* Always display the message */ + (*cinfo->err->output_message) (cinfo); + + /* Let the memory manager delete any temp files before we die */ + jpeg_destroy(cinfo); + + exit(EXIT_FAILURE); +} + + +/* + * Actual output of an error or trace message. + * Applications may override this method to send JPEG messages somewhere + * other than stderr. + * + * On Windows, printing to stderr is generally completely useless, + * so we provide optional code to produce an error-dialog popup. + * Most Windows applications will still prefer to override this routine, + * but if they don't, it'll do something at least marginally useful. + * + * NOTE: to use the library in an environment that doesn't support the + * C stdio library, you may have to delete the call to fprintf() entirely, + * not just not use this routine. + */ + +METHODDEF(void) +output_message (j_common_ptr cinfo) +{ + char buffer[JMSG_LENGTH_MAX]; + + /* Create the message */ + (*cinfo->err->format_message) (cinfo, buffer); + +#ifdef USE_WINDOWS_MESSAGEBOX + /* Display it in a message dialog box */ + MessageBox(GetActiveWindow(), buffer, "JPEG Library Error", + MB_OK | MB_ICONERROR); +#else + /* Send it to stderr, adding a newline */ + fprintf(stderr, "%s\n", buffer); +#endif +} + + +/* + * Decide whether to emit a trace or warning message. + * msg_level is one of: + * -1: recoverable corrupt-data warning, may want to abort. + * 0: important advisory messages (always display to user). + * 1: first level of tracing detail. + * 2,3,...: successively more detailed tracing messages. + * An application might override this method if it wanted to abort on warnings + * or change the policy about which messages to display. + */ + +METHODDEF(void) +emit_message (j_common_ptr cinfo, int msg_level) +{ + struct jpeg_error_mgr * err = cinfo->err; + + if (msg_level < 0) { + /* It's a warning message. Since corrupt files may generate many warnings, + * the policy implemented here is to show only the first warning, + * unless trace_level >= 3. + */ + if (err->num_warnings == 0 || err->trace_level >= 3) + (*err->output_message) (cinfo); + /* Always count warnings in num_warnings. */ + err->num_warnings++; + } else { + /* It's a trace message. Show it if trace_level >= msg_level. */ + if (err->trace_level >= msg_level) + (*err->output_message) (cinfo); + } +} + + +/* + * Format a message string for the most recent JPEG error or message. + * The message is stored into buffer, which should be at least JMSG_LENGTH_MAX + * characters. Note that no '\n' character is added to the string. + * Few applications should need to override this method. + */ + +METHODDEF(void) +format_message (j_common_ptr cinfo, char * buffer) +{ + struct jpeg_error_mgr * err = cinfo->err; + int msg_code = err->msg_code; + const char * msgtext = NULL; + const char * msgptr; + char ch; + boolean isstring; + + /* Look up message string in proper table */ + if (msg_code > 0 && msg_code <= err->last_jpeg_message) { + msgtext = err->jpeg_message_table[msg_code]; + } else if (err->addon_message_table != NULL && + msg_code >= err->first_addon_message && + msg_code <= err->last_addon_message) { + msgtext = err->addon_message_table[msg_code - err->first_addon_message]; + } + + /* Defend against bogus message number */ + if (msgtext == NULL) { + err->msg_parm.i[0] = msg_code; + msgtext = err->jpeg_message_table[0]; + } + + /* Check for string parameter, as indicated by %s in the message text */ + isstring = FALSE; + msgptr = msgtext; + while ((ch = *msgptr++) != '\0') { + if (ch == '%') { + if (*msgptr == 's') isstring = TRUE; + break; + } + } + + /* Format the message into the passed buffer */ + if (isstring) + sprintf(buffer, msgtext, err->msg_parm.s); + else + sprintf(buffer, msgtext, + err->msg_parm.i[0], err->msg_parm.i[1], + err->msg_parm.i[2], err->msg_parm.i[3], + err->msg_parm.i[4], err->msg_parm.i[5], + err->msg_parm.i[6], err->msg_parm.i[7]); +} + + +/* + * Reset error state variables at start of a new image. + * This is called during compression startup to reset trace/error + * processing to default state, without losing any application-specific + * method pointers. An application might possibly want to override + * this method if it has additional error processing state. + */ + +METHODDEF(void) +reset_error_mgr (j_common_ptr cinfo) +{ + cinfo->err->num_warnings = 0; + /* trace_level is not reset since it is an application-supplied parameter */ + cinfo->err->msg_code = 0; /* may be useful as a flag for "no error" */ +} + + +/* + * Fill in the standard error-handling methods in a jpeg_error_mgr object. + * Typical call is: + * struct jpeg_compress_struct cinfo; + * struct jpeg_error_mgr err; + * + * cinfo.err = jpeg_std_error(&err); + * after which the application may override some of the methods. + */ + +GLOBAL(struct jpeg_error_mgr *) +jpeg_std_error (struct jpeg_error_mgr * err) +{ + err->error_exit = error_exit; + err->emit_message = emit_message; + err->output_message = output_message; + err->format_message = format_message; + err->reset_error_mgr = reset_error_mgr; + + err->trace_level = 0; /* default = no tracing */ + err->num_warnings = 0; /* no warnings emitted yet */ + err->msg_code = 0; /* may be useful as a flag for "no error" */ + + /* Initialize message table pointers */ + err->jpeg_message_table = jpeg_std_message_table; + err->last_jpeg_message = (int) JMSG_LASTMSGCODE - 1; + + err->addon_message_table = NULL; + err->first_addon_message = 0; /* for safety */ + err->last_addon_message = 0; + + return err; +} diff --git a/test/monniaux/jpeg-6b/jmemmgr.c b/test/monniaux/jpeg-6b/jmemmgr.c index b636f1be..d801b322 100644 --- a/test/monniaux/jpeg-6b/jmemmgr.c +++ b/test/monniaux/jpeg-6b/jmemmgr.c @@ -1,1118 +1,1118 @@ -/* - * jmemmgr.c - * - * Copyright (C) 1991-1997, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains the JPEG system-independent memory management - * routines. This code is usable across a wide variety of machines; most - * of the system dependencies have been isolated in a separate file. - * The major functions provided here are: - * * pool-based allocation and freeing of memory; - * * policy decisions about how to divide available memory among the - * virtual arrays; - * * control logic for swapping virtual arrays between main memory and - * backing storage. - * The separate system-dependent file provides the actual backing-storage - * access code, and it contains the policy decision about how much total - * main memory to use. - * This file is system-dependent in the sense that some of its functions - * are unnecessary in some systems. For example, if there is enough virtual - * memory so that backing storage will never be used, much of the virtual - * array control logic could be removed. (Of course, if you have that much - * memory then you shouldn't care about a little bit of unused code...) - */ - -#define JPEG_INTERNALS -#define AM_MEMORY_MANAGER /* we define jvirt_Xarray_control structs */ -#include "jinclude.h" -#include "jpeglib.h" -#include "jmemsys.h" /* import the system-dependent declarations */ - -#ifndef NO_GETENV -#ifndef HAVE_STDLIB_H /* should declare getenv() */ -extern char * getenv JPP((const char * name)); -#endif -#endif - - -/* - * Some important notes: - * The allocation routines provided here must never return NULL. - * They should exit to error_exit if unsuccessful. - * - * It's not a good idea to try to merge the sarray and barray routines, - * even though they are textually almost the same, because samples are - * usually stored as bytes while coefficients are shorts or ints. Thus, - * in machines where byte pointers have a different representation from - * word pointers, the resulting machine code could not be the same. - */ - - -/* - * Many machines require storage alignment: longs must start on 4-byte - * boundaries, doubles on 8-byte boundaries, etc. On such machines, malloc() - * always returns pointers that are multiples of the worst-case alignment - * requirement, and we had better do so too. - * There isn't any really portable way to determine the worst-case alignment - * requirement. This module assumes that the alignment requirement is - * multiples of sizeof(ALIGN_TYPE). - * By default, we define ALIGN_TYPE as double. This is necessary on some - * workstations (where doubles really do need 8-byte alignment) and will work - * fine on nearly everything. If your machine has lesser alignment needs, - * you can save a few bytes by making ALIGN_TYPE smaller. - * The only place I know of where this will NOT work is certain Macintosh - * 680x0 compilers that define double as a 10-byte IEEE extended float. - * Doing 10-byte alignment is counterproductive because longwords won't be - * aligned well. Put "#define ALIGN_TYPE long" in jconfig.h if you have - * such a compiler. - */ - -#ifndef ALIGN_TYPE /* so can override from jconfig.h */ -#define ALIGN_TYPE double -#endif - - -/* - * We allocate objects from "pools", where each pool is gotten with a single - * request to jpeg_get_small() or jpeg_get_large(). There is no per-object - * overhead within a pool, except for alignment padding. Each pool has a - * header with a link to the next pool of the same class. - * Small and large pool headers are identical except that the latter's - * link pointer must be FAR on 80x86 machines. - * Notice that the "real" header fields are union'ed with a dummy ALIGN_TYPE - * field. This forces the compiler to make SIZEOF(small_pool_hdr) a multiple - * of the alignment requirement of ALIGN_TYPE. - */ - -typedef union small_pool_struct * small_pool_ptr; - -typedef union small_pool_struct { - struct { - small_pool_ptr next; /* next in list of pools */ - size_t bytes_used; /* how many bytes already used within pool */ - size_t bytes_left; /* bytes still available in this pool */ - } hdr; - ALIGN_TYPE dummy; /* included in union to ensure alignment */ -} small_pool_hdr; - -typedef union large_pool_struct FAR * large_pool_ptr; - -typedef union large_pool_struct { - struct { - large_pool_ptr next; /* next in list of pools */ - size_t bytes_used; /* how many bytes already used within pool */ - size_t bytes_left; /* bytes still available in this pool */ - } hdr; - ALIGN_TYPE dummy; /* included in union to ensure alignment */ -} large_pool_hdr; - - -/* - * Here is the full definition of a memory manager object. - */ - -typedef struct { - struct jpeg_memory_mgr pub; /* public fields */ - - /* Each pool identifier (lifetime class) names a linked list of pools. */ - small_pool_ptr small_list[JPOOL_NUMPOOLS]; - large_pool_ptr large_list[JPOOL_NUMPOOLS]; - - /* Since we only have one lifetime class of virtual arrays, only one - * linked list is necessary (for each datatype). Note that the virtual - * array control blocks being linked together are actually stored somewhere - * in the small-pool list. - */ - jvirt_sarray_ptr virt_sarray_list; - jvirt_barray_ptr virt_barray_list; - - /* This counts total space obtained from jpeg_get_small/large */ - long total_space_allocated; - - /* alloc_sarray and alloc_barray set this value for use by virtual - * array routines. - */ - JDIMENSION last_rowsperchunk; /* from most recent alloc_sarray/barray */ -} my_memory_mgr; - -typedef my_memory_mgr * my_mem_ptr; - - -/* - * The control blocks for virtual arrays. - * Note that these blocks are allocated in the "small" pool area. - * System-dependent info for the associated backing store (if any) is hidden - * inside the backing_store_info struct. - */ - -struct jvirt_sarray_control { - JSAMPARRAY mem_buffer; /* => the in-memory buffer */ - JDIMENSION rows_in_array; /* total virtual array height */ - JDIMENSION samplesperrow; /* width of array (and of memory buffer) */ - JDIMENSION maxaccess; /* max rows accessed by access_virt_sarray */ - JDIMENSION rows_in_mem; /* height of memory buffer */ - JDIMENSION rowsperchunk; /* allocation chunk size in mem_buffer */ - JDIMENSION cur_start_row; /* first logical row # in the buffer */ - JDIMENSION first_undef_row; /* row # of first uninitialized row */ - boolean pre_zero; /* pre-zero mode requested? */ - boolean dirty; /* do current buffer contents need written? */ - boolean b_s_open; /* is backing-store data valid? */ - jvirt_sarray_ptr next; /* link to next virtual sarray control block */ - backing_store_info b_s_info; /* System-dependent control info */ -}; - -struct jvirt_barray_control { - JBLOCKARRAY mem_buffer; /* => the in-memory buffer */ - JDIMENSION rows_in_array; /* total virtual array height */ - JDIMENSION blocksperrow; /* width of array (and of memory buffer) */ - JDIMENSION maxaccess; /* max rows accessed by access_virt_barray */ - JDIMENSION rows_in_mem; /* height of memory buffer */ - JDIMENSION rowsperchunk; /* allocation chunk size in mem_buffer */ - JDIMENSION cur_start_row; /* first logical row # in the buffer */ - JDIMENSION first_undef_row; /* row # of first uninitialized row */ - boolean pre_zero; /* pre-zero mode requested? */ - boolean dirty; /* do current buffer contents need written? */ - boolean b_s_open; /* is backing-store data valid? */ - jvirt_barray_ptr next; /* link to next virtual barray control block */ - backing_store_info b_s_info; /* System-dependent control info */ -}; - - -#ifdef MEM_STATS /* optional extra stuff for statistics */ - -LOCAL(void) -print_mem_stats (j_common_ptr cinfo, int pool_id) -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - small_pool_ptr shdr_ptr; - large_pool_ptr lhdr_ptr; - - /* Since this is only a debugging stub, we can cheat a little by using - * fprintf directly rather than going through the trace message code. - * This is helpful because message parm array can't handle longs. - */ - fprintf(stderr, "Freeing pool %d, total space = %ld\n", - pool_id, mem->total_space_allocated); - - for (lhdr_ptr = mem->large_list[pool_id]; lhdr_ptr != NULL; - lhdr_ptr = lhdr_ptr->hdr.next) { - fprintf(stderr, " Large chunk used %ld\n", - (long) lhdr_ptr->hdr.bytes_used); - } - - for (shdr_ptr = mem->small_list[pool_id]; shdr_ptr != NULL; - shdr_ptr = shdr_ptr->hdr.next) { - fprintf(stderr, " Small chunk used %ld free %ld\n", - (long) shdr_ptr->hdr.bytes_used, - (long) shdr_ptr->hdr.bytes_left); - } -} - -#endif /* MEM_STATS */ - - -LOCAL(void) -out_of_memory (j_common_ptr cinfo, int which) -/* Report an out-of-memory error and stop execution */ -/* If we compiled MEM_STATS support, report alloc requests before dying */ -{ -#ifdef MEM_STATS - cinfo->err->trace_level = 2; /* force self_destruct to report stats */ -#endif - ERREXIT1(cinfo, JERR_OUT_OF_MEMORY, which); -} - - -/* - * Allocation of "small" objects. - * - * For these, we use pooled storage. When a new pool must be created, - * we try to get enough space for the current request plus a "slop" factor, - * where the slop will be the amount of leftover space in the new pool. - * The speed vs. space tradeoff is largely determined by the slop values. - * A different slop value is provided for each pool class (lifetime), - * and we also distinguish the first pool of a class from later ones. - * NOTE: the values given work fairly well on both 16- and 32-bit-int - * machines, but may be too small if longs are 64 bits or more. - */ - -static const size_t first_pool_slop[JPOOL_NUMPOOLS] = -{ - 1600, /* first PERMANENT pool */ - 16000 /* first IMAGE pool */ -}; - -static const size_t extra_pool_slop[JPOOL_NUMPOOLS] = -{ - 0, /* additional PERMANENT pools */ - 5000 /* additional IMAGE pools */ -}; - -#define MIN_SLOP 50 /* greater than 0 to avoid futile looping */ - - -METHODDEF(void *) -alloc_small (j_common_ptr cinfo, int pool_id, size_t sizeofobject) -/* Allocate a "small" object */ -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - small_pool_ptr hdr_ptr, prev_hdr_ptr; - char * data_ptr; - size_t odd_bytes, min_request, slop; - - /* Check for unsatisfiable request (do now to ensure no overflow below) */ - if (sizeofobject > (size_t) (MAX_ALLOC_CHUNK-SIZEOF(small_pool_hdr))) - out_of_memory(cinfo, 1); /* request exceeds malloc's ability */ - - /* Round up the requested size to a multiple of SIZEOF(ALIGN_TYPE) */ - odd_bytes = sizeofobject % SIZEOF(ALIGN_TYPE); - if (odd_bytes > 0) - sizeofobject += SIZEOF(ALIGN_TYPE) - odd_bytes; - - /* See if space is available in any existing pool */ - if (pool_id < 0 || pool_id >= JPOOL_NUMPOOLS) - ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ - prev_hdr_ptr = NULL; - hdr_ptr = mem->small_list[pool_id]; - while (hdr_ptr != NULL) { - if (hdr_ptr->hdr.bytes_left >= sizeofobject) - break; /* found pool with enough space */ - prev_hdr_ptr = hdr_ptr; - hdr_ptr = hdr_ptr->hdr.next; - } - - /* Time to make a new pool? */ - if (hdr_ptr == NULL) { - /* min_request is what we need now, slop is what will be leftover */ - min_request = sizeofobject + SIZEOF(small_pool_hdr); - if (prev_hdr_ptr == NULL) /* first pool in class? */ - slop = first_pool_slop[pool_id]; - else - slop = extra_pool_slop[pool_id]; - /* Don't ask for more than MAX_ALLOC_CHUNK */ - if (slop > (size_t) (MAX_ALLOC_CHUNK-min_request)) - slop = (size_t) (MAX_ALLOC_CHUNK-min_request); - /* Try to get space, if fail reduce slop and try again */ - for (;;) { - hdr_ptr = (small_pool_ptr) jpeg_get_small(cinfo, min_request + slop); - if (hdr_ptr != NULL) - break; - slop /= 2; - if (slop < MIN_SLOP) /* give up when it gets real small */ - out_of_memory(cinfo, 2); /* jpeg_get_small failed */ - } - mem->total_space_allocated += min_request + slop; - /* Success, initialize the new pool header and add to end of list */ - hdr_ptr->hdr.next = NULL; - hdr_ptr->hdr.bytes_used = 0; - hdr_ptr->hdr.bytes_left = sizeofobject + slop; - if (prev_hdr_ptr == NULL) /* first pool in class? */ - mem->small_list[pool_id] = hdr_ptr; - else - prev_hdr_ptr->hdr.next = hdr_ptr; - } - - /* OK, allocate the object from the current pool */ - data_ptr = (char *) (hdr_ptr + 1); /* point to first data byte in pool */ - data_ptr += hdr_ptr->hdr.bytes_used; /* point to place for object */ - hdr_ptr->hdr.bytes_used += sizeofobject; - hdr_ptr->hdr.bytes_left -= sizeofobject; - - return (void *) data_ptr; -} - - -/* - * Allocation of "large" objects. - * - * The external semantics of these are the same as "small" objects, - * except that FAR pointers are used on 80x86. However the pool - * management heuristics are quite different. We assume that each - * request is large enough that it may as well be passed directly to - * jpeg_get_large; the pool management just links everything together - * so that we can free it all on demand. - * Note: the major use of "large" objects is in JSAMPARRAY and JBLOCKARRAY - * structures. The routines that create these structures (see below) - * deliberately bunch rows together to ensure a large request size. - */ - -METHODDEF(void FAR *) -alloc_large (j_common_ptr cinfo, int pool_id, size_t sizeofobject) -/* Allocate a "large" object */ -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - large_pool_ptr hdr_ptr; - size_t odd_bytes; - - /* Check for unsatisfiable request (do now to ensure no overflow below) */ - if (sizeofobject > (size_t) (MAX_ALLOC_CHUNK-SIZEOF(large_pool_hdr))) - out_of_memory(cinfo, 3); /* request exceeds malloc's ability */ - - /* Round up the requested size to a multiple of SIZEOF(ALIGN_TYPE) */ - odd_bytes = sizeofobject % SIZEOF(ALIGN_TYPE); - if (odd_bytes > 0) - sizeofobject += SIZEOF(ALIGN_TYPE) - odd_bytes; - - /* Always make a new pool */ - if (pool_id < 0 || pool_id >= JPOOL_NUMPOOLS) - ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ - - hdr_ptr = (large_pool_ptr) jpeg_get_large(cinfo, sizeofobject + - SIZEOF(large_pool_hdr)); - if (hdr_ptr == NULL) - out_of_memory(cinfo, 4); /* jpeg_get_large failed */ - mem->total_space_allocated += sizeofobject + SIZEOF(large_pool_hdr); - - /* Success, initialize the new pool header and add to list */ - hdr_ptr->hdr.next = mem->large_list[pool_id]; - /* We maintain space counts in each pool header for statistical purposes, - * even though they are not needed for allocation. - */ - hdr_ptr->hdr.bytes_used = sizeofobject; - hdr_ptr->hdr.bytes_left = 0; - mem->large_list[pool_id] = hdr_ptr; - - return (void FAR *) (hdr_ptr + 1); /* point to first data byte in pool */ -} - - -/* - * Creation of 2-D sample arrays. - * The pointers are in near heap, the samples themselves in FAR heap. - * - * To minimize allocation overhead and to allow I/O of large contiguous - * blocks, we allocate the sample rows in groups of as many rows as possible - * without exceeding MAX_ALLOC_CHUNK total bytes per allocation request. - * NB: the virtual array control routines, later in this file, know about - * this chunking of rows. The rowsperchunk value is left in the mem manager - * object so that it can be saved away if this sarray is the workspace for - * a virtual array. - */ - -METHODDEF(JSAMPARRAY) -alloc_sarray (j_common_ptr cinfo, int pool_id, - JDIMENSION samplesperrow, JDIMENSION numrows) -/* Allocate a 2-D sample array */ -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - JSAMPARRAY result; - JSAMPROW workspace; - JDIMENSION rowsperchunk, currow, i; - long ltemp; - - /* Calculate max # of rows allowed in one allocation chunk */ - ltemp = (MAX_ALLOC_CHUNK-SIZEOF(large_pool_hdr)) / - ((long) samplesperrow * SIZEOF(JSAMPLE)); - if (ltemp <= 0) - ERREXIT(cinfo, JERR_WIDTH_OVERFLOW); - if (ltemp < (long) numrows) - rowsperchunk = (JDIMENSION) ltemp; - else - rowsperchunk = numrows; - mem->last_rowsperchunk = rowsperchunk; - - /* Get space for row pointers (small object) */ - result = (JSAMPARRAY) alloc_small(cinfo, pool_id, - (size_t) (numrows * SIZEOF(JSAMPROW))); - - /* Get the rows themselves (large objects) */ - currow = 0; - while (currow < numrows) { - rowsperchunk = MIN(rowsperchunk, numrows - currow); - workspace = (JSAMPROW) alloc_large(cinfo, pool_id, - (size_t) ((size_t) rowsperchunk * (size_t) samplesperrow - * SIZEOF(JSAMPLE))); - for (i = rowsperchunk; i > 0; i--) { - result[currow++] = workspace; - workspace += samplesperrow; - } - } - - return result; -} - - -/* - * Creation of 2-D coefficient-block arrays. - * This is essentially the same as the code for sample arrays, above. - */ - -METHODDEF(JBLOCKARRAY) -alloc_barray (j_common_ptr cinfo, int pool_id, - JDIMENSION blocksperrow, JDIMENSION numrows) -/* Allocate a 2-D coefficient-block array */ -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - JBLOCKARRAY result; - JBLOCKROW workspace; - JDIMENSION rowsperchunk, currow, i; - long ltemp; - - /* Calculate max # of rows allowed in one allocation chunk */ - ltemp = (MAX_ALLOC_CHUNK-SIZEOF(large_pool_hdr)) / - ((long) blocksperrow * SIZEOF(JBLOCK)); - if (ltemp <= 0) - ERREXIT(cinfo, JERR_WIDTH_OVERFLOW); - if (ltemp < (long) numrows) - rowsperchunk = (JDIMENSION) ltemp; - else - rowsperchunk = numrows; - mem->last_rowsperchunk = rowsperchunk; - - /* Get space for row pointers (small object) */ - result = (JBLOCKARRAY) alloc_small(cinfo, pool_id, - (size_t) (numrows * SIZEOF(JBLOCKROW))); - - /* Get the rows themselves (large objects) */ - currow = 0; - while (currow < numrows) { - rowsperchunk = MIN(rowsperchunk, numrows - currow); - workspace = (JBLOCKROW) alloc_large(cinfo, pool_id, - (size_t) ((size_t) rowsperchunk * (size_t) blocksperrow - * SIZEOF(JBLOCK))); - for (i = rowsperchunk; i > 0; i--) { - result[currow++] = workspace; - workspace += blocksperrow; - } - } - - return result; -} - - -/* - * About virtual array management: - * - * The above "normal" array routines are only used to allocate strip buffers - * (as wide as the image, but just a few rows high). Full-image-sized buffers - * are handled as "virtual" arrays. The array is still accessed a strip at a - * time, but the memory manager must save the whole array for repeated - * accesses. The intended implementation is that there is a strip buffer in - * memory (as high as is possible given the desired memory limit), plus a - * backing file that holds the rest of the array. - * - * The request_virt_array routines are told the total size of the image and - * the maximum number of rows that will be accessed at once. The in-memory - * buffer must be at least as large as the maxaccess value. - * - * The request routines create control blocks but not the in-memory buffers. - * That is postponed until realize_virt_arrays is called. At that time the - * total amount of space needed is known (approximately, anyway), so free - * memory can be divided up fairly. - * - * The access_virt_array routines are responsible for making a specific strip - * area accessible (after reading or writing the backing file, if necessary). - * Note that the access routines are told whether the caller intends to modify - * the accessed strip; during a read-only pass this saves having to rewrite - * data to disk. The access routines are also responsible for pre-zeroing - * any newly accessed rows, if pre-zeroing was requested. - * - * In current usage, the access requests are usually for nonoverlapping - * strips; that is, successive access start_row numbers differ by exactly - * num_rows = maxaccess. This means we can get good performance with simple - * buffer dump/reload logic, by making the in-memory buffer be a multiple - * of the access height; then there will never be accesses across bufferload - * boundaries. The code will still work with overlapping access requests, - * but it doesn't handle bufferload overlaps very efficiently. - */ - - -METHODDEF(jvirt_sarray_ptr) -request_virt_sarray (j_common_ptr cinfo, int pool_id, boolean pre_zero, - JDIMENSION samplesperrow, JDIMENSION numrows, - JDIMENSION maxaccess) -/* Request a virtual 2-D sample array */ -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - jvirt_sarray_ptr result; - - /* Only IMAGE-lifetime virtual arrays are currently supported */ - if (pool_id != JPOOL_IMAGE) - ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ - - /* get control block */ - result = (jvirt_sarray_ptr) alloc_small(cinfo, pool_id, - SIZEOF(struct jvirt_sarray_control)); - - result->mem_buffer = NULL; /* marks array not yet realized */ - result->rows_in_array = numrows; - result->samplesperrow = samplesperrow; - result->maxaccess = maxaccess; - result->pre_zero = pre_zero; - result->b_s_open = FALSE; /* no associated backing-store object */ - result->next = mem->virt_sarray_list; /* add to list of virtual arrays */ - mem->virt_sarray_list = result; - - return result; -} - - -METHODDEF(jvirt_barray_ptr) -request_virt_barray (j_common_ptr cinfo, int pool_id, boolean pre_zero, - JDIMENSION blocksperrow, JDIMENSION numrows, - JDIMENSION maxaccess) -/* Request a virtual 2-D coefficient-block array */ -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - jvirt_barray_ptr result; - - /* Only IMAGE-lifetime virtual arrays are currently supported */ - if (pool_id != JPOOL_IMAGE) - ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ - - /* get control block */ - result = (jvirt_barray_ptr) alloc_small(cinfo, pool_id, - SIZEOF(struct jvirt_barray_control)); - - result->mem_buffer = NULL; /* marks array not yet realized */ - result->rows_in_array = numrows; - result->blocksperrow = blocksperrow; - result->maxaccess = maxaccess; - result->pre_zero = pre_zero; - result->b_s_open = FALSE; /* no associated backing-store object */ - result->next = mem->virt_barray_list; /* add to list of virtual arrays */ - mem->virt_barray_list = result; - - return result; -} - - -METHODDEF(void) -realize_virt_arrays (j_common_ptr cinfo) -/* Allocate the in-memory buffers for any unrealized virtual arrays */ -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - long space_per_minheight, maximum_space, avail_mem; - long minheights, max_minheights; - jvirt_sarray_ptr sptr; - jvirt_barray_ptr bptr; - - /* Compute the minimum space needed (maxaccess rows in each buffer) - * and the maximum space needed (full image height in each buffer). - * These may be of use to the system-dependent jpeg_mem_available routine. - */ - space_per_minheight = 0; - maximum_space = 0; - for (sptr = mem->virt_sarray_list; sptr != NULL; sptr = sptr->next) { - if (sptr->mem_buffer == NULL) { /* if not realized yet */ - space_per_minheight += (long) sptr->maxaccess * - (long) sptr->samplesperrow * SIZEOF(JSAMPLE); - maximum_space += (long) sptr->rows_in_array * - (long) sptr->samplesperrow * SIZEOF(JSAMPLE); - } - } - for (bptr = mem->virt_barray_list; bptr != NULL; bptr = bptr->next) { - if (bptr->mem_buffer == NULL) { /* if not realized yet */ - space_per_minheight += (long) bptr->maxaccess * - (long) bptr->blocksperrow * SIZEOF(JBLOCK); - maximum_space += (long) bptr->rows_in_array * - (long) bptr->blocksperrow * SIZEOF(JBLOCK); - } - } - - if (space_per_minheight <= 0) - return; /* no unrealized arrays, no work */ - - /* Determine amount of memory to actually use; this is system-dependent. */ - avail_mem = jpeg_mem_available(cinfo, space_per_minheight, maximum_space, - mem->total_space_allocated); - - /* If the maximum space needed is available, make all the buffers full - * height; otherwise parcel it out with the same number of minheights - * in each buffer. - */ - if (avail_mem >= maximum_space) - max_minheights = 1000000000L; - else { - max_minheights = avail_mem / space_per_minheight; - /* If there doesn't seem to be enough space, try to get the minimum - * anyway. This allows a "stub" implementation of jpeg_mem_available(). - */ - if (max_minheights <= 0) - max_minheights = 1; - } - - /* Allocate the in-memory buffers and initialize backing store as needed. */ - - for (sptr = mem->virt_sarray_list; sptr != NULL; sptr = sptr->next) { - if (sptr->mem_buffer == NULL) { /* if not realized yet */ - minheights = ((long) sptr->rows_in_array - 1L) / sptr->maxaccess + 1L; - if (minheights <= max_minheights) { - /* This buffer fits in memory */ - sptr->rows_in_mem = sptr->rows_in_array; - } else { - /* It doesn't fit in memory, create backing store. */ - sptr->rows_in_mem = (JDIMENSION) (max_minheights * sptr->maxaccess); - jpeg_open_backing_store(cinfo, & sptr->b_s_info, - (long) sptr->rows_in_array * - (long) sptr->samplesperrow * - (long) SIZEOF(JSAMPLE)); - sptr->b_s_open = TRUE; - } - sptr->mem_buffer = alloc_sarray(cinfo, JPOOL_IMAGE, - sptr->samplesperrow, sptr->rows_in_mem); - sptr->rowsperchunk = mem->last_rowsperchunk; - sptr->cur_start_row = 0; - sptr->first_undef_row = 0; - sptr->dirty = FALSE; - } - } - - for (bptr = mem->virt_barray_list; bptr != NULL; bptr = bptr->next) { - if (bptr->mem_buffer == NULL) { /* if not realized yet */ - minheights = ((long) bptr->rows_in_array - 1L) / bptr->maxaccess + 1L; - if (minheights <= max_minheights) { - /* This buffer fits in memory */ - bptr->rows_in_mem = bptr->rows_in_array; - } else { - /* It doesn't fit in memory, create backing store. */ - bptr->rows_in_mem = (JDIMENSION) (max_minheights * bptr->maxaccess); - jpeg_open_backing_store(cinfo, & bptr->b_s_info, - (long) bptr->rows_in_array * - (long) bptr->blocksperrow * - (long) SIZEOF(JBLOCK)); - bptr->b_s_open = TRUE; - } - bptr->mem_buffer = alloc_barray(cinfo, JPOOL_IMAGE, - bptr->blocksperrow, bptr->rows_in_mem); - bptr->rowsperchunk = mem->last_rowsperchunk; - bptr->cur_start_row = 0; - bptr->first_undef_row = 0; - bptr->dirty = FALSE; - } - } -} - - -LOCAL(void) -do_sarray_io (j_common_ptr cinfo, jvirt_sarray_ptr ptr, boolean writing) -/* Do backing store read or write of a virtual sample array */ -{ - long bytesperrow, file_offset, byte_count, rows, thisrow, i; - - bytesperrow = (long) ptr->samplesperrow * SIZEOF(JSAMPLE); - file_offset = ptr->cur_start_row * bytesperrow; - /* Loop to read or write each allocation chunk in mem_buffer */ - for (i = 0; i < (long) ptr->rows_in_mem; i += ptr->rowsperchunk) { - /* One chunk, but check for short chunk at end of buffer */ - rows = MIN((long) ptr->rowsperchunk, (long) ptr->rows_in_mem - i); - /* Transfer no more than is currently defined */ - thisrow = (long) ptr->cur_start_row + i; - rows = MIN(rows, (long) ptr->first_undef_row - thisrow); - /* Transfer no more than fits in file */ - rows = MIN(rows, (long) ptr->rows_in_array - thisrow); - if (rows <= 0) /* this chunk might be past end of file! */ - break; - byte_count = rows * bytesperrow; - if (writing) - (*ptr->b_s_info.write_backing_store) (cinfo, & ptr->b_s_info, - (void FAR *) ptr->mem_buffer[i], - file_offset, byte_count); - else - (*ptr->b_s_info.read_backing_store) (cinfo, & ptr->b_s_info, - (void FAR *) ptr->mem_buffer[i], - file_offset, byte_count); - file_offset += byte_count; - } -} - - -LOCAL(void) -do_barray_io (j_common_ptr cinfo, jvirt_barray_ptr ptr, boolean writing) -/* Do backing store read or write of a virtual coefficient-block array */ -{ - long bytesperrow, file_offset, byte_count, rows, thisrow, i; - - bytesperrow = (long) ptr->blocksperrow * SIZEOF(JBLOCK); - file_offset = ptr->cur_start_row * bytesperrow; - /* Loop to read or write each allocation chunk in mem_buffer */ - for (i = 0; i < (long) ptr->rows_in_mem; i += ptr->rowsperchunk) { - /* One chunk, but check for short chunk at end of buffer */ - rows = MIN((long) ptr->rowsperchunk, (long) ptr->rows_in_mem - i); - /* Transfer no more than is currently defined */ - thisrow = (long) ptr->cur_start_row + i; - rows = MIN(rows, (long) ptr->first_undef_row - thisrow); - /* Transfer no more than fits in file */ - rows = MIN(rows, (long) ptr->rows_in_array - thisrow); - if (rows <= 0) /* this chunk might be past end of file! */ - break; - byte_count = rows * bytesperrow; - if (writing) - (*ptr->b_s_info.write_backing_store) (cinfo, & ptr->b_s_info, - (void FAR *) ptr->mem_buffer[i], - file_offset, byte_count); - else - (*ptr->b_s_info.read_backing_store) (cinfo, & ptr->b_s_info, - (void FAR *) ptr->mem_buffer[i], - file_offset, byte_count); - file_offset += byte_count; - } -} - - -METHODDEF(JSAMPARRAY) -access_virt_sarray (j_common_ptr cinfo, jvirt_sarray_ptr ptr, - JDIMENSION start_row, JDIMENSION num_rows, - boolean writable) -/* Access the part of a virtual sample array starting at start_row */ -/* and extending for num_rows rows. writable is true if */ -/* caller intends to modify the accessed area. */ -{ - JDIMENSION end_row = start_row + num_rows; - JDIMENSION undef_row; - - /* debugging check */ - if (end_row > ptr->rows_in_array || num_rows > ptr->maxaccess || - ptr->mem_buffer == NULL) - ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); - - /* Make the desired part of the virtual array accessible */ - if (start_row < ptr->cur_start_row || - end_row > ptr->cur_start_row+ptr->rows_in_mem) { - if (! ptr->b_s_open) - ERREXIT(cinfo, JERR_VIRTUAL_BUG); - /* Flush old buffer contents if necessary */ - if (ptr->dirty) { - do_sarray_io(cinfo, ptr, TRUE); - ptr->dirty = FALSE; - } - /* Decide what part of virtual array to access. - * Algorithm: if target address > current window, assume forward scan, - * load starting at target address. If target address < current window, - * assume backward scan, load so that target area is top of window. - * Note that when switching from forward write to forward read, will have - * start_row = 0, so the limiting case applies and we load from 0 anyway. - */ - if (start_row > ptr->cur_start_row) { - ptr->cur_start_row = start_row; - } else { - /* use long arithmetic here to avoid overflow & unsigned problems */ - long ltemp; - - ltemp = (long) end_row - (long) ptr->rows_in_mem; - if (ltemp < 0) - ltemp = 0; /* don't fall off front end of file */ - ptr->cur_start_row = (JDIMENSION) ltemp; - } - /* Read in the selected part of the array. - * During the initial write pass, we will do no actual read - * because the selected part is all undefined. - */ - do_sarray_io(cinfo, ptr, FALSE); - } - /* Ensure the accessed part of the array is defined; prezero if needed. - * To improve locality of access, we only prezero the part of the array - * that the caller is about to access, not the entire in-memory array. - */ - if (ptr->first_undef_row < end_row) { - if (ptr->first_undef_row < start_row) { - if (writable) /* writer skipped over a section of array */ - ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); - undef_row = start_row; /* but reader is allowed to read ahead */ - } else { - undef_row = ptr->first_undef_row; - } - if (writable) - ptr->first_undef_row = end_row; - if (ptr->pre_zero) { - size_t bytesperrow = (size_t) ptr->samplesperrow * SIZEOF(JSAMPLE); - undef_row -= ptr->cur_start_row; /* make indexes relative to buffer */ - end_row -= ptr->cur_start_row; - while (undef_row < end_row) { - jzero_far((void FAR *) ptr->mem_buffer[undef_row], bytesperrow); - undef_row++; - } - } else { - if (! writable) /* reader looking at undefined data */ - ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); - } - } - /* Flag the buffer dirty if caller will write in it */ - if (writable) - ptr->dirty = TRUE; - /* Return address of proper part of the buffer */ - return ptr->mem_buffer + (start_row - ptr->cur_start_row); -} - - -METHODDEF(JBLOCKARRAY) -access_virt_barray (j_common_ptr cinfo, jvirt_barray_ptr ptr, - JDIMENSION start_row, JDIMENSION num_rows, - boolean writable) -/* Access the part of a virtual block array starting at start_row */ -/* and extending for num_rows rows. writable is true if */ -/* caller intends to modify the accessed area. */ -{ - JDIMENSION end_row = start_row + num_rows; - JDIMENSION undef_row; - - /* debugging check */ - if (end_row > ptr->rows_in_array || num_rows > ptr->maxaccess || - ptr->mem_buffer == NULL) - ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); - - /* Make the desired part of the virtual array accessible */ - if (start_row < ptr->cur_start_row || - end_row > ptr->cur_start_row+ptr->rows_in_mem) { - if (! ptr->b_s_open) - ERREXIT(cinfo, JERR_VIRTUAL_BUG); - /* Flush old buffer contents if necessary */ - if (ptr->dirty) { - do_barray_io(cinfo, ptr, TRUE); - ptr->dirty = FALSE; - } - /* Decide what part of virtual array to access. - * Algorithm: if target address > current window, assume forward scan, - * load starting at target address. If target address < current window, - * assume backward scan, load so that target area is top of window. - * Note that when switching from forward write to forward read, will have - * start_row = 0, so the limiting case applies and we load from 0 anyway. - */ - if (start_row > ptr->cur_start_row) { - ptr->cur_start_row = start_row; - } else { - /* use long arithmetic here to avoid overflow & unsigned problems */ - long ltemp; - - ltemp = (long) end_row - (long) ptr->rows_in_mem; - if (ltemp < 0) - ltemp = 0; /* don't fall off front end of file */ - ptr->cur_start_row = (JDIMENSION) ltemp; - } - /* Read in the selected part of the array. - * During the initial write pass, we will do no actual read - * because the selected part is all undefined. - */ - do_barray_io(cinfo, ptr, FALSE); - } - /* Ensure the accessed part of the array is defined; prezero if needed. - * To improve locality of access, we only prezero the part of the array - * that the caller is about to access, not the entire in-memory array. - */ - if (ptr->first_undef_row < end_row) { - if (ptr->first_undef_row < start_row) { - if (writable) /* writer skipped over a section of array */ - ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); - undef_row = start_row; /* but reader is allowed to read ahead */ - } else { - undef_row = ptr->first_undef_row; - } - if (writable) - ptr->first_undef_row = end_row; - if (ptr->pre_zero) { - size_t bytesperrow = (size_t) ptr->blocksperrow * SIZEOF(JBLOCK); - undef_row -= ptr->cur_start_row; /* make indexes relative to buffer */ - end_row -= ptr->cur_start_row; - while (undef_row < end_row) { - jzero_far((void FAR *) ptr->mem_buffer[undef_row], bytesperrow); - undef_row++; - } - } else { - if (! writable) /* reader looking at undefined data */ - ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); - } - } - /* Flag the buffer dirty if caller will write in it */ - if (writable) - ptr->dirty = TRUE; - /* Return address of proper part of the buffer */ - return ptr->mem_buffer + (start_row - ptr->cur_start_row); -} - - -/* - * Release all objects belonging to a specified pool. - */ - -METHODDEF(void) -free_pool (j_common_ptr cinfo, int pool_id) -{ - my_mem_ptr mem = (my_mem_ptr) cinfo->mem; - small_pool_ptr shdr_ptr; - large_pool_ptr lhdr_ptr; - size_t space_freed; - - if (pool_id < 0 || pool_id >= JPOOL_NUMPOOLS) - ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ - -#ifdef MEM_STATS - if (cinfo->err->trace_level > 1) - print_mem_stats(cinfo, pool_id); /* print pool's memory usage statistics */ -#endif - - /* If freeing IMAGE pool, close any virtual arrays first */ - if (pool_id == JPOOL_IMAGE) { - jvirt_sarray_ptr sptr; - jvirt_barray_ptr bptr; - - for (sptr = mem->virt_sarray_list; sptr != NULL; sptr = sptr->next) { - if (sptr->b_s_open) { /* there may be no backing store */ - sptr->b_s_open = FALSE; /* prevent recursive close if error */ - (*sptr->b_s_info.close_backing_store) (cinfo, & sptr->b_s_info); - } - } - mem->virt_sarray_list = NULL; - for (bptr = mem->virt_barray_list; bptr != NULL; bptr = bptr->next) { - if (bptr->b_s_open) { /* there may be no backing store */ - bptr->b_s_open = FALSE; /* prevent recursive close if error */ - (*bptr->b_s_info.close_backing_store) (cinfo, & bptr->b_s_info); - } - } - mem->virt_barray_list = NULL; - } - - /* Release large objects */ - lhdr_ptr = mem->large_list[pool_id]; - mem->large_list[pool_id] = NULL; - - while (lhdr_ptr != NULL) { - large_pool_ptr next_lhdr_ptr = lhdr_ptr->hdr.next; - space_freed = lhdr_ptr->hdr.bytes_used + - lhdr_ptr->hdr.bytes_left + - SIZEOF(large_pool_hdr); - jpeg_free_large(cinfo, (void FAR *) lhdr_ptr, space_freed); - mem->total_space_allocated -= space_freed; - lhdr_ptr = next_lhdr_ptr; - } - - /* Release small objects */ - shdr_ptr = mem->small_list[pool_id]; - mem->small_list[pool_id] = NULL; - - while (shdr_ptr != NULL) { - small_pool_ptr next_shdr_ptr = shdr_ptr->hdr.next; - space_freed = shdr_ptr->hdr.bytes_used + - shdr_ptr->hdr.bytes_left + - SIZEOF(small_pool_hdr); - jpeg_free_small(cinfo, (void *) shdr_ptr, space_freed); - mem->total_space_allocated -= space_freed; - shdr_ptr = next_shdr_ptr; - } -} - - -/* - * Close up shop entirely. - * Note that this cannot be called unless cinfo->mem is non-NULL. - */ - -METHODDEF(void) -self_destruct (j_common_ptr cinfo) -{ - int pool; - - /* Close all backing store, release all memory. - * Releasing pools in reverse order might help avoid fragmentation - * with some (brain-damaged) malloc libraries. - */ - for (pool = JPOOL_NUMPOOLS-1; pool >= JPOOL_PERMANENT; pool--) { - free_pool(cinfo, pool); - } - - /* Release the memory manager control block too. */ - jpeg_free_small(cinfo, (void *) cinfo->mem, SIZEOF(my_memory_mgr)); - cinfo->mem = NULL; /* ensures I will be called only once */ - - jpeg_mem_term(cinfo); /* system-dependent cleanup */ -} - - -/* - * Memory manager initialization. - * When this is called, only the error manager pointer is valid in cinfo! - */ - -GLOBAL(void) -jinit_memory_mgr (j_common_ptr cinfo) -{ - my_mem_ptr mem; - long max_to_use; - int pool; - size_t test_mac; - - cinfo->mem = NULL; /* for safety if init fails */ - - /* Check for configuration errors. - * SIZEOF(ALIGN_TYPE) should be a power of 2; otherwise, it probably - * doesn't reflect any real hardware alignment requirement. - * The test is a little tricky: for X>0, X and X-1 have no one-bits - * in common if and only if X is a power of 2, ie has only one one-bit. - * Some compilers may give an "unreachable code" warning here; ignore it. - */ - if ((SIZEOF(ALIGN_TYPE) & (SIZEOF(ALIGN_TYPE)-1)) != 0) - ERREXIT(cinfo, JERR_BAD_ALIGN_TYPE); - /* MAX_ALLOC_CHUNK must be representable as type size_t, and must be - * a multiple of SIZEOF(ALIGN_TYPE). - * Again, an "unreachable code" warning may be ignored here. - * But a "constant too large" warning means you need to fix MAX_ALLOC_CHUNK. - */ - test_mac = (size_t) MAX_ALLOC_CHUNK; - if ((long) test_mac != MAX_ALLOC_CHUNK || - (MAX_ALLOC_CHUNK % SIZEOF(ALIGN_TYPE)) != 0) - ERREXIT(cinfo, JERR_BAD_ALLOC_CHUNK); - - max_to_use = jpeg_mem_init(cinfo); /* system-dependent initialization */ - - /* Attempt to allocate memory manager's control block */ - mem = (my_mem_ptr) jpeg_get_small(cinfo, SIZEOF(my_memory_mgr)); - - if (mem == NULL) { - jpeg_mem_term(cinfo); /* system-dependent cleanup */ - ERREXIT1(cinfo, JERR_OUT_OF_MEMORY, 0); - } - - /* OK, fill in the method pointers */ - mem->pub.alloc_small = alloc_small; - mem->pub.alloc_large = alloc_large; - mem->pub.alloc_sarray = alloc_sarray; - mem->pub.alloc_barray = alloc_barray; - mem->pub.request_virt_sarray = request_virt_sarray; - mem->pub.request_virt_barray = request_virt_barray; - mem->pub.realize_virt_arrays = realize_virt_arrays; - mem->pub.access_virt_sarray = access_virt_sarray; - mem->pub.access_virt_barray = access_virt_barray; - mem->pub.free_pool = free_pool; - mem->pub.self_destruct = self_destruct; - - /* Make MAX_ALLOC_CHUNK accessible to other modules */ - mem->pub.max_alloc_chunk = MAX_ALLOC_CHUNK; - - /* Initialize working state */ - mem->pub.max_memory_to_use = max_to_use; - - for (pool = JPOOL_NUMPOOLS-1; pool >= JPOOL_PERMANENT; pool--) { - mem->small_list[pool] = NULL; - mem->large_list[pool] = NULL; - } - mem->virt_sarray_list = NULL; - mem->virt_barray_list = NULL; - - mem->total_space_allocated = SIZEOF(my_memory_mgr); - - /* Declare ourselves open for business */ - cinfo->mem = & mem->pub; - - /* Check for an environment variable JPEGMEM; if found, override the - * default max_memory setting from jpeg_mem_init. Note that the - * surrounding application may again override this value. - * If your system doesn't support getenv(), define NO_GETENV to disable - * this feature. - */ -#ifndef NO_GETENV - { char * memenv; - - if ((memenv = getenv("JPEGMEM")) != NULL) { - char ch = 'x'; - - if (sscanf(memenv, "%ld%c", &max_to_use, &ch) > 0) { - if (ch == 'm' || ch == 'M') - max_to_use *= 1000L; - mem->pub.max_memory_to_use = max_to_use * 1000L; - } - } - } -#endif - -} +/* + * jmemmgr.c + * + * Copyright (C) 1991-1997, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains the JPEG system-independent memory management + * routines. This code is usable across a wide variety of machines; most + * of the system dependencies have been isolated in a separate file. + * The major functions provided here are: + * * pool-based allocation and freeing of memory; + * * policy decisions about how to divide available memory among the + * virtual arrays; + * * control logic for swapping virtual arrays between main memory and + * backing storage. + * The separate system-dependent file provides the actual backing-storage + * access code, and it contains the policy decision about how much total + * main memory to use. + * This file is system-dependent in the sense that some of its functions + * are unnecessary in some systems. For example, if there is enough virtual + * memory so that backing storage will never be used, much of the virtual + * array control logic could be removed. (Of course, if you have that much + * memory then you shouldn't care about a little bit of unused code...) + */ + +#define JPEG_INTERNALS +#define AM_MEMORY_MANAGER /* we define jvirt_Xarray_control structs */ +#include "jinclude.h" +#include "jpeglib.h" +#include "jmemsys.h" /* import the system-dependent declarations */ + +#ifndef NO_GETENV +#ifndef HAVE_STDLIB_H /* should declare getenv() */ +extern char * getenv JPP((const char * name)); +#endif +#endif + + +/* + * Some important notes: + * The allocation routines provided here must never return NULL. + * They should exit to error_exit if unsuccessful. + * + * It's not a good idea to try to merge the sarray and barray routines, + * even though they are textually almost the same, because samples are + * usually stored as bytes while coefficients are shorts or ints. Thus, + * in machines where byte pointers have a different representation from + * word pointers, the resulting machine code could not be the same. + */ + + +/* + * Many machines require storage alignment: longs must start on 4-byte + * boundaries, doubles on 8-byte boundaries, etc. On such machines, malloc() + * always returns pointers that are multiples of the worst-case alignment + * requirement, and we had better do so too. + * There isn't any really portable way to determine the worst-case alignment + * requirement. This module assumes that the alignment requirement is + * multiples of sizeof(ALIGN_TYPE). + * By default, we define ALIGN_TYPE as double. This is necessary on some + * workstations (where doubles really do need 8-byte alignment) and will work + * fine on nearly everything. If your machine has lesser alignment needs, + * you can save a few bytes by making ALIGN_TYPE smaller. + * The only place I know of where this will NOT work is certain Macintosh + * 680x0 compilers that define double as a 10-byte IEEE extended float. + * Doing 10-byte alignment is counterproductive because longwords won't be + * aligned well. Put "#define ALIGN_TYPE long" in jconfig.h if you have + * such a compiler. + */ + +#ifndef ALIGN_TYPE /* so can override from jconfig.h */ +#define ALIGN_TYPE double +#endif + + +/* + * We allocate objects from "pools", where each pool is gotten with a single + * request to jpeg_get_small() or jpeg_get_large(). There is no per-object + * overhead within a pool, except for alignment padding. Each pool has a + * header with a link to the next pool of the same class. + * Small and large pool headers are identical except that the latter's + * link pointer must be FAR on 80x86 machines. + * Notice that the "real" header fields are union'ed with a dummy ALIGN_TYPE + * field. This forces the compiler to make SIZEOF(small_pool_hdr) a multiple + * of the alignment requirement of ALIGN_TYPE. + */ + +typedef union small_pool_struct * small_pool_ptr; + +typedef union small_pool_struct { + struct { + small_pool_ptr next; /* next in list of pools */ + size_t bytes_used; /* how many bytes already used within pool */ + size_t bytes_left; /* bytes still available in this pool */ + } hdr; + ALIGN_TYPE dummy; /* included in union to ensure alignment */ +} small_pool_hdr; + +typedef union large_pool_struct FAR * large_pool_ptr; + +typedef union large_pool_struct { + struct { + large_pool_ptr next; /* next in list of pools */ + size_t bytes_used; /* how many bytes already used within pool */ + size_t bytes_left; /* bytes still available in this pool */ + } hdr; + ALIGN_TYPE dummy; /* included in union to ensure alignment */ +} large_pool_hdr; + + +/* + * Here is the full definition of a memory manager object. + */ + +typedef struct { + struct jpeg_memory_mgr pub; /* public fields */ + + /* Each pool identifier (lifetime class) names a linked list of pools. */ + small_pool_ptr small_list[JPOOL_NUMPOOLS]; + large_pool_ptr large_list[JPOOL_NUMPOOLS]; + + /* Since we only have one lifetime class of virtual arrays, only one + * linked list is necessary (for each datatype). Note that the virtual + * array control blocks being linked together are actually stored somewhere + * in the small-pool list. + */ + jvirt_sarray_ptr virt_sarray_list; + jvirt_barray_ptr virt_barray_list; + + /* This counts total space obtained from jpeg_get_small/large */ + long total_space_allocated; + + /* alloc_sarray and alloc_barray set this value for use by virtual + * array routines. + */ + JDIMENSION last_rowsperchunk; /* from most recent alloc_sarray/barray */ +} my_memory_mgr; + +typedef my_memory_mgr * my_mem_ptr; + + +/* + * The control blocks for virtual arrays. + * Note that these blocks are allocated in the "small" pool area. + * System-dependent info for the associated backing store (if any) is hidden + * inside the backing_store_info struct. + */ + +struct jvirt_sarray_control { + JSAMPARRAY mem_buffer; /* => the in-memory buffer */ + JDIMENSION rows_in_array; /* total virtual array height */ + JDIMENSION samplesperrow; /* width of array (and of memory buffer) */ + JDIMENSION maxaccess; /* max rows accessed by access_virt_sarray */ + JDIMENSION rows_in_mem; /* height of memory buffer */ + JDIMENSION rowsperchunk; /* allocation chunk size in mem_buffer */ + JDIMENSION cur_start_row; /* first logical row # in the buffer */ + JDIMENSION first_undef_row; /* row # of first uninitialized row */ + boolean pre_zero; /* pre-zero mode requested? */ + boolean dirty; /* do current buffer contents need written? */ + boolean b_s_open; /* is backing-store data valid? */ + jvirt_sarray_ptr next; /* link to next virtual sarray control block */ + backing_store_info b_s_info; /* System-dependent control info */ +}; + +struct jvirt_barray_control { + JBLOCKARRAY mem_buffer; /* => the in-memory buffer */ + JDIMENSION rows_in_array; /* total virtual array height */ + JDIMENSION blocksperrow; /* width of array (and of memory buffer) */ + JDIMENSION maxaccess; /* max rows accessed by access_virt_barray */ + JDIMENSION rows_in_mem; /* height of memory buffer */ + JDIMENSION rowsperchunk; /* allocation chunk size in mem_buffer */ + JDIMENSION cur_start_row; /* first logical row # in the buffer */ + JDIMENSION first_undef_row; /* row # of first uninitialized row */ + boolean pre_zero; /* pre-zero mode requested? */ + boolean dirty; /* do current buffer contents need written? */ + boolean b_s_open; /* is backing-store data valid? */ + jvirt_barray_ptr next; /* link to next virtual barray control block */ + backing_store_info b_s_info; /* System-dependent control info */ +}; + + +#ifdef MEM_STATS /* optional extra stuff for statistics */ + +LOCAL(void) +print_mem_stats (j_common_ptr cinfo, int pool_id) +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + small_pool_ptr shdr_ptr; + large_pool_ptr lhdr_ptr; + + /* Since this is only a debugging stub, we can cheat a little by using + * fprintf directly rather than going through the trace message code. + * This is helpful because message parm array can't handle longs. + */ + fprintf(stderr, "Freeing pool %d, total space = %ld\n", + pool_id, mem->total_space_allocated); + + for (lhdr_ptr = mem->large_list[pool_id]; lhdr_ptr != NULL; + lhdr_ptr = lhdr_ptr->hdr.next) { + fprintf(stderr, " Large chunk used %ld\n", + (long) lhdr_ptr->hdr.bytes_used); + } + + for (shdr_ptr = mem->small_list[pool_id]; shdr_ptr != NULL; + shdr_ptr = shdr_ptr->hdr.next) { + fprintf(stderr, " Small chunk used %ld free %ld\n", + (long) shdr_ptr->hdr.bytes_used, + (long) shdr_ptr->hdr.bytes_left); + } +} + +#endif /* MEM_STATS */ + + +LOCAL(void) +out_of_memory (j_common_ptr cinfo, int which) +/* Report an out-of-memory error and stop execution */ +/* If we compiled MEM_STATS support, report alloc requests before dying */ +{ +#ifdef MEM_STATS + cinfo->err->trace_level = 2; /* force self_destruct to report stats */ +#endif + ERREXIT1(cinfo, JERR_OUT_OF_MEMORY, which); +} + + +/* + * Allocation of "small" objects. + * + * For these, we use pooled storage. When a new pool must be created, + * we try to get enough space for the current request plus a "slop" factor, + * where the slop will be the amount of leftover space in the new pool. + * The speed vs. space tradeoff is largely determined by the slop values. + * A different slop value is provided for each pool class (lifetime), + * and we also distinguish the first pool of a class from later ones. + * NOTE: the values given work fairly well on both 16- and 32-bit-int + * machines, but may be too small if longs are 64 bits or more. + */ + +static const size_t first_pool_slop[JPOOL_NUMPOOLS] = +{ + 1600, /* first PERMANENT pool */ + 16000 /* first IMAGE pool */ +}; + +static const size_t extra_pool_slop[JPOOL_NUMPOOLS] = +{ + 0, /* additional PERMANENT pools */ + 5000 /* additional IMAGE pools */ +}; + +#define MIN_SLOP 50 /* greater than 0 to avoid futile looping */ + + +METHODDEF(void *) +alloc_small (j_common_ptr cinfo, int pool_id, size_t sizeofobject) +/* Allocate a "small" object */ +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + small_pool_ptr hdr_ptr, prev_hdr_ptr; + char * data_ptr; + size_t odd_bytes, min_request, slop; + + /* Check for unsatisfiable request (do now to ensure no overflow below) */ + if (sizeofobject > (size_t) (MAX_ALLOC_CHUNK-SIZEOF(small_pool_hdr))) + out_of_memory(cinfo, 1); /* request exceeds malloc's ability */ + + /* Round up the requested size to a multiple of SIZEOF(ALIGN_TYPE) */ + odd_bytes = sizeofobject % SIZEOF(ALIGN_TYPE); + if (odd_bytes > 0) + sizeofobject += SIZEOF(ALIGN_TYPE) - odd_bytes; + + /* See if space is available in any existing pool */ + if (pool_id < 0 || pool_id >= JPOOL_NUMPOOLS) + ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ + prev_hdr_ptr = NULL; + hdr_ptr = mem->small_list[pool_id]; + while (hdr_ptr != NULL) { + if (hdr_ptr->hdr.bytes_left >= sizeofobject) + break; /* found pool with enough space */ + prev_hdr_ptr = hdr_ptr; + hdr_ptr = hdr_ptr->hdr.next; + } + + /* Time to make a new pool? */ + if (hdr_ptr == NULL) { + /* min_request is what we need now, slop is what will be leftover */ + min_request = sizeofobject + SIZEOF(small_pool_hdr); + if (prev_hdr_ptr == NULL) /* first pool in class? */ + slop = first_pool_slop[pool_id]; + else + slop = extra_pool_slop[pool_id]; + /* Don't ask for more than MAX_ALLOC_CHUNK */ + if (slop > (size_t) (MAX_ALLOC_CHUNK-min_request)) + slop = (size_t) (MAX_ALLOC_CHUNK-min_request); + /* Try to get space, if fail reduce slop and try again */ + for (;;) { + hdr_ptr = (small_pool_ptr) jpeg_get_small(cinfo, min_request + slop); + if (hdr_ptr != NULL) + break; + slop /= 2; + if (slop < MIN_SLOP) /* give up when it gets real small */ + out_of_memory(cinfo, 2); /* jpeg_get_small failed */ + } + mem->total_space_allocated += min_request + slop; + /* Success, initialize the new pool header and add to end of list */ + hdr_ptr->hdr.next = NULL; + hdr_ptr->hdr.bytes_used = 0; + hdr_ptr->hdr.bytes_left = sizeofobject + slop; + if (prev_hdr_ptr == NULL) /* first pool in class? */ + mem->small_list[pool_id] = hdr_ptr; + else + prev_hdr_ptr->hdr.next = hdr_ptr; + } + + /* OK, allocate the object from the current pool */ + data_ptr = (char *) (hdr_ptr + 1); /* point to first data byte in pool */ + data_ptr += hdr_ptr->hdr.bytes_used; /* point to place for object */ + hdr_ptr->hdr.bytes_used += sizeofobject; + hdr_ptr->hdr.bytes_left -= sizeofobject; + + return (void *) data_ptr; +} + + +/* + * Allocation of "large" objects. + * + * The external semantics of these are the same as "small" objects, + * except that FAR pointers are used on 80x86. However the pool + * management heuristics are quite different. We assume that each + * request is large enough that it may as well be passed directly to + * jpeg_get_large; the pool management just links everything together + * so that we can free it all on demand. + * Note: the major use of "large" objects is in JSAMPARRAY and JBLOCKARRAY + * structures. The routines that create these structures (see below) + * deliberately bunch rows together to ensure a large request size. + */ + +METHODDEF(void FAR *) +alloc_large (j_common_ptr cinfo, int pool_id, size_t sizeofobject) +/* Allocate a "large" object */ +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + large_pool_ptr hdr_ptr; + size_t odd_bytes; + + /* Check for unsatisfiable request (do now to ensure no overflow below) */ + if (sizeofobject > (size_t) (MAX_ALLOC_CHUNK-SIZEOF(large_pool_hdr))) + out_of_memory(cinfo, 3); /* request exceeds malloc's ability */ + + /* Round up the requested size to a multiple of SIZEOF(ALIGN_TYPE) */ + odd_bytes = sizeofobject % SIZEOF(ALIGN_TYPE); + if (odd_bytes > 0) + sizeofobject += SIZEOF(ALIGN_TYPE) - odd_bytes; + + /* Always make a new pool */ + if (pool_id < 0 || pool_id >= JPOOL_NUMPOOLS) + ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ + + hdr_ptr = (large_pool_ptr) jpeg_get_large(cinfo, sizeofobject + + SIZEOF(large_pool_hdr)); + if (hdr_ptr == NULL) + out_of_memory(cinfo, 4); /* jpeg_get_large failed */ + mem->total_space_allocated += sizeofobject + SIZEOF(large_pool_hdr); + + /* Success, initialize the new pool header and add to list */ + hdr_ptr->hdr.next = mem->large_list[pool_id]; + /* We maintain space counts in each pool header for statistical purposes, + * even though they are not needed for allocation. + */ + hdr_ptr->hdr.bytes_used = sizeofobject; + hdr_ptr->hdr.bytes_left = 0; + mem->large_list[pool_id] = hdr_ptr; + + return (void FAR *) (hdr_ptr + 1); /* point to first data byte in pool */ +} + + +/* + * Creation of 2-D sample arrays. + * The pointers are in near heap, the samples themselves in FAR heap. + * + * To minimize allocation overhead and to allow I/O of large contiguous + * blocks, we allocate the sample rows in groups of as many rows as possible + * without exceeding MAX_ALLOC_CHUNK total bytes per allocation request. + * NB: the virtual array control routines, later in this file, know about + * this chunking of rows. The rowsperchunk value is left in the mem manager + * object so that it can be saved away if this sarray is the workspace for + * a virtual array. + */ + +METHODDEF(JSAMPARRAY) +alloc_sarray (j_common_ptr cinfo, int pool_id, + JDIMENSION samplesperrow, JDIMENSION numrows) +/* Allocate a 2-D sample array */ +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + JSAMPARRAY result; + JSAMPROW workspace; + JDIMENSION rowsperchunk, currow, i; + long ltemp; + + /* Calculate max # of rows allowed in one allocation chunk */ + ltemp = (MAX_ALLOC_CHUNK-SIZEOF(large_pool_hdr)) / + ((long) samplesperrow * SIZEOF(JSAMPLE)); + if (ltemp <= 0) + ERREXIT(cinfo, JERR_WIDTH_OVERFLOW); + if (ltemp < (long) numrows) + rowsperchunk = (JDIMENSION) ltemp; + else + rowsperchunk = numrows; + mem->last_rowsperchunk = rowsperchunk; + + /* Get space for row pointers (small object) */ + result = (JSAMPARRAY) alloc_small(cinfo, pool_id, + (size_t) (numrows * SIZEOF(JSAMPROW))); + + /* Get the rows themselves (large objects) */ + currow = 0; + while (currow < numrows) { + rowsperchunk = MIN(rowsperchunk, numrows - currow); + workspace = (JSAMPROW) alloc_large(cinfo, pool_id, + (size_t) ((size_t) rowsperchunk * (size_t) samplesperrow + * SIZEOF(JSAMPLE))); + for (i = rowsperchunk; i > 0; i--) { + result[currow++] = workspace; + workspace += samplesperrow; + } + } + + return result; +} + + +/* + * Creation of 2-D coefficient-block arrays. + * This is essentially the same as the code for sample arrays, above. + */ + +METHODDEF(JBLOCKARRAY) +alloc_barray (j_common_ptr cinfo, int pool_id, + JDIMENSION blocksperrow, JDIMENSION numrows) +/* Allocate a 2-D coefficient-block array */ +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + JBLOCKARRAY result; + JBLOCKROW workspace; + JDIMENSION rowsperchunk, currow, i; + long ltemp; + + /* Calculate max # of rows allowed in one allocation chunk */ + ltemp = (MAX_ALLOC_CHUNK-SIZEOF(large_pool_hdr)) / + ((long) blocksperrow * SIZEOF(JBLOCK)); + if (ltemp <= 0) + ERREXIT(cinfo, JERR_WIDTH_OVERFLOW); + if (ltemp < (long) numrows) + rowsperchunk = (JDIMENSION) ltemp; + else + rowsperchunk = numrows; + mem->last_rowsperchunk = rowsperchunk; + + /* Get space for row pointers (small object) */ + result = (JBLOCKARRAY) alloc_small(cinfo, pool_id, + (size_t) (numrows * SIZEOF(JBLOCKROW))); + + /* Get the rows themselves (large objects) */ + currow = 0; + while (currow < numrows) { + rowsperchunk = MIN(rowsperchunk, numrows - currow); + workspace = (JBLOCKROW) alloc_large(cinfo, pool_id, + (size_t) ((size_t) rowsperchunk * (size_t) blocksperrow + * SIZEOF(JBLOCK))); + for (i = rowsperchunk; i > 0; i--) { + result[currow++] = workspace; + workspace += blocksperrow; + } + } + + return result; +} + + +/* + * About virtual array management: + * + * The above "normal" array routines are only used to allocate strip buffers + * (as wide as the image, but just a few rows high). Full-image-sized buffers + * are handled as "virtual" arrays. The array is still accessed a strip at a + * time, but the memory manager must save the whole array for repeated + * accesses. The intended implementation is that there is a strip buffer in + * memory (as high as is possible given the desired memory limit), plus a + * backing file that holds the rest of the array. + * + * The request_virt_array routines are told the total size of the image and + * the maximum number of rows that will be accessed at once. The in-memory + * buffer must be at least as large as the maxaccess value. + * + * The request routines create control blocks but not the in-memory buffers. + * That is postponed until realize_virt_arrays is called. At that time the + * total amount of space needed is known (approximately, anyway), so free + * memory can be divided up fairly. + * + * The access_virt_array routines are responsible for making a specific strip + * area accessible (after reading or writing the backing file, if necessary). + * Note that the access routines are told whether the caller intends to modify + * the accessed strip; during a read-only pass this saves having to rewrite + * data to disk. The access routines are also responsible for pre-zeroing + * any newly accessed rows, if pre-zeroing was requested. + * + * In current usage, the access requests are usually for nonoverlapping + * strips; that is, successive access start_row numbers differ by exactly + * num_rows = maxaccess. This means we can get good performance with simple + * buffer dump/reload logic, by making the in-memory buffer be a multiple + * of the access height; then there will never be accesses across bufferload + * boundaries. The code will still work with overlapping access requests, + * but it doesn't handle bufferload overlaps very efficiently. + */ + + +METHODDEF(jvirt_sarray_ptr) +request_virt_sarray (j_common_ptr cinfo, int pool_id, boolean pre_zero, + JDIMENSION samplesperrow, JDIMENSION numrows, + JDIMENSION maxaccess) +/* Request a virtual 2-D sample array */ +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + jvirt_sarray_ptr result; + + /* Only IMAGE-lifetime virtual arrays are currently supported */ + if (pool_id != JPOOL_IMAGE) + ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ + + /* get control block */ + result = (jvirt_sarray_ptr) alloc_small(cinfo, pool_id, + SIZEOF(struct jvirt_sarray_control)); + + result->mem_buffer = NULL; /* marks array not yet realized */ + result->rows_in_array = numrows; + result->samplesperrow = samplesperrow; + result->maxaccess = maxaccess; + result->pre_zero = pre_zero; + result->b_s_open = FALSE; /* no associated backing-store object */ + result->next = mem->virt_sarray_list; /* add to list of virtual arrays */ + mem->virt_sarray_list = result; + + return result; +} + + +METHODDEF(jvirt_barray_ptr) +request_virt_barray (j_common_ptr cinfo, int pool_id, boolean pre_zero, + JDIMENSION blocksperrow, JDIMENSION numrows, + JDIMENSION maxaccess) +/* Request a virtual 2-D coefficient-block array */ +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + jvirt_barray_ptr result; + + /* Only IMAGE-lifetime virtual arrays are currently supported */ + if (pool_id != JPOOL_IMAGE) + ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ + + /* get control block */ + result = (jvirt_barray_ptr) alloc_small(cinfo, pool_id, + SIZEOF(struct jvirt_barray_control)); + + result->mem_buffer = NULL; /* marks array not yet realized */ + result->rows_in_array = numrows; + result->blocksperrow = blocksperrow; + result->maxaccess = maxaccess; + result->pre_zero = pre_zero; + result->b_s_open = FALSE; /* no associated backing-store object */ + result->next = mem->virt_barray_list; /* add to list of virtual arrays */ + mem->virt_barray_list = result; + + return result; +} + + +METHODDEF(void) +realize_virt_arrays (j_common_ptr cinfo) +/* Allocate the in-memory buffers for any unrealized virtual arrays */ +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + long space_per_minheight, maximum_space, avail_mem; + long minheights, max_minheights; + jvirt_sarray_ptr sptr; + jvirt_barray_ptr bptr; + + /* Compute the minimum space needed (maxaccess rows in each buffer) + * and the maximum space needed (full image height in each buffer). + * These may be of use to the system-dependent jpeg_mem_available routine. + */ + space_per_minheight = 0; + maximum_space = 0; + for (sptr = mem->virt_sarray_list; sptr != NULL; sptr = sptr->next) { + if (sptr->mem_buffer == NULL) { /* if not realized yet */ + space_per_minheight += (long) sptr->maxaccess * + (long) sptr->samplesperrow * SIZEOF(JSAMPLE); + maximum_space += (long) sptr->rows_in_array * + (long) sptr->samplesperrow * SIZEOF(JSAMPLE); + } + } + for (bptr = mem->virt_barray_list; bptr != NULL; bptr = bptr->next) { + if (bptr->mem_buffer == NULL) { /* if not realized yet */ + space_per_minheight += (long) bptr->maxaccess * + (long) bptr->blocksperrow * SIZEOF(JBLOCK); + maximum_space += (long) bptr->rows_in_array * + (long) bptr->blocksperrow * SIZEOF(JBLOCK); + } + } + + if (space_per_minheight <= 0) + return; /* no unrealized arrays, no work */ + + /* Determine amount of memory to actually use; this is system-dependent. */ + avail_mem = jpeg_mem_available(cinfo, space_per_minheight, maximum_space, + mem->total_space_allocated); + + /* If the maximum space needed is available, make all the buffers full + * height; otherwise parcel it out with the same number of minheights + * in each buffer. + */ + if (avail_mem >= maximum_space) + max_minheights = 1000000000L; + else { + max_minheights = avail_mem / space_per_minheight; + /* If there doesn't seem to be enough space, try to get the minimum + * anyway. This allows a "stub" implementation of jpeg_mem_available(). + */ + if (max_minheights <= 0) + max_minheights = 1; + } + + /* Allocate the in-memory buffers and initialize backing store as needed. */ + + for (sptr = mem->virt_sarray_list; sptr != NULL; sptr = sptr->next) { + if (sptr->mem_buffer == NULL) { /* if not realized yet */ + minheights = ((long) sptr->rows_in_array - 1L) / sptr->maxaccess + 1L; + if (minheights <= max_minheights) { + /* This buffer fits in memory */ + sptr->rows_in_mem = sptr->rows_in_array; + } else { + /* It doesn't fit in memory, create backing store. */ + sptr->rows_in_mem = (JDIMENSION) (max_minheights * sptr->maxaccess); + jpeg_open_backing_store(cinfo, & sptr->b_s_info, + (long) sptr->rows_in_array * + (long) sptr->samplesperrow * + (long) SIZEOF(JSAMPLE)); + sptr->b_s_open = TRUE; + } + sptr->mem_buffer = alloc_sarray(cinfo, JPOOL_IMAGE, + sptr->samplesperrow, sptr->rows_in_mem); + sptr->rowsperchunk = mem->last_rowsperchunk; + sptr->cur_start_row = 0; + sptr->first_undef_row = 0; + sptr->dirty = FALSE; + } + } + + for (bptr = mem->virt_barray_list; bptr != NULL; bptr = bptr->next) { + if (bptr->mem_buffer == NULL) { /* if not realized yet */ + minheights = ((long) bptr->rows_in_array - 1L) / bptr->maxaccess + 1L; + if (minheights <= max_minheights) { + /* This buffer fits in memory */ + bptr->rows_in_mem = bptr->rows_in_array; + } else { + /* It doesn't fit in memory, create backing store. */ + bptr->rows_in_mem = (JDIMENSION) (max_minheights * bptr->maxaccess); + jpeg_open_backing_store(cinfo, & bptr->b_s_info, + (long) bptr->rows_in_array * + (long) bptr->blocksperrow * + (long) SIZEOF(JBLOCK)); + bptr->b_s_open = TRUE; + } + bptr->mem_buffer = alloc_barray(cinfo, JPOOL_IMAGE, + bptr->blocksperrow, bptr->rows_in_mem); + bptr->rowsperchunk = mem->last_rowsperchunk; + bptr->cur_start_row = 0; + bptr->first_undef_row = 0; + bptr->dirty = FALSE; + } + } +} + + +LOCAL(void) +do_sarray_io (j_common_ptr cinfo, jvirt_sarray_ptr ptr, boolean writing) +/* Do backing store read or write of a virtual sample array */ +{ + long bytesperrow, file_offset, byte_count, rows, thisrow, i; + + bytesperrow = (long) ptr->samplesperrow * SIZEOF(JSAMPLE); + file_offset = ptr->cur_start_row * bytesperrow; + /* Loop to read or write each allocation chunk in mem_buffer */ + for (i = 0; i < (long) ptr->rows_in_mem; i += ptr->rowsperchunk) { + /* One chunk, but check for short chunk at end of buffer */ + rows = MIN((long) ptr->rowsperchunk, (long) ptr->rows_in_mem - i); + /* Transfer no more than is currently defined */ + thisrow = (long) ptr->cur_start_row + i; + rows = MIN(rows, (long) ptr->first_undef_row - thisrow); + /* Transfer no more than fits in file */ + rows = MIN(rows, (long) ptr->rows_in_array - thisrow); + if (rows <= 0) /* this chunk might be past end of file! */ + break; + byte_count = rows * bytesperrow; + if (writing) + (*ptr->b_s_info.write_backing_store) (cinfo, & ptr->b_s_info, + (void FAR *) ptr->mem_buffer[i], + file_offset, byte_count); + else + (*ptr->b_s_info.read_backing_store) (cinfo, & ptr->b_s_info, + (void FAR *) ptr->mem_buffer[i], + file_offset, byte_count); + file_offset += byte_count; + } +} + + +LOCAL(void) +do_barray_io (j_common_ptr cinfo, jvirt_barray_ptr ptr, boolean writing) +/* Do backing store read or write of a virtual coefficient-block array */ +{ + long bytesperrow, file_offset, byte_count, rows, thisrow, i; + + bytesperrow = (long) ptr->blocksperrow * SIZEOF(JBLOCK); + file_offset = ptr->cur_start_row * bytesperrow; + /* Loop to read or write each allocation chunk in mem_buffer */ + for (i = 0; i < (long) ptr->rows_in_mem; i += ptr->rowsperchunk) { + /* One chunk, but check for short chunk at end of buffer */ + rows = MIN((long) ptr->rowsperchunk, (long) ptr->rows_in_mem - i); + /* Transfer no more than is currently defined */ + thisrow = (long) ptr->cur_start_row + i; + rows = MIN(rows, (long) ptr->first_undef_row - thisrow); + /* Transfer no more than fits in file */ + rows = MIN(rows, (long) ptr->rows_in_array - thisrow); + if (rows <= 0) /* this chunk might be past end of file! */ + break; + byte_count = rows * bytesperrow; + if (writing) + (*ptr->b_s_info.write_backing_store) (cinfo, & ptr->b_s_info, + (void FAR *) ptr->mem_buffer[i], + file_offset, byte_count); + else + (*ptr->b_s_info.read_backing_store) (cinfo, & ptr->b_s_info, + (void FAR *) ptr->mem_buffer[i], + file_offset, byte_count); + file_offset += byte_count; + } +} + + +METHODDEF(JSAMPARRAY) +access_virt_sarray (j_common_ptr cinfo, jvirt_sarray_ptr ptr, + JDIMENSION start_row, JDIMENSION num_rows, + boolean writable) +/* Access the part of a virtual sample array starting at start_row */ +/* and extending for num_rows rows. writable is true if */ +/* caller intends to modify the accessed area. */ +{ + JDIMENSION end_row = start_row + num_rows; + JDIMENSION undef_row; + + /* debugging check */ + if (end_row > ptr->rows_in_array || num_rows > ptr->maxaccess || + ptr->mem_buffer == NULL) + ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); + + /* Make the desired part of the virtual array accessible */ + if (start_row < ptr->cur_start_row || + end_row > ptr->cur_start_row+ptr->rows_in_mem) { + if (! ptr->b_s_open) + ERREXIT(cinfo, JERR_VIRTUAL_BUG); + /* Flush old buffer contents if necessary */ + if (ptr->dirty) { + do_sarray_io(cinfo, ptr, TRUE); + ptr->dirty = FALSE; + } + /* Decide what part of virtual array to access. + * Algorithm: if target address > current window, assume forward scan, + * load starting at target address. If target address < current window, + * assume backward scan, load so that target area is top of window. + * Note that when switching from forward write to forward read, will have + * start_row = 0, so the limiting case applies and we load from 0 anyway. + */ + if (start_row > ptr->cur_start_row) { + ptr->cur_start_row = start_row; + } else { + /* use long arithmetic here to avoid overflow & unsigned problems */ + long ltemp; + + ltemp = (long) end_row - (long) ptr->rows_in_mem; + if (ltemp < 0) + ltemp = 0; /* don't fall off front end of file */ + ptr->cur_start_row = (JDIMENSION) ltemp; + } + /* Read in the selected part of the array. + * During the initial write pass, we will do no actual read + * because the selected part is all undefined. + */ + do_sarray_io(cinfo, ptr, FALSE); + } + /* Ensure the accessed part of the array is defined; prezero if needed. + * To improve locality of access, we only prezero the part of the array + * that the caller is about to access, not the entire in-memory array. + */ + if (ptr->first_undef_row < end_row) { + if (ptr->first_undef_row < start_row) { + if (writable) /* writer skipped over a section of array */ + ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); + undef_row = start_row; /* but reader is allowed to read ahead */ + } else { + undef_row = ptr->first_undef_row; + } + if (writable) + ptr->first_undef_row = end_row; + if (ptr->pre_zero) { + size_t bytesperrow = (size_t) ptr->samplesperrow * SIZEOF(JSAMPLE); + undef_row -= ptr->cur_start_row; /* make indexes relative to buffer */ + end_row -= ptr->cur_start_row; + while (undef_row < end_row) { + jzero_far((void FAR *) ptr->mem_buffer[undef_row], bytesperrow); + undef_row++; + } + } else { + if (! writable) /* reader looking at undefined data */ + ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); + } + } + /* Flag the buffer dirty if caller will write in it */ + if (writable) + ptr->dirty = TRUE; + /* Return address of proper part of the buffer */ + return ptr->mem_buffer + (start_row - ptr->cur_start_row); +} + + +METHODDEF(JBLOCKARRAY) +access_virt_barray (j_common_ptr cinfo, jvirt_barray_ptr ptr, + JDIMENSION start_row, JDIMENSION num_rows, + boolean writable) +/* Access the part of a virtual block array starting at start_row */ +/* and extending for num_rows rows. writable is true if */ +/* caller intends to modify the accessed area. */ +{ + JDIMENSION end_row = start_row + num_rows; + JDIMENSION undef_row; + + /* debugging check */ + if (end_row > ptr->rows_in_array || num_rows > ptr->maxaccess || + ptr->mem_buffer == NULL) + ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); + + /* Make the desired part of the virtual array accessible */ + if (start_row < ptr->cur_start_row || + end_row > ptr->cur_start_row+ptr->rows_in_mem) { + if (! ptr->b_s_open) + ERREXIT(cinfo, JERR_VIRTUAL_BUG); + /* Flush old buffer contents if necessary */ + if (ptr->dirty) { + do_barray_io(cinfo, ptr, TRUE); + ptr->dirty = FALSE; + } + /* Decide what part of virtual array to access. + * Algorithm: if target address > current window, assume forward scan, + * load starting at target address. If target address < current window, + * assume backward scan, load so that target area is top of window. + * Note that when switching from forward write to forward read, will have + * start_row = 0, so the limiting case applies and we load from 0 anyway. + */ + if (start_row > ptr->cur_start_row) { + ptr->cur_start_row = start_row; + } else { + /* use long arithmetic here to avoid overflow & unsigned problems */ + long ltemp; + + ltemp = (long) end_row - (long) ptr->rows_in_mem; + if (ltemp < 0) + ltemp = 0; /* don't fall off front end of file */ + ptr->cur_start_row = (JDIMENSION) ltemp; + } + /* Read in the selected part of the array. + * During the initial write pass, we will do no actual read + * because the selected part is all undefined. + */ + do_barray_io(cinfo, ptr, FALSE); + } + /* Ensure the accessed part of the array is defined; prezero if needed. + * To improve locality of access, we only prezero the part of the array + * that the caller is about to access, not the entire in-memory array. + */ + if (ptr->first_undef_row < end_row) { + if (ptr->first_undef_row < start_row) { + if (writable) /* writer skipped over a section of array */ + ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); + undef_row = start_row; /* but reader is allowed to read ahead */ + } else { + undef_row = ptr->first_undef_row; + } + if (writable) + ptr->first_undef_row = end_row; + if (ptr->pre_zero) { + size_t bytesperrow = (size_t) ptr->blocksperrow * SIZEOF(JBLOCK); + undef_row -= ptr->cur_start_row; /* make indexes relative to buffer */ + end_row -= ptr->cur_start_row; + while (undef_row < end_row) { + jzero_far((void FAR *) ptr->mem_buffer[undef_row], bytesperrow); + undef_row++; + } + } else { + if (! writable) /* reader looking at undefined data */ + ERREXIT(cinfo, JERR_BAD_VIRTUAL_ACCESS); + } + } + /* Flag the buffer dirty if caller will write in it */ + if (writable) + ptr->dirty = TRUE; + /* Return address of proper part of the buffer */ + return ptr->mem_buffer + (start_row - ptr->cur_start_row); +} + + +/* + * Release all objects belonging to a specified pool. + */ + +METHODDEF(void) +free_pool (j_common_ptr cinfo, int pool_id) +{ + my_mem_ptr mem = (my_mem_ptr) cinfo->mem; + small_pool_ptr shdr_ptr; + large_pool_ptr lhdr_ptr; + size_t space_freed; + + if (pool_id < 0 || pool_id >= JPOOL_NUMPOOLS) + ERREXIT1(cinfo, JERR_BAD_POOL_ID, pool_id); /* safety check */ + +#ifdef MEM_STATS + if (cinfo->err->trace_level > 1) + print_mem_stats(cinfo, pool_id); /* print pool's memory usage statistics */ +#endif + + /* If freeing IMAGE pool, close any virtual arrays first */ + if (pool_id == JPOOL_IMAGE) { + jvirt_sarray_ptr sptr; + jvirt_barray_ptr bptr; + + for (sptr = mem->virt_sarray_list; sptr != NULL; sptr = sptr->next) { + if (sptr->b_s_open) { /* there may be no backing store */ + sptr->b_s_open = FALSE; /* prevent recursive close if error */ + (*sptr->b_s_info.close_backing_store) (cinfo, & sptr->b_s_info); + } + } + mem->virt_sarray_list = NULL; + for (bptr = mem->virt_barray_list; bptr != NULL; bptr = bptr->next) { + if (bptr->b_s_open) { /* there may be no backing store */ + bptr->b_s_open = FALSE; /* prevent recursive close if error */ + (*bptr->b_s_info.close_backing_store) (cinfo, & bptr->b_s_info); + } + } + mem->virt_barray_list = NULL; + } + + /* Release large objects */ + lhdr_ptr = mem->large_list[pool_id]; + mem->large_list[pool_id] = NULL; + + while (lhdr_ptr != NULL) { + large_pool_ptr next_lhdr_ptr = lhdr_ptr->hdr.next; + space_freed = lhdr_ptr->hdr.bytes_used + + lhdr_ptr->hdr.bytes_left + + SIZEOF(large_pool_hdr); + jpeg_free_large(cinfo, (void FAR *) lhdr_ptr, space_freed); + mem->total_space_allocated -= space_freed; + lhdr_ptr = next_lhdr_ptr; + } + + /* Release small objects */ + shdr_ptr = mem->small_list[pool_id]; + mem->small_list[pool_id] = NULL; + + while (shdr_ptr != NULL) { + small_pool_ptr next_shdr_ptr = shdr_ptr->hdr.next; + space_freed = shdr_ptr->hdr.bytes_used + + shdr_ptr->hdr.bytes_left + + SIZEOF(small_pool_hdr); + jpeg_free_small(cinfo, (void *) shdr_ptr, space_freed); + mem->total_space_allocated -= space_freed; + shdr_ptr = next_shdr_ptr; + } +} + + +/* + * Close up shop entirely. + * Note that this cannot be called unless cinfo->mem is non-NULL. + */ + +METHODDEF(void) +self_destruct (j_common_ptr cinfo) +{ + int pool; + + /* Close all backing store, release all memory. + * Releasing pools in reverse order might help avoid fragmentation + * with some (brain-damaged) malloc libraries. + */ + for (pool = JPOOL_NUMPOOLS-1; pool >= JPOOL_PERMANENT; pool--) { + free_pool(cinfo, pool); + } + + /* Release the memory manager control block too. */ + jpeg_free_small(cinfo, (void *) cinfo->mem, SIZEOF(my_memory_mgr)); + cinfo->mem = NULL; /* ensures I will be called only once */ + + jpeg_mem_term(cinfo); /* system-dependent cleanup */ +} + + +/* + * Memory manager initialization. + * When this is called, only the error manager pointer is valid in cinfo! + */ + +GLOBAL(void) +jinit_memory_mgr (j_common_ptr cinfo) +{ + my_mem_ptr mem; + long max_to_use; + int pool; + size_t test_mac; + + cinfo->mem = NULL; /* for safety if init fails */ + + /* Check for configuration errors. + * SIZEOF(ALIGN_TYPE) should be a power of 2; otherwise, it probably + * doesn't reflect any real hardware alignment requirement. + * The test is a little tricky: for X>0, X and X-1 have no one-bits + * in common if and only if X is a power of 2, ie has only one one-bit. + * Some compilers may give an "unreachable code" warning here; ignore it. + */ + if ((SIZEOF(ALIGN_TYPE) & (SIZEOF(ALIGN_TYPE)-1)) != 0) + ERREXIT(cinfo, JERR_BAD_ALIGN_TYPE); + /* MAX_ALLOC_CHUNK must be representable as type size_t, and must be + * a multiple of SIZEOF(ALIGN_TYPE). + * Again, an "unreachable code" warning may be ignored here. + * But a "constant too large" warning means you need to fix MAX_ALLOC_CHUNK. + */ + test_mac = (size_t) MAX_ALLOC_CHUNK; + if ((long) test_mac != MAX_ALLOC_CHUNK || + (MAX_ALLOC_CHUNK % SIZEOF(ALIGN_TYPE)) != 0) + ERREXIT(cinfo, JERR_BAD_ALLOC_CHUNK); + + max_to_use = jpeg_mem_init(cinfo); /* system-dependent initialization */ + + /* Attempt to allocate memory manager's control block */ + mem = (my_mem_ptr) jpeg_get_small(cinfo, SIZEOF(my_memory_mgr)); + + if (mem == NULL) { + jpeg_mem_term(cinfo); /* system-dependent cleanup */ + ERREXIT1(cinfo, JERR_OUT_OF_MEMORY, 0); + } + + /* OK, fill in the method pointers */ + mem->pub.alloc_small = alloc_small; + mem->pub.alloc_large = alloc_large; + mem->pub.alloc_sarray = alloc_sarray; + mem->pub.alloc_barray = alloc_barray; + mem->pub.request_virt_sarray = request_virt_sarray; + mem->pub.request_virt_barray = request_virt_barray; + mem->pub.realize_virt_arrays = realize_virt_arrays; + mem->pub.access_virt_sarray = access_virt_sarray; + mem->pub.access_virt_barray = access_virt_barray; + mem->pub.free_pool = free_pool; + mem->pub.self_destruct = self_destruct; + + /* Make MAX_ALLOC_CHUNK accessible to other modules */ + mem->pub.max_alloc_chunk = MAX_ALLOC_CHUNK; + + /* Initialize working state */ + mem->pub.max_memory_to_use = max_to_use; + + for (pool = JPOOL_NUMPOOLS-1; pool >= JPOOL_PERMANENT; pool--) { + mem->small_list[pool] = NULL; + mem->large_list[pool] = NULL; + } + mem->virt_sarray_list = NULL; + mem->virt_barray_list = NULL; + + mem->total_space_allocated = SIZEOF(my_memory_mgr); + + /* Declare ourselves open for business */ + cinfo->mem = & mem->pub; + + /* Check for an environment variable JPEGMEM; if found, override the + * default max_memory setting from jpeg_mem_init. Note that the + * surrounding application may again override this value. + * If your system doesn't support getenv(), define NO_GETENV to disable + * this feature. + */ +#ifndef NO_GETENV + { char * memenv; + + if ((memenv = getenv("JPEGMEM")) != NULL) { + char ch = 'x'; + + if (sscanf(memenv, "%ld%c", &max_to_use, &ch) > 0) { + if (ch == 'm' || ch == 'M') + max_to_use *= 1000L; + mem->pub.max_memory_to_use = max_to_use * 1000L; + } + } + } +#endif + +} diff --git a/test/monniaux/jpeg-6b/jmorecfg.h b/test/monniaux/jpeg-6b/jmorecfg.h index c274ca3a..54a7d1c4 100644 --- a/test/monniaux/jpeg-6b/jmorecfg.h +++ b/test/monniaux/jpeg-6b/jmorecfg.h @@ -262,9 +262,7 @@ typedef int boolean; #define DCT_ISLOW_SUPPORTED /* slow but accurate integer algorithm */ #define DCT_IFAST_SUPPORTED /* faster, less accurate integer method */ - -/* DM */ -#undef DCT_FLOAT_SUPPORTED /* floating-point: accurate, fast on fast HW */ +#define DCT_FLOAT_SUPPORTED /* floating-point: accurate, fast on fast HW */ /* Encoder capability options: */ diff --git a/test/monniaux/jpeg-6b/jpegtran.c b/test/monniaux/jpeg-6b/jpegtran.c index 719aaa73..20ef111b 100644 --- a/test/monniaux/jpeg-6b/jpegtran.c +++ b/test/monniaux/jpeg-6b/jpegtran.c @@ -1,504 +1,504 @@ -/* - * jpegtran.c - * - * Copyright (C) 1995-1997, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains a command-line user interface for JPEG transcoding. - * It is very similar to cjpeg.c, but provides lossless transcoding between - * different JPEG file formats. It also provides some lossless and sort-of- - * lossless transformations of JPEG data. - */ - -#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ -#include "transupp.h" /* Support routines for jpegtran */ -#include "jversion.h" /* for version message */ - -#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ -#ifdef __MWERKS__ -#include /* Metrowerks needs this */ -#include /* ... and this */ -#endif -#ifdef THINK_C -#include /* Think declares it here */ -#endif -#endif - - -/* - * Argument-parsing code. - * The switch parser is designed to be useful with DOS-style command line - * syntax, ie, intermixed switches and file names, where only the switches - * to the left of a given file name affect processing of that file. - * The main program in this file doesn't actually use this capability... - */ - - -static const char * progname; /* program name for error messages */ -static char * outfilename; /* for -outfile switch */ -static JCOPY_OPTION copyoption; /* -copy switch */ -static jpeg_transform_info transformoption; /* image transformation options */ - - -LOCAL(void) -usage (void) -/* complain about bad command line */ -{ - fprintf(stderr, "usage: %s [switches] ", progname); -#ifdef TWO_FILE_COMMANDLINE - fprintf(stderr, "inputfile outputfile\n"); -#else - fprintf(stderr, "[inputfile]\n"); -#endif - - fprintf(stderr, "Switches (names may be abbreviated):\n"); - fprintf(stderr, " -copy none Copy no extra markers from source file\n"); - fprintf(stderr, " -copy comments Copy only comment markers (default)\n"); - fprintf(stderr, " -copy all Copy all extra markers\n"); -#ifdef ENTROPY_OPT_SUPPORTED - fprintf(stderr, " -optimize Optimize Huffman table (smaller file, but slow compression)\n"); -#endif -#ifdef C_PROGRESSIVE_SUPPORTED - fprintf(stderr, " -progressive Create progressive JPEG file\n"); -#endif -#if TRANSFORMS_SUPPORTED - fprintf(stderr, "Switches for modifying the image:\n"); - fprintf(stderr, " -grayscale Reduce to grayscale (omit color data)\n"); - fprintf(stderr, " -flip [horizontal|vertical] Mirror image (left-right or top-bottom)\n"); - fprintf(stderr, " -rotate [90|180|270] Rotate image (degrees clockwise)\n"); - fprintf(stderr, " -transpose Transpose image\n"); - fprintf(stderr, " -transverse Transverse transpose image\n"); - fprintf(stderr, " -trim Drop non-transformable edge blocks\n"); -#endif /* TRANSFORMS_SUPPORTED */ - fprintf(stderr, "Switches for advanced users:\n"); - fprintf(stderr, " -restart N Set restart interval in rows, or in blocks with B\n"); - fprintf(stderr, " -maxmemory N Maximum memory to use (in kbytes)\n"); - fprintf(stderr, " -outfile name Specify name for output file\n"); - fprintf(stderr, " -verbose or -debug Emit debug output\n"); - fprintf(stderr, "Switches for wizards:\n"); -#ifdef C_ARITH_CODING_SUPPORTED - fprintf(stderr, " -arithmetic Use arithmetic coding\n"); -#endif -#ifdef C_MULTISCAN_FILES_SUPPORTED - fprintf(stderr, " -scans file Create multi-scan JPEG per script file\n"); -#endif - exit(EXIT_FAILURE); -} - - -LOCAL(void) -select_transform (JXFORM_CODE transform) -/* Silly little routine to detect multiple transform options, - * which we can't handle. - */ -{ -#if TRANSFORMS_SUPPORTED - if (transformoption.transform == JXFORM_NONE || - transformoption.transform == transform) { - transformoption.transform = transform; - } else { - fprintf(stderr, "%s: can only do one image transformation at a time\n", - progname); - usage(); - } -#else - fprintf(stderr, "%s: sorry, image transformation was not compiled\n", - progname); - exit(EXIT_FAILURE); -#endif -} - - -LOCAL(int) -parse_switches (j_compress_ptr cinfo, int argc, char **argv, - int last_file_arg_seen, boolean for_real) -/* Parse optional switches. - * Returns argv[] index of first file-name argument (== argc if none). - * Any file names with indexes <= last_file_arg_seen are ignored; - * they have presumably been processed in a previous iteration. - * (Pass 0 for last_file_arg_seen on the first or only iteration.) - * for_real is FALSE on the first (dummy) pass; we may skip any expensive - * processing. - */ -{ - int argn; - char * arg; - boolean simple_progressive; - char * scansarg = NULL; /* saves -scans parm if any */ - - /* Set up default JPEG parameters. */ - simple_progressive = FALSE; - outfilename = NULL; - copyoption = JCOPYOPT_DEFAULT; - transformoption.transform = JXFORM_NONE; - transformoption.trim = FALSE; - transformoption.force_grayscale = FALSE; - cinfo->err->trace_level = 0; - - /* Scan command line options, adjust parameters */ - - for (argn = 1; argn < argc; argn++) { - arg = argv[argn]; - if (*arg != '-') { - /* Not a switch, must be a file name argument */ - if (argn <= last_file_arg_seen) { - outfilename = NULL; /* -outfile applies to just one input file */ - continue; /* ignore this name if previously processed */ - } - break; /* else done parsing switches */ - } - arg++; /* advance past switch marker character */ - - if (keymatch(arg, "arithmetic", 1)) { - /* Use arithmetic coding. */ -#ifdef C_ARITH_CODING_SUPPORTED - cinfo->arith_code = TRUE; -#else - fprintf(stderr, "%s: sorry, arithmetic coding not supported\n", - progname); - exit(EXIT_FAILURE); -#endif - - } else if (keymatch(arg, "copy", 1)) { - /* Select which extra markers to copy. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (keymatch(argv[argn], "none", 1)) { - copyoption = JCOPYOPT_NONE; - } else if (keymatch(argv[argn], "comments", 1)) { - copyoption = JCOPYOPT_COMMENTS; - } else if (keymatch(argv[argn], "all", 1)) { - copyoption = JCOPYOPT_ALL; - } else - usage(); - - } else if (keymatch(arg, "debug", 1) || keymatch(arg, "verbose", 1)) { - /* Enable debug printouts. */ - /* On first -d, print version identification */ - static boolean printed_version = FALSE; - - if (! printed_version) { - fprintf(stderr, "Independent JPEG Group's JPEGTRAN, version %s\n%s\n", - JVERSION, JCOPYRIGHT); - printed_version = TRUE; - } - cinfo->err->trace_level++; - - } else if (keymatch(arg, "flip", 1)) { - /* Mirror left-right or top-bottom. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (keymatch(argv[argn], "horizontal", 1)) - select_transform(JXFORM_FLIP_H); - else if (keymatch(argv[argn], "vertical", 1)) - select_transform(JXFORM_FLIP_V); - else - usage(); - - } else if (keymatch(arg, "grayscale", 1) || keymatch(arg, "greyscale",1)) { - /* Force to grayscale. */ -#if TRANSFORMS_SUPPORTED - transformoption.force_grayscale = TRUE; -#else - select_transform(JXFORM_NONE); /* force an error */ -#endif - - } else if (keymatch(arg, "maxmemory", 3)) { - /* Maximum memory in Kb (or Mb with 'm'). */ - long lval; - char ch = 'x'; - - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) - usage(); - if (ch == 'm' || ch == 'M') - lval *= 1000L; - cinfo->mem->max_memory_to_use = lval * 1000L; - - } else if (keymatch(arg, "optimize", 1) || keymatch(arg, "optimise", 1)) { - /* Enable entropy parm optimization. */ -#ifdef ENTROPY_OPT_SUPPORTED - cinfo->optimize_coding = TRUE; -#else - fprintf(stderr, "%s: sorry, entropy optimization was not compiled\n", - progname); - exit(EXIT_FAILURE); -#endif - - } else if (keymatch(arg, "outfile", 4)) { - /* Set output file name. */ - if (++argn >= argc) /* advance to next argument */ - usage(); - outfilename = argv[argn]; /* save it away for later use */ - - } else if (keymatch(arg, "progressive", 1)) { - /* Select simple progressive mode. */ -#ifdef C_PROGRESSIVE_SUPPORTED - simple_progressive = TRUE; - /* We must postpone execution until num_components is known. */ -#else - fprintf(stderr, "%s: sorry, progressive output was not compiled\n", - progname); - exit(EXIT_FAILURE); -#endif - - } else if (keymatch(arg, "restart", 1)) { - /* Restart interval in MCU rows (or in MCUs with 'b'). */ - long lval; - char ch = 'x'; - - if (++argn >= argc) /* advance to next argument */ - usage(); - if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) - usage(); - if (lval < 0 || lval > 65535L) - usage(); - if (ch == 'b' || ch == 'B') { - cinfo->restart_interval = (unsigned int) lval; - cinfo->restart_in_rows = 0; /* else prior '-restart n' overrides me */ - } else { - cinfo->restart_in_rows = (int) lval; - /* restart_interval will be computed during startup */ - } - - } else if (keymatch(arg, "rotate", 2)) { - /* Rotate 90, 180, or 270 degrees (measured clockwise). */ - if (++argn >= argc) /* advance to next argument */ - usage(); - if (keymatch(argv[argn], "90", 2)) - select_transform(JXFORM_ROT_90); - else if (keymatch(argv[argn], "180", 3)) - select_transform(JXFORM_ROT_180); - else if (keymatch(argv[argn], "270", 3)) - select_transform(JXFORM_ROT_270); - else - usage(); - - } else if (keymatch(arg, "scans", 1)) { - /* Set scan script. */ -#ifdef C_MULTISCAN_FILES_SUPPORTED - if (++argn >= argc) /* advance to next argument */ - usage(); - scansarg = argv[argn]; - /* We must postpone reading the file in case -progressive appears. */ -#else - fprintf(stderr, "%s: sorry, multi-scan output was not compiled\n", - progname); - exit(EXIT_FAILURE); -#endif - - } else if (keymatch(arg, "transpose", 1)) { - /* Transpose (across UL-to-LR axis). */ - select_transform(JXFORM_TRANSPOSE); - - } else if (keymatch(arg, "transverse", 6)) { - /* Transverse transpose (across UR-to-LL axis). */ - select_transform(JXFORM_TRANSVERSE); - - } else if (keymatch(arg, "trim", 3)) { - /* Trim off any partial edge MCUs that the transform can't handle. */ - transformoption.trim = TRUE; - - } else { - usage(); /* bogus switch */ - } - } - - /* Post-switch-scanning cleanup */ - - if (for_real) { - -#ifdef C_PROGRESSIVE_SUPPORTED - if (simple_progressive) /* process -progressive; -scans can override */ - jpeg_simple_progression(cinfo); -#endif - -#ifdef C_MULTISCAN_FILES_SUPPORTED - if (scansarg != NULL) /* process -scans if it was present */ - if (! read_scan_script(cinfo, scansarg)) - usage(); -#endif - } - - return argn; /* return index of next arg (file name) */ -} - - -/* - * The main program. - */ - -int -main (int argc, char **argv) -{ - struct jpeg_decompress_struct srcinfo; - struct jpeg_compress_struct dstinfo; - struct jpeg_error_mgr jsrcerr, jdsterr; -#ifdef PROGRESS_REPORT - struct cdjpeg_progress_mgr progress; -#endif - jvirt_barray_ptr * src_coef_arrays; - jvirt_barray_ptr * dst_coef_arrays; - int file_index; - FILE * input_file; - FILE * output_file; - - /* On Mac, fetch a command line. */ -#ifdef USE_CCOMMAND - argc = ccommand(&argv); -#endif - - progname = argv[0]; - if (progname == NULL || progname[0] == 0) - progname = "jpegtran"; /* in case C library doesn't provide it */ - - /* Initialize the JPEG decompression object with default error handling. */ - srcinfo.err = jpeg_std_error(&jsrcerr); - jpeg_create_decompress(&srcinfo); - /* Initialize the JPEG compression object with default error handling. */ - dstinfo.err = jpeg_std_error(&jdsterr); - jpeg_create_compress(&dstinfo); - - /* Now safe to enable signal catcher. - * Note: we assume only the decompression object will have virtual arrays. - */ -#ifdef NEED_SIGNAL_CATCHER - enable_signal_catcher((j_common_ptr) &srcinfo); -#endif - - /* Scan command line to find file names. - * It is convenient to use just one switch-parsing routine, but the switch - * values read here are mostly ignored; we will rescan the switches after - * opening the input file. Also note that most of the switches affect the - * destination JPEG object, so we parse into that and then copy over what - * needs to affects the source too. - */ - - file_index = parse_switches(&dstinfo, argc, argv, 0, FALSE); - jsrcerr.trace_level = jdsterr.trace_level; - srcinfo.mem->max_memory_to_use = dstinfo.mem->max_memory_to_use; - -#ifdef TWO_FILE_COMMANDLINE - /* Must have either -outfile switch or explicit output file name */ - if (outfilename == NULL) { - if (file_index != argc-2) { - fprintf(stderr, "%s: must name one input and one output file\n", - progname); - usage(); - } - outfilename = argv[file_index+1]; - } else { - if (file_index != argc-1) { - fprintf(stderr, "%s: must name one input and one output file\n", - progname); - usage(); - } - } -#else - /* Unix style: expect zero or one file name */ - if (file_index < argc-1) { - fprintf(stderr, "%s: only one input file\n", progname); - usage(); - } -#endif /* TWO_FILE_COMMANDLINE */ - - /* Open the input file. */ - if (file_index < argc) { - if ((input_file = fopen(argv[file_index], READ_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, argv[file_index]); - exit(EXIT_FAILURE); - } - } else { - /* default input file is stdin */ - input_file = read_stdin(); - } - - /* Open the output file. */ - if (outfilename != NULL) { - if ((output_file = fopen(outfilename, WRITE_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, outfilename); - exit(EXIT_FAILURE); - } - } else { - /* default output file is stdout */ - output_file = write_stdout(); - } - -#ifdef PROGRESS_REPORT - start_progress_monitor((j_common_ptr) &dstinfo, &progress); -#endif - - /* Specify data source for decompression */ - jpeg_stdio_src(&srcinfo, input_file); - - /* Enable saving of extra markers that we want to copy */ - jcopy_markers_setup(&srcinfo, copyoption); - - /* Read file header */ - (void) jpeg_read_header(&srcinfo, TRUE); - - /* Any space needed by a transform option must be requested before - * jpeg_read_coefficients so that memory allocation will be done right. - */ -#if TRANSFORMS_SUPPORTED - jtransform_request_workspace(&srcinfo, &transformoption); -#endif - - /* Read source file as DCT coefficients */ - src_coef_arrays = jpeg_read_coefficients(&srcinfo); - - /* Initialize destination compression parameters from source values */ - jpeg_copy_critical_parameters(&srcinfo, &dstinfo); - - /* Adjust destination parameters if required by transform options; - * also find out which set of coefficient arrays will hold the output. - */ -#if TRANSFORMS_SUPPORTED - dst_coef_arrays = jtransform_adjust_parameters(&srcinfo, &dstinfo, - src_coef_arrays, - &transformoption); -#else - dst_coef_arrays = src_coef_arrays; -#endif - - /* Adjust default compression parameters by re-parsing the options */ - file_index = parse_switches(&dstinfo, argc, argv, 0, TRUE); - - /* Specify data destination for compression */ - jpeg_stdio_dest(&dstinfo, output_file); - - /* Start compressor (note no image data is actually written here) */ - jpeg_write_coefficients(&dstinfo, dst_coef_arrays); - - /* Copy to the output file any extra markers that we want to preserve */ - jcopy_markers_execute(&srcinfo, &dstinfo, copyoption); - - /* Execute image transformation, if any */ -#if TRANSFORMS_SUPPORTED - jtransform_execute_transformation(&srcinfo, &dstinfo, - src_coef_arrays, - &transformoption); -#endif - - /* Finish compression and release memory */ - jpeg_finish_compress(&dstinfo); - jpeg_destroy_compress(&dstinfo); - (void) jpeg_finish_decompress(&srcinfo); - jpeg_destroy_decompress(&srcinfo); - - /* Close files, if we opened them */ - if (input_file != stdin) - fclose(input_file); - if (output_file != stdout) - fclose(output_file); - -#ifdef PROGRESS_REPORT - end_progress_monitor((j_common_ptr) &dstinfo); -#endif - - /* All done. */ - exit(jsrcerr.num_warnings + jdsterr.num_warnings ?EXIT_WARNING:EXIT_SUCCESS); - return 0; /* suppress no-return-value warnings */ -} +/* + * jpegtran.c + * + * Copyright (C) 1995-1997, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains a command-line user interface for JPEG transcoding. + * It is very similar to cjpeg.c, but provides lossless transcoding between + * different JPEG file formats. It also provides some lossless and sort-of- + * lossless transformations of JPEG data. + */ + +#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ +#include "transupp.h" /* Support routines for jpegtran */ +#include "jversion.h" /* for version message */ + +#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ +#ifdef __MWERKS__ +#include /* Metrowerks needs this */ +#include /* ... and this */ +#endif +#ifdef THINK_C +#include /* Think declares it here */ +#endif +#endif + + +/* + * Argument-parsing code. + * The switch parser is designed to be useful with DOS-style command line + * syntax, ie, intermixed switches and file names, where only the switches + * to the left of a given file name affect processing of that file. + * The main program in this file doesn't actually use this capability... + */ + + +static const char * progname; /* program name for error messages */ +static char * outfilename; /* for -outfile switch */ +static JCOPY_OPTION copyoption; /* -copy switch */ +static jpeg_transform_info transformoption; /* image transformation options */ + + +LOCAL(void) +usage (void) +/* complain about bad command line */ +{ + fprintf(stderr, "usage: %s [switches] ", progname); +#ifdef TWO_FILE_COMMANDLINE + fprintf(stderr, "inputfile outputfile\n"); +#else + fprintf(stderr, "[inputfile]\n"); +#endif + + fprintf(stderr, "Switches (names may be abbreviated):\n"); + fprintf(stderr, " -copy none Copy no extra markers from source file\n"); + fprintf(stderr, " -copy comments Copy only comment markers (default)\n"); + fprintf(stderr, " -copy all Copy all extra markers\n"); +#ifdef ENTROPY_OPT_SUPPORTED + fprintf(stderr, " -optimize Optimize Huffman table (smaller file, but slow compression)\n"); +#endif +#ifdef C_PROGRESSIVE_SUPPORTED + fprintf(stderr, " -progressive Create progressive JPEG file\n"); +#endif +#if TRANSFORMS_SUPPORTED + fprintf(stderr, "Switches for modifying the image:\n"); + fprintf(stderr, " -grayscale Reduce to grayscale (omit color data)\n"); + fprintf(stderr, " -flip [horizontal|vertical] Mirror image (left-right or top-bottom)\n"); + fprintf(stderr, " -rotate [90|180|270] Rotate image (degrees clockwise)\n"); + fprintf(stderr, " -transpose Transpose image\n"); + fprintf(stderr, " -transverse Transverse transpose image\n"); + fprintf(stderr, " -trim Drop non-transformable edge blocks\n"); +#endif /* TRANSFORMS_SUPPORTED */ + fprintf(stderr, "Switches for advanced users:\n"); + fprintf(stderr, " -restart N Set restart interval in rows, or in blocks with B\n"); + fprintf(stderr, " -maxmemory N Maximum memory to use (in kbytes)\n"); + fprintf(stderr, " -outfile name Specify name for output file\n"); + fprintf(stderr, " -verbose or -debug Emit debug output\n"); + fprintf(stderr, "Switches for wizards:\n"); +#ifdef C_ARITH_CODING_SUPPORTED + fprintf(stderr, " -arithmetic Use arithmetic coding\n"); +#endif +#ifdef C_MULTISCAN_FILES_SUPPORTED + fprintf(stderr, " -scans file Create multi-scan JPEG per script file\n"); +#endif + exit(EXIT_FAILURE); +} + + +LOCAL(void) +select_transform (JXFORM_CODE transform) +/* Silly little routine to detect multiple transform options, + * which we can't handle. + */ +{ +#if TRANSFORMS_SUPPORTED + if (transformoption.transform == JXFORM_NONE || + transformoption.transform == transform) { + transformoption.transform = transform; + } else { + fprintf(stderr, "%s: can only do one image transformation at a time\n", + progname); + usage(); + } +#else + fprintf(stderr, "%s: sorry, image transformation was not compiled\n", + progname); + exit(EXIT_FAILURE); +#endif +} + + +LOCAL(int) +parse_switches (j_compress_ptr cinfo, int argc, char **argv, + int last_file_arg_seen, boolean for_real) +/* Parse optional switches. + * Returns argv[] index of first file-name argument (== argc if none). + * Any file names with indexes <= last_file_arg_seen are ignored; + * they have presumably been processed in a previous iteration. + * (Pass 0 for last_file_arg_seen on the first or only iteration.) + * for_real is FALSE on the first (dummy) pass; we may skip any expensive + * processing. + */ +{ + int argn; + char * arg; + boolean simple_progressive; + char * scansarg = NULL; /* saves -scans parm if any */ + + /* Set up default JPEG parameters. */ + simple_progressive = FALSE; + outfilename = NULL; + copyoption = JCOPYOPT_DEFAULT; + transformoption.transform = JXFORM_NONE; + transformoption.trim = FALSE; + transformoption.force_grayscale = FALSE; + cinfo->err->trace_level = 0; + + /* Scan command line options, adjust parameters */ + + for (argn = 1; argn < argc; argn++) { + arg = argv[argn]; + if (*arg != '-') { + /* Not a switch, must be a file name argument */ + if (argn <= last_file_arg_seen) { + outfilename = NULL; /* -outfile applies to just one input file */ + continue; /* ignore this name if previously processed */ + } + break; /* else done parsing switches */ + } + arg++; /* advance past switch marker character */ + + if (keymatch(arg, "arithmetic", 1)) { + /* Use arithmetic coding. */ +#ifdef C_ARITH_CODING_SUPPORTED + cinfo->arith_code = TRUE; +#else + fprintf(stderr, "%s: sorry, arithmetic coding not supported\n", + progname); + exit(EXIT_FAILURE); +#endif + + } else if (keymatch(arg, "copy", 1)) { + /* Select which extra markers to copy. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (keymatch(argv[argn], "none", 1)) { + copyoption = JCOPYOPT_NONE; + } else if (keymatch(argv[argn], "comments", 1)) { + copyoption = JCOPYOPT_COMMENTS; + } else if (keymatch(argv[argn], "all", 1)) { + copyoption = JCOPYOPT_ALL; + } else + usage(); + + } else if (keymatch(arg, "debug", 1) || keymatch(arg, "verbose", 1)) { + /* Enable debug printouts. */ + /* On first -d, print version identification */ + static boolean printed_version = FALSE; + + if (! printed_version) { + fprintf(stderr, "Independent JPEG Group's JPEGTRAN, version %s\n%s\n", + JVERSION, JCOPYRIGHT); + printed_version = TRUE; + } + cinfo->err->trace_level++; + + } else if (keymatch(arg, "flip", 1)) { + /* Mirror left-right or top-bottom. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (keymatch(argv[argn], "horizontal", 1)) + select_transform(JXFORM_FLIP_H); + else if (keymatch(argv[argn], "vertical", 1)) + select_transform(JXFORM_FLIP_V); + else + usage(); + + } else if (keymatch(arg, "grayscale", 1) || keymatch(arg, "greyscale",1)) { + /* Force to grayscale. */ +#if TRANSFORMS_SUPPORTED + transformoption.force_grayscale = TRUE; +#else + select_transform(JXFORM_NONE); /* force an error */ +#endif + + } else if (keymatch(arg, "maxmemory", 3)) { + /* Maximum memory in Kb (or Mb with 'm'). */ + long lval; + char ch = 'x'; + + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) + usage(); + if (ch == 'm' || ch == 'M') + lval *= 1000L; + cinfo->mem->max_memory_to_use = lval * 1000L; + + } else if (keymatch(arg, "optimize", 1) || keymatch(arg, "optimise", 1)) { + /* Enable entropy parm optimization. */ +#ifdef ENTROPY_OPT_SUPPORTED + cinfo->optimize_coding = TRUE; +#else + fprintf(stderr, "%s: sorry, entropy optimization was not compiled\n", + progname); + exit(EXIT_FAILURE); +#endif + + } else if (keymatch(arg, "outfile", 4)) { + /* Set output file name. */ + if (++argn >= argc) /* advance to next argument */ + usage(); + outfilename = argv[argn]; /* save it away for later use */ + + } else if (keymatch(arg, "progressive", 1)) { + /* Select simple progressive mode. */ +#ifdef C_PROGRESSIVE_SUPPORTED + simple_progressive = TRUE; + /* We must postpone execution until num_components is known. */ +#else + fprintf(stderr, "%s: sorry, progressive output was not compiled\n", + progname); + exit(EXIT_FAILURE); +#endif + + } else if (keymatch(arg, "restart", 1)) { + /* Restart interval in MCU rows (or in MCUs with 'b'). */ + long lval; + char ch = 'x'; + + if (++argn >= argc) /* advance to next argument */ + usage(); + if (sscanf(argv[argn], "%ld%c", &lval, &ch) < 1) + usage(); + if (lval < 0 || lval > 65535L) + usage(); + if (ch == 'b' || ch == 'B') { + cinfo->restart_interval = (unsigned int) lval; + cinfo->restart_in_rows = 0; /* else prior '-restart n' overrides me */ + } else { + cinfo->restart_in_rows = (int) lval; + /* restart_interval will be computed during startup */ + } + + } else if (keymatch(arg, "rotate", 2)) { + /* Rotate 90, 180, or 270 degrees (measured clockwise). */ + if (++argn >= argc) /* advance to next argument */ + usage(); + if (keymatch(argv[argn], "90", 2)) + select_transform(JXFORM_ROT_90); + else if (keymatch(argv[argn], "180", 3)) + select_transform(JXFORM_ROT_180); + else if (keymatch(argv[argn], "270", 3)) + select_transform(JXFORM_ROT_270); + else + usage(); + + } else if (keymatch(arg, "scans", 1)) { + /* Set scan script. */ +#ifdef C_MULTISCAN_FILES_SUPPORTED + if (++argn >= argc) /* advance to next argument */ + usage(); + scansarg = argv[argn]; + /* We must postpone reading the file in case -progressive appears. */ +#else + fprintf(stderr, "%s: sorry, multi-scan output was not compiled\n", + progname); + exit(EXIT_FAILURE); +#endif + + } else if (keymatch(arg, "transpose", 1)) { + /* Transpose (across UL-to-LR axis). */ + select_transform(JXFORM_TRANSPOSE); + + } else if (keymatch(arg, "transverse", 6)) { + /* Transverse transpose (across UR-to-LL axis). */ + select_transform(JXFORM_TRANSVERSE); + + } else if (keymatch(arg, "trim", 3)) { + /* Trim off any partial edge MCUs that the transform can't handle. */ + transformoption.trim = TRUE; + + } else { + usage(); /* bogus switch */ + } + } + + /* Post-switch-scanning cleanup */ + + if (for_real) { + +#ifdef C_PROGRESSIVE_SUPPORTED + if (simple_progressive) /* process -progressive; -scans can override */ + jpeg_simple_progression(cinfo); +#endif + +#ifdef C_MULTISCAN_FILES_SUPPORTED + if (scansarg != NULL) /* process -scans if it was present */ + if (! read_scan_script(cinfo, scansarg)) + usage(); +#endif + } + + return argn; /* return index of next arg (file name) */ +} + + +/* + * The main program. + */ + +int +main (int argc, char **argv) +{ + struct jpeg_decompress_struct srcinfo; + struct jpeg_compress_struct dstinfo; + struct jpeg_error_mgr jsrcerr, jdsterr; +#ifdef PROGRESS_REPORT + struct cdjpeg_progress_mgr progress; +#endif + jvirt_barray_ptr * src_coef_arrays; + jvirt_barray_ptr * dst_coef_arrays; + int file_index; + FILE * input_file; + FILE * output_file; + + /* On Mac, fetch a command line. */ +#ifdef USE_CCOMMAND + argc = ccommand(&argv); +#endif + + progname = argv[0]; + if (progname == NULL || progname[0] == 0) + progname = "jpegtran"; /* in case C library doesn't provide it */ + + /* Initialize the JPEG decompression object with default error handling. */ + srcinfo.err = jpeg_std_error(&jsrcerr); + jpeg_create_decompress(&srcinfo); + /* Initialize the JPEG compression object with default error handling. */ + dstinfo.err = jpeg_std_error(&jdsterr); + jpeg_create_compress(&dstinfo); + + /* Now safe to enable signal catcher. + * Note: we assume only the decompression object will have virtual arrays. + */ +#ifdef NEED_SIGNAL_CATCHER + enable_signal_catcher((j_common_ptr) &srcinfo); +#endif + + /* Scan command line to find file names. + * It is convenient to use just one switch-parsing routine, but the switch + * values read here are mostly ignored; we will rescan the switches after + * opening the input file. Also note that most of the switches affect the + * destination JPEG object, so we parse into that and then copy over what + * needs to affects the source too. + */ + + file_index = parse_switches(&dstinfo, argc, argv, 0, FALSE); + jsrcerr.trace_level = jdsterr.trace_level; + srcinfo.mem->max_memory_to_use = dstinfo.mem->max_memory_to_use; + +#ifdef TWO_FILE_COMMANDLINE + /* Must have either -outfile switch or explicit output file name */ + if (outfilename == NULL) { + if (file_index != argc-2) { + fprintf(stderr, "%s: must name one input and one output file\n", + progname); + usage(); + } + outfilename = argv[file_index+1]; + } else { + if (file_index != argc-1) { + fprintf(stderr, "%s: must name one input and one output file\n", + progname); + usage(); + } + } +#else + /* Unix style: expect zero or one file name */ + if (file_index < argc-1) { + fprintf(stderr, "%s: only one input file\n", progname); + usage(); + } +#endif /* TWO_FILE_COMMANDLINE */ + + /* Open the input file. */ + if (file_index < argc) { + if ((input_file = fopen(argv[file_index], READ_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, argv[file_index]); + exit(EXIT_FAILURE); + } + } else { + /* default input file is stdin */ + input_file = read_stdin(); + } + + /* Open the output file. */ + if (outfilename != NULL) { + if ((output_file = fopen(outfilename, WRITE_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, outfilename); + exit(EXIT_FAILURE); + } + } else { + /* default output file is stdout */ + output_file = write_stdout(); + } + +#ifdef PROGRESS_REPORT + start_progress_monitor((j_common_ptr) &dstinfo, &progress); +#endif + + /* Specify data source for decompression */ + jpeg_stdio_src(&srcinfo, input_file); + + /* Enable saving of extra markers that we want to copy */ + jcopy_markers_setup(&srcinfo, copyoption); + + /* Read file header */ + (void) jpeg_read_header(&srcinfo, TRUE); + + /* Any space needed by a transform option must be requested before + * jpeg_read_coefficients so that memory allocation will be done right. + */ +#if TRANSFORMS_SUPPORTED + jtransform_request_workspace(&srcinfo, &transformoption); +#endif + + /* Read source file as DCT coefficients */ + src_coef_arrays = jpeg_read_coefficients(&srcinfo); + + /* Initialize destination compression parameters from source values */ + jpeg_copy_critical_parameters(&srcinfo, &dstinfo); + + /* Adjust destination parameters if required by transform options; + * also find out which set of coefficient arrays will hold the output. + */ +#if TRANSFORMS_SUPPORTED + dst_coef_arrays = jtransform_adjust_parameters(&srcinfo, &dstinfo, + src_coef_arrays, + &transformoption); +#else + dst_coef_arrays = src_coef_arrays; +#endif + + /* Adjust default compression parameters by re-parsing the options */ + file_index = parse_switches(&dstinfo, argc, argv, 0, TRUE); + + /* Specify data destination for compression */ + jpeg_stdio_dest(&dstinfo, output_file); + + /* Start compressor (note no image data is actually written here) */ + jpeg_write_coefficients(&dstinfo, dst_coef_arrays); + + /* Copy to the output file any extra markers that we want to preserve */ + jcopy_markers_execute(&srcinfo, &dstinfo, copyoption); + + /* Execute image transformation, if any */ +#if TRANSFORMS_SUPPORTED + jtransform_execute_transformation(&srcinfo, &dstinfo, + src_coef_arrays, + &transformoption); +#endif + + /* Finish compression and release memory */ + jpeg_finish_compress(&dstinfo); + jpeg_destroy_compress(&dstinfo); + (void) jpeg_finish_decompress(&srcinfo); + jpeg_destroy_decompress(&srcinfo); + + /* Close files, if we opened them */ + if (input_file != stdin) + fclose(input_file); + if (output_file != stdout) + fclose(output_file); + +#ifdef PROGRESS_REPORT + end_progress_monitor((j_common_ptr) &dstinfo); +#endif + + /* All done. */ + exit(jsrcerr.num_warnings + jdsterr.num_warnings ?EXIT_WARNING:EXIT_SUCCESS); + return 0; /* suppress no-return-value warnings */ +} diff --git a/test/monniaux/jpeg-6b/jquant1.c b/test/monniaux/jpeg-6b/jquant1.c index 698e4846..b2f96aa1 100644 --- a/test/monniaux/jpeg-6b/jquant1.c +++ b/test/monniaux/jpeg-6b/jquant1.c @@ -227,7 +227,7 @@ select_ncolors (j_decompress_ptr cinfo, int Ncolors[]) for (i = 0; i < nc; i++) { j = (cinfo->out_color_space == JCS_RGB ? RGB_order[i] : i); /* calculate new total_colors if Ncolors[j] is incremented */ - temp = INT_DIV(total_colors, Ncolors[j]); + temp = total_colors / Ncolors[j]; temp *= Ncolors[j]+1; /* done in long arith to avoid oflo */ if (temp > (long) max_colors) break; /* won't fit, done with this pass */ @@ -251,7 +251,7 @@ output_value (j_decompress_ptr cinfo, int ci, int j, int maxj) * (Forcing the upper and lower values to the limits ensures that * dithering can't produce a color outside the selected gamut.) */ - return (int) INT_DIV(((INT32) j * MAXJSAMPLE + maxj/2), maxj); + return (int) (((INT32) j * MAXJSAMPLE + maxj/2) / maxj); } @@ -261,7 +261,7 @@ largest_input_value (j_decompress_ptr cinfo, int ci, int j, int maxj) /* Must have largest(j=0) >= 0, and largest(j=maxj) >= MAXJSAMPLE */ { /* Breakpoints are halfway between values returned by output_value */ - return (int) INT_DIV(((INT32) (2*j + 1) * MAXJSAMPLE + maxj), (2*maxj)); + return (int) (((INT32) (2*j + 1) * MAXJSAMPLE + maxj) / (2*maxj)); } @@ -303,7 +303,7 @@ create_colormap (j_decompress_ptr cinfo) for (i = 0; i < cinfo->out_color_components; i++) { /* fill in colormap entries for i'th color component */ nci = cquantize->Ncolors[i]; /* # of distinct values for this color */ - blksize = INT_DIV(blkdist, nci); + blksize = blkdist / nci; for (j = 0; j < nci; j++) { /* Compute j'th output value (out of nci) for component */ val = output_value(cinfo, i, j, nci-1); @@ -360,7 +360,7 @@ create_colorindex (j_decompress_ptr cinfo) for (i = 0; i < cinfo->out_color_components; i++) { /* fill in colorindex entries for i'th color component */ nci = cquantize->Ncolors[i]; /* # of distinct values for this color */ - blksize = INT_DIV(blksize, nci); + blksize = blksize / nci; /* adjust colorindex pointers to provide padding at negative indexes. */ if (pad) @@ -415,7 +415,7 @@ make_odither_array (j_decompress_ptr cinfo, int ncolors) /* Ensure round towards zero despite C's lack of consistency * about rounding negative values in integer division... */ - odither[j][k] = (int) (num<0 ? -INT_DIV((-num),den) : INT_DIV(num,den)); + odither[j][k] = (int) (num<0 ? -((-num)/den) : num/den); } } return odither; diff --git a/test/monniaux/jpeg-6b/jquant2.c b/test/monniaux/jpeg-6b/jquant2.c index 6eb3bb17..af601e33 100644 --- a/test/monniaux/jpeg-6b/jquant2.c +++ b/test/monniaux/jpeg-6b/jquant2.c @@ -529,9 +529,9 @@ compute_color (j_decompress_ptr cinfo, boxptr boxp, int icolor) } } - cinfo->colormap[0][icolor] = (JSAMPLE) INT_DIV((c0total + (total>>1)), total); - cinfo->colormap[1][icolor] = (JSAMPLE) INT_DIV((c1total + (total>>1)), total); - cinfo->colormap[2][icolor] = (JSAMPLE) INT_DIV((c2total + (total>>1)), total); + cinfo->colormap[0][icolor] = (JSAMPLE) ((c0total + (total>>1)) / total); + cinfo->colormap[1][icolor] = (JSAMPLE) ((c1total + (total>>1)) / total); + cinfo->colormap[2][icolor] = (JSAMPLE) ((c2total + (total>>1)) / total); } diff --git a/test/monniaux/jpeg-6b/rdbmp.c b/test/monniaux/jpeg-6b/rdbmp.c index 29b1d484..b05fe2ac 100644 --- a/test/monniaux/jpeg-6b/rdbmp.c +++ b/test/monniaux/jpeg-6b/rdbmp.c @@ -229,8 +229,7 @@ preload_image (j_compress_ptr cinfo, cjpeg_source_ptr sinfo) source->source_row = cinfo->image_height; /* And read the first row */ - int ret= (*source->pub.get_pixel_rows) (cinfo, sinfo); - return ret; + return (*source->pub.get_pixel_rows) (cinfo, sinfo); } @@ -337,8 +336,8 @@ start_input_bmp (j_compress_ptr cinfo, cjpeg_source_ptr sinfo) if (biXPelsPerMeter > 0 && biYPelsPerMeter > 0) { /* Set JFIF density parameters from the BMP data */ - cinfo->X_density = (UINT16) INT_DIV(biXPelsPerMeter, 100); /* 100 cm per meter */ - cinfo->Y_density = (UINT16) INT_DIV(biYPelsPerMeter, 100); + cinfo->X_density = (UINT16) (biXPelsPerMeter/100); /* 100 cm per meter */ + cinfo->Y_density = (UINT16) (biYPelsPerMeter/100); cinfo->density_unit = 2; /* dots/cm */ } break; diff --git a/test/monniaux/jpeg-6b/rdgif.c b/test/monniaux/jpeg-6b/rdgif.c index b0757e71..b27c1675 100644 --- a/test/monniaux/jpeg-6b/rdgif.c +++ b/test/monniaux/jpeg-6b/rdgif.c @@ -1,38 +1,38 @@ -/* - * rdgif.c - * - * Copyright (C) 1991-1997, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains routines to read input images in GIF format. - * - ***************************************************************************** - * NOTE: to avoid entanglements with Unisys' patent on LZW compression, * - * the ability to read GIF files has been removed from the IJG distribution. * - * Sorry about that. * - ***************************************************************************** - * - * We are required to state that - * "The Graphics Interchange Format(c) is the Copyright property of - * CompuServe Incorporated. GIF(sm) is a Service Mark property of - * CompuServe Incorporated." - */ - -#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ - -#ifdef GIF_SUPPORTED - -/* - * The module selection routine for GIF format input. - */ - -GLOBAL(cjpeg_source_ptr) -jinit_read_gif (j_compress_ptr cinfo) -{ - fprintf(stderr, "GIF input is unsupported for legal reasons. Sorry.\n"); - exit(EXIT_FAILURE); - return NULL; /* keep compiler happy */ -} - -#endif /* GIF_SUPPORTED */ +/* + * rdgif.c + * + * Copyright (C) 1991-1997, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains routines to read input images in GIF format. + * + ***************************************************************************** + * NOTE: to avoid entanglements with Unisys' patent on LZW compression, * + * the ability to read GIF files has been removed from the IJG distribution. * + * Sorry about that. * + ***************************************************************************** + * + * We are required to state that + * "The Graphics Interchange Format(c) is the Copyright property of + * CompuServe Incorporated. GIF(sm) is a Service Mark property of + * CompuServe Incorporated." + */ + +#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ + +#ifdef GIF_SUPPORTED + +/* + * The module selection routine for GIF format input. + */ + +GLOBAL(cjpeg_source_ptr) +jinit_read_gif (j_compress_ptr cinfo) +{ + fprintf(stderr, "GIF input is unsupported for legal reasons. Sorry.\n"); + exit(EXIT_FAILURE); + return NULL; /* keep compiler happy */ +} + +#endif /* GIF_SUPPORTED */ diff --git a/test/monniaux/jpeg-6b/rdjpgcom.c b/test/monniaux/jpeg-6b/rdjpgcom.c index c6cb47ce..ffe6fc62 100644 --- a/test/monniaux/jpeg-6b/rdjpgcom.c +++ b/test/monniaux/jpeg-6b/rdjpgcom.c @@ -1,513 +1,496 @@ -/* - * rdjpgcom.c - * - * Copyright (C) 1994-1997, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains a very simple stand-alone application that displays - * the text in COM (comment) markers in a JFIF file. - * This may be useful as an example of the minimum logic needed to parse - * JPEG markers. - */ - -#define JPEG_CJPEG_DJPEG /* to get the command-line config symbols */ -#include "jinclude.h" /* get auto-config symbols, */ - -#include /* to declare isupper(), tolower() */ -#ifdef USE_SETMODE -#include /* to declare setmode()'s parameter macros */ -/* If you have setmode() but not , just delete this line: */ -#include /* to declare setmode() */ -#endif - -#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ -#ifdef __MWERKS__ -#include /* Metrowerks needs this */ -#include /* ... and this */ -#endif -#ifdef THINK_C -#include /* Think declares it here */ -#endif -#endif - -#ifdef DONT_USE_B_MODE /* define mode parameters for fopen() */ -#define READ_BINARY "r" -#else -#ifdef VMS /* VMS is very nonstandard */ -#define READ_BINARY "rb", "ctx=stm" -#else /* standard ANSI-compliant case */ -#define READ_BINARY "rb" -#endif -#endif - -#ifndef EXIT_FAILURE /* define exit() codes if not provided */ -#define EXIT_FAILURE 1 -#endif -#ifndef EXIT_SUCCESS -#ifdef VMS -#define EXIT_SUCCESS 1 /* VMS is very nonstandard */ -#else -#define EXIT_SUCCESS 0 -#endif -#endif - - -/* - * These macros are used to read the input file. - * To reuse this code in another application, you might need to change these. - */ - -static FILE * infile; /* input JPEG file */ - -/* Return next input byte, or EOF if no more */ -#define NEXTBYTE() getc(infile) - - -/* Error exit handler */ -#define ERREXIT(msg) (fprintf(stderr, "%s\n", msg), exit(EXIT_FAILURE)) - - -/* Read one byte, testing for EOF */ -static int -read_1_byte (void) -{ - int c; - - c = NEXTBYTE(); - if (c == EOF) - ERREXIT("Premature EOF in JPEG file"); - return c; -} - -/* Read 2 bytes, convert to unsigned int */ -/* All 2-byte quantities in JPEG markers are MSB first */ -static unsigned int -read_2_bytes (void) -{ - int c1, c2; - - c1 = NEXTBYTE(); - if (c1 == EOF) - ERREXIT("Premature EOF in JPEG file"); - c2 = NEXTBYTE(); - if (c2 == EOF) - ERREXIT("Premature EOF in JPEG file"); - return (((unsigned int) c1) << 8) + ((unsigned int) c2); -} - - -/* - * JPEG markers consist of one or more 0xFF bytes, followed by a marker - * code byte (which is not an FF). Here are the marker codes of interest - * in this program. (See jdmarker.c for a more complete list.) - */ - -#define M_SOF0 0xC0 /* Start Of Frame N */ -#define M_SOF1 0xC1 /* N indicates which compression process */ -#define M_SOF2 0xC2 /* Only SOF0-SOF2 are now in common use */ -#define M_SOF3 0xC3 -#define M_SOF5 0xC5 /* NB: codes C4 and CC are NOT SOF markers */ -#define M_SOF6 0xC6 -#define M_SOF7 0xC7 -#define M_SOF9 0xC9 -#define M_SOF10 0xCA -#define M_SOF11 0xCB -#define M_SOF13 0xCD -#define M_SOF14 0xCE -#define M_SOF15 0xCF -#define M_SOI 0xD8 /* Start Of Image (beginning of datastream) */ -#define M_EOI 0xD9 /* End Of Image (end of datastream) */ -#define M_SOS 0xDA /* Start Of Scan (begins compressed data) */ -#define M_APP0 0xE0 /* Application-specific marker, type N */ -#define M_APP12 0xEC /* (we don't bother to list all 16 APPn's) */ -#define M_COM 0xFE /* COMment */ - - -/* - * Find the next JPEG marker and return its marker code. - * We expect at least one FF byte, possibly more if the compressor used FFs - * to pad the file. - * There could also be non-FF garbage between markers. The treatment of such - * garbage is unspecified; we choose to skip over it but emit a warning msg. - * NB: this routine must not be used after seeing SOS marker, since it will - * not deal correctly with FF/00 sequences in the compressed image data... - */ - -static int -next_marker (void) -{ - int c; - int discarded_bytes = 0; - - /* Find 0xFF byte; count and skip any non-FFs. */ - c = read_1_byte(); - while (c != 0xFF) { - discarded_bytes++; - c = read_1_byte(); - } - /* Get marker code byte, swallowing any duplicate FF bytes. Extra FFs - * are legal as pad bytes, so don't count them in discarded_bytes. - */ - do { - c = read_1_byte(); - } while (c == 0xFF); - - if (discarded_bytes != 0) { - fprintf(stderr, "Warning: garbage data found in JPEG file\n"); - } - - return c; -} - - -/* - * Read the initial marker, which should be SOI. - * For a JFIF file, the first two bytes of the file should be literally - * 0xFF M_SOI. To be more general, we could use next_marker, but if the - * input file weren't actually JPEG at all, next_marker might read the whole - * file and then return a misleading error message... - */ - -static int -first_marker (void) -{ - int c1, c2; - - c1 = NEXTBYTE(); - c2 = NEXTBYTE(); - if (c1 != 0xFF || c2 != M_SOI) - ERREXIT("Not a JPEG file"); - return c2; -} - - -/* - * Most types of marker are followed by a variable-length parameter segment. - * This routine skips over the parameters for any marker we don't otherwise - * want to process. - * Note that we MUST skip the parameter segment explicitly in order not to - * be fooled by 0xFF bytes that might appear within the parameter segment; - * such bytes do NOT introduce new markers. - */ - -static void -skip_variable (void) -/* Skip over an unknown or uninteresting variable-length marker */ -{ - unsigned int length; - - /* Get the marker parameter length count */ - length = read_2_bytes(); - /* Length includes itself, so must be at least 2 */ - if (length < 2) - ERREXIT("Erroneous JPEG marker length"); - length -= 2; - /* Skip over the remaining bytes */ - while (length > 0) { - (void) read_1_byte(); - length--; - } -} - - -/* - * Process a COM marker. - * We want to print out the marker contents as legible text; - * we must guard against non-text junk and varying newline representations. - */ - -static void -process_COM (void) -{ - unsigned int length; - int ch; - int lastch = 0; - - /* Get the marker parameter length count */ - length = read_2_bytes(); - /* Length includes itself, so must be at least 2 */ - if (length < 2) - ERREXIT("Erroneous JPEG marker length"); - length -= 2; - - while (length > 0) { - ch = read_1_byte(); - /* Emit the character in a readable form. - * Nonprintables are converted to \nnn form, - * while \ is converted to \\. - * Newlines in CR, CR/LF, or LF form will be printed as one newline. - */ - if (ch == '\r') { - printf("\n"); - } else if (ch == '\n') { - if (lastch != '\r') - printf("\n"); - } else if (ch == '\\') { - printf("\\\\"); - } else if (isprint(ch)) { - putc(ch, stdout); - } else { - printf("\\%03o", ch); - } - lastch = ch; - length--; - } - printf("\n"); -} - - -/* - * Process a SOFn marker. - * This code is only needed if you want to know the image dimensions... - */ - -static void -process_SOFn (int marker) -{ - unsigned int length; - unsigned int image_height, image_width; - int data_precision, num_components; - const char * process; - int ci; - - length = read_2_bytes(); /* usual parameter length count */ - - data_precision = read_1_byte(); - image_height = read_2_bytes(); - image_width = read_2_bytes(); - num_components = read_1_byte(); - -#ifdef NO_SWITCH - if (marker==M_SOF0) process = "Baseline"; - else if (marker==M_SOF1) process = "Extended sequential"; - else if (marker==M_SOF2) process = "Progressive"; - else if (marker==M_SOF3) process = "Lossless"; - else if (marker==M_SOF5) process = "Differential sequential"; - else if (marker==M_SOF6) process = "Differential progressive"; - else if (marker==M_SOF7) process = "Differential lossless"; - else if (marker==M_SOF9) process = "Extended sequential, arithmetic coding"; - else if (marker==M_SOF10) process = "Progressive, arithmetic coding"; - else if (marker==M_SOF11) process = "Lossless, arithmetic coding"; - else if (marker==M_SOF13) process = "Differential sequential, arithmetic coding"; - else if (marker==M_SOF14) process = "Differential progressive, arithmetic coding"; - else if (marker==M_SOF15) process = "Differential lossless, arithmetic coding"; - else process = "Unknown"; -#else - switch (marker) { - case M_SOF0: process = "Baseline"; break; - case M_SOF1: process = "Extended sequential"; break; - case M_SOF2: process = "Progressive"; break; - case M_SOF3: process = "Lossless"; break; - case M_SOF5: process = "Differential sequential"; break; - case M_SOF6: process = "Differential progressive"; break; - case M_SOF7: process = "Differential lossless"; break; - case M_SOF9: process = "Extended sequential, arithmetic coding"; break; - case M_SOF10: process = "Progressive, arithmetic coding"; break; - case M_SOF11: process = "Lossless, arithmetic coding"; break; - case M_SOF13: process = "Differential sequential, arithmetic coding"; break; - case M_SOF14: process = "Differential progressive, arithmetic coding"; break; - case M_SOF15: process = "Differential lossless, arithmetic coding"; break; - default: process = "Unknown"; break; - } -#endif - - printf("JPEG image is %uw * %uh, %d color components, %d bits per sample\n", - image_width, image_height, num_components, data_precision); - printf("JPEG process: %s\n", process); - - if (length != (unsigned int) (8 + num_components * 3)) - ERREXIT("Bogus SOF marker length"); - - for (ci = 0; ci < num_components; ci++) { - (void) read_1_byte(); /* Component ID code */ - (void) read_1_byte(); /* H, V sampling factors */ - (void) read_1_byte(); /* Quantization table number */ - } -} - - -/* - * Parse the marker stream until SOS or EOI is seen; - * display any COM markers. - * While the companion program wrjpgcom will always insert COM markers before - * SOFn, other implementations might not, so we scan to SOS before stopping. - * If we were only interested in the image dimensions, we would stop at SOFn. - * (Conversely, if we only cared about COM markers, there would be no need - * for special code to handle SOFn; we could treat it like other markers.) - */ - -static int -scan_JPEG_header (int verbose) -{ - int marker; - - /* Expect SOI at start of file */ - if (first_marker() != M_SOI) - ERREXIT("Expected SOI marker first"); - - /* Scan miscellaneous markers until we reach SOS. */ - for (;;) { - marker = next_marker(); - switch (marker) { - /* Note that marker codes 0xC4, 0xC8, 0xCC are not, and must not be, - * treated as SOFn. C4 in particular is actually DHT. - */ - case M_SOF0: /* Baseline */ - case M_SOF1: /* Extended sequential, Huffman */ - case M_SOF2: /* Progressive, Huffman */ - case M_SOF3: /* Lossless, Huffman */ - case M_SOF5: /* Differential sequential, Huffman */ - case M_SOF6: /* Differential progressive, Huffman */ - case M_SOF7: /* Differential lossless, Huffman */ - case M_SOF9: /* Extended sequential, arithmetic */ - case M_SOF10: /* Progressive, arithmetic */ - case M_SOF11: /* Lossless, arithmetic */ - case M_SOF13: /* Differential sequential, arithmetic */ - case M_SOF14: /* Differential progressive, arithmetic */ - case M_SOF15: /* Differential lossless, arithmetic */ - if (verbose) - process_SOFn(marker); - else - skip_variable(); - break; - - case M_SOS: /* stop before hitting compressed data */ - return marker; - - case M_EOI: /* in case it's a tables-only JPEG stream */ - return marker; - - case M_COM: - process_COM(); - break; - - case M_APP12: - /* Some digital camera makers put useful textual information into - * APP12 markers, so we print those out too when in -verbose mode. - */ - if (verbose) { - printf("APP12 contains:\n"); - process_COM(); - } else - skip_variable(); - break; - - default: /* Anything else just gets skipped */ - skip_variable(); /* we assume it has a parameter count... */ - break; - } - } /* end loop */ -} - - -/* Command line parsing code */ - -static const char * progname; /* program name for error messages */ - - -static void -usage (void) -/* complain about bad command line */ -{ - fprintf(stderr, "rdjpgcom displays any textual comments in a JPEG file.\n"); - - fprintf(stderr, "Usage: %s [switches] [inputfile]\n", progname); - - fprintf(stderr, "Switches (names may be abbreviated):\n"); - fprintf(stderr, " -verbose Also display dimensions of JPEG image\n"); - - exit(EXIT_FAILURE); -} - - -static int -keymatch (char * arg, const char * keyword, int minchars) -/* Case-insensitive matching of (possibly abbreviated) keyword switches. */ -/* keyword is the constant keyword (must be lower case already), */ -/* minchars is length of minimum legal abbreviation. */ -{ - register int ca, ck; - register int nmatched = 0; - - while ((ca = *arg++) != '\0') { - if ((ck = *keyword++) == '\0') - return 0; /* arg longer than keyword, no good */ - if (isupper(ca)) /* force arg to lcase (assume ck is already) */ - ca = tolower(ca); - if (ca != ck) - return 0; /* no good */ - nmatched++; /* count matched characters */ - } - /* reached end of argument; fail if it's too short for unique abbrev */ - if (nmatched < minchars) - return 0; - return 1; /* A-OK */ -} - - -/* - * The main program. - */ - -int -main (int argc, char **argv) -{ - int argn; - char * arg; - int verbose = 0; - - /* On Mac, fetch a command line. */ -#ifdef USE_CCOMMAND - argc = ccommand(&argv); -#endif - - progname = argv[0]; - if (progname == NULL || progname[0] == 0) - progname = "rdjpgcom"; /* in case C library doesn't provide it */ - - /* Parse switches, if any */ - for (argn = 1; argn < argc; argn++) { - arg = argv[argn]; - if (arg[0] != '-') - break; /* not switch, must be file name */ - arg++; /* advance over '-' */ - if (keymatch(arg, "verbose", 1)) { - verbose++; - } else - usage(); - } - - /* Open the input file. */ - /* Unix style: expect zero or one file name */ - if (argn < argc-1) { - fprintf(stderr, "%s: only one input file\n", progname); - usage(); - } - if (argn < argc) { - if ((infile = fopen(argv[argn], READ_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, argv[argn]); - exit(EXIT_FAILURE); - } - } else { - /* default input file is stdin */ -#ifdef USE_SETMODE /* need to hack file mode? */ - setmode(fileno(stdin), O_BINARY); -#endif -#ifdef USE_FDOPEN /* need to re-open in binary mode? */ - if ((infile = fdopen(fileno(stdin), READ_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open stdin\n", progname); - exit(EXIT_FAILURE); - } -#else - infile = stdin; -#endif - } - - /* Scan the JPEG headers. */ - (void) scan_JPEG_header(verbose); - - /* All done. */ - exit(EXIT_SUCCESS); - return 0; /* suppress no-return-value warnings */ -} +/* + * rdjpgcom.c + * + * Copyright (C) 1994-1997, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains a very simple stand-alone application that displays + * the text in COM (comment) markers in a JFIF file. + * This may be useful as an example of the minimum logic needed to parse + * JPEG markers. + */ + +#define JPEG_CJPEG_DJPEG /* to get the command-line config symbols */ +#include "jinclude.h" /* get auto-config symbols, */ + +#include /* to declare isupper(), tolower() */ +#ifdef USE_SETMODE +#include /* to declare setmode()'s parameter macros */ +/* If you have setmode() but not , just delete this line: */ +#include /* to declare setmode() */ +#endif + +#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ +#ifdef __MWERKS__ +#include /* Metrowerks needs this */ +#include /* ... and this */ +#endif +#ifdef THINK_C +#include /* Think declares it here */ +#endif +#endif + +#ifdef DONT_USE_B_MODE /* define mode parameters for fopen() */ +#define READ_BINARY "r" +#else +#ifdef VMS /* VMS is very nonstandard */ +#define READ_BINARY "rb", "ctx=stm" +#else /* standard ANSI-compliant case */ +#define READ_BINARY "rb" +#endif +#endif + +#ifndef EXIT_FAILURE /* define exit() codes if not provided */ +#define EXIT_FAILURE 1 +#endif +#ifndef EXIT_SUCCESS +#ifdef VMS +#define EXIT_SUCCESS 1 /* VMS is very nonstandard */ +#else +#define EXIT_SUCCESS 0 +#endif +#endif + + +/* + * These macros are used to read the input file. + * To reuse this code in another application, you might need to change these. + */ + +static FILE * infile; /* input JPEG file */ + +/* Return next input byte, or EOF if no more */ +#define NEXTBYTE() getc(infile) + + +/* Error exit handler */ +#define ERREXIT(msg) (fprintf(stderr, "%s\n", msg), exit(EXIT_FAILURE)) + + +/* Read one byte, testing for EOF */ +static int +read_1_byte (void) +{ + int c; + + c = NEXTBYTE(); + if (c == EOF) + ERREXIT("Premature EOF in JPEG file"); + return c; +} + +/* Read 2 bytes, convert to unsigned int */ +/* All 2-byte quantities in JPEG markers are MSB first */ +static unsigned int +read_2_bytes (void) +{ + int c1, c2; + + c1 = NEXTBYTE(); + if (c1 == EOF) + ERREXIT("Premature EOF in JPEG file"); + c2 = NEXTBYTE(); + if (c2 == EOF) + ERREXIT("Premature EOF in JPEG file"); + return (((unsigned int) c1) << 8) + ((unsigned int) c2); +} + + +/* + * JPEG markers consist of one or more 0xFF bytes, followed by a marker + * code byte (which is not an FF). Here are the marker codes of interest + * in this program. (See jdmarker.c for a more complete list.) + */ + +#define M_SOF0 0xC0 /* Start Of Frame N */ +#define M_SOF1 0xC1 /* N indicates which compression process */ +#define M_SOF2 0xC2 /* Only SOF0-SOF2 are now in common use */ +#define M_SOF3 0xC3 +#define M_SOF5 0xC5 /* NB: codes C4 and CC are NOT SOF markers */ +#define M_SOF6 0xC6 +#define M_SOF7 0xC7 +#define M_SOF9 0xC9 +#define M_SOF10 0xCA +#define M_SOF11 0xCB +#define M_SOF13 0xCD +#define M_SOF14 0xCE +#define M_SOF15 0xCF +#define M_SOI 0xD8 /* Start Of Image (beginning of datastream) */ +#define M_EOI 0xD9 /* End Of Image (end of datastream) */ +#define M_SOS 0xDA /* Start Of Scan (begins compressed data) */ +#define M_APP0 0xE0 /* Application-specific marker, type N */ +#define M_APP12 0xEC /* (we don't bother to list all 16 APPn's) */ +#define M_COM 0xFE /* COMment */ + + +/* + * Find the next JPEG marker and return its marker code. + * We expect at least one FF byte, possibly more if the compressor used FFs + * to pad the file. + * There could also be non-FF garbage between markers. The treatment of such + * garbage is unspecified; we choose to skip over it but emit a warning msg. + * NB: this routine must not be used after seeing SOS marker, since it will + * not deal correctly with FF/00 sequences in the compressed image data... + */ + +static int +next_marker (void) +{ + int c; + int discarded_bytes = 0; + + /* Find 0xFF byte; count and skip any non-FFs. */ + c = read_1_byte(); + while (c != 0xFF) { + discarded_bytes++; + c = read_1_byte(); + } + /* Get marker code byte, swallowing any duplicate FF bytes. Extra FFs + * are legal as pad bytes, so don't count them in discarded_bytes. + */ + do { + c = read_1_byte(); + } while (c == 0xFF); + + if (discarded_bytes != 0) { + fprintf(stderr, "Warning: garbage data found in JPEG file\n"); + } + + return c; +} + + +/* + * Read the initial marker, which should be SOI. + * For a JFIF file, the first two bytes of the file should be literally + * 0xFF M_SOI. To be more general, we could use next_marker, but if the + * input file weren't actually JPEG at all, next_marker might read the whole + * file and then return a misleading error message... + */ + +static int +first_marker (void) +{ + int c1, c2; + + c1 = NEXTBYTE(); + c2 = NEXTBYTE(); + if (c1 != 0xFF || c2 != M_SOI) + ERREXIT("Not a JPEG file"); + return c2; +} + + +/* + * Most types of marker are followed by a variable-length parameter segment. + * This routine skips over the parameters for any marker we don't otherwise + * want to process. + * Note that we MUST skip the parameter segment explicitly in order not to + * be fooled by 0xFF bytes that might appear within the parameter segment; + * such bytes do NOT introduce new markers. + */ + +static void +skip_variable (void) +/* Skip over an unknown or uninteresting variable-length marker */ +{ + unsigned int length; + + /* Get the marker parameter length count */ + length = read_2_bytes(); + /* Length includes itself, so must be at least 2 */ + if (length < 2) + ERREXIT("Erroneous JPEG marker length"); + length -= 2; + /* Skip over the remaining bytes */ + while (length > 0) { + (void) read_1_byte(); + length--; + } +} + + +/* + * Process a COM marker. + * We want to print out the marker contents as legible text; + * we must guard against non-text junk and varying newline representations. + */ + +static void +process_COM (void) +{ + unsigned int length; + int ch; + int lastch = 0; + + /* Get the marker parameter length count */ + length = read_2_bytes(); + /* Length includes itself, so must be at least 2 */ + if (length < 2) + ERREXIT("Erroneous JPEG marker length"); + length -= 2; + + while (length > 0) { + ch = read_1_byte(); + /* Emit the character in a readable form. + * Nonprintables are converted to \nnn form, + * while \ is converted to \\. + * Newlines in CR, CR/LF, or LF form will be printed as one newline. + */ + if (ch == '\r') { + printf("\n"); + } else if (ch == '\n') { + if (lastch != '\r') + printf("\n"); + } else if (ch == '\\') { + printf("\\\\"); + } else if (isprint(ch)) { + putc(ch, stdout); + } else { + printf("\\%03o", ch); + } + lastch = ch; + length--; + } + printf("\n"); +} + + +/* + * Process a SOFn marker. + * This code is only needed if you want to know the image dimensions... + */ + +static void +process_SOFn (int marker) +{ + unsigned int length; + unsigned int image_height, image_width; + int data_precision, num_components; + const char * process; + int ci; + + length = read_2_bytes(); /* usual parameter length count */ + + data_precision = read_1_byte(); + image_height = read_2_bytes(); + image_width = read_2_bytes(); + num_components = read_1_byte(); + + switch (marker) { + case M_SOF0: process = "Baseline"; break; + case M_SOF1: process = "Extended sequential"; break; + case M_SOF2: process = "Progressive"; break; + case M_SOF3: process = "Lossless"; break; + case M_SOF5: process = "Differential sequential"; break; + case M_SOF6: process = "Differential progressive"; break; + case M_SOF7: process = "Differential lossless"; break; + case M_SOF9: process = "Extended sequential, arithmetic coding"; break; + case M_SOF10: process = "Progressive, arithmetic coding"; break; + case M_SOF11: process = "Lossless, arithmetic coding"; break; + case M_SOF13: process = "Differential sequential, arithmetic coding"; break; + case M_SOF14: process = "Differential progressive, arithmetic coding"; break; + case M_SOF15: process = "Differential lossless, arithmetic coding"; break; + default: process = "Unknown"; break; + } + + printf("JPEG image is %uw * %uh, %d color components, %d bits per sample\n", + image_width, image_height, num_components, data_precision); + printf("JPEG process: %s\n", process); + + if (length != (unsigned int) (8 + num_components * 3)) + ERREXIT("Bogus SOF marker length"); + + for (ci = 0; ci < num_components; ci++) { + (void) read_1_byte(); /* Component ID code */ + (void) read_1_byte(); /* H, V sampling factors */ + (void) read_1_byte(); /* Quantization table number */ + } +} + + +/* + * Parse the marker stream until SOS or EOI is seen; + * display any COM markers. + * While the companion program wrjpgcom will always insert COM markers before + * SOFn, other implementations might not, so we scan to SOS before stopping. + * If we were only interested in the image dimensions, we would stop at SOFn. + * (Conversely, if we only cared about COM markers, there would be no need + * for special code to handle SOFn; we could treat it like other markers.) + */ + +static int +scan_JPEG_header (int verbose) +{ + int marker; + + /* Expect SOI at start of file */ + if (first_marker() != M_SOI) + ERREXIT("Expected SOI marker first"); + + /* Scan miscellaneous markers until we reach SOS. */ + for (;;) { + marker = next_marker(); + switch (marker) { + /* Note that marker codes 0xC4, 0xC8, 0xCC are not, and must not be, + * treated as SOFn. C4 in particular is actually DHT. + */ + case M_SOF0: /* Baseline */ + case M_SOF1: /* Extended sequential, Huffman */ + case M_SOF2: /* Progressive, Huffman */ + case M_SOF3: /* Lossless, Huffman */ + case M_SOF5: /* Differential sequential, Huffman */ + case M_SOF6: /* Differential progressive, Huffman */ + case M_SOF7: /* Differential lossless, Huffman */ + case M_SOF9: /* Extended sequential, arithmetic */ + case M_SOF10: /* Progressive, arithmetic */ + case M_SOF11: /* Lossless, arithmetic */ + case M_SOF13: /* Differential sequential, arithmetic */ + case M_SOF14: /* Differential progressive, arithmetic */ + case M_SOF15: /* Differential lossless, arithmetic */ + if (verbose) + process_SOFn(marker); + else + skip_variable(); + break; + + case M_SOS: /* stop before hitting compressed data */ + return marker; + + case M_EOI: /* in case it's a tables-only JPEG stream */ + return marker; + + case M_COM: + process_COM(); + break; + + case M_APP12: + /* Some digital camera makers put useful textual information into + * APP12 markers, so we print those out too when in -verbose mode. + */ + if (verbose) { + printf("APP12 contains:\n"); + process_COM(); + } else + skip_variable(); + break; + + default: /* Anything else just gets skipped */ + skip_variable(); /* we assume it has a parameter count... */ + break; + } + } /* end loop */ +} + + +/* Command line parsing code */ + +static const char * progname; /* program name for error messages */ + + +static void +usage (void) +/* complain about bad command line */ +{ + fprintf(stderr, "rdjpgcom displays any textual comments in a JPEG file.\n"); + + fprintf(stderr, "Usage: %s [switches] [inputfile]\n", progname); + + fprintf(stderr, "Switches (names may be abbreviated):\n"); + fprintf(stderr, " -verbose Also display dimensions of JPEG image\n"); + + exit(EXIT_FAILURE); +} + + +static int +keymatch (char * arg, const char * keyword, int minchars) +/* Case-insensitive matching of (possibly abbreviated) keyword switches. */ +/* keyword is the constant keyword (must be lower case already), */ +/* minchars is length of minimum legal abbreviation. */ +{ + register int ca, ck; + register int nmatched = 0; + + while ((ca = *arg++) != '\0') { + if ((ck = *keyword++) == '\0') + return 0; /* arg longer than keyword, no good */ + if (isupper(ca)) /* force arg to lcase (assume ck is already) */ + ca = tolower(ca); + if (ca != ck) + return 0; /* no good */ + nmatched++; /* count matched characters */ + } + /* reached end of argument; fail if it's too short for unique abbrev */ + if (nmatched < minchars) + return 0; + return 1; /* A-OK */ +} + + +/* + * The main program. + */ + +int +main (int argc, char **argv) +{ + int argn; + char * arg; + int verbose = 0; + + /* On Mac, fetch a command line. */ +#ifdef USE_CCOMMAND + argc = ccommand(&argv); +#endif + + progname = argv[0]; + if (progname == NULL || progname[0] == 0) + progname = "rdjpgcom"; /* in case C library doesn't provide it */ + + /* Parse switches, if any */ + for (argn = 1; argn < argc; argn++) { + arg = argv[argn]; + if (arg[0] != '-') + break; /* not switch, must be file name */ + arg++; /* advance over '-' */ + if (keymatch(arg, "verbose", 1)) { + verbose++; + } else + usage(); + } + + /* Open the input file. */ + /* Unix style: expect zero or one file name */ + if (argn < argc-1) { + fprintf(stderr, "%s: only one input file\n", progname); + usage(); + } + if (argn < argc) { + if ((infile = fopen(argv[argn], READ_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, argv[argn]); + exit(EXIT_FAILURE); + } + } else { + /* default input file is stdin */ +#ifdef USE_SETMODE /* need to hack file mode? */ + setmode(fileno(stdin), O_BINARY); +#endif +#ifdef USE_FDOPEN /* need to re-open in binary mode? */ + if ((infile = fdopen(fileno(stdin), READ_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open stdin\n", progname); + exit(EXIT_FAILURE); + } +#else + infile = stdin; +#endif + } + + /* Scan the JPEG headers. */ + (void) scan_JPEG_header(verbose); + + /* All done. */ + exit(EXIT_SUCCESS); + return 0; /* suppress no-return-value warnings */ +} diff --git a/test/monniaux/jpeg-6b/rdswitch.c b/test/monniaux/jpeg-6b/rdswitch.c index b915ad9d..4f4bb4f5 100644 --- a/test/monniaux/jpeg-6b/rdswitch.c +++ b/test/monniaux/jpeg-6b/rdswitch.c @@ -1,332 +1,332 @@ -/* - * rdswitch.c - * - * Copyright (C) 1991-1996, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains routines to process some of cjpeg's more complicated - * command-line switches. Switches processed here are: - * -qtables file Read quantization tables from text file - * -scans file Read scan script from text file - * -qslots N[,N,...] Set component quantization table selectors - * -sample HxV[,HxV,...] Set component sampling factors - */ - -#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ -#include /* to declare isdigit(), isspace() */ - - -LOCAL(int) -text_getc (FILE * file) -/* Read next char, skipping over any comments (# to end of line) */ -/* A comment/newline sequence is returned as a newline */ -{ - register int ch; - - ch = getc(file); - if (ch == '#') { - do { - ch = getc(file); - } while (ch != '\n' && ch != EOF); - } - return ch; -} - - -LOCAL(boolean) -read_text_integer (FILE * file, long * result, int * termchar) -/* Read an unsigned decimal integer from a file, store it in result */ -/* Reads one trailing character after the integer; returns it in termchar */ -{ - register int ch; - register long val; - - /* Skip any leading whitespace, detect EOF */ - do { - ch = text_getc(file); - if (ch == EOF) { - *termchar = ch; - return FALSE; - } - } while (isspace(ch)); - - if (! isdigit(ch)) { - *termchar = ch; - return FALSE; - } - - val = ch - '0'; - while ((ch = text_getc(file)) != EOF) { - if (! isdigit(ch)) - break; - val *= 10; - val += ch - '0'; - } - *result = val; - *termchar = ch; - return TRUE; -} - - -GLOBAL(boolean) -read_quant_tables (j_compress_ptr cinfo, char * filename, - int scale_factor, boolean force_baseline) -/* Read a set of quantization tables from the specified file. - * The file is plain ASCII text: decimal numbers with whitespace between. - * Comments preceded by '#' may be included in the file. - * There may be one to NUM_QUANT_TBLS tables in the file, each of 64 values. - * The tables are implicitly numbered 0,1,etc. - * NOTE: does not affect the qslots mapping, which will default to selecting - * table 0 for luminance (or primary) components, 1 for chrominance components. - * You must use -qslots if you want a different component->table mapping. - */ -{ - FILE * fp; - int tblno, i, termchar; - long val; - unsigned int table[DCTSIZE2]; - - if ((fp = fopen(filename, "r")) == NULL) { - fprintf(stderr, "Can't open table file %s\n", filename); - return FALSE; - } - tblno = 0; - - while (read_text_integer(fp, &val, &termchar)) { /* read 1st element of table */ - if (tblno >= NUM_QUANT_TBLS) { - fprintf(stderr, "Too many tables in file %s\n", filename); - fclose(fp); - return FALSE; - } - table[0] = (unsigned int) val; - for (i = 1; i < DCTSIZE2; i++) { - if (! read_text_integer(fp, &val, &termchar)) { - fprintf(stderr, "Invalid table data in file %s\n", filename); - fclose(fp); - return FALSE; - } - table[i] = (unsigned int) val; - } - jpeg_add_quant_table(cinfo, tblno, table, scale_factor, force_baseline); - tblno++; - } - - if (termchar != EOF) { - fprintf(stderr, "Non-numeric data in file %s\n", filename); - fclose(fp); - return FALSE; - } - - fclose(fp); - return TRUE; -} - - -#ifdef C_MULTISCAN_FILES_SUPPORTED - -LOCAL(boolean) -read_scan_integer (FILE * file, long * result, int * termchar) -/* Variant of read_text_integer that always looks for a non-space termchar; - * this simplifies parsing of punctuation in scan scripts. - */ -{ - register int ch; - - if (! read_text_integer(file, result, termchar)) - return FALSE; - ch = *termchar; - while (ch != EOF && isspace(ch)) - ch = text_getc(file); - if (isdigit(ch)) { /* oops, put it back */ - if (ungetc(ch, file) == EOF) - return FALSE; - ch = ' '; - } else { - /* Any separators other than ';' and ':' are ignored; - * this allows user to insert commas, etc, if desired. - */ - if (ch != EOF && ch != ';' && ch != ':') - ch = ' '; - } - *termchar = ch; - return TRUE; -} - - -GLOBAL(boolean) -read_scan_script (j_compress_ptr cinfo, char * filename) -/* Read a scan script from the specified text file. - * Each entry in the file defines one scan to be emitted. - * Entries are separated by semicolons ';'. - * An entry contains one to four component indexes, - * optionally followed by a colon ':' and four progressive-JPEG parameters. - * The component indexes denote which component(s) are to be transmitted - * in the current scan. The first component has index 0. - * Sequential JPEG is used if the progressive-JPEG parameters are omitted. - * The file is free format text: any whitespace may appear between numbers - * and the ':' and ';' punctuation marks. Also, other punctuation (such - * as commas or dashes) can be placed between numbers if desired. - * Comments preceded by '#' may be included in the file. - * Note: we do very little validity checking here; - * jcmaster.c will validate the script parameters. - */ -{ - FILE * fp; - int scanno, ncomps, termchar; - long val; - jpeg_scan_info * scanptr; -#define MAX_SCANS 100 /* quite arbitrary limit */ - jpeg_scan_info scans[MAX_SCANS]; - - if ((fp = fopen(filename, "r")) == NULL) { - fprintf(stderr, "Can't open scan definition file %s\n", filename); - return FALSE; - } - scanptr = scans; - scanno = 0; - - while (read_scan_integer(fp, &val, &termchar)) { - if (scanno >= MAX_SCANS) { - fprintf(stderr, "Too many scans defined in file %s\n", filename); - fclose(fp); - return FALSE; - } - scanptr->component_index[0] = (int) val; - ncomps = 1; - while (termchar == ' ') { - if (ncomps >= MAX_COMPS_IN_SCAN) { - fprintf(stderr, "Too many components in one scan in file %s\n", - filename); - fclose(fp); - return FALSE; - } - if (! read_scan_integer(fp, &val, &termchar)) - goto bogus; - scanptr->component_index[ncomps] = (int) val; - ncomps++; - } - scanptr->comps_in_scan = ncomps; - if (termchar == ':') { - if (! read_scan_integer(fp, &val, &termchar) || termchar != ' ') - goto bogus; - scanptr->Ss = (int) val; - if (! read_scan_integer(fp, &val, &termchar) || termchar != ' ') - goto bogus; - scanptr->Se = (int) val; - if (! read_scan_integer(fp, &val, &termchar) || termchar != ' ') - goto bogus; - scanptr->Ah = (int) val; - if (! read_scan_integer(fp, &val, &termchar)) - goto bogus; - scanptr->Al = (int) val; - } else { - /* set non-progressive parameters */ - scanptr->Ss = 0; - scanptr->Se = DCTSIZE2-1; - scanptr->Ah = 0; - scanptr->Al = 0; - } - if (termchar != ';' && termchar != EOF) { -bogus: - fprintf(stderr, "Invalid scan entry format in file %s\n", filename); - fclose(fp); - return FALSE; - } - scanptr++, scanno++; - } - - if (termchar != EOF) { - fprintf(stderr, "Non-numeric data in file %s\n", filename); - fclose(fp); - return FALSE; - } - - if (scanno > 0) { - /* Stash completed scan list in cinfo structure. - * NOTE: for cjpeg's use, JPOOL_IMAGE is the right lifetime for this data, - * but if you want to compress multiple images you'd want JPOOL_PERMANENT. - */ - scanptr = (jpeg_scan_info *) - (*cinfo->mem->alloc_small) ((j_common_ptr) cinfo, JPOOL_IMAGE, - scanno * SIZEOF(jpeg_scan_info)); - MEMCOPY(scanptr, scans, scanno * SIZEOF(jpeg_scan_info)); - cinfo->scan_info = scanptr; - cinfo->num_scans = scanno; - } - - fclose(fp); - return TRUE; -} - -#endif /* C_MULTISCAN_FILES_SUPPORTED */ - - -GLOBAL(boolean) -set_quant_slots (j_compress_ptr cinfo, char *arg) -/* Process a quantization-table-selectors parameter string, of the form - * N[,N,...] - * If there are more components than parameters, the last value is replicated. - */ -{ - int val = 0; /* default table # */ - int ci; - char ch; - - for (ci = 0; ci < MAX_COMPONENTS; ci++) { - if (*arg) { - ch = ','; /* if not set by sscanf, will be ',' */ - if (sscanf(arg, "%d%c", &val, &ch) < 1) - return FALSE; - if (ch != ',') /* syntax check */ - return FALSE; - if (val < 0 || val >= NUM_QUANT_TBLS) { - fprintf(stderr, "JPEG quantization tables are numbered 0..%d\n", - NUM_QUANT_TBLS-1); - return FALSE; - } - cinfo->comp_info[ci].quant_tbl_no = val; - while (*arg && *arg++ != ',') /* advance to next segment of arg string */ - ; - } else { - /* reached end of parameter, set remaining components to last table */ - cinfo->comp_info[ci].quant_tbl_no = val; - } - } - return TRUE; -} - - -GLOBAL(boolean) -set_sample_factors (j_compress_ptr cinfo, char *arg) -/* Process a sample-factors parameter string, of the form - * HxV[,HxV,...] - * If there are more components than parameters, "1x1" is assumed for the rest. - */ -{ - int ci, val1, val2; - char ch1, ch2; - - for (ci = 0; ci < MAX_COMPONENTS; ci++) { - if (*arg) { - ch2 = ','; /* if not set by sscanf, will be ',' */ - if (sscanf(arg, "%d%c%d%c", &val1, &ch1, &val2, &ch2) < 3) - return FALSE; - if ((ch1 != 'x' && ch1 != 'X') || ch2 != ',') /* syntax check */ - return FALSE; - if (val1 <= 0 || val1 > 4 || val2 <= 0 || val2 > 4) { - fprintf(stderr, "JPEG sampling factors must be 1..4\n"); - return FALSE; - } - cinfo->comp_info[ci].h_samp_factor = val1; - cinfo->comp_info[ci].v_samp_factor = val2; - while (*arg && *arg++ != ',') /* advance to next segment of arg string */ - ; - } else { - /* reached end of parameter, set remaining components to 1x1 sampling */ - cinfo->comp_info[ci].h_samp_factor = 1; - cinfo->comp_info[ci].v_samp_factor = 1; - } - } - return TRUE; -} +/* + * rdswitch.c + * + * Copyright (C) 1991-1996, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains routines to process some of cjpeg's more complicated + * command-line switches. Switches processed here are: + * -qtables file Read quantization tables from text file + * -scans file Read scan script from text file + * -qslots N[,N,...] Set component quantization table selectors + * -sample HxV[,HxV,...] Set component sampling factors + */ + +#include "cdjpeg.h" /* Common decls for cjpeg/djpeg applications */ +#include /* to declare isdigit(), isspace() */ + + +LOCAL(int) +text_getc (FILE * file) +/* Read next char, skipping over any comments (# to end of line) */ +/* A comment/newline sequence is returned as a newline */ +{ + register int ch; + + ch = getc(file); + if (ch == '#') { + do { + ch = getc(file); + } while (ch != '\n' && ch != EOF); + } + return ch; +} + + +LOCAL(boolean) +read_text_integer (FILE * file, long * result, int * termchar) +/* Read an unsigned decimal integer from a file, store it in result */ +/* Reads one trailing character after the integer; returns it in termchar */ +{ + register int ch; + register long val; + + /* Skip any leading whitespace, detect EOF */ + do { + ch = text_getc(file); + if (ch == EOF) { + *termchar = ch; + return FALSE; + } + } while (isspace(ch)); + + if (! isdigit(ch)) { + *termchar = ch; + return FALSE; + } + + val = ch - '0'; + while ((ch = text_getc(file)) != EOF) { + if (! isdigit(ch)) + break; + val *= 10; + val += ch - '0'; + } + *result = val; + *termchar = ch; + return TRUE; +} + + +GLOBAL(boolean) +read_quant_tables (j_compress_ptr cinfo, char * filename, + int scale_factor, boolean force_baseline) +/* Read a set of quantization tables from the specified file. + * The file is plain ASCII text: decimal numbers with whitespace between. + * Comments preceded by '#' may be included in the file. + * There may be one to NUM_QUANT_TBLS tables in the file, each of 64 values. + * The tables are implicitly numbered 0,1,etc. + * NOTE: does not affect the qslots mapping, which will default to selecting + * table 0 for luminance (or primary) components, 1 for chrominance components. + * You must use -qslots if you want a different component->table mapping. + */ +{ + FILE * fp; + int tblno, i, termchar; + long val; + unsigned int table[DCTSIZE2]; + + if ((fp = fopen(filename, "r")) == NULL) { + fprintf(stderr, "Can't open table file %s\n", filename); + return FALSE; + } + tblno = 0; + + while (read_text_integer(fp, &val, &termchar)) { /* read 1st element of table */ + if (tblno >= NUM_QUANT_TBLS) { + fprintf(stderr, "Too many tables in file %s\n", filename); + fclose(fp); + return FALSE; + } + table[0] = (unsigned int) val; + for (i = 1; i < DCTSIZE2; i++) { + if (! read_text_integer(fp, &val, &termchar)) { + fprintf(stderr, "Invalid table data in file %s\n", filename); + fclose(fp); + return FALSE; + } + table[i] = (unsigned int) val; + } + jpeg_add_quant_table(cinfo, tblno, table, scale_factor, force_baseline); + tblno++; + } + + if (termchar != EOF) { + fprintf(stderr, "Non-numeric data in file %s\n", filename); + fclose(fp); + return FALSE; + } + + fclose(fp); + return TRUE; +} + + +#ifdef C_MULTISCAN_FILES_SUPPORTED + +LOCAL(boolean) +read_scan_integer (FILE * file, long * result, int * termchar) +/* Variant of read_text_integer that always looks for a non-space termchar; + * this simplifies parsing of punctuation in scan scripts. + */ +{ + register int ch; + + if (! read_text_integer(file, result, termchar)) + return FALSE; + ch = *termchar; + while (ch != EOF && isspace(ch)) + ch = text_getc(file); + if (isdigit(ch)) { /* oops, put it back */ + if (ungetc(ch, file) == EOF) + return FALSE; + ch = ' '; + } else { + /* Any separators other than ';' and ':' are ignored; + * this allows user to insert commas, etc, if desired. + */ + if (ch != EOF && ch != ';' && ch != ':') + ch = ' '; + } + *termchar = ch; + return TRUE; +} + + +GLOBAL(boolean) +read_scan_script (j_compress_ptr cinfo, char * filename) +/* Read a scan script from the specified text file. + * Each entry in the file defines one scan to be emitted. + * Entries are separated by semicolons ';'. + * An entry contains one to four component indexes, + * optionally followed by a colon ':' and four progressive-JPEG parameters. + * The component indexes denote which component(s) are to be transmitted + * in the current scan. The first component has index 0. + * Sequential JPEG is used if the progressive-JPEG parameters are omitted. + * The file is free format text: any whitespace may appear between numbers + * and the ':' and ';' punctuation marks. Also, other punctuation (such + * as commas or dashes) can be placed between numbers if desired. + * Comments preceded by '#' may be included in the file. + * Note: we do very little validity checking here; + * jcmaster.c will validate the script parameters. + */ +{ + FILE * fp; + int scanno, ncomps, termchar; + long val; + jpeg_scan_info * scanptr; +#define MAX_SCANS 100 /* quite arbitrary limit */ + jpeg_scan_info scans[MAX_SCANS]; + + if ((fp = fopen(filename, "r")) == NULL) { + fprintf(stderr, "Can't open scan definition file %s\n", filename); + return FALSE; + } + scanptr = scans; + scanno = 0; + + while (read_scan_integer(fp, &val, &termchar)) { + if (scanno >= MAX_SCANS) { + fprintf(stderr, "Too many scans defined in file %s\n", filename); + fclose(fp); + return FALSE; + } + scanptr->component_index[0] = (int) val; + ncomps = 1; + while (termchar == ' ') { + if (ncomps >= MAX_COMPS_IN_SCAN) { + fprintf(stderr, "Too many components in one scan in file %s\n", + filename); + fclose(fp); + return FALSE; + } + if (! read_scan_integer(fp, &val, &termchar)) + goto bogus; + scanptr->component_index[ncomps] = (int) val; + ncomps++; + } + scanptr->comps_in_scan = ncomps; + if (termchar == ':') { + if (! read_scan_integer(fp, &val, &termchar) || termchar != ' ') + goto bogus; + scanptr->Ss = (int) val; + if (! read_scan_integer(fp, &val, &termchar) || termchar != ' ') + goto bogus; + scanptr->Se = (int) val; + if (! read_scan_integer(fp, &val, &termchar) || termchar != ' ') + goto bogus; + scanptr->Ah = (int) val; + if (! read_scan_integer(fp, &val, &termchar)) + goto bogus; + scanptr->Al = (int) val; + } else { + /* set non-progressive parameters */ + scanptr->Ss = 0; + scanptr->Se = DCTSIZE2-1; + scanptr->Ah = 0; + scanptr->Al = 0; + } + if (termchar != ';' && termchar != EOF) { +bogus: + fprintf(stderr, "Invalid scan entry format in file %s\n", filename); + fclose(fp); + return FALSE; + } + scanptr++, scanno++; + } + + if (termchar != EOF) { + fprintf(stderr, "Non-numeric data in file %s\n", filename); + fclose(fp); + return FALSE; + } + + if (scanno > 0) { + /* Stash completed scan list in cinfo structure. + * NOTE: for cjpeg's use, JPOOL_IMAGE is the right lifetime for this data, + * but if you want to compress multiple images you'd want JPOOL_PERMANENT. + */ + scanptr = (jpeg_scan_info *) + (*cinfo->mem->alloc_small) ((j_common_ptr) cinfo, JPOOL_IMAGE, + scanno * SIZEOF(jpeg_scan_info)); + MEMCOPY(scanptr, scans, scanno * SIZEOF(jpeg_scan_info)); + cinfo->scan_info = scanptr; + cinfo->num_scans = scanno; + } + + fclose(fp); + return TRUE; +} + +#endif /* C_MULTISCAN_FILES_SUPPORTED */ + + +GLOBAL(boolean) +set_quant_slots (j_compress_ptr cinfo, char *arg) +/* Process a quantization-table-selectors parameter string, of the form + * N[,N,...] + * If there are more components than parameters, the last value is replicated. + */ +{ + int val = 0; /* default table # */ + int ci; + char ch; + + for (ci = 0; ci < MAX_COMPONENTS; ci++) { + if (*arg) { + ch = ','; /* if not set by sscanf, will be ',' */ + if (sscanf(arg, "%d%c", &val, &ch) < 1) + return FALSE; + if (ch != ',') /* syntax check */ + return FALSE; + if (val < 0 || val >= NUM_QUANT_TBLS) { + fprintf(stderr, "JPEG quantization tables are numbered 0..%d\n", + NUM_QUANT_TBLS-1); + return FALSE; + } + cinfo->comp_info[ci].quant_tbl_no = val; + while (*arg && *arg++ != ',') /* advance to next segment of arg string */ + ; + } else { + /* reached end of parameter, set remaining components to last table */ + cinfo->comp_info[ci].quant_tbl_no = val; + } + } + return TRUE; +} + + +GLOBAL(boolean) +set_sample_factors (j_compress_ptr cinfo, char *arg) +/* Process a sample-factors parameter string, of the form + * HxV[,HxV,...] + * If there are more components than parameters, "1x1" is assumed for the rest. + */ +{ + int ci, val1, val2; + char ch1, ch2; + + for (ci = 0; ci < MAX_COMPONENTS; ci++) { + if (*arg) { + ch2 = ','; /* if not set by sscanf, will be ',' */ + if (sscanf(arg, "%d%c%d%c", &val1, &ch1, &val2, &ch2) < 3) + return FALSE; + if ((ch1 != 'x' && ch1 != 'X') || ch2 != ',') /* syntax check */ + return FALSE; + if (val1 <= 0 || val1 > 4 || val2 <= 0 || val2 > 4) { + fprintf(stderr, "JPEG sampling factors must be 1..4\n"); + return FALSE; + } + cinfo->comp_info[ci].h_samp_factor = val1; + cinfo->comp_info[ci].v_samp_factor = val2; + while (*arg && *arg++ != ',') /* advance to next segment of arg string */ + ; + } else { + /* reached end of parameter, set remaining components to 1x1 sampling */ + cinfo->comp_info[ci].h_samp_factor = 1; + cinfo->comp_info[ci].v_samp_factor = 1; + } + } + return TRUE; +} diff --git a/test/monniaux/jpeg-6b/transupp.c b/test/monniaux/jpeg-6b/transupp.c index 9f6e57f4..e5ec5642 100644 --- a/test/monniaux/jpeg-6b/transupp.c +++ b/test/monniaux/jpeg-6b/transupp.c @@ -79,7 +79,7 @@ do_flip_h (j_decompress_ptr srcinfo, j_compress_ptr dstinfo, * mirroring by changing the signs of odd-numbered columns. * Partial iMCUs at the right edge are left untouched. */ - MCU_cols = INT_DIV(dstinfo->image_width, (dstinfo->max_h_samp_factor * DCTSIZE)); + MCU_cols = dstinfo->image_width / (dstinfo->max_h_samp_factor * DCTSIZE); for (ci = 0; ci < dstinfo->num_components; ci++) { compptr = dstinfo->comp_info + ci; @@ -131,7 +131,7 @@ do_flip_v (j_decompress_ptr srcinfo, j_compress_ptr dstinfo, * of odd-numbered rows. * Partial iMCUs at the bottom edge are copied verbatim. */ - MCU_rows = INT_DIV(dstinfo->image_height, (dstinfo->max_v_samp_factor * DCTSIZE)); + MCU_rows = dstinfo->image_height / (dstinfo->max_v_samp_factor * DCTSIZE); for (ci = 0; ci < dstinfo->num_components; ci++) { compptr = dstinfo->comp_info + ci; @@ -246,7 +246,7 @@ do_rot_90 (j_decompress_ptr srcinfo, j_compress_ptr dstinfo, * at the (output) right edge properly. They just get transposed and * not mirrored. */ - MCU_cols = INT_DIV(dstinfo->image_width, (dstinfo->max_h_samp_factor * DCTSIZE)); + MCU_cols = dstinfo->image_width / (dstinfo->max_h_samp_factor * DCTSIZE); for (ci = 0; ci < dstinfo->num_components; ci++) { compptr = dstinfo->comp_info + ci; @@ -310,7 +310,7 @@ do_rot_270 (j_decompress_ptr srcinfo, j_compress_ptr dstinfo, * at the (output) bottom edge properly. They just get transposed and * not mirrored. */ - MCU_rows = INT_DIV(dstinfo->image_height, (dstinfo->max_v_samp_factor * DCTSIZE)); + MCU_rows = dstinfo->image_height / (dstinfo->max_v_samp_factor * DCTSIZE); for (ci = 0; ci < dstinfo->num_components; ci++) { compptr = dstinfo->comp_info + ci; @@ -371,8 +371,8 @@ do_rot_180 (j_decompress_ptr srcinfo, j_compress_ptr dstinfo, JCOEFPTR src_ptr, dst_ptr; jpeg_component_info *compptr; - MCU_cols = INT_DIV(dstinfo->image_width, (dstinfo->max_h_samp_factor * DCTSIZE)); - MCU_rows = INT_DIV(dstinfo->image_height, (dstinfo->max_v_samp_factor * DCTSIZE)); + MCU_cols = dstinfo->image_width / (dstinfo->max_h_samp_factor * DCTSIZE); + MCU_rows = dstinfo->image_height / (dstinfo->max_v_samp_factor * DCTSIZE); for (ci = 0; ci < dstinfo->num_components; ci++) { compptr = dstinfo->comp_info + ci; @@ -475,8 +475,8 @@ do_transverse (j_decompress_ptr srcinfo, j_compress_ptr dstinfo, JCOEFPTR src_ptr, dst_ptr; jpeg_component_info *compptr; - MCU_cols = INT_DIV(dstinfo->image_width, (dstinfo->max_h_samp_factor * DCTSIZE)); - MCU_rows = INT_DIV(dstinfo->image_height, (dstinfo->max_v_samp_factor * DCTSIZE)); + MCU_cols = dstinfo->image_width / (dstinfo->max_h_samp_factor * DCTSIZE); + MCU_rows = dstinfo->image_height / (dstinfo->max_v_samp_factor * DCTSIZE); for (ci = 0; ci < dstinfo->num_components; ci++) { compptr = dstinfo->comp_info + ci; @@ -582,52 +582,6 @@ jtransform_request_workspace (j_decompress_ptr srcinfo, info->num_components = srcinfo->num_components; } -#ifdef NO_SWITCH - int choice = info->transform; - if (choice == JXFORM_NONE || - choice == JXFORM_FLIP_H) { - } else if (choice == JXFORM_FLIP_V || - choice == JXFORM_ROT_180) { - /* Need workspace arrays having same dimensions as source image. - * Note that we allocate arrays padded out to the next iMCU boundary, - * so that transform routines need not worry about missing edge blocks. - */ - coef_arrays = (jvirt_barray_ptr *) - (*srcinfo->mem->alloc_small) ((j_common_ptr) srcinfo, JPOOL_IMAGE, - SIZEOF(jvirt_barray_ptr) * info->num_components); - for (ci = 0; ci < info->num_components; ci++) { - compptr = srcinfo->comp_info + ci; - coef_arrays[ci] = (*srcinfo->mem->request_virt_barray) - ((j_common_ptr) srcinfo, JPOOL_IMAGE, FALSE, - (JDIMENSION) jround_up((long) compptr->width_in_blocks, - (long) compptr->h_samp_factor), - (JDIMENSION) jround_up((long) compptr->height_in_blocks, - (long) compptr->v_samp_factor), - (JDIMENSION) compptr->v_samp_factor); - } - } else if (choice == JXFORM_TRANSPOSE || - choice == JXFORM_TRANSVERSE || - choice == JXFORM_ROT_90 || - choice == JXFORM_ROT_270) { - /* Need workspace arrays having transposed dimensions. - * Note that we allocate arrays padded out to the next iMCU boundary, - * so that transform routines need not worry about missing edge blocks. - */ - coef_arrays = (jvirt_barray_ptr *) - (*srcinfo->mem->alloc_small) ((j_common_ptr) srcinfo, JPOOL_IMAGE, - SIZEOF(jvirt_barray_ptr) * info->num_components); - for (ci = 0; ci < info->num_components; ci++) { - compptr = srcinfo->comp_info + ci; - coef_arrays[ci] = (*srcinfo->mem->request_virt_barray) - ((j_common_ptr) srcinfo, JPOOL_IMAGE, FALSE, - (JDIMENSION) jround_up((long) compptr->height_in_blocks, - (long) compptr->v_samp_factor), - (JDIMENSION) jround_up((long) compptr->width_in_blocks, - (long) compptr->h_samp_factor), - (JDIMENSION) compptr->h_samp_factor); - } - } -#else switch (info->transform) { case JXFORM_NONE: case JXFORM_FLIP_H: @@ -676,7 +630,6 @@ jtransform_request_workspace (j_decompress_ptr srcinfo, } break; } -#endif info->workspace_coef_arrays = coef_arrays; } @@ -738,7 +691,7 @@ trim_right_edge (j_compress_ptr dstinfo) int h_samp_factor = dstinfo->comp_info[ci].h_samp_factor; max_h_samp_factor = MAX(max_h_samp_factor, h_samp_factor); } - MCU_cols = INT_DIV(dstinfo->image_width, (max_h_samp_factor * DCTSIZE)); + MCU_cols = dstinfo->image_width / (max_h_samp_factor * DCTSIZE); if (MCU_cols > 0) /* can't trim to 0 pixels */ dstinfo->image_width = MCU_cols * (max_h_samp_factor * DCTSIZE); } @@ -758,7 +711,7 @@ trim_bottom_edge (j_compress_ptr dstinfo) int v_samp_factor = dstinfo->comp_info[ci].v_samp_factor; max_v_samp_factor = MAX(max_v_samp_factor, v_samp_factor); } - MCU_rows = INT_DIV(dstinfo->image_height, (max_v_samp_factor * DCTSIZE)); + MCU_rows = dstinfo->image_height / (max_v_samp_factor * DCTSIZE); if (MCU_rows > 0) /* can't trim to 0 pixels */ dstinfo->image_height = MCU_rows * (max_v_samp_factor * DCTSIZE); } @@ -805,39 +758,6 @@ jtransform_adjust_parameters (j_decompress_ptr srcinfo, } /* Correct the destination's image dimensions etc if necessary */ -#ifdef NO_SWITCH - int choice = info->transform; - if (choice == JXFORM_NONE) { - } else if (choice == JXFORM_FLIP_H) { - if (info->trim) - trim_right_edge(dstinfo); - } else if (choice == JXFORM_FLIP_V) { - if (info->trim) - trim_bottom_edge(dstinfo); - } else if (choice == JXFORM_TRANSPOSE) { - transpose_critical_parameters(dstinfo); - /* transpose does NOT have to trim anything */ - } else if (choice == JXFORM_TRANSVERSE) { - transpose_critical_parameters(dstinfo); - if (info->trim) { - trim_right_edge(dstinfo); - trim_bottom_edge(dstinfo); - } - } else if (choice == JXFORM_ROT_90) { - transpose_critical_parameters(dstinfo); - if (info->trim) - trim_right_edge(dstinfo); - } else if (choice == JXFORM_ROT_180) { - if (info->trim) { - trim_right_edge(dstinfo); - trim_bottom_edge(dstinfo); - } - } else if (choice == JXFORM_ROT_270) { - transpose_critical_parameters(dstinfo); - if (info->trim) - trim_bottom_edge(dstinfo); - } -#else switch (info->transform) { case JXFORM_NONE: /* Nothing to do */ @@ -878,7 +798,7 @@ jtransform_adjust_parameters (j_decompress_ptr srcinfo, trim_bottom_edge(dstinfo); break; } -#endif + /* Return the appropriate output data set */ if (info->workspace_coef_arrays != NULL) return info->workspace_coef_arrays; @@ -903,25 +823,6 @@ jtransform_execute_transformation (j_decompress_ptr srcinfo, { jvirt_barray_ptr *dst_coef_arrays = info->workspace_coef_arrays; -#ifdef NO_SWITCH - int choice = info->transform; - if (choice == JXFORM_NONE) { - } else if (choice == JXFORM_FLIP_H) { - do_flip_h(srcinfo, dstinfo, src_coef_arrays); - } else if (choice == JXFORM_FLIP_V) { - do_flip_v(srcinfo, dstinfo, src_coef_arrays, dst_coef_arrays); - } else if (choice == JXFORM_TRANSPOSE) { - do_transpose(srcinfo, dstinfo, src_coef_arrays, dst_coef_arrays); - } else if (choice == JXFORM_TRANSVERSE) { - do_transverse(srcinfo, dstinfo, src_coef_arrays, dst_coef_arrays); - } else if (choice == JXFORM_ROT_90) { - do_rot_90(srcinfo, dstinfo, src_coef_arrays, dst_coef_arrays); - } else if (choice == JXFORM_ROT_180) { - do_rot_180(srcinfo, dstinfo, src_coef_arrays, dst_coef_arrays); - } else if (choice == JXFORM_ROT_270) { - do_rot_270(srcinfo, dstinfo, src_coef_arrays, dst_coef_arrays); - } -#else switch (info->transform) { case JXFORM_NONE: break; @@ -947,7 +848,6 @@ jtransform_execute_transformation (j_decompress_ptr srcinfo, do_rot_270(srcinfo, dstinfo, src_coef_arrays, dst_coef_arrays); break; } -#endif } #endif /* TRANSFORMS_SUPPORTED */ diff --git a/test/monniaux/jpeg-6b/wrgif.c b/test/monniaux/jpeg-6b/wrgif.c index 65acd437..5fe83283 100644 --- a/test/monniaux/jpeg-6b/wrgif.c +++ b/test/monniaux/jpeg-6b/wrgif.c @@ -271,7 +271,7 @@ emit_header (gif_dest_ptr dinfo, int num_colors, JSAMPARRAY colormap) } } else { /* Create a gray-scale map of num_colors values, range 0..255 */ - put_3bytes(dinfo, INT_DIV((i * 255 + (num_colors-1)/2), (num_colors-1))); + put_3bytes(dinfo, (i * 255 + (num_colors-1)/2) / (num_colors-1)); } } else { /* fill out the map to a power of 2 */ diff --git a/test/monniaux/jpeg-6b/wrjpgcom.c b/test/monniaux/jpeg-6b/wrjpgcom.c index 7d10ee66..8c04b055 100644 --- a/test/monniaux/jpeg-6b/wrjpgcom.c +++ b/test/monniaux/jpeg-6b/wrjpgcom.c @@ -1,583 +1,583 @@ -/* - * wrjpgcom.c - * - * Copyright (C) 1994-1997, Thomas G. Lane. - * This file is part of the Independent JPEG Group's software. - * For conditions of distribution and use, see the accompanying README file. - * - * This file contains a very simple stand-alone application that inserts - * user-supplied text as a COM (comment) marker in a JFIF file. - * This may be useful as an example of the minimum logic needed to parse - * JPEG markers. - */ - -#define JPEG_CJPEG_DJPEG /* to get the command-line config symbols */ -#include "jinclude.h" /* get auto-config symbols, */ - -#ifndef HAVE_STDLIB_H /* should declare malloc() */ -extern void * malloc (); -#endif -#include /* to declare isupper(), tolower() */ -#ifdef USE_SETMODE -#include /* to declare setmode()'s parameter macros */ -/* If you have setmode() but not , just delete this line: */ -#include /* to declare setmode() */ -#endif - -#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ -#ifdef __MWERKS__ -#include /* Metrowerks needs this */ -#include /* ... and this */ -#endif -#ifdef THINK_C -#include /* Think declares it here */ -#endif -#endif - -#ifdef DONT_USE_B_MODE /* define mode parameters for fopen() */ -#define READ_BINARY "r" -#define WRITE_BINARY "w" -#else -#ifdef VMS /* VMS is very nonstandard */ -#define READ_BINARY "rb", "ctx=stm" -#define WRITE_BINARY "wb", "ctx=stm" -#else /* standard ANSI-compliant case */ -#define READ_BINARY "rb" -#define WRITE_BINARY "wb" -#endif -#endif - -#ifndef EXIT_FAILURE /* define exit() codes if not provided */ -#define EXIT_FAILURE 1 -#endif -#ifndef EXIT_SUCCESS -#ifdef VMS -#define EXIT_SUCCESS 1 /* VMS is very nonstandard */ -#else -#define EXIT_SUCCESS 0 -#endif -#endif - -/* Reduce this value if your malloc() can't allocate blocks up to 64K. - * On DOS, compiling in large model is usually a better solution. - */ - -#ifndef MAX_COM_LENGTH -#define MAX_COM_LENGTH 65000L /* must be <= 65533 in any case */ -#endif - - -/* - * These macros are used to read the input file and write the output file. - * To reuse this code in another application, you might need to change these. - */ - -static FILE * infile; /* input JPEG file */ - -/* Return next input byte, or EOF if no more */ -#define NEXTBYTE() getc(infile) - -static FILE * outfile; /* output JPEG file */ - -/* Emit an output byte */ -#define PUTBYTE(x) putc((x), outfile) - - -/* Error exit handler */ -#define ERREXIT(msg) (fprintf(stderr, "%s\n", msg), exit(EXIT_FAILURE)) - - -/* Read one byte, testing for EOF */ -static int -read_1_byte (void) -{ - int c; - - c = NEXTBYTE(); - if (c == EOF) - ERREXIT("Premature EOF in JPEG file"); - return c; -} - -/* Read 2 bytes, convert to unsigned int */ -/* All 2-byte quantities in JPEG markers are MSB first */ -static unsigned int -read_2_bytes (void) -{ - int c1, c2; - - c1 = NEXTBYTE(); - if (c1 == EOF) - ERREXIT("Premature EOF in JPEG file"); - c2 = NEXTBYTE(); - if (c2 == EOF) - ERREXIT("Premature EOF in JPEG file"); - return (((unsigned int) c1) << 8) + ((unsigned int) c2); -} - - -/* Routines to write data to output file */ - -static void -write_1_byte (int c) -{ - PUTBYTE(c); -} - -static void -write_2_bytes (unsigned int val) -{ - PUTBYTE((val >> 8) & 0xFF); - PUTBYTE(val & 0xFF); -} - -static void -write_marker (int marker) -{ - PUTBYTE(0xFF); - PUTBYTE(marker); -} - -static void -copy_rest_of_file (void) -{ - int c; - - while ((c = NEXTBYTE()) != EOF) - PUTBYTE(c); -} - - -/* - * JPEG markers consist of one or more 0xFF bytes, followed by a marker - * code byte (which is not an FF). Here are the marker codes of interest - * in this program. (See jdmarker.c for a more complete list.) - */ - -#define M_SOF0 0xC0 /* Start Of Frame N */ -#define M_SOF1 0xC1 /* N indicates which compression process */ -#define M_SOF2 0xC2 /* Only SOF0-SOF2 are now in common use */ -#define M_SOF3 0xC3 -#define M_SOF5 0xC5 /* NB: codes C4 and CC are NOT SOF markers */ -#define M_SOF6 0xC6 -#define M_SOF7 0xC7 -#define M_SOF9 0xC9 -#define M_SOF10 0xCA -#define M_SOF11 0xCB -#define M_SOF13 0xCD -#define M_SOF14 0xCE -#define M_SOF15 0xCF -#define M_SOI 0xD8 /* Start Of Image (beginning of datastream) */ -#define M_EOI 0xD9 /* End Of Image (end of datastream) */ -#define M_SOS 0xDA /* Start Of Scan (begins compressed data) */ -#define M_COM 0xFE /* COMment */ - - -/* - * Find the next JPEG marker and return its marker code. - * We expect at least one FF byte, possibly more if the compressor used FFs - * to pad the file. (Padding FFs will NOT be replicated in the output file.) - * There could also be non-FF garbage between markers. The treatment of such - * garbage is unspecified; we choose to skip over it but emit a warning msg. - * NB: this routine must not be used after seeing SOS marker, since it will - * not deal correctly with FF/00 sequences in the compressed image data... - */ - -static int -next_marker (void) -{ - int c; - int discarded_bytes = 0; - - /* Find 0xFF byte; count and skip any non-FFs. */ - c = read_1_byte(); - while (c != 0xFF) { - discarded_bytes++; - c = read_1_byte(); - } - /* Get marker code byte, swallowing any duplicate FF bytes. Extra FFs - * are legal as pad bytes, so don't count them in discarded_bytes. - */ - do { - c = read_1_byte(); - } while (c == 0xFF); - - if (discarded_bytes != 0) { - fprintf(stderr, "Warning: garbage data found in JPEG file\n"); - } - - return c; -} - - -/* - * Read the initial marker, which should be SOI. - * For a JFIF file, the first two bytes of the file should be literally - * 0xFF M_SOI. To be more general, we could use next_marker, but if the - * input file weren't actually JPEG at all, next_marker might read the whole - * file and then return a misleading error message... - */ - -static int -first_marker (void) -{ - int c1, c2; - - c1 = NEXTBYTE(); - c2 = NEXTBYTE(); - if (c1 != 0xFF || c2 != M_SOI) - ERREXIT("Not a JPEG file"); - return c2; -} - - -/* - * Most types of marker are followed by a variable-length parameter segment. - * This routine skips over the parameters for any marker we don't otherwise - * want to process. - * Note that we MUST skip the parameter segment explicitly in order not to - * be fooled by 0xFF bytes that might appear within the parameter segment; - * such bytes do NOT introduce new markers. - */ - -static void -copy_variable (void) -/* Copy an unknown or uninteresting variable-length marker */ -{ - unsigned int length; - - /* Get the marker parameter length count */ - length = read_2_bytes(); - write_2_bytes(length); - /* Length includes itself, so must be at least 2 */ - if (length < 2) - ERREXIT("Erroneous JPEG marker length"); - length -= 2; - /* Skip over the remaining bytes */ - while (length > 0) { - write_1_byte(read_1_byte()); - length--; - } -} - -static void -skip_variable (void) -/* Skip over an unknown or uninteresting variable-length marker */ -{ - unsigned int length; - - /* Get the marker parameter length count */ - length = read_2_bytes(); - /* Length includes itself, so must be at least 2 */ - if (length < 2) - ERREXIT("Erroneous JPEG marker length"); - length -= 2; - /* Skip over the remaining bytes */ - while (length > 0) { - (void) read_1_byte(); - length--; - } -} - - -/* - * Parse the marker stream until SOFn or EOI is seen; - * copy data to output, but discard COM markers unless keep_COM is true. - */ - -static int -scan_JPEG_header (int keep_COM) -{ - int marker; - - /* Expect SOI at start of file */ - if (first_marker() != M_SOI) - ERREXIT("Expected SOI marker first"); - write_marker(M_SOI); - - /* Scan miscellaneous markers until we reach SOFn. */ - for (;;) { - marker = next_marker(); - switch (marker) { - /* Note that marker codes 0xC4, 0xC8, 0xCC are not, and must not be, - * treated as SOFn. C4 in particular is actually DHT. - */ - case M_SOF0: /* Baseline */ - case M_SOF1: /* Extended sequential, Huffman */ - case M_SOF2: /* Progressive, Huffman */ - case M_SOF3: /* Lossless, Huffman */ - case M_SOF5: /* Differential sequential, Huffman */ - case M_SOF6: /* Differential progressive, Huffman */ - case M_SOF7: /* Differential lossless, Huffman */ - case M_SOF9: /* Extended sequential, arithmetic */ - case M_SOF10: /* Progressive, arithmetic */ - case M_SOF11: /* Lossless, arithmetic */ - case M_SOF13: /* Differential sequential, arithmetic */ - case M_SOF14: /* Differential progressive, arithmetic */ - case M_SOF15: /* Differential lossless, arithmetic */ - return marker; - - case M_SOS: /* should not see compressed data before SOF */ - ERREXIT("SOS without prior SOFn"); - break; - - case M_EOI: /* in case it's a tables-only JPEG stream */ - return marker; - - case M_COM: /* Existing COM: conditionally discard */ - if (keep_COM) { - write_marker(marker); - copy_variable(); - } else { - skip_variable(); - } - break; - - default: /* Anything else just gets copied */ - write_marker(marker); - copy_variable(); /* we assume it has a parameter count... */ - break; - } - } /* end loop */ -} - - -/* Command line parsing code */ - -static const char * progname; /* program name for error messages */ - - -static void -usage (void) -/* complain about bad command line */ -{ - fprintf(stderr, "wrjpgcom inserts a textual comment in a JPEG file.\n"); - fprintf(stderr, "You can add to or replace any existing comment(s).\n"); - - fprintf(stderr, "Usage: %s [switches] ", progname); -#ifdef TWO_FILE_COMMANDLINE - fprintf(stderr, "inputfile outputfile\n"); -#else - fprintf(stderr, "[inputfile]\n"); -#endif - - fprintf(stderr, "Switches (names may be abbreviated):\n"); - fprintf(stderr, " -replace Delete any existing comments\n"); - fprintf(stderr, " -comment \"text\" Insert comment with given text\n"); - fprintf(stderr, " -cfile name Read comment from named file\n"); - fprintf(stderr, "Notice that you must put quotes around the comment text\n"); - fprintf(stderr, "when you use -comment.\n"); - fprintf(stderr, "If you do not give either -comment or -cfile on the command line,\n"); - fprintf(stderr, "then the comment text is read from standard input.\n"); - fprintf(stderr, "It can be multiple lines, up to %u characters total.\n", - (unsigned int) MAX_COM_LENGTH); -#ifndef TWO_FILE_COMMANDLINE - fprintf(stderr, "You must specify an input JPEG file name when supplying\n"); - fprintf(stderr, "comment text from standard input.\n"); -#endif - - exit(EXIT_FAILURE); -} - - -static int -keymatch (char * arg, const char * keyword, int minchars) -/* Case-insensitive matching of (possibly abbreviated) keyword switches. */ -/* keyword is the constant keyword (must be lower case already), */ -/* minchars is length of minimum legal abbreviation. */ -{ - register int ca, ck; - register int nmatched = 0; - - while ((ca = *arg++) != '\0') { - if ((ck = *keyword++) == '\0') - return 0; /* arg longer than keyword, no good */ - if (isupper(ca)) /* force arg to lcase (assume ck is already) */ - ca = tolower(ca); - if (ca != ck) - return 0; /* no good */ - nmatched++; /* count matched characters */ - } - /* reached end of argument; fail if it's too short for unique abbrev */ - if (nmatched < minchars) - return 0; - return 1; /* A-OK */ -} - - -/* - * The main program. - */ - -int -main (int argc, char **argv) -{ - int argn; - char * arg; - int keep_COM = 1; - char * comment_arg = NULL; - FILE * comment_file = NULL; - unsigned int comment_length = 0; - int marker; - - /* On Mac, fetch a command line. */ -#ifdef USE_CCOMMAND - argc = ccommand(&argv); -#endif - - progname = argv[0]; - if (progname == NULL || progname[0] == 0) - progname = "wrjpgcom"; /* in case C library doesn't provide it */ - - /* Parse switches, if any */ - for (argn = 1; argn < argc; argn++) { - arg = argv[argn]; - if (arg[0] != '-') - break; /* not switch, must be file name */ - arg++; /* advance over '-' */ - if (keymatch(arg, "replace", 1)) { - keep_COM = 0; - } else if (keymatch(arg, "cfile", 2)) { - if (++argn >= argc) usage(); - if ((comment_file = fopen(argv[argn], "r")) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, argv[argn]); - exit(EXIT_FAILURE); - } - } else if (keymatch(arg, "comment", 1)) { - if (++argn >= argc) usage(); - comment_arg = argv[argn]; - /* If the comment text starts with '"', then we are probably running - * under MS-DOG and must parse out the quoted string ourselves. Sigh. - */ - if (comment_arg[0] == '"') { - comment_arg = (char *) malloc((size_t) MAX_COM_LENGTH); - if (comment_arg == NULL) - ERREXIT("Insufficient memory"); - strcpy(comment_arg, argv[argn]+1); - for (;;) { - comment_length = (unsigned int) strlen(comment_arg); - if (comment_length > 0 && comment_arg[comment_length-1] == '"') { - comment_arg[comment_length-1] = '\0'; /* zap terminating quote */ - break; - } - if (++argn >= argc) - ERREXIT("Missing ending quote mark"); - strcat(comment_arg, " "); - strcat(comment_arg, argv[argn]); - } - } - comment_length = (unsigned int) strlen(comment_arg); - } else - usage(); - } - - /* Cannot use both -comment and -cfile. */ - if (comment_arg != NULL && comment_file != NULL) - usage(); - /* If there is neither -comment nor -cfile, we will read the comment text - * from stdin; in this case there MUST be an input JPEG file name. - */ - if (comment_arg == NULL && comment_file == NULL && argn >= argc) - usage(); - - /* Open the input file. */ - if (argn < argc) { - if ((infile = fopen(argv[argn], READ_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, argv[argn]); - exit(EXIT_FAILURE); - } - } else { - /* default input file is stdin */ -#ifdef USE_SETMODE /* need to hack file mode? */ - setmode(fileno(stdin), O_BINARY); -#endif -#ifdef USE_FDOPEN /* need to re-open in binary mode? */ - if ((infile = fdopen(fileno(stdin), READ_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open stdin\n", progname); - exit(EXIT_FAILURE); - } -#else - infile = stdin; -#endif - } - - /* Open the output file. */ -#ifdef TWO_FILE_COMMANDLINE - /* Must have explicit output file name */ - if (argn != argc-2) { - fprintf(stderr, "%s: must name one input and one output file\n", - progname); - usage(); - } - if ((outfile = fopen(argv[argn+1], WRITE_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open %s\n", progname, argv[argn+1]); - exit(EXIT_FAILURE); - } -#else - /* Unix style: expect zero or one file name */ - if (argn < argc-1) { - fprintf(stderr, "%s: only one input file\n", progname); - usage(); - } - /* default output file is stdout */ -#ifdef USE_SETMODE /* need to hack file mode? */ - setmode(fileno(stdout), O_BINARY); -#endif -#ifdef USE_FDOPEN /* need to re-open in binary mode? */ - if ((outfile = fdopen(fileno(stdout), WRITE_BINARY)) == NULL) { - fprintf(stderr, "%s: can't open stdout\n", progname); - exit(EXIT_FAILURE); - } -#else - outfile = stdout; -#endif -#endif /* TWO_FILE_COMMANDLINE */ - - /* Collect comment text from comment_file or stdin, if necessary */ - if (comment_arg == NULL) { - FILE * src_file; - int c; - - comment_arg = (char *) malloc((size_t) MAX_COM_LENGTH); - if (comment_arg == NULL) - ERREXIT("Insufficient memory"); - comment_length = 0; - src_file = (comment_file != NULL ? comment_file : stdin); - while ((c = getc(src_file)) != EOF) { - if (comment_length >= (unsigned int) MAX_COM_LENGTH) { - fprintf(stderr, "Comment text may not exceed %u bytes\n", - (unsigned int) MAX_COM_LENGTH); - exit(EXIT_FAILURE); - } - comment_arg[comment_length++] = (char) c; - } - if (comment_file != NULL) - fclose(comment_file); - } - - /* Copy JPEG headers until SOFn marker; - * we will insert the new comment marker just before SOFn. - * This (a) causes the new comment to appear after, rather than before, - * existing comments; and (b) ensures that comments come after any JFIF - * or JFXX markers, as required by the JFIF specification. - */ - marker = scan_JPEG_header(keep_COM); - /* Insert the new COM marker, but only if nonempty text has been supplied */ - if (comment_length > 0) { - write_marker(M_COM); - write_2_bytes(comment_length + 2); - while (comment_length > 0) { - write_1_byte(*comment_arg++); - comment_length--; - } - } - /* Duplicate the remainder of the source file. - * Note that any COM markers occuring after SOF will not be touched. - */ - write_marker(marker); - copy_rest_of_file(); - - /* All done. */ - exit(EXIT_SUCCESS); - return 0; /* suppress no-return-value warnings */ -} +/* + * wrjpgcom.c + * + * Copyright (C) 1994-1997, Thomas G. Lane. + * This file is part of the Independent JPEG Group's software. + * For conditions of distribution and use, see the accompanying README file. + * + * This file contains a very simple stand-alone application that inserts + * user-supplied text as a COM (comment) marker in a JFIF file. + * This may be useful as an example of the minimum logic needed to parse + * JPEG markers. + */ + +#define JPEG_CJPEG_DJPEG /* to get the command-line config symbols */ +#include "jinclude.h" /* get auto-config symbols, */ + +#ifndef HAVE_STDLIB_H /* should declare malloc() */ +extern void * malloc (); +#endif +#include /* to declare isupper(), tolower() */ +#ifdef USE_SETMODE +#include /* to declare setmode()'s parameter macros */ +/* If you have setmode() but not , just delete this line: */ +#include /* to declare setmode() */ +#endif + +#ifdef USE_CCOMMAND /* command-line reader for Macintosh */ +#ifdef __MWERKS__ +#include /* Metrowerks needs this */ +#include /* ... and this */ +#endif +#ifdef THINK_C +#include /* Think declares it here */ +#endif +#endif + +#ifdef DONT_USE_B_MODE /* define mode parameters for fopen() */ +#define READ_BINARY "r" +#define WRITE_BINARY "w" +#else +#ifdef VMS /* VMS is very nonstandard */ +#define READ_BINARY "rb", "ctx=stm" +#define WRITE_BINARY "wb", "ctx=stm" +#else /* standard ANSI-compliant case */ +#define READ_BINARY "rb" +#define WRITE_BINARY "wb" +#endif +#endif + +#ifndef EXIT_FAILURE /* define exit() codes if not provided */ +#define EXIT_FAILURE 1 +#endif +#ifndef EXIT_SUCCESS +#ifdef VMS +#define EXIT_SUCCESS 1 /* VMS is very nonstandard */ +#else +#define EXIT_SUCCESS 0 +#endif +#endif + +/* Reduce this value if your malloc() can't allocate blocks up to 64K. + * On DOS, compiling in large model is usually a better solution. + */ + +#ifndef MAX_COM_LENGTH +#define MAX_COM_LENGTH 65000L /* must be <= 65533 in any case */ +#endif + + +/* + * These macros are used to read the input file and write the output file. + * To reuse this code in another application, you might need to change these. + */ + +static FILE * infile; /* input JPEG file */ + +/* Return next input byte, or EOF if no more */ +#define NEXTBYTE() getc(infile) + +static FILE * outfile; /* output JPEG file */ + +/* Emit an output byte */ +#define PUTBYTE(x) putc((x), outfile) + + +/* Error exit handler */ +#define ERREXIT(msg) (fprintf(stderr, "%s\n", msg), exit(EXIT_FAILURE)) + + +/* Read one byte, testing for EOF */ +static int +read_1_byte (void) +{ + int c; + + c = NEXTBYTE(); + if (c == EOF) + ERREXIT("Premature EOF in JPEG file"); + return c; +} + +/* Read 2 bytes, convert to unsigned int */ +/* All 2-byte quantities in JPEG markers are MSB first */ +static unsigned int +read_2_bytes (void) +{ + int c1, c2; + + c1 = NEXTBYTE(); + if (c1 == EOF) + ERREXIT("Premature EOF in JPEG file"); + c2 = NEXTBYTE(); + if (c2 == EOF) + ERREXIT("Premature EOF in JPEG file"); + return (((unsigned int) c1) << 8) + ((unsigned int) c2); +} + + +/* Routines to write data to output file */ + +static void +write_1_byte (int c) +{ + PUTBYTE(c); +} + +static void +write_2_bytes (unsigned int val) +{ + PUTBYTE((val >> 8) & 0xFF); + PUTBYTE(val & 0xFF); +} + +static void +write_marker (int marker) +{ + PUTBYTE(0xFF); + PUTBYTE(marker); +} + +static void +copy_rest_of_file (void) +{ + int c; + + while ((c = NEXTBYTE()) != EOF) + PUTBYTE(c); +} + + +/* + * JPEG markers consist of one or more 0xFF bytes, followed by a marker + * code byte (which is not an FF). Here are the marker codes of interest + * in this program. (See jdmarker.c for a more complete list.) + */ + +#define M_SOF0 0xC0 /* Start Of Frame N */ +#define M_SOF1 0xC1 /* N indicates which compression process */ +#define M_SOF2 0xC2 /* Only SOF0-SOF2 are now in common use */ +#define M_SOF3 0xC3 +#define M_SOF5 0xC5 /* NB: codes C4 and CC are NOT SOF markers */ +#define M_SOF6 0xC6 +#define M_SOF7 0xC7 +#define M_SOF9 0xC9 +#define M_SOF10 0xCA +#define M_SOF11 0xCB +#define M_SOF13 0xCD +#define M_SOF14 0xCE +#define M_SOF15 0xCF +#define M_SOI 0xD8 /* Start Of Image (beginning of datastream) */ +#define M_EOI 0xD9 /* End Of Image (end of datastream) */ +#define M_SOS 0xDA /* Start Of Scan (begins compressed data) */ +#define M_COM 0xFE /* COMment */ + + +/* + * Find the next JPEG marker and return its marker code. + * We expect at least one FF byte, possibly more if the compressor used FFs + * to pad the file. (Padding FFs will NOT be replicated in the output file.) + * There could also be non-FF garbage between markers. The treatment of such + * garbage is unspecified; we choose to skip over it but emit a warning msg. + * NB: this routine must not be used after seeing SOS marker, since it will + * not deal correctly with FF/00 sequences in the compressed image data... + */ + +static int +next_marker (void) +{ + int c; + int discarded_bytes = 0; + + /* Find 0xFF byte; count and skip any non-FFs. */ + c = read_1_byte(); + while (c != 0xFF) { + discarded_bytes++; + c = read_1_byte(); + } + /* Get marker code byte, swallowing any duplicate FF bytes. Extra FFs + * are legal as pad bytes, so don't count them in discarded_bytes. + */ + do { + c = read_1_byte(); + } while (c == 0xFF); + + if (discarded_bytes != 0) { + fprintf(stderr, "Warning: garbage data found in JPEG file\n"); + } + + return c; +} + + +/* + * Read the initial marker, which should be SOI. + * For a JFIF file, the first two bytes of the file should be literally + * 0xFF M_SOI. To be more general, we could use next_marker, but if the + * input file weren't actually JPEG at all, next_marker might read the whole + * file and then return a misleading error message... + */ + +static int +first_marker (void) +{ + int c1, c2; + + c1 = NEXTBYTE(); + c2 = NEXTBYTE(); + if (c1 != 0xFF || c2 != M_SOI) + ERREXIT("Not a JPEG file"); + return c2; +} + + +/* + * Most types of marker are followed by a variable-length parameter segment. + * This routine skips over the parameters for any marker we don't otherwise + * want to process. + * Note that we MUST skip the parameter segment explicitly in order not to + * be fooled by 0xFF bytes that might appear within the parameter segment; + * such bytes do NOT introduce new markers. + */ + +static void +copy_variable (void) +/* Copy an unknown or uninteresting variable-length marker */ +{ + unsigned int length; + + /* Get the marker parameter length count */ + length = read_2_bytes(); + write_2_bytes(length); + /* Length includes itself, so must be at least 2 */ + if (length < 2) + ERREXIT("Erroneous JPEG marker length"); + length -= 2; + /* Skip over the remaining bytes */ + while (length > 0) { + write_1_byte(read_1_byte()); + length--; + } +} + +static void +skip_variable (void) +/* Skip over an unknown or uninteresting variable-length marker */ +{ + unsigned int length; + + /* Get the marker parameter length count */ + length = read_2_bytes(); + /* Length includes itself, so must be at least 2 */ + if (length < 2) + ERREXIT("Erroneous JPEG marker length"); + length -= 2; + /* Skip over the remaining bytes */ + while (length > 0) { + (void) read_1_byte(); + length--; + } +} + + +/* + * Parse the marker stream until SOFn or EOI is seen; + * copy data to output, but discard COM markers unless keep_COM is true. + */ + +static int +scan_JPEG_header (int keep_COM) +{ + int marker; + + /* Expect SOI at start of file */ + if (first_marker() != M_SOI) + ERREXIT("Expected SOI marker first"); + write_marker(M_SOI); + + /* Scan miscellaneous markers until we reach SOFn. */ + for (;;) { + marker = next_marker(); + switch (marker) { + /* Note that marker codes 0xC4, 0xC8, 0xCC are not, and must not be, + * treated as SOFn. C4 in particular is actually DHT. + */ + case M_SOF0: /* Baseline */ + case M_SOF1: /* Extended sequential, Huffman */ + case M_SOF2: /* Progressive, Huffman */ + case M_SOF3: /* Lossless, Huffman */ + case M_SOF5: /* Differential sequential, Huffman */ + case M_SOF6: /* Differential progressive, Huffman */ + case M_SOF7: /* Differential lossless, Huffman */ + case M_SOF9: /* Extended sequential, arithmetic */ + case M_SOF10: /* Progressive, arithmetic */ + case M_SOF11: /* Lossless, arithmetic */ + case M_SOF13: /* Differential sequential, arithmetic */ + case M_SOF14: /* Differential progressive, arithmetic */ + case M_SOF15: /* Differential lossless, arithmetic */ + return marker; + + case M_SOS: /* should not see compressed data before SOF */ + ERREXIT("SOS without prior SOFn"); + break; + + case M_EOI: /* in case it's a tables-only JPEG stream */ + return marker; + + case M_COM: /* Existing COM: conditionally discard */ + if (keep_COM) { + write_marker(marker); + copy_variable(); + } else { + skip_variable(); + } + break; + + default: /* Anything else just gets copied */ + write_marker(marker); + copy_variable(); /* we assume it has a parameter count... */ + break; + } + } /* end loop */ +} + + +/* Command line parsing code */ + +static const char * progname; /* program name for error messages */ + + +static void +usage (void) +/* complain about bad command line */ +{ + fprintf(stderr, "wrjpgcom inserts a textual comment in a JPEG file.\n"); + fprintf(stderr, "You can add to or replace any existing comment(s).\n"); + + fprintf(stderr, "Usage: %s [switches] ", progname); +#ifdef TWO_FILE_COMMANDLINE + fprintf(stderr, "inputfile outputfile\n"); +#else + fprintf(stderr, "[inputfile]\n"); +#endif + + fprintf(stderr, "Switches (names may be abbreviated):\n"); + fprintf(stderr, " -replace Delete any existing comments\n"); + fprintf(stderr, " -comment \"text\" Insert comment with given text\n"); + fprintf(stderr, " -cfile name Read comment from named file\n"); + fprintf(stderr, "Notice that you must put quotes around the comment text\n"); + fprintf(stderr, "when you use -comment.\n"); + fprintf(stderr, "If you do not give either -comment or -cfile on the command line,\n"); + fprintf(stderr, "then the comment text is read from standard input.\n"); + fprintf(stderr, "It can be multiple lines, up to %u characters total.\n", + (unsigned int) MAX_COM_LENGTH); +#ifndef TWO_FILE_COMMANDLINE + fprintf(stderr, "You must specify an input JPEG file name when supplying\n"); + fprintf(stderr, "comment text from standard input.\n"); +#endif + + exit(EXIT_FAILURE); +} + + +static int +keymatch (char * arg, const char * keyword, int minchars) +/* Case-insensitive matching of (possibly abbreviated) keyword switches. */ +/* keyword is the constant keyword (must be lower case already), */ +/* minchars is length of minimum legal abbreviation. */ +{ + register int ca, ck; + register int nmatched = 0; + + while ((ca = *arg++) != '\0') { + if ((ck = *keyword++) == '\0') + return 0; /* arg longer than keyword, no good */ + if (isupper(ca)) /* force arg to lcase (assume ck is already) */ + ca = tolower(ca); + if (ca != ck) + return 0; /* no good */ + nmatched++; /* count matched characters */ + } + /* reached end of argument; fail if it's too short for unique abbrev */ + if (nmatched < minchars) + return 0; + return 1; /* A-OK */ +} + + +/* + * The main program. + */ + +int +main (int argc, char **argv) +{ + int argn; + char * arg; + int keep_COM = 1; + char * comment_arg = NULL; + FILE * comment_file = NULL; + unsigned int comment_length = 0; + int marker; + + /* On Mac, fetch a command line. */ +#ifdef USE_CCOMMAND + argc = ccommand(&argv); +#endif + + progname = argv[0]; + if (progname == NULL || progname[0] == 0) + progname = "wrjpgcom"; /* in case C library doesn't provide it */ + + /* Parse switches, if any */ + for (argn = 1; argn < argc; argn++) { + arg = argv[argn]; + if (arg[0] != '-') + break; /* not switch, must be file name */ + arg++; /* advance over '-' */ + if (keymatch(arg, "replace", 1)) { + keep_COM = 0; + } else if (keymatch(arg, "cfile", 2)) { + if (++argn >= argc) usage(); + if ((comment_file = fopen(argv[argn], "r")) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, argv[argn]); + exit(EXIT_FAILURE); + } + } else if (keymatch(arg, "comment", 1)) { + if (++argn >= argc) usage(); + comment_arg = argv[argn]; + /* If the comment text starts with '"', then we are probably running + * under MS-DOG and must parse out the quoted string ourselves. Sigh. + */ + if (comment_arg[0] == '"') { + comment_arg = (char *) malloc((size_t) MAX_COM_LENGTH); + if (comment_arg == NULL) + ERREXIT("Insufficient memory"); + strcpy(comment_arg, argv[argn]+1); + for (;;) { + comment_length = (unsigned int) strlen(comment_arg); + if (comment_length > 0 && comment_arg[comment_length-1] == '"') { + comment_arg[comment_length-1] = '\0'; /* zap terminating quote */ + break; + } + if (++argn >= argc) + ERREXIT("Missing ending quote mark"); + strcat(comment_arg, " "); + strcat(comment_arg, argv[argn]); + } + } + comment_length = (unsigned int) strlen(comment_arg); + } else + usage(); + } + + /* Cannot use both -comment and -cfile. */ + if (comment_arg != NULL && comment_file != NULL) + usage(); + /* If there is neither -comment nor -cfile, we will read the comment text + * from stdin; in this case there MUST be an input JPEG file name. + */ + if (comment_arg == NULL && comment_file == NULL && argn >= argc) + usage(); + + /* Open the input file. */ + if (argn < argc) { + if ((infile = fopen(argv[argn], READ_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, argv[argn]); + exit(EXIT_FAILURE); + } + } else { + /* default input file is stdin */ +#ifdef USE_SETMODE /* need to hack file mode? */ + setmode(fileno(stdin), O_BINARY); +#endif +#ifdef USE_FDOPEN /* need to re-open in binary mode? */ + if ((infile = fdopen(fileno(stdin), READ_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open stdin\n", progname); + exit(EXIT_FAILURE); + } +#else + infile = stdin; +#endif + } + + /* Open the output file. */ +#ifdef TWO_FILE_COMMANDLINE + /* Must have explicit output file name */ + if (argn != argc-2) { + fprintf(stderr, "%s: must name one input and one output file\n", + progname); + usage(); + } + if ((outfile = fopen(argv[argn+1], WRITE_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open %s\n", progname, argv[argn+1]); + exit(EXIT_FAILURE); + } +#else + /* Unix style: expect zero or one file name */ + if (argn < argc-1) { + fprintf(stderr, "%s: only one input file\n", progname); + usage(); + } + /* default output file is stdout */ +#ifdef USE_SETMODE /* need to hack file mode? */ + setmode(fileno(stdout), O_BINARY); +#endif +#ifdef USE_FDOPEN /* need to re-open in binary mode? */ + if ((outfile = fdopen(fileno(stdout), WRITE_BINARY)) == NULL) { + fprintf(stderr, "%s: can't open stdout\n", progname); + exit(EXIT_FAILURE); + } +#else + outfile = stdout; +#endif +#endif /* TWO_FILE_COMMANDLINE */ + + /* Collect comment text from comment_file or stdin, if necessary */ + if (comment_arg == NULL) { + FILE * src_file; + int c; + + comment_arg = (char *) malloc((size_t) MAX_COM_LENGTH); + if (comment_arg == NULL) + ERREXIT("Insufficient memory"); + comment_length = 0; + src_file = (comment_file != NULL ? comment_file : stdin); + while ((c = getc(src_file)) != EOF) { + if (comment_length >= (unsigned int) MAX_COM_LENGTH) { + fprintf(stderr, "Comment text may not exceed %u bytes\n", + (unsigned int) MAX_COM_LENGTH); + exit(EXIT_FAILURE); + } + comment_arg[comment_length++] = (char) c; + } + if (comment_file != NULL) + fclose(comment_file); + } + + /* Copy JPEG headers until SOFn marker; + * we will insert the new comment marker just before SOFn. + * This (a) causes the new comment to appear after, rather than before, + * existing comments; and (b) ensures that comments come after any JFIF + * or JFXX markers, as required by the JFIF specification. + */ + marker = scan_JPEG_header(keep_COM); + /* Insert the new COM marker, but only if nonempty text has been supplied */ + if (comment_length > 0) { + write_marker(M_COM); + write_2_bytes(comment_length + 2); + while (comment_length > 0) { + write_1_byte(*comment_arg++); + comment_length--; + } + } + /* Duplicate the remainder of the source file. + * Note that any COM markers occuring after SOF will not be touched. + */ + write_marker(marker); + copy_rest_of_file(); + + /* All done. */ + exit(EXIT_SUCCESS); + return 0; /* suppress no-return-value warnings */ +} -- cgit From 85dc3bff6792d0cc2e1f7b3cc2446cd01428907b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 20 Mar 2019 13:33:45 +0100 Subject: Proof of div32/mod32/divf32/divf64 lemmas --- backend/Selectionproof.v | 6 +++--- mppa_k1c/SelectOpproof.v | 27 +++++++++++++++++++++------ 2 files changed, 24 insertions(+), 9 deletions(-) diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v index dc01ad20..5a0fafaf 100644 --- a/backend/Selectionproof.v +++ b/backend/Selectionproof.v @@ -86,7 +86,7 @@ Lemma get_helpers_correct: forall p hf, get_helpers (prog_defmap p) = OK hf -> helper_functions_declared p hf. Proof. - intros. monadInv H. red; simpl. auto 20 using lookup_helper_correct. + intros. monadInv H. red; simpl. auto 22 using lookup_helper_correct. Qed. Theorem transf_program_match: @@ -106,7 +106,7 @@ Proof. { unfold helper_declared; intros. destruct (prog_defmap_linkorder _ _ _ _ H0 H1) as (gd & P & Q). inv Q. inv H3. auto. } - red in H. decompose [Logic.and] H; clear H. red; auto 20. + red in H. decompose [Logic.and] H; clear H. red; auto 22. Qed. (** * Correctness of the instruction selection functions for expressions *) @@ -164,7 +164,7 @@ Proof. generalize (match_program_defmap _ _ _ _ _ TRANSF id). unfold Cminor.fundef; rewrite H; intros R; inv R. inv H2. destruct H4 as (cu & A & B). monadInv B. auto. } - unfold helper_functions_declared; intros. decompose [Logic.and] H; clear H. auto 20. + unfold helper_functions_declared; intros. decompose [Logic.and] H; clear H. auto 22. Qed. Section CMCONSTR. diff --git a/mppa_k1c/SelectOpproof.v b/mppa_k1c/SelectOpproof.v index ca6c342a..a4a5f72b 100644 --- a/mppa_k1c/SelectOpproof.v +++ b/mppa_k1c/SelectOpproof.v @@ -69,6 +69,8 @@ Axiom i64_helpers_correct : /\ (forall x y z, Val.divu x y = Some z -> external_implements "__compcert_i32_udiv" sig_ii_i (x::y::nil) z) /\ (forall x y z, Val.mods x y = Some z -> external_implements "__compcert_i32_smod" sig_ii_i (x::y::nil) z) /\ (forall x y z, Val.modu x y = Some z -> external_implements "__compcert_i32_umod" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.divfs x y = z -> external_implements "__compcert_f32_div" sig_ss_s (x::y::nil) z) + /\ (forall x y z, Val.divf x y = z -> external_implements "__compcert_f64_div" sig_ff_f (x::y::nil) z) . Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop := @@ -93,7 +95,10 @@ Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F) /\ helper_declared p i32_sdiv "__compcert_i32_sdiv" sig_ii_i /\ helper_declared p i32_udiv "__compcert_i32_udiv" sig_ii_i /\ helper_declared p i32_smod "__compcert_i32_smod" sig_ii_i - /\ helper_declared p i32_umod "__compcert_i32_umod" sig_ii_i. + /\ helper_declared p i32_umod "__compcert_i32_umod" sig_ii_i + /\ helper_declared p f32_div "__compcert_f32_div" sig_ss_s + /\ helper_declared p f64_div "__compcert_f64_div" sig_ff_f +. (** * Useful lemmas and tactics *) @@ -672,7 +677,9 @@ Theorem eval_mods_base: Val.mods x y = Some z -> exists v, eval_expr ge sp e m le (mods_base a b) v /\ Val.lessdef z v. Proof. -Admitted. + intros; unfold mods_base. + econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto. +Qed. Theorem eval_divu_base: forall le a b x y z, @@ -681,7 +688,9 @@ Theorem eval_divu_base: Val.divu x y = Some z -> exists v, eval_expr ge sp e m le (divu_base a b) v /\ Val.lessdef z v. Proof. -Admitted. + intros; unfold divu_base. + econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto. +Qed. Theorem eval_modu_base: forall le a b x y z, @@ -690,7 +699,9 @@ Theorem eval_modu_base: Val.modu x y = Some z -> exists v, eval_expr ge sp e m le (modu_base a b) v /\ Val.lessdef z v. Proof. -Admitted. + intros; unfold modu_base. + econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto. +Qed. Theorem eval_shrximm: forall le a n x z, @@ -1077,7 +1088,9 @@ Theorem eval_divf_base: eval_expr ge sp e m le b y -> exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v. Proof. -Admitted. + intros; unfold divf_base. + econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto. +Qed. Theorem eval_divfs_base: forall le a b x y, @@ -1085,5 +1098,7 @@ Theorem eval_divfs_base: eval_expr ge sp e m le b y -> exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v. Proof. -Admitted. + intros; unfold divfs_base. + econstructor; split. eapply eval_helper_2; eauto. DeclHelper. UseHelper. auto. +Qed. End CMCONSTR. -- cgit From 9fe81a5bc203db0ad7e7a88fa27d522922a8057a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 13:35:29 +0100 Subject: rm need for gcc --- test/monniaux/BearSSL/conf/KalrayCompCert.mk | 1 - test/monniaux/BearSSL/mk/Rules.mk | 5 ++--- 2 files changed, 2 insertions(+), 4 deletions(-) diff --git a/test/monniaux/BearSSL/conf/KalrayCompCert.mk b/test/monniaux/BearSSL/conf/KalrayCompCert.mk index ad5bd2da..9c0e951c 100644 --- a/test/monniaux/BearSSL/conf/KalrayCompCert.mk +++ b/test/monniaux/BearSSL/conf/KalrayCompCert.mk @@ -38,7 +38,6 @@ MKDIR = mkdir -p # C compiler and flags. CC = ../../../ccomp -fstruct-passing -GCC=k1-mbr-gcc CFLAGS = -W -Wall -Wno-c11-extensions -O3 -D_POSIX_C_SOURCE=200909L CCOUT = -c -o diff --git a/test/monniaux/BearSSL/mk/Rules.mk b/test/monniaux/BearSSL/mk/Rules.mk index 97338c38..0da4558b 100644 --- a/test/monniaux/BearSSL/mk/Rules.mk +++ b/test/monniaux/BearSSL/mk/Rules.mk @@ -1248,16 +1248,15 @@ $(OBJDIR)$Pencode_rsa_rawder$O: src$Px509$Pencode_rsa_rawder.c $(HEADERSPRIV) $(OBJDIR)$Pskey_decoder$O: src$Px509$Pskey_decoder.c $(HEADERSPRIV) $(CC) $(CFLAGS) $(INCFLAGS) $(CCOUT)$(OBJDIR)$Pskey_decoder$O src$Px509$Pskey_decoder.c -# Omod: 32-bits modulo not supported yet. Please use 64-bits. $(OBJDIR)$Px509_decoder$O: src$Px509$Px509_decoder.c $(HEADERSPRIV) - $(GCC) $(CFLAGS) $(INCFLAGS) $(CCOUT)$(OBJDIR)$Px509_decoder$O src$Px509$Px509_decoder.c + $(CC) $(CFLAGS) $(INCFLAGS) $(CCOUT)$(OBJDIR)$Px509_decoder$O src$Px509$Px509_decoder.c $(OBJDIR)$Px509_knownkey$O: src$Px509$Px509_knownkey.c $(HEADERSPRIV) $(CC) $(CFLAGS) $(INCFLAGS) $(CCOUT)$(OBJDIR)$Px509_knownkey$O src$Px509$Px509_knownkey.c # nsl_op: Omod: 32-bits modulo not supported yet. Please use 64-bits. $(OBJDIR)$Px509_minimal$O: src$Px509$Px509_minimal.c $(HEADERSPRIV) - $(GCC) $(CFLAGS) $(INCFLAGS) $(CCOUT)$(OBJDIR)$Px509_minimal$O src$Px509$Px509_minimal.c + $(CC) $(CFLAGS) $(INCFLAGS) $(CCOUT)$(OBJDIR)$Px509_minimal$O src$Px509$Px509_minimal.c $(OBJDIR)$Px509_minimal_full$O: src$Px509$Px509_minimal_full.c $(HEADERSPRIV) $(CC) $(CFLAGS) $(INCFLAGS) $(CCOUT)$(OBJDIR)$Px509_minimal_full$O src$Px509$Px509_minimal_full.c -- cgit From 4c9c95b6a0ac8aa31abb1f7ab48c3f645c059bd6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 17:32:00 +0100 Subject: ocaml byterunner example --- test/monniaux/ocaml/Makefile.common | 30 + test/monniaux/ocaml/VERSION | 4 + test/monniaux/ocaml/byterun/.depend | 803 ++ test/monniaux/ocaml/byterun/Makefile | 234 + test/monniaux/ocaml/byterun/afl.c | 162 + test/monniaux/ocaml/byterun/alloc.c | 247 + test/monniaux/ocaml/byterun/array.c | 586 + test/monniaux/ocaml/byterun/backtrace.c | 349 + test/monniaux/ocaml/byterun/backtrace_prim.c | 457 + test/monniaux/ocaml/byterun/bigarray.c | 1223 ++ test/monniaux/ocaml/byterun/callback.c | 262 + test/monniaux/ocaml/byterun/caml/address_class.h | 85 + test/monniaux/ocaml/byterun/caml/alloc.h | 82 + test/monniaux/ocaml/byterun/caml/backtrace.h | 136 + test/monniaux/ocaml/byterun/caml/backtrace_prim.h | 91 + test/monniaux/ocaml/byterun/caml/bigarray.h | 134 + test/monniaux/ocaml/byterun/caml/callback.h | 63 + test/monniaux/ocaml/byterun/caml/compact.h | 31 + test/monniaux/ocaml/byterun/caml/compare.h | 25 + test/monniaux/ocaml/byterun/caml/compatibility.h | 375 + test/monniaux/ocaml/byterun/caml/config.h | 208 + test/monniaux/ocaml/byterun/caml/custom.h | 73 + test/monniaux/ocaml/byterun/caml/debugger.h | 117 + test/monniaux/ocaml/byterun/caml/dynlink.h | 46 + test/monniaux/ocaml/byterun/caml/exec.h | 65 + test/monniaux/ocaml/byterun/caml/fail.h | 144 + test/monniaux/ocaml/byterun/caml/finalise.h | 36 + test/monniaux/ocaml/byterun/caml/fix_code.h | 45 + test/monniaux/ocaml/byterun/caml/freelist.h | 38 + test/monniaux/ocaml/byterun/caml/gc.h | 79 + test/monniaux/ocaml/byterun/caml/gc_ctrl.h | 58 + test/monniaux/ocaml/byterun/caml/globroots.h | 31 + test/monniaux/ocaml/byterun/caml/hash.h | 39 + test/monniaux/ocaml/byterun/caml/hooks.h | 42 + test/monniaux/ocaml/byterun/caml/instrtrace.h | 35 + test/monniaux/ocaml/byterun/caml/instruct.h | 68 + test/monniaux/ocaml/byterun/caml/int64_emul.h | 293 + test/monniaux/ocaml/byterun/caml/int64_format.h | 111 + test/monniaux/ocaml/byterun/caml/int64_native.h | 67 + test/monniaux/ocaml/byterun/caml/interp.h | 37 + test/monniaux/ocaml/byterun/caml/intext.h | 207 + test/monniaux/ocaml/byterun/caml/io.h | 126 + test/monniaux/ocaml/byterun/caml/m.h | 86 + test/monniaux/ocaml/byterun/caml/major_gc.h | 95 + test/monniaux/ocaml/byterun/caml/md5.h | 47 + test/monniaux/ocaml/byterun/caml/memory.h | 595 + test/monniaux/ocaml/byterun/caml/minor_gc.h | 120 + test/monniaux/ocaml/byterun/caml/misc.h | 525 + test/monniaux/ocaml/byterun/caml/mlvalues.h | 377 + test/monniaux/ocaml/byterun/caml/osdeps.h | 152 + test/monniaux/ocaml/byterun/caml/prims.h | 40 + test/monniaux/ocaml/byterun/caml/printexc.h | 35 + test/monniaux/ocaml/byterun/caml/reverse.h | 92 + test/monniaux/ocaml/byterun/caml/roots.h | 44 + test/monniaux/ocaml/byterun/caml/s.h | 217 + test/monniaux/ocaml/byterun/caml/signals.h | 59 + test/monniaux/ocaml/byterun/caml/signals_machdep.h | 74 + test/monniaux/ocaml/byterun/caml/spacetime.h | 203 + test/monniaux/ocaml/byterun/caml/stack.h | 124 + test/monniaux/ocaml/byterun/caml/stacks.h | 46 + test/monniaux/ocaml/byterun/caml/startup.h | 52 + test/monniaux/ocaml/byterun/caml/startup_aux.h | 44 + test/monniaux/ocaml/byterun/caml/sys.h | 45 + test/monniaux/ocaml/byterun/caml/ui.h | 32 + test/monniaux/ocaml/byterun/caml/weak.h | 93 + test/monniaux/ocaml/byterun/compact.c | 562 + test/monniaux/ocaml/byterun/compare.c | 363 + test/monniaux/ocaml/byterun/custom.c | 124 + test/monniaux/ocaml/byterun/debugger.c | 454 + test/monniaux/ocaml/byterun/dynlink.c | 300 + test/monniaux/ocaml/byterun/extern.c | 925 ++ test/monniaux/ocaml/byterun/fail.c | 204 + test/monniaux/ocaml/byterun/finalise.c | 445 + test/monniaux/ocaml/byterun/fix_code.c | 195 + test/monniaux/ocaml/byterun/floats.c | 720 + test/monniaux/ocaml/byterun/freelist.c | 621 + test/monniaux/ocaml/byterun/gc_ctrl.c | 691 + test/monniaux/ocaml/byterun/globroots.c | 291 + test/monniaux/ocaml/byterun/hash.c | 419 + test/monniaux/ocaml/byterun/instrtrace.c | 269 + test/monniaux/ocaml/byterun/intern.c | 1048 ++ test/monniaux/ocaml/byterun/interp.c | 1172 ++ test/monniaux/ocaml/byterun/ints.c | 830 ++ test/monniaux/ocaml/byterun/io.c | 828 ++ test/monniaux/ocaml/byterun/lexing.c | 233 + test/monniaux/ocaml/byterun/main.c | 47 + test/monniaux/ocaml/byterun/major_gc.c | 943 ++ test/monniaux/ocaml/byterun/make.log | 17 + test/monniaux/ocaml/byterun/make2.log | 321 + test/monniaux/ocaml/byterun/md5.c | 325 + test/monniaux/ocaml/byterun/memory.c | 1016 ++ test/monniaux/ocaml/byterun/meta.c | 234 + test/monniaux/ocaml/byterun/minor_gc.c | 558 + test/monniaux/ocaml/byterun/misc.c | 276 + test/monniaux/ocaml/byterun/obj.c | 390 + test/monniaux/ocaml/byterun/parsing.c | 304 + test/monniaux/ocaml/byterun/printexc.c | 155 + test/monniaux/ocaml/byterun/roots.c | 120 + test/monniaux/ocaml/byterun/signals.c | 398 + test/monniaux/ocaml/byterun/signals_byt.c | 101 + test/monniaux/ocaml/byterun/spacetime.c | 38 + test/monniaux/ocaml/byterun/stacks.c | 119 + test/monniaux/ocaml/byterun/startup.c | 531 + test/monniaux/ocaml/byterun/startup_aux.c | 168 + test/monniaux/ocaml/byterun/str.c | 474 + test/monniaux/ocaml/byterun/sys.c | 723 + test/monniaux/ocaml/byterun/toto | 13996 +++++++++++++++++++ test/monniaux/ocaml/byterun/unix.c | 459 + test/monniaux/ocaml/byterun/weak.c | 427 + test/monniaux/ocaml/byterun/win32.c | 1019 ++ test/monniaux/ocaml/config/Makefile | 109 + test/monniaux/ocaml/config/Makefile-templ | 202 + test/monniaux/ocaml/config/m-nt.h | 65 + test/monniaux/ocaml/config/m-templ.h | 83 + test/monniaux/ocaml/config/s-nt.h | 41 + test/monniaux/ocaml/config/s-templ.h | 214 + test/monniaux/ocaml/examples/quicksort | Bin 0 -> 146138 bytes test/monniaux/ocaml/examples/quicksort.ml | 11 + test/monniaux/ocaml/tools/make-version-header.sh | 55 + 119 files changed, 45174 insertions(+) create mode 100644 test/monniaux/ocaml/Makefile.common create mode 100644 test/monniaux/ocaml/VERSION create mode 100644 test/monniaux/ocaml/byterun/.depend create mode 100644 test/monniaux/ocaml/byterun/Makefile create mode 100644 test/monniaux/ocaml/byterun/afl.c create mode 100644 test/monniaux/ocaml/byterun/alloc.c create mode 100644 test/monniaux/ocaml/byterun/array.c create mode 100644 test/monniaux/ocaml/byterun/backtrace.c create mode 100644 test/monniaux/ocaml/byterun/backtrace_prim.c create mode 100644 test/monniaux/ocaml/byterun/bigarray.c create mode 100644 test/monniaux/ocaml/byterun/callback.c create mode 100644 test/monniaux/ocaml/byterun/caml/address_class.h create mode 100644 test/monniaux/ocaml/byterun/caml/alloc.h create mode 100644 test/monniaux/ocaml/byterun/caml/backtrace.h create mode 100644 test/monniaux/ocaml/byterun/caml/backtrace_prim.h create mode 100644 test/monniaux/ocaml/byterun/caml/bigarray.h create mode 100644 test/monniaux/ocaml/byterun/caml/callback.h create mode 100644 test/monniaux/ocaml/byterun/caml/compact.h create mode 100644 test/monniaux/ocaml/byterun/caml/compare.h create mode 100644 test/monniaux/ocaml/byterun/caml/compatibility.h create mode 100644 test/monniaux/ocaml/byterun/caml/config.h create mode 100644 test/monniaux/ocaml/byterun/caml/custom.h create mode 100644 test/monniaux/ocaml/byterun/caml/debugger.h create mode 100644 test/monniaux/ocaml/byterun/caml/dynlink.h create mode 100644 test/monniaux/ocaml/byterun/caml/exec.h create mode 100644 test/monniaux/ocaml/byterun/caml/fail.h create mode 100644 test/monniaux/ocaml/byterun/caml/finalise.h create mode 100644 test/monniaux/ocaml/byterun/caml/fix_code.h create mode 100644 test/monniaux/ocaml/byterun/caml/freelist.h create mode 100644 test/monniaux/ocaml/byterun/caml/gc.h create mode 100644 test/monniaux/ocaml/byterun/caml/gc_ctrl.h create mode 100644 test/monniaux/ocaml/byterun/caml/globroots.h create mode 100644 test/monniaux/ocaml/byterun/caml/hash.h create mode 100644 test/monniaux/ocaml/byterun/caml/hooks.h create mode 100644 test/monniaux/ocaml/byterun/caml/instrtrace.h create mode 100644 test/monniaux/ocaml/byterun/caml/instruct.h create mode 100644 test/monniaux/ocaml/byterun/caml/int64_emul.h create mode 100644 test/monniaux/ocaml/byterun/caml/int64_format.h create mode 100644 test/monniaux/ocaml/byterun/caml/int64_native.h create mode 100644 test/monniaux/ocaml/byterun/caml/interp.h create mode 100644 test/monniaux/ocaml/byterun/caml/intext.h create mode 100644 test/monniaux/ocaml/byterun/caml/io.h create mode 100644 test/monniaux/ocaml/byterun/caml/m.h create mode 100644 test/monniaux/ocaml/byterun/caml/major_gc.h create mode 100644 test/monniaux/ocaml/byterun/caml/md5.h create mode 100644 test/monniaux/ocaml/byterun/caml/memory.h create mode 100644 test/monniaux/ocaml/byterun/caml/minor_gc.h create mode 100644 test/monniaux/ocaml/byterun/caml/misc.h create mode 100644 test/monniaux/ocaml/byterun/caml/mlvalues.h create mode 100644 test/monniaux/ocaml/byterun/caml/osdeps.h create mode 100644 test/monniaux/ocaml/byterun/caml/prims.h create mode 100644 test/monniaux/ocaml/byterun/caml/printexc.h create mode 100644 test/monniaux/ocaml/byterun/caml/reverse.h create mode 100644 test/monniaux/ocaml/byterun/caml/roots.h create mode 100644 test/monniaux/ocaml/byterun/caml/s.h create mode 100644 test/monniaux/ocaml/byterun/caml/signals.h create mode 100644 test/monniaux/ocaml/byterun/caml/signals_machdep.h create mode 100644 test/monniaux/ocaml/byterun/caml/spacetime.h create mode 100644 test/monniaux/ocaml/byterun/caml/stack.h create mode 100644 test/monniaux/ocaml/byterun/caml/stacks.h create mode 100644 test/monniaux/ocaml/byterun/caml/startup.h create mode 100644 test/monniaux/ocaml/byterun/caml/startup_aux.h create mode 100644 test/monniaux/ocaml/byterun/caml/sys.h create mode 100644 test/monniaux/ocaml/byterun/caml/ui.h create mode 100644 test/monniaux/ocaml/byterun/caml/weak.h create mode 100644 test/monniaux/ocaml/byterun/compact.c create mode 100644 test/monniaux/ocaml/byterun/compare.c create mode 100644 test/monniaux/ocaml/byterun/custom.c create mode 100644 test/monniaux/ocaml/byterun/debugger.c create mode 100644 test/monniaux/ocaml/byterun/dynlink.c create mode 100644 test/monniaux/ocaml/byterun/extern.c create mode 100644 test/monniaux/ocaml/byterun/fail.c create mode 100644 test/monniaux/ocaml/byterun/finalise.c create mode 100644 test/monniaux/ocaml/byterun/fix_code.c create mode 100644 test/monniaux/ocaml/byterun/floats.c create mode 100644 test/monniaux/ocaml/byterun/freelist.c create mode 100644 test/monniaux/ocaml/byterun/gc_ctrl.c create mode 100644 test/monniaux/ocaml/byterun/globroots.c create mode 100644 test/monniaux/ocaml/byterun/hash.c create mode 100644 test/monniaux/ocaml/byterun/instrtrace.c create mode 100644 test/monniaux/ocaml/byterun/intern.c create mode 100644 test/monniaux/ocaml/byterun/interp.c create mode 100644 test/monniaux/ocaml/byterun/ints.c create mode 100644 test/monniaux/ocaml/byterun/io.c create mode 100644 test/monniaux/ocaml/byterun/lexing.c create mode 100644 test/monniaux/ocaml/byterun/main.c create mode 100644 test/monniaux/ocaml/byterun/major_gc.c create mode 100644 test/monniaux/ocaml/byterun/make.log create mode 100644 test/monniaux/ocaml/byterun/make2.log create mode 100644 test/monniaux/ocaml/byterun/md5.c create mode 100644 test/monniaux/ocaml/byterun/memory.c create mode 100644 test/monniaux/ocaml/byterun/meta.c create mode 100644 test/monniaux/ocaml/byterun/minor_gc.c create mode 100644 test/monniaux/ocaml/byterun/misc.c create mode 100644 test/monniaux/ocaml/byterun/obj.c create mode 100644 test/monniaux/ocaml/byterun/parsing.c create mode 100644 test/monniaux/ocaml/byterun/printexc.c create mode 100644 test/monniaux/ocaml/byterun/roots.c create mode 100644 test/monniaux/ocaml/byterun/signals.c create mode 100644 test/monniaux/ocaml/byterun/signals_byt.c create mode 100644 test/monniaux/ocaml/byterun/spacetime.c create mode 100644 test/monniaux/ocaml/byterun/stacks.c create mode 100644 test/monniaux/ocaml/byterun/startup.c create mode 100644 test/monniaux/ocaml/byterun/startup_aux.c create mode 100644 test/monniaux/ocaml/byterun/str.c create mode 100644 test/monniaux/ocaml/byterun/sys.c create mode 100644 test/monniaux/ocaml/byterun/toto create mode 100644 test/monniaux/ocaml/byterun/unix.c create mode 100644 test/monniaux/ocaml/byterun/weak.c create mode 100644 test/monniaux/ocaml/byterun/win32.c create mode 100644 test/monniaux/ocaml/config/Makefile create mode 100644 test/monniaux/ocaml/config/Makefile-templ create mode 100644 test/monniaux/ocaml/config/m-nt.h create mode 100644 test/monniaux/ocaml/config/m-templ.h create mode 100644 test/monniaux/ocaml/config/s-nt.h create mode 100644 test/monniaux/ocaml/config/s-templ.h create mode 100755 test/monniaux/ocaml/examples/quicksort create mode 100644 test/monniaux/ocaml/examples/quicksort.ml create mode 100755 test/monniaux/ocaml/tools/make-version-header.sh diff --git a/test/monniaux/ocaml/Makefile.common b/test/monniaux/ocaml/Makefile.common new file mode 100644 index 00000000..cd9f676b --- /dev/null +++ b/test/monniaux/ocaml/Makefile.common @@ -0,0 +1,30 @@ +#************************************************************************** +#* * +#* OCaml * +#* * +#* Gabriel Scherer, projet Parsifal, INRIA Saclay * +#* * +#* Copyright 2018 Institut National de Recherche en Informatique et * +#* en Automatique. * +#* * +#* All rights reserved. This file is distributed under the terms of * +#* the GNU Lesser General Public License version 2.1, with the * +#* special exception on linking described in the file LICENSE. * +#* * +#************************************************************************** + +# This makefile contains common definitions shared by other Makefiles +# We assume that config/Makefile has already been included + +INSTALL ?= install +INSTALL_DATA ?= $(INSTALL) -m u=rw,g=rw,o=r +INSTALL_PROG ?= $(INSTALL) -m u=rwx,g=rwx,o=rx + +# note: these are defined by lazy expansions +# as some parts of the makefiles change BINDIR, etc. +# and expect INSTALL_BINDIR, etc. to stay in synch +# (see `shellquote` in tools/Makefile) +INSTALL_BINDIR = $(DESTDIR)$(BINDIR) +INSTALL_LIBDIR = $(DESTDIR)$(LIBDIR) +INSTALL_STUBLIBDIR = $(DESTDIR)$(STUBLIBDIR) +INSTALL_MANDIR = $(DESTDIR)$(MANDIR) diff --git a/test/monniaux/ocaml/VERSION b/test/monniaux/ocaml/VERSION new file mode 100644 index 00000000..0e48c0f1 --- /dev/null +++ b/test/monniaux/ocaml/VERSION @@ -0,0 +1,4 @@ +4.07.1 + +# The version string is the first line of this file. +# It must be in the format described in stdlib/sys.mli diff --git a/test/monniaux/ocaml/byterun/.depend b/test/monniaux/ocaml/byterun/.depend new file mode 100644 index 00000000..c0f81615 --- /dev/null +++ b/test/monniaux/ocaml/byterun/.depend @@ -0,0 +1,803 @@ +afl.$(O): afl.c caml/config.h caml/m.h caml/s.h caml/misc.h caml/mlvalues.h \ + caml/osdeps.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h +alloc.$(O): alloc.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/custom.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/stacks.h +array.$(O): array.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/signals.h \ + caml/spacetime.h caml/io.h caml/stack.h +backtrace.$(O): backtrace.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/backtrace.h \ + caml/exec.h caml/backtrace_prim.h caml/fail.h +backtrace_prim.$(O): backtrace_prim.c caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/misc.h caml/alloc.h caml/custom.h caml/io.h \ + caml/instruct.h caml/intext.h caml/exec.h caml/fix_code.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/startup.h caml/stacks.h \ + caml/sys.h caml/backtrace.h caml/fail.h caml/backtrace_prim.h +bigarray.$(O): bigarray.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/bigarray.h caml/custom.h caml/fail.h \ + caml/intext.h caml/io.h caml/hash.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/signals.h +callback.$(O): callback.c caml/callback.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/interp.h caml/instruct.h caml/fix_code.h caml/stacks.h +compact.$(O): compact.c caml/address_class.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/finalise.h caml/roots.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/gc_ctrl.h caml/weak.h caml/compact.h +compare.$(O): compare.c caml/custom.h caml/mlvalues.h caml/config.h caml/m.h \ + caml/s.h caml/misc.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +custom.$(O): custom.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/custom.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/signals.h +debugger.$(O): debugger.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/debugger.h caml/osdeps.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/fail.h caml/fix_code.h caml/instruct.h \ + caml/intext.h caml/io.h caml/stacks.h caml/sys.h +dynlink.$(O): dynlink.c caml/config.h caml/m.h caml/s.h caml/alloc.h \ + caml/misc.h caml/mlvalues.h caml/dynlink.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/osdeps.h caml/prims.h caml/signals.h +extern.$(O): extern.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/custom.h caml/fail.h caml/gc.h \ + caml/intext.h caml/io.h caml/md5.h caml/memory.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/reverse.h +fail.$(O): fail.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/io.h caml/gc.h caml/memory.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/printexc.h caml/signals.h caml/stacks.h +finalise.$(O): finalise.c caml/callback.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/compact.h caml/fail.h \ + caml/finalise.h caml/roots.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/signals.h +fix_code.$(O): fix_code.c caml/config.h caml/m.h caml/s.h caml/debugger.h \ + caml/misc.h caml/mlvalues.h caml/fix_code.h caml/instruct.h \ + caml/intext.h caml/io.h caml/md5.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h +floats.$(O): floats.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h caml/stacks.h +freelist.$(O): freelist.c caml/config.h caml/m.h caml/s.h caml/freelist.h \ + caml/misc.h caml/mlvalues.h caml/gc.h caml/gc_ctrl.h caml/memory.h \ + caml/major_gc.h caml/minor_gc.h caml/address_class.h +gc_ctrl.$(O): gc_ctrl.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/backtrace.h caml/exec.h caml/compact.h \ + caml/custom.h caml/fail.h caml/finalise.h caml/roots.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/gc_ctrl.h caml/signals.h caml/stacks.h \ + caml/startup_aux.h +globroots.$(O): globroots.c caml/memory.h caml/config.h caml/m.h caml/s.h \ + caml/gc.h caml/mlvalues.h caml/misc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/roots.h caml/globroots.h +hash.$(O): hash.c caml/mlvalues.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/custom.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/hash.h +instrtrace.$(O): instrtrace.c +intern.$(O): intern.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/callback.h caml/custom.h caml/fail.h \ + caml/gc.h caml/intext.h caml/io.h caml/md5.h caml/memory.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h +interp.$(O): interp.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/backtrace.h caml/exec.h caml/callback.h \ + caml/debugger.h caml/fail.h caml/fix_code.h caml/instrtrace.h \ + caml/instruct.h caml/interp.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/prims.h caml/signals.h caml/stacks.h caml/startup_aux.h \ + caml/jumptbl.h +ints.$(O): ints.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/custom.h caml/fail.h caml/intext.h caml/io.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h +io.$(O): io.c caml/config.h caml/m.h caml/s.h caml/alloc.h caml/misc.h \ + caml/mlvalues.h caml/custom.h caml/fail.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/osdeps.h caml/signals.h caml/sys.h +lexing.$(O): lexing.c caml/fail.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +main.$(O): main.c caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/sys.h caml/osdeps.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +major_gc.$(O): major_gc.c caml/compact.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/custom.h caml/fail.h caml/finalise.h \ + caml/roots.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/gc_ctrl.h caml/signals.h \ + caml/weak.h +md5.$(O): md5.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/md5.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/reverse.h +memory.$(O): memory.c caml/address_class.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/fail.h caml/freelist.h caml/gc.h \ + caml/gc_ctrl.h caml/major_gc.h caml/memory.h caml/minor_gc.h \ + caml/signals.h +meta.$(O): meta.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/fix_code.h caml/interp.h \ + caml/intext.h caml/io.h caml/major_gc.h caml/freelist.h caml/memory.h \ + caml/gc.h caml/minor_gc.h caml/address_class.h caml/prims.h \ + caml/stacks.h +minor_gc.$(O): minor_gc.c caml/custom.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/fail.h caml/finalise.h caml/roots.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/gc_ctrl.h caml/signals.h \ + caml/weak.h +misc.$(O): misc.c caml/config.h caml/m.h caml/s.h caml/misc.h caml/memory.h \ + caml/gc.h caml/mlvalues.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/osdeps.h caml/version.h +obj.$(O): obj.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/gc.h caml/interp.h caml/major_gc.h \ + caml/freelist.h caml/memory.h caml/minor_gc.h caml/address_class.h \ + caml/prims.h caml/spacetime.h caml/io.h caml/stack.h +parsing.$(O): parsing.c caml/config.h caml/m.h caml/s.h caml/mlvalues.h \ + caml/misc.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/alloc.h +prims.$(O): prims.c caml/mlvalues.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/prims.h +printexc.$(O): printexc.c caml/backtrace.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/exec.h caml/callback.h \ + caml/debugger.h caml/fail.h caml/printexc.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +roots.$(O): roots.c caml/finalise.h caml/roots.h caml/misc.h caml/config.h \ + caml/m.h caml/s.h caml/memory.h caml/gc.h caml/mlvalues.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/globroots.h caml/stacks.h +signals.$(O): signals.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/callback.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/roots.h caml/signals.h \ + caml/signals_machdep.h caml/sys.h +signals_byt.$(O): signals_byt.c caml/config.h caml/m.h caml/s.h \ + caml/memory.h caml/gc.h caml/mlvalues.h caml/misc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/osdeps.h \ + caml/signals.h caml/signals_machdep.h +spacetime.$(O): spacetime.c caml/fail.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h +stacks.$(O): stacks.c caml/config.h caml/m.h caml/s.h caml/fail.h \ + caml/misc.h caml/mlvalues.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +startup.$(O): startup.c caml/config.h caml/m.h caml/s.h caml/alloc.h \ + caml/misc.h caml/mlvalues.h caml/backtrace.h caml/exec.h \ + caml/callback.h caml/custom.h caml/debugger.h caml/dynlink.h \ + caml/fail.h caml/fix_code.h caml/freelist.h caml/gc_ctrl.h \ + caml/instrtrace.h caml/interp.h caml/intext.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/minor_gc.h caml/address_class.h \ + caml/osdeps.h caml/prims.h caml/printexc.h caml/reverse.h \ + caml/signals.h caml/stacks.h caml/sys.h caml/startup.h \ + caml/startup_aux.h caml/version.h +startup_aux.$(O): startup_aux.c caml/backtrace.h caml/mlvalues.h \ + caml/config.h caml/m.h caml/s.h caml/misc.h caml/exec.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/callback.h caml/dynlink.h caml/osdeps.h \ + caml/startup_aux.h +str.$(O): str.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h +sys.$(O): sys.c caml/config.h caml/m.h caml/s.h caml/alloc.h caml/misc.h \ + caml/mlvalues.h caml/debugger.h caml/fail.h caml/gc_ctrl.h caml/io.h \ + caml/osdeps.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/signals.h caml/stacks.h \ + caml/sys.h caml/version.h caml/callback.h caml/startup_aux.h +unix.$(O): unix.c caml/config.h caml/m.h caml/s.h caml/fail.h caml/misc.h \ + caml/mlvalues.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/osdeps.h \ + caml/signals.h caml/sys.h caml/io.h caml/alloc.h +weak.$(O): weak.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/weak.h +afl.d.$(O): afl.c caml/config.h caml/m.h caml/s.h caml/misc.h caml/mlvalues.h \ + caml/osdeps.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h +alloc.d.$(O): alloc.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/custom.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/stacks.h +array.d.$(O): array.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/signals.h \ + caml/spacetime.h caml/io.h caml/stack.h +backtrace.d.$(O): backtrace.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/backtrace.h \ + caml/exec.h caml/backtrace_prim.h caml/fail.h +backtrace_prim.d.$(O): backtrace_prim.c caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/misc.h caml/alloc.h caml/custom.h caml/io.h \ + caml/instruct.h caml/intext.h caml/exec.h caml/fix_code.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/startup.h caml/stacks.h \ + caml/sys.h caml/backtrace.h caml/fail.h caml/backtrace_prim.h +bigarray.d.$(O): bigarray.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/bigarray.h caml/custom.h caml/fail.h \ + caml/intext.h caml/io.h caml/hash.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/signals.h +callback.d.$(O): callback.c caml/callback.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/interp.h caml/instruct.h caml/fix_code.h caml/stacks.h +compact.d.$(O): compact.c caml/address_class.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/finalise.h caml/roots.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/gc_ctrl.h caml/weak.h caml/compact.h +compare.d.$(O): compare.c caml/custom.h caml/mlvalues.h caml/config.h caml/m.h \ + caml/s.h caml/misc.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +custom.d.$(O): custom.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/custom.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/signals.h +debugger.d.$(O): debugger.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/debugger.h caml/osdeps.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/fail.h caml/fix_code.h caml/instruct.h \ + caml/intext.h caml/io.h caml/stacks.h caml/sys.h +dynlink.d.$(O): dynlink.c caml/config.h caml/m.h caml/s.h caml/alloc.h \ + caml/misc.h caml/mlvalues.h caml/dynlink.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/osdeps.h caml/prims.h caml/signals.h +extern.d.$(O): extern.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/custom.h caml/fail.h caml/gc.h \ + caml/intext.h caml/io.h caml/md5.h caml/memory.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/reverse.h +fail.d.$(O): fail.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/io.h caml/gc.h caml/memory.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/printexc.h caml/signals.h caml/stacks.h +finalise.d.$(O): finalise.c caml/callback.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/compact.h caml/fail.h \ + caml/finalise.h caml/roots.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/signals.h +fix_code.d.$(O): fix_code.c caml/config.h caml/m.h caml/s.h caml/debugger.h \ + caml/misc.h caml/mlvalues.h caml/fix_code.h caml/instruct.h \ + caml/intext.h caml/io.h caml/md5.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h +floats.d.$(O): floats.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h caml/stacks.h +freelist.d.$(O): freelist.c caml/config.h caml/m.h caml/s.h caml/freelist.h \ + caml/misc.h caml/mlvalues.h caml/gc.h caml/gc_ctrl.h caml/memory.h \ + caml/major_gc.h caml/minor_gc.h caml/address_class.h +gc_ctrl.d.$(O): gc_ctrl.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/backtrace.h caml/exec.h caml/compact.h \ + caml/custom.h caml/fail.h caml/finalise.h caml/roots.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/gc_ctrl.h caml/signals.h caml/stacks.h \ + caml/startup_aux.h +globroots.d.$(O): globroots.c caml/memory.h caml/config.h caml/m.h caml/s.h \ + caml/gc.h caml/mlvalues.h caml/misc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/roots.h caml/globroots.h +hash.d.$(O): hash.c caml/mlvalues.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/custom.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/hash.h +instrtrace.d.$(O): instrtrace.c caml/instrtrace.h caml/mlvalues.h \ + caml/config.h caml/m.h caml/s.h caml/misc.h caml/instruct.h \ + caml/opnames.h caml/prims.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/startup_aux.h +intern.d.$(O): intern.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/callback.h caml/custom.h caml/fail.h \ + caml/gc.h caml/intext.h caml/io.h caml/md5.h caml/memory.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h +interp.d.$(O): interp.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/backtrace.h caml/exec.h caml/callback.h \ + caml/debugger.h caml/fail.h caml/fix_code.h caml/instrtrace.h \ + caml/instruct.h caml/interp.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/prims.h caml/signals.h caml/stacks.h caml/startup_aux.h +ints.d.$(O): ints.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/custom.h caml/fail.h caml/intext.h caml/io.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h +io.d.$(O): io.c caml/config.h caml/m.h caml/s.h caml/alloc.h caml/misc.h \ + caml/mlvalues.h caml/custom.h caml/fail.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/osdeps.h caml/signals.h caml/sys.h +lexing.d.$(O): lexing.c caml/fail.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +main.d.$(O): main.c caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/sys.h caml/osdeps.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +major_gc.d.$(O): major_gc.c caml/compact.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/custom.h caml/fail.h caml/finalise.h \ + caml/roots.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/gc_ctrl.h caml/signals.h \ + caml/weak.h +md5.d.$(O): md5.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/md5.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/reverse.h +memory.d.$(O): memory.c caml/address_class.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/fail.h caml/freelist.h caml/gc.h \ + caml/gc_ctrl.h caml/major_gc.h caml/memory.h caml/minor_gc.h \ + caml/signals.h +meta.d.$(O): meta.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/fix_code.h caml/interp.h \ + caml/intext.h caml/io.h caml/major_gc.h caml/freelist.h caml/memory.h \ + caml/gc.h caml/minor_gc.h caml/address_class.h caml/prims.h \ + caml/stacks.h +minor_gc.d.$(O): minor_gc.c caml/custom.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/fail.h caml/finalise.h caml/roots.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/gc_ctrl.h caml/signals.h \ + caml/weak.h +misc.d.$(O): misc.c caml/config.h caml/m.h caml/s.h caml/misc.h caml/memory.h \ + caml/gc.h caml/mlvalues.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/osdeps.h caml/version.h +obj.d.$(O): obj.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/gc.h caml/interp.h caml/major_gc.h \ + caml/freelist.h caml/memory.h caml/minor_gc.h caml/address_class.h \ + caml/prims.h caml/spacetime.h caml/io.h caml/stack.h +parsing.d.$(O): parsing.c caml/config.h caml/m.h caml/s.h caml/mlvalues.h \ + caml/misc.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/alloc.h +prims.d.$(O): prims.c caml/mlvalues.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/prims.h +printexc.d.$(O): printexc.c caml/backtrace.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/exec.h caml/callback.h \ + caml/debugger.h caml/fail.h caml/printexc.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +roots.d.$(O): roots.c caml/finalise.h caml/roots.h caml/misc.h caml/config.h \ + caml/m.h caml/s.h caml/memory.h caml/gc.h caml/mlvalues.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/globroots.h caml/stacks.h +signals.d.$(O): signals.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/callback.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/roots.h caml/signals.h \ + caml/signals_machdep.h caml/sys.h +signals_byt.d.$(O): signals_byt.c caml/config.h caml/m.h caml/s.h \ + caml/memory.h caml/gc.h caml/mlvalues.h caml/misc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/osdeps.h \ + caml/signals.h caml/signals_machdep.h +spacetime.d.$(O): spacetime.c caml/fail.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h +stacks.d.$(O): stacks.c caml/config.h caml/m.h caml/s.h caml/fail.h \ + caml/misc.h caml/mlvalues.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +startup.d.$(O): startup.c caml/config.h caml/m.h caml/s.h caml/alloc.h \ + caml/misc.h caml/mlvalues.h caml/backtrace.h caml/exec.h \ + caml/callback.h caml/custom.h caml/debugger.h caml/dynlink.h \ + caml/fail.h caml/fix_code.h caml/freelist.h caml/gc_ctrl.h \ + caml/instrtrace.h caml/interp.h caml/intext.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/minor_gc.h caml/address_class.h \ + caml/osdeps.h caml/prims.h caml/printexc.h caml/reverse.h \ + caml/signals.h caml/stacks.h caml/sys.h caml/startup.h \ + caml/startup_aux.h caml/version.h +startup_aux.d.$(O): startup_aux.c caml/backtrace.h caml/mlvalues.h \ + caml/config.h caml/m.h caml/s.h caml/misc.h caml/exec.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/callback.h caml/dynlink.h caml/osdeps.h \ + caml/startup_aux.h +str.d.$(O): str.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h +sys.d.$(O): sys.c caml/config.h caml/m.h caml/s.h caml/alloc.h caml/misc.h \ + caml/mlvalues.h caml/debugger.h caml/fail.h caml/gc_ctrl.h caml/io.h \ + caml/osdeps.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/signals.h caml/stacks.h \ + caml/sys.h caml/version.h caml/callback.h caml/startup_aux.h +unix.d.$(O): unix.c caml/config.h caml/m.h caml/s.h caml/fail.h caml/misc.h \ + caml/mlvalues.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/osdeps.h \ + caml/signals.h caml/sys.h caml/io.h caml/alloc.h +weak.d.$(O): weak.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/weak.h +afl.i.$(O): afl.c caml/config.h caml/m.h caml/s.h caml/misc.h caml/mlvalues.h \ + caml/osdeps.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h +alloc.i.$(O): alloc.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/custom.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/stacks.h +array.i.$(O): array.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/signals.h \ + caml/spacetime.h caml/io.h caml/stack.h +backtrace.i.$(O): backtrace.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/backtrace.h \ + caml/exec.h caml/backtrace_prim.h caml/fail.h +backtrace_prim.i.$(O): backtrace_prim.c caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/misc.h caml/alloc.h caml/custom.h caml/io.h \ + caml/instruct.h caml/intext.h caml/exec.h caml/fix_code.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/startup.h caml/stacks.h \ + caml/sys.h caml/backtrace.h caml/fail.h caml/backtrace_prim.h +bigarray.i.$(O): bigarray.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/bigarray.h caml/custom.h caml/fail.h \ + caml/intext.h caml/io.h caml/hash.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/signals.h +callback.i.$(O): callback.c caml/callback.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/interp.h caml/instruct.h caml/fix_code.h caml/stacks.h +compact.i.$(O): compact.c caml/address_class.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/finalise.h caml/roots.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/gc_ctrl.h caml/weak.h caml/compact.h +compare.i.$(O): compare.c caml/custom.h caml/mlvalues.h caml/config.h caml/m.h \ + caml/s.h caml/misc.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +custom.i.$(O): custom.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/custom.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/signals.h +debugger.i.$(O): debugger.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/debugger.h caml/osdeps.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/fail.h caml/fix_code.h caml/instruct.h \ + caml/intext.h caml/io.h caml/stacks.h caml/sys.h +dynlink.i.$(O): dynlink.c caml/config.h caml/m.h caml/s.h caml/alloc.h \ + caml/misc.h caml/mlvalues.h caml/dynlink.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/osdeps.h caml/prims.h caml/signals.h +extern.i.$(O): extern.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/custom.h caml/fail.h caml/gc.h \ + caml/intext.h caml/io.h caml/md5.h caml/memory.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/reverse.h +fail.i.$(O): fail.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/io.h caml/gc.h caml/memory.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/printexc.h caml/signals.h caml/stacks.h +finalise.i.$(O): finalise.c caml/callback.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/compact.h caml/fail.h \ + caml/finalise.h caml/roots.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/signals.h +fix_code.i.$(O): fix_code.c caml/config.h caml/m.h caml/s.h caml/debugger.h \ + caml/misc.h caml/mlvalues.h caml/fix_code.h caml/instruct.h \ + caml/intext.h caml/io.h caml/md5.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h +floats.i.$(O): floats.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h caml/stacks.h +freelist.i.$(O): freelist.c caml/config.h caml/m.h caml/s.h caml/freelist.h \ + caml/misc.h caml/mlvalues.h caml/gc.h caml/gc_ctrl.h caml/memory.h \ + caml/major_gc.h caml/minor_gc.h caml/address_class.h +gc_ctrl.i.$(O): gc_ctrl.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/backtrace.h caml/exec.h caml/compact.h \ + caml/custom.h caml/fail.h caml/finalise.h caml/roots.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/gc_ctrl.h caml/signals.h caml/stacks.h \ + caml/startup_aux.h +globroots.i.$(O): globroots.c caml/memory.h caml/config.h caml/m.h caml/s.h \ + caml/gc.h caml/mlvalues.h caml/misc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/roots.h caml/globroots.h +hash.i.$(O): hash.c caml/mlvalues.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/custom.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/hash.h +instrtrace.i.$(O): instrtrace.c +intern.i.$(O): intern.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/callback.h caml/custom.h caml/fail.h \ + caml/gc.h caml/intext.h caml/io.h caml/md5.h caml/memory.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h +interp.i.$(O): interp.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/backtrace.h caml/exec.h caml/callback.h \ + caml/debugger.h caml/fail.h caml/fix_code.h caml/instrtrace.h \ + caml/instruct.h caml/interp.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/prims.h caml/signals.h caml/stacks.h caml/startup_aux.h \ + caml/jumptbl.h +ints.i.$(O): ints.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/custom.h caml/fail.h caml/intext.h caml/io.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h +io.i.$(O): io.c caml/config.h caml/m.h caml/s.h caml/alloc.h caml/misc.h \ + caml/mlvalues.h caml/custom.h caml/fail.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/osdeps.h caml/signals.h caml/sys.h +lexing.i.$(O): lexing.c caml/fail.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +main.i.$(O): main.c caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/sys.h caml/osdeps.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +major_gc.i.$(O): major_gc.c caml/compact.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/custom.h caml/fail.h caml/finalise.h \ + caml/roots.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/gc_ctrl.h caml/signals.h \ + caml/weak.h +md5.i.$(O): md5.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/md5.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/reverse.h +memory.i.$(O): memory.c caml/address_class.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/fail.h caml/freelist.h caml/gc.h \ + caml/gc_ctrl.h caml/major_gc.h caml/memory.h caml/minor_gc.h \ + caml/signals.h +meta.i.$(O): meta.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/fix_code.h caml/interp.h \ + caml/intext.h caml/io.h caml/major_gc.h caml/freelist.h caml/memory.h \ + caml/gc.h caml/minor_gc.h caml/address_class.h caml/prims.h \ + caml/stacks.h +minor_gc.i.$(O): minor_gc.c caml/custom.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/fail.h caml/finalise.h caml/roots.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/gc_ctrl.h caml/signals.h \ + caml/weak.h +misc.i.$(O): misc.c caml/config.h caml/m.h caml/s.h caml/misc.h caml/memory.h \ + caml/gc.h caml/mlvalues.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/osdeps.h caml/version.h +obj.i.$(O): obj.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/gc.h caml/interp.h caml/major_gc.h \ + caml/freelist.h caml/memory.h caml/minor_gc.h caml/address_class.h \ + caml/prims.h caml/spacetime.h caml/io.h caml/stack.h +parsing.i.$(O): parsing.c caml/config.h caml/m.h caml/s.h caml/mlvalues.h \ + caml/misc.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/alloc.h +prims.i.$(O): prims.c caml/mlvalues.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/prims.h +printexc.i.$(O): printexc.c caml/backtrace.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/exec.h caml/callback.h \ + caml/debugger.h caml/fail.h caml/printexc.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +roots.i.$(O): roots.c caml/finalise.h caml/roots.h caml/misc.h caml/config.h \ + caml/m.h caml/s.h caml/memory.h caml/gc.h caml/mlvalues.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/globroots.h caml/stacks.h +signals.i.$(O): signals.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/callback.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/roots.h caml/signals.h \ + caml/signals_machdep.h caml/sys.h +signals_byt.i.$(O): signals_byt.c caml/config.h caml/m.h caml/s.h \ + caml/memory.h caml/gc.h caml/mlvalues.h caml/misc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/osdeps.h \ + caml/signals.h caml/signals_machdep.h +spacetime.i.$(O): spacetime.c caml/fail.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h +stacks.i.$(O): stacks.c caml/config.h caml/m.h caml/s.h caml/fail.h \ + caml/misc.h caml/mlvalues.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +startup.i.$(O): startup.c caml/config.h caml/m.h caml/s.h caml/alloc.h \ + caml/misc.h caml/mlvalues.h caml/backtrace.h caml/exec.h \ + caml/callback.h caml/custom.h caml/debugger.h caml/dynlink.h \ + caml/fail.h caml/fix_code.h caml/freelist.h caml/gc_ctrl.h \ + caml/instrtrace.h caml/interp.h caml/intext.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/minor_gc.h caml/address_class.h \ + caml/osdeps.h caml/prims.h caml/printexc.h caml/reverse.h \ + caml/signals.h caml/stacks.h caml/sys.h caml/startup.h \ + caml/startup_aux.h caml/version.h +startup_aux.i.$(O): startup_aux.c caml/backtrace.h caml/mlvalues.h \ + caml/config.h caml/m.h caml/s.h caml/misc.h caml/exec.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/callback.h caml/dynlink.h caml/osdeps.h \ + caml/startup_aux.h +str.i.$(O): str.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h +sys.i.$(O): sys.c caml/config.h caml/m.h caml/s.h caml/alloc.h caml/misc.h \ + caml/mlvalues.h caml/debugger.h caml/fail.h caml/gc_ctrl.h caml/io.h \ + caml/osdeps.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/signals.h caml/stacks.h \ + caml/sys.h caml/version.h caml/callback.h caml/startup_aux.h +unix.i.$(O): unix.c caml/config.h caml/m.h caml/s.h caml/fail.h caml/misc.h \ + caml/mlvalues.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/osdeps.h \ + caml/signals.h caml/sys.h caml/io.h caml/alloc.h +weak.i.$(O): weak.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/weak.h +afl.pic.$(O): afl.c caml/config.h caml/m.h caml/s.h caml/misc.h caml/mlvalues.h \ + caml/osdeps.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h +alloc.pic.$(O): alloc.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/custom.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/stacks.h +array.pic.$(O): array.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/signals.h \ + caml/spacetime.h caml/io.h caml/stack.h +backtrace.pic.$(O): backtrace.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/backtrace.h \ + caml/exec.h caml/backtrace_prim.h caml/fail.h +backtrace_prim.pic.$(O): backtrace_prim.c caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/misc.h caml/alloc.h caml/custom.h caml/io.h \ + caml/instruct.h caml/intext.h caml/exec.h caml/fix_code.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/startup.h caml/stacks.h \ + caml/sys.h caml/backtrace.h caml/fail.h caml/backtrace_prim.h +bigarray.pic.$(O): bigarray.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/bigarray.h caml/custom.h caml/fail.h \ + caml/intext.h caml/io.h caml/hash.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/signals.h +callback.pic.$(O): callback.c caml/callback.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/interp.h caml/instruct.h caml/fix_code.h caml/stacks.h +compact.pic.$(O): compact.c caml/address_class.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/finalise.h caml/roots.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/gc_ctrl.h caml/weak.h caml/compact.h +compare.pic.$(O): compare.c caml/custom.h caml/mlvalues.h caml/config.h caml/m.h \ + caml/s.h caml/misc.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +custom.pic.$(O): custom.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/custom.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/signals.h +debugger.pic.$(O): debugger.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/debugger.h caml/osdeps.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/fail.h caml/fix_code.h caml/instruct.h \ + caml/intext.h caml/io.h caml/stacks.h caml/sys.h +dynlink.pic.$(O): dynlink.c caml/config.h caml/m.h caml/s.h caml/alloc.h \ + caml/misc.h caml/mlvalues.h caml/dynlink.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/osdeps.h caml/prims.h caml/signals.h +extern.pic.$(O): extern.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/custom.h caml/fail.h caml/gc.h \ + caml/intext.h caml/io.h caml/md5.h caml/memory.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/reverse.h +fail.pic.$(O): fail.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/io.h caml/gc.h caml/memory.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/printexc.h caml/signals.h caml/stacks.h +finalise.pic.$(O): finalise.c caml/callback.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/compact.h caml/fail.h \ + caml/finalise.h caml/roots.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/signals.h +fix_code.pic.$(O): fix_code.c caml/config.h caml/m.h caml/s.h caml/debugger.h \ + caml/misc.h caml/mlvalues.h caml/fix_code.h caml/instruct.h \ + caml/intext.h caml/io.h caml/md5.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h +floats.pic.$(O): floats.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h caml/stacks.h +freelist.pic.$(O): freelist.c caml/config.h caml/m.h caml/s.h caml/freelist.h \ + caml/misc.h caml/mlvalues.h caml/gc.h caml/gc_ctrl.h caml/memory.h \ + caml/major_gc.h caml/minor_gc.h caml/address_class.h +gc_ctrl.pic.$(O): gc_ctrl.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/backtrace.h caml/exec.h caml/compact.h \ + caml/custom.h caml/fail.h caml/finalise.h caml/roots.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/gc_ctrl.h caml/signals.h caml/stacks.h \ + caml/startup_aux.h +globroots.pic.$(O): globroots.c caml/memory.h caml/config.h caml/m.h caml/s.h \ + caml/gc.h caml/mlvalues.h caml/misc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/roots.h caml/globroots.h +hash.pic.$(O): hash.c caml/mlvalues.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/custom.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/hash.h +instrtrace.pic.$(O): instrtrace.c +intern.pic.$(O): intern.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/callback.h caml/custom.h caml/fail.h \ + caml/gc.h caml/intext.h caml/io.h caml/md5.h caml/memory.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/reverse.h +interp.pic.$(O): interp.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/backtrace.h caml/exec.h caml/callback.h \ + caml/debugger.h caml/fail.h caml/fix_code.h caml/instrtrace.h \ + caml/instruct.h caml/interp.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/prims.h caml/signals.h caml/stacks.h caml/startup_aux.h \ + caml/jumptbl.h +ints.pic.$(O): ints.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/custom.h caml/fail.h caml/intext.h caml/io.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h +io.pic.$(O): io.c caml/config.h caml/m.h caml/s.h caml/alloc.h caml/misc.h \ + caml/mlvalues.h caml/custom.h caml/fail.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/osdeps.h caml/signals.h caml/sys.h +lexing.pic.$(O): lexing.c caml/fail.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +main.pic.$(O): main.c caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/sys.h caml/osdeps.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +major_gc.pic.$(O): major_gc.c caml/compact.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/custom.h caml/fail.h caml/finalise.h \ + caml/roots.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/gc_ctrl.h caml/signals.h \ + caml/weak.h +md5.pic.$(O): md5.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/md5.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/reverse.h +memory.pic.$(O): memory.c caml/address_class.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/mlvalues.h caml/fail.h caml/freelist.h caml/gc.h \ + caml/gc_ctrl.h caml/major_gc.h caml/memory.h caml/minor_gc.h \ + caml/signals.h +meta.pic.$(O): meta.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/fix_code.h caml/interp.h \ + caml/intext.h caml/io.h caml/major_gc.h caml/freelist.h caml/memory.h \ + caml/gc.h caml/minor_gc.h caml/address_class.h caml/prims.h \ + caml/stacks.h +minor_gc.pic.$(O): minor_gc.c caml/custom.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/fail.h caml/finalise.h caml/roots.h \ + caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/gc_ctrl.h caml/signals.h \ + caml/weak.h +misc.pic.$(O): misc.c caml/config.h caml/m.h caml/s.h caml/misc.h caml/memory.h \ + caml/gc.h caml/mlvalues.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/osdeps.h caml/version.h +obj.pic.$(O): obj.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/gc.h caml/interp.h caml/major_gc.h \ + caml/freelist.h caml/memory.h caml/minor_gc.h caml/address_class.h \ + caml/prims.h caml/spacetime.h caml/io.h caml/stack.h +parsing.pic.$(O): parsing.c caml/config.h caml/m.h caml/s.h caml/mlvalues.h \ + caml/misc.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/alloc.h +prims.pic.$(O): prims.c caml/mlvalues.h caml/config.h caml/m.h caml/s.h \ + caml/misc.h caml/prims.h +printexc.pic.$(O): printexc.c caml/backtrace.h caml/mlvalues.h caml/config.h \ + caml/m.h caml/s.h caml/misc.h caml/exec.h caml/callback.h \ + caml/debugger.h caml/fail.h caml/printexc.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +roots.pic.$(O): roots.c caml/finalise.h caml/roots.h caml/misc.h caml/config.h \ + caml/m.h caml/s.h caml/memory.h caml/gc.h caml/mlvalues.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h \ + caml/globroots.h caml/stacks.h +signals.pic.$(O): signals.c caml/alloc.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h caml/callback.h caml/fail.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/roots.h caml/signals.h \ + caml/signals_machdep.h caml/sys.h +signals_byt.pic.$(O): signals_byt.c caml/config.h caml/m.h caml/s.h \ + caml/memory.h caml/gc.h caml/mlvalues.h caml/misc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/osdeps.h \ + caml/signals.h caml/signals_machdep.h +spacetime.pic.$(O): spacetime.c caml/fail.h caml/misc.h caml/config.h caml/m.h \ + caml/s.h caml/mlvalues.h +stacks.pic.$(O): stacks.c caml/config.h caml/m.h caml/s.h caml/fail.h \ + caml/misc.h caml/mlvalues.h caml/stacks.h caml/memory.h caml/gc.h \ + caml/major_gc.h caml/freelist.h caml/minor_gc.h caml/address_class.h +startup.pic.$(O): startup.c caml/config.h caml/m.h caml/s.h caml/alloc.h \ + caml/misc.h caml/mlvalues.h caml/backtrace.h caml/exec.h \ + caml/callback.h caml/custom.h caml/debugger.h caml/dynlink.h \ + caml/fail.h caml/fix_code.h caml/freelist.h caml/gc_ctrl.h \ + caml/instrtrace.h caml/interp.h caml/intext.h caml/io.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/minor_gc.h caml/address_class.h \ + caml/osdeps.h caml/prims.h caml/printexc.h caml/reverse.h \ + caml/signals.h caml/stacks.h caml/sys.h caml/startup.h \ + caml/startup_aux.h caml/version.h +startup_aux.pic.$(O): startup_aux.c caml/backtrace.h caml/mlvalues.h \ + caml/config.h caml/m.h caml/s.h caml/misc.h caml/exec.h caml/memory.h \ + caml/gc.h caml/major_gc.h caml/freelist.h caml/minor_gc.h \ + caml/address_class.h caml/callback.h caml/dynlink.h caml/osdeps.h \ + caml/startup_aux.h +str.pic.$(O): str.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h +sys.pic.$(O): sys.c caml/config.h caml/m.h caml/s.h caml/alloc.h caml/misc.h \ + caml/mlvalues.h caml/debugger.h caml/fail.h caml/gc_ctrl.h caml/io.h \ + caml/osdeps.h caml/memory.h caml/gc.h caml/major_gc.h caml/freelist.h \ + caml/minor_gc.h caml/address_class.h caml/signals.h caml/stacks.h \ + caml/sys.h caml/version.h caml/callback.h caml/startup_aux.h +unix.pic.$(O): unix.c caml/config.h caml/m.h caml/s.h caml/fail.h caml/misc.h \ + caml/mlvalues.h caml/memory.h caml/gc.h caml/major_gc.h \ + caml/freelist.h caml/minor_gc.h caml/address_class.h caml/osdeps.h \ + caml/signals.h caml/sys.h caml/io.h caml/alloc.h +weak.pic.$(O): weak.c caml/alloc.h caml/misc.h caml/config.h caml/m.h caml/s.h \ + caml/mlvalues.h caml/fail.h caml/major_gc.h caml/freelist.h \ + caml/memory.h caml/gc.h caml/minor_gc.h caml/address_class.h \ + caml/weak.h diff --git a/test/monniaux/ocaml/byterun/Makefile b/test/monniaux/ocaml/byterun/Makefile new file mode 100644 index 00000000..7b9f71ea --- /dev/null +++ b/test/monniaux/ocaml/byterun/Makefile @@ -0,0 +1,234 @@ +#************************************************************************** +#* * +#* OCaml * +#* * +#* Xavier Leroy, projet Cristal, INRIA Rocquencourt * +#* * +#* Copyright 1999 Institut National de Recherche en Informatique et * +#* en Automatique. * +#* * +#* All rights reserved. This file is distributed under the terms of * +#* the GNU Lesser General Public License version 2.1, with the * +#* special exception on linking described in the file LICENSE. * +#* * +#************************************************************************** + +include ../config/Makefile +include ../Makefile.common + +# The PROGRAMS (resp. LIBRARIES) variable list the files to build and +# install as programs in $(INSTALL_BINDIR) (resp. libraries in +# $(INSTALL_LIBDIR)) + +PROGRAMS = ocamlrun$(EXE) +LIBRARIES = ld.conf libcamlrun.$(A) +DYNLIBRARIES= + +ifeq "$(RUNTIMED)" "true" +PROGRAMS += ocamlrund$(EXE) +LIBRARIES += libcamlrund.$(A) +endif + +ifeq "$(RUNTIMEI)" "true" +PROGRAMS += ocamlruni$(EXE) +LIBRARIES += libcamlruni.$(A) +endif + +ifeq "$(UNIX_OR_WIN32)" "unix" +ifeq "$(SUPPORTS_SHARED_LIBRARIES)" "true" +LIBRARIES += libcamlrun_pic.$(A) +DYNLIBRARIES += libcamlrun_shared.$(SO) +endif +endif + +ifdef BOOTSTRAPPING_FLEXLINK +CFLAGS += -DBOOTSTRAPPING_FLEXLINK +endif + +# On Windows, OCAML_STDLIB_DIR needs to be defined dynamically + +ifeq "$(UNIX_OR_WIN32)" "win32" +# OCAML_STDLIB_DIR needs to arrive in dynlink.c as a string which both gcc and +# msvc are willing parse without warning. This means we can't pass UTF-8 +# directly since, as far as I can tell, cl can cope, but the pre-processor +# can't. So the string needs to be directly translated to L"" form. To do this, +# we take advantage of the fact that Cygwin uses GNU libiconv which includes a +# Java pseudo-encoding which translates any UTF-8 sequences to \uXXXX (and, +# unlike the C99 pseudo-encoding, emits two surrogate values when needed, rather +# than \UXXXXXXXX). The \u is then translated to \x in order to accommodate +# pre-Visual Studio 2013 compilers where \x is a non-standard alias for \u. +OCAML_STDLIB_DIR = $(shell echo $(LIBDIR)| iconv -t JAVA | sed -e 's/\\u/\\x/g') +CFLAGS += -DOCAML_STDLIB_DIR='L"$(OCAML_STDLIB_DIR)"' +endif + +CFLAGS += $(IFLEXDIR) + +ifneq "$(CCOMPTYPE)" "msvc" +CFLAGS += -g +endif + +DFLAGS=$(CFLAGS) -DDEBUG +IFLAGS=$(CFLAGS) +PICFLAGS=$(CFLAGS) $(SHAREDCCCOMPOPTS) + +DBGO=d.$(O) + +ifeq "$(UNIX_OR_WIN32)" "win32" +LIBS = $(BYTECCLIBS) $(EXTRALIBS) +ifdef BOOTSTRAPPING_FLEXLINK +MAKE_OCAMLRUN=$(MKEXE_BOOT) +else +MAKE_OCAMLRUN = $(MKEXE) -o $(1) $(2) +endif +else +LIBS = $(BYTECCLIBS) +MAKE_OCAMLRUN = $(MKEXE) $(LDFLAGS) -o $(1) $(2) +endif + +PRIMS=\ + alloc.c array.c compare.c extern.c floats.c gc_ctrl.c hash.c \ + intern.c interp.c ints.c io.c lexing.c md5.c meta.c obj.c parsing.c \ + signals.c str.c sys.c callback.c weak.c finalise.c stacks.c \ + dynlink.c backtrace_prim.c backtrace.c spacetime.c afl.c \ + bigarray.c + +OBJS=$(addsuffix .$(O), \ + interp misc stacks fix_code startup_aux startup \ + freelist major_gc minor_gc memory alloc roots globroots \ + fail signals signals_byt printexc backtrace_prim backtrace \ + compare ints floats str array io extern intern \ + hash sys meta parsing gc_ctrl md5 obj \ + lexing callback debugger weak compact finalise custom \ + dynlink spacetime afl $(UNIX_OR_WIN32) bigarray main) + +DOBJS=$(OBJS:.$(O)=.$(DBGO)) instrtrace.$(DBGO) +IOBJS=$(OBJS:.$(O)=.i.$(O)) +PICOBJS=$(OBJS:.$(O)=.pic.$(O)) + +.PHONY: all +all: $(LIBRARIES) $(DYNLIBRARIES) $(PROGRAMS) + +ld.conf: ../config/Makefile + echo "$(STUBLIBDIR)" > $@ + echo "$(LIBDIR)" >> $@ + +INSTALL_INCDIR=$(INSTALL_LIBDIR)/caml + +.PHONY: install +install: + $(INSTALL_PROG) $(PROGRAMS) "$(INSTALL_BINDIR)" + $(INSTALL_DATA) $(LIBRARIES) "$(INSTALL_LIBDIR)" + if test -n "$(DYNLIBRARIES)"; then \ + $(INSTALL_PROG) $(DYNLIBRARIES) "$(INSTALL_LIBDIR)"; \ + fi + mkdir -p "$(INSTALL_INCDIR)" + $(INSTALL_DATA) caml/*.h "$(INSTALL_INCDIR)" + +# If primitives contain duplicated lines (e.g. because the code is defined +# like +# #ifdef X +# CAMLprim value caml_foo() ... +# #else +# CAMLprim value caml_foo() ... +# end), horrible things will happen (duplicated entries in Runtimedef -> +# double registration in Symtable -> empty entry in the PRIM table -> +# the bytecode interpreter is confused). +# We sort the primitive file and remove duplicates to avoid this problem. + +# Warning: we use "sort | uniq" instead of "sort -u" because in the MSVC +# port, the "sort" program in the path is Microsoft's and not cygwin's + +# Warning: POSIX sort is locale dependent, that's why we set LC_ALL explicitly. +# Sort is unstable for "is_directory" and "isatty" +# see http://pubs.opengroup.org/onlinepubs/9699919799/utilities/sort.html: +# "using sort to process pathnames, it is recommended that LC_ALL .. set to C" + + +primitives : $(PRIMS) + sed -n -e "s/CAMLprim value \([a-z0-9_][a-z0-9_]*\).*/\1/p" $(PRIMS) \ + | LC_ALL=C sort | uniq > primitives + +prims.c : primitives + (echo '#define CAML_INTERNALS'; \ + echo '#include "caml/mlvalues.h"'; \ + echo '#include "caml/prims.h"'; \ + sed -e 's/.*/extern value &();/' primitives; \ + echo 'c_primitive caml_builtin_cprim[] = {'; \ + sed -e 's/.*/ &,/' primitives; \ + echo ' 0 };'; \ + echo 'char * caml_names_of_builtin_cprim[] = {'; \ + sed -e 's/.*/ "&",/' primitives; \ + echo ' 0 };') > prims.c + +caml/opnames.h : caml/instruct.h + cat $^ | tr -d '\r' | \ + sed -e '/\/\*/d' \ + -e '/^#/d' \ + -e 's/enum /char * names_of_/' \ + -e 's/{$$/[] = {/' \ + -e 's/\([[:upper:]][[:upper:]_0-9]*\)/"\1"/g' > $@ + +# caml/jumptbl.h is required only if you have GCC 2.0 or later +caml/jumptbl.h : caml/instruct.h + cat $^ | tr -d '\r' | \ + sed -n -e '/^ /s/ \([A-Z]\)/ \&\&lbl_\1/gp' \ + -e '/^}/q' > $@ + +caml/version.h : ../VERSION ../tools/make-version-header.sh + ../tools/make-version-header.sh ../VERSION > caml/version.h + +.PHONY: clean +clean: + rm -f $(LIBRARIES) $(DYNLIBRARIES) $(PROGRAMS) *.$(O) *.$(A) *.$(SO) + rm -f primitives prims.c caml/opnames.h caml/jumptbl.h + rm -f caml/version.h + +ocamlrun$(EXE): prims.$(O) libcamlrun.$(A) + $(call MAKE_OCAMLRUN,$@,$^ $(LIBS)) + +libcamlrun.$(A): $(OBJS) + $(call MKLIB,$@, $^) + +ocamlrund$(EXE): prims.$(O) libcamlrund.$(A) + $(MKEXE) $(MKEXEDEBUGFLAG) -o $@ $^ $(LIBS) + +libcamlrund.$(A): $(DOBJS) + $(call MKLIB,$@, $^) + +ocamlruni$(EXE): prims.$(O) libcamlruni.$(A) + $(MKEXE) -o $@ $^ $(LIBS) + +libcamlruni.$(A): $(IOBJS) + $(call MKLIB,$@, $^) + +libcamlrun_pic.$(A): $(PICOBJS) + $(call MKLIB,$@, $^) + +libcamlrun_shared.$(SO): $(PICOBJS) + $(MKDLL) -o $@ $^ $(BYTECCLIBS) + +%.$(O): %.c + $(CC) -c $(CFLAGS) $(CPPFLAGS) $< + +%.$(DBGO): %.c + $(CC) -c $(DFLAGS) $(CPPFLAGS) $(OUTPUTOBJ)$@ $< + +%.i.$(O): %.c + $(CC) -c $(IFLAGS) $(CPPFLAGS) $(OUTPUTOBJ)$@ $< + +%.pic.$(O): %.c + $(CC) -c $(PICFLAGS) $(CPPFLAGS) $(OUTPUTOBJ)$@ $< + +.PHONY: depend +ifeq "$(TOOLCHAIN)" "msvc" +depend: + $(error Dependencies cannot be regenerated using the MSVC ports) +else +depend: prims.c caml/opnames.h caml/jumptbl.h caml/version.h + $(CC) -MM $(CFLAGS) $(CPPFLAGS) *.c | sed -e 's/\.o/.$$(O)/'>.$@ + $(CC) -MM $(DFLAGS) $(CPPFLAGS) *.c | sed -e 's/\.o/.d.$$(O)/'>>.$@ + $(CC) -MM $(IFLAGS) $(CPPFLAGS) *.c | sed -e 's/\.o/.i.$$(O)/'>>.$@ + $(CC) -MM $(PICFLAGS) $(CPPFLAGS) *.c | sed -e 's/\.o/.pic.$$(O)/'>>.$@ +endif + +include .depend diff --git a/test/monniaux/ocaml/byterun/afl.c b/test/monniaux/ocaml/byterun/afl.c new file mode 100644 index 00000000..324a3c34 --- /dev/null +++ b/test/monniaux/ocaml/byterun/afl.c @@ -0,0 +1,162 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Stephen Dolan, University of Cambridge */ +/* */ +/* Copyright 2016 Stephen Dolan. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Runtime support for afl-fuzz */ +#include "caml/config.h" + +#if !defined(HAS_SYS_SHM_H) + +#include "caml/mlvalues.h" + +CAMLprim value caml_setup_afl (value unit) +{ + return Val_unit; +} + +CAMLprim value caml_reset_afl_instrumentation(value unused) +{ + return Val_unit; +} + +#else + +#include +#include +#include +#include +#include +#include +#include + +#define CAML_INTERNALS +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/osdeps.h" + +static int afl_initialised = 0; + +/* afl uses abnormal termination (SIGABRT) to check whether + to count a testcase as "crashing" */ +extern int caml_abort_on_uncaught_exn; + +/* Values used by the instrumentation logic (see cmmgen.ml) */ +static unsigned char afl_area_initial[1 << 16]; +unsigned char* caml_afl_area_ptr = afl_area_initial; +uintnat caml_afl_prev_loc; + +/* File descriptors used to synchronise with afl-fuzz */ +#define FORKSRV_FD_READ 198 +#define FORKSRV_FD_WRITE 199 + +static void afl_write(uint32_t msg) +{ + if (write(FORKSRV_FD_WRITE, &msg, 4) != 4) + caml_fatal_error("writing to afl-fuzz"); +} + +static uint32_t afl_read() +{ + uint32_t msg; + if (read(FORKSRV_FD_READ, &msg, 4) != 4) + caml_fatal_error("reading from afl-fuzz"); + return msg; +} + +CAMLprim value caml_setup_afl(value unit) +{ + if (afl_initialised) return Val_unit; + afl_initialised = 1; + + char* shm_id_str = caml_secure_getenv("__AFL_SHM_ID"); + if (shm_id_str == NULL) { + /* Not running under afl-fuzz, continue as normal */ + return Val_unit; + } + + /* if afl-fuzz is attached, we want it to know about uncaught exceptions */ + caml_abort_on_uncaught_exn = 1; + + char* shm_id_end; + long int shm_id = strtol(shm_id_str, &shm_id_end, 10); + if (!(*shm_id_str != '\0' && *shm_id_end == '\0')) + caml_fatal_error("afl-fuzz: bad shm id"); + + caml_afl_area_ptr = shmat((int)shm_id, NULL, 0); + if (caml_afl_area_ptr == (void*)-1) + caml_fatal_error("afl-fuzz: could not attach shm area"); + + /* poke the bitmap so that afl-fuzz knows we exist, even if the + application has sparse instrumentation */ + caml_afl_area_ptr[0] = 1; + + /* synchronise with afl-fuzz */ + uint32_t startup_msg = 0; + if (write(FORKSRV_FD_WRITE, &startup_msg, 4) != 4) { + /* initial write failed, so assume we're not meant to fork. + afl-tmin uses this mode. */ + return Val_unit; + } + afl_read(); + + while (1) { + int child_pid = fork(); + if (child_pid < 0) caml_fatal_error("afl-fuzz: could not fork"); + else if (child_pid == 0) { + /* Run the program */ + close(FORKSRV_FD_READ); + close(FORKSRV_FD_WRITE); + return Val_unit; + } + + /* As long as the child keeps raising SIGSTOP, we re-use the same process */ + while (1) { + afl_write((uint32_t)child_pid); + + int status; + /* WUNTRACED means wait until termination or SIGSTOP */ + if (waitpid(child_pid, &status, WUNTRACED) < 0) + caml_fatal_error("afl-fuzz: waitpid failed"); + + afl_write((uint32_t)status); + + uint32_t was_killed = afl_read(); + if (WIFSTOPPED(status)) { + /* child stopped, waiting for another test case */ + if (was_killed) { + /* we saw the child stop, but since then afl-fuzz killed it. + we should wait for it before forking another child */ + if (waitpid(child_pid, &status, 0) < 0) + caml_fatal_error("afl-fuzz: waitpid failed"); + break; + } else { + kill(child_pid, SIGCONT); + } + } else { + /* child died */ + break; + } + } + } +} + +CAMLprim value caml_reset_afl_instrumentation(value full) +{ + if (full != Val_int(0)) { + memset(caml_afl_area_ptr, 0, sizeof(afl_area_initial)); + } + caml_afl_prev_loc = 0; + return Val_unit; +} + +#endif /* HAS_SYS_SHM_H */ diff --git a/test/monniaux/ocaml/byterun/alloc.c b/test/monniaux/ocaml/byterun/alloc.c new file mode 100644 index 00000000..8924dbc0 --- /dev/null +++ b/test/monniaux/ocaml/byterun/alloc.c @@ -0,0 +1,247 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* 1. Allocation functions doing the same work as the macros in the + case where [Setup_for_gc] and [Restore_after_gc] are no-ops. + 2. Convenience functions related to allocation. +*/ + +#include +#include "caml/alloc.h" +#include "caml/custom.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/stacks.h" + +#define Setup_for_gc +#define Restore_after_gc + +CAMLexport value caml_alloc (mlsize_t wosize, tag_t tag) +{ + value result; + mlsize_t i; + + CAMLassert (tag < 256); + CAMLassert (tag != Infix_tag); + if (wosize <= Max_young_wosize){ + if (wosize == 0){ + result = Atom (tag); + }else{ + Alloc_small (result, wosize, tag); + if (tag < No_scan_tag){ + for (i = 0; i < wosize; i++) Field (result, i) = Val_unit; + } + } + }else{ + result = caml_alloc_shr (wosize, tag); + if (tag < No_scan_tag){ + for (i = 0; i < wosize; i++) Field (result, i) = Val_unit; + } + result = caml_check_urgent_gc (result); + } + return result; +} + +CAMLexport value caml_alloc_small (mlsize_t wosize, tag_t tag) +{ + value result; + + CAMLassert (wosize > 0); + CAMLassert (wosize <= Max_young_wosize); + CAMLassert (tag < 256); + Alloc_small (result, wosize, tag); + return result; +} + +CAMLexport value caml_alloc_small_with_my_or_given_profinfo (mlsize_t wosize, + tag_t tag, uintnat profinfo) +{ + if (profinfo == 0) { + return caml_alloc_small(wosize, tag); + } + else { + value result; + + CAMLassert (wosize > 0); + CAMLassert (wosize <= Max_young_wosize); + CAMLassert (tag < 256); + Alloc_small_with_profinfo (result, wosize, tag, profinfo); + return result; + } +} + +/* [n] is a number of words (fields) */ +CAMLexport value caml_alloc_tuple(mlsize_t n) +{ + return caml_alloc(n, 0); +} + +/* [len] is a number of bytes (chars) */ +CAMLexport value caml_alloc_string (mlsize_t len) +{ + value result; + mlsize_t offset_index; + mlsize_t wosize = (len + sizeof (value)) / sizeof (value); + + if (wosize <= Max_young_wosize) { + Alloc_small (result, wosize, String_tag); + }else{ + result = caml_alloc_shr (wosize, String_tag); + result = caml_check_urgent_gc (result); + } + Field (result, wosize - 1) = 0; + offset_index = Bsize_wsize (wosize) - 1; + Byte (result, offset_index) = offset_index - len; + return result; +} + +/* [len] is a number of bytes (chars) */ +CAMLexport value caml_alloc_initialized_string (mlsize_t len, const char *p) +{ + value result = caml_alloc_string (len); + memcpy((char *)String_val(result), p, len); + return result; +} + +/* [len] is a number of words. + [mem] and [max] are relative (without unit). +*/ +CAMLexport value caml_alloc_final (mlsize_t len, final_fun fun, + mlsize_t mem, mlsize_t max) +{ + return caml_alloc_custom(caml_final_custom_operations(fun), + len * sizeof(value), mem, max); +} + +CAMLexport value caml_copy_string(char const *s) +{ + int len; + value res; + + len = strlen(s); + res = caml_alloc_initialized_string(len, s); + return res; +} + +CAMLexport value caml_alloc_array(value (*funct)(char const *), + char const ** arr) +{ + CAMLparam0 (); + mlsize_t nbr, n; + CAMLlocal2 (v, result); + + nbr = 0; + while (arr[nbr] != 0) nbr++; + result = caml_alloc (nbr, 0); + for (n = 0; n < nbr; n++) { + /* The two statements below must be separate because of evaluation + order (don't take the address &Field(result, n) before + calling funct, which may cause a GC and move result). */ + v = funct(arr[n]); + caml_modify(&Field(result, n), v); + } + CAMLreturn (result); +} + +/* [len] is a number of floats */ +value caml_alloc_float_array(mlsize_t len) +{ +#ifdef FLAT_FLOAT_ARRAY + mlsize_t wosize = len * Double_wosize; + value result; + /* For consistency with [caml_make_vect], which can't tell whether it should + create a float array or not when the size is zero, the tag is set to + zero when the size is zero. */ + if (wosize <= Max_young_wosize){ + if (wosize == 0) + return Atom(0); + else + Alloc_small (result, wosize, Double_array_tag); + }else { + result = caml_alloc_shr (wosize, Double_array_tag); + result = caml_check_urgent_gc (result); + } + return result; +#else + return caml_alloc (len, 0); +#endif +} + + +CAMLexport value caml_copy_string_array(char const ** arr) +{ + return caml_alloc_array(caml_copy_string, arr); +} + +CAMLexport int caml_convert_flag_list(value list, int *flags) +{ + int res; + res = 0; + while (list != Val_int(0)) { + res |= flags[Int_val(Field(list, 0))]; + list = Field(list, 1); + } + return res; +} + +/* For compiling let rec over values */ + +/* [size] is a [value] representing number of words (fields) */ +CAMLprim value caml_alloc_dummy(value size) +{ + mlsize_t wosize = Long_val(size); + return caml_alloc (wosize, 0); +} + +/* [size] is a [value] representing number of words (fields) */ +CAMLprim value caml_alloc_dummy_function(value size,value arity) +{ + /* the arity argument is used by the js_of_ocaml runtime */ + return caml_alloc_dummy(size); +} + +/* [size] is a [value] representing number of floats. */ +CAMLprim value caml_alloc_dummy_float (value size) +{ + mlsize_t wosize = Long_val(size) * Double_wosize; + return caml_alloc (wosize, 0); +} + +CAMLprim value caml_update_dummy(value dummy, value newval) +{ + mlsize_t size, i; + tag_t tag; + + size = Wosize_val(newval); + tag = Tag_val (newval); + CAMLassert (size == Wosize_val(dummy)); + CAMLassert (tag < No_scan_tag || tag == Double_array_tag); + + Tag_val(dummy) = tag; + if (tag == Double_array_tag){ + size = Wosize_val (newval) / Double_wosize; + for (i = 0; i < size; i++){ + Store_double_flat_field (dummy, i, Double_flat_field (newval, i)); + } + }else{ + for (i = 0; i < size; i++){ + caml_modify (&Field(dummy, i), Field(newval, i)); + } + } + return Val_unit; +} diff --git a/test/monniaux/ocaml/byterun/array.c b/test/monniaux/ocaml/byterun/array.c new file mode 100644 index 00000000..5367532b --- /dev/null +++ b/test/monniaux/ocaml/byterun/array.c @@ -0,0 +1,586 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Operations on arrays */ +#include +#include "caml/alloc.h" +#include "caml/fail.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/signals.h" +/* Why is caml/spacetime.h included conditionnally sometimes and not here ? */ +#include "caml/spacetime.h" + +static const mlsize_t mlsize_t_max = -1; + +/* returns number of elements (either fields or floats) */ +/* [ 'a array -> int ] */ +CAMLexport mlsize_t caml_array_length(value array) +{ +#ifdef FLAT_FLOAT_ARRAY + if (Tag_val(array) == Double_array_tag) + return Wosize_val(array) / Double_wosize; + else +#endif + return Wosize_val(array); +} + +CAMLexport int caml_is_double_array(value array) +{ + return (Tag_val(array) == Double_array_tag); +} + +/* Note: the OCaml types on the following primitives will work both with + and without the -no-flat-float-array configure-time option. If you + respect them, your C code should work in both configurations. +*/ + +/* [ 'a array -> int -> 'a ] where 'a != float */ +CAMLprim value caml_array_get_addr(value array, value index) +{ + intnat idx = Long_val(index); + if (idx < 0 || idx >= Wosize_val(array)) caml_array_bound_error(); + return Field(array, idx); +} + +/* [ float array -> int -> float ] */ +CAMLprim value caml_array_get_float(value array, value index) +{ + intnat idx = Long_val(index); +#ifdef FLAT_FLOAT_ARRAY + double d; + value res; + + if (idx < 0 || idx >= Wosize_val(array) / Double_wosize) + caml_array_bound_error(); + d = Double_flat_field(array, idx); +#define Setup_for_gc +#define Restore_after_gc + Alloc_small(res, Double_wosize, Double_tag); +#undef Setup_for_gc +#undef Restore_after_gc + Store_double_val(res, d); + return res; +#else + CAMLassert (Tag_val (array) != Double_array_tag); + if (idx < 0 || idx >= Wosize_val(array)) caml_array_bound_error(); + return Field(array, idx); +#endif /* FLAT_FLOAT_ARRAY */ +} + +/* [ 'a array -> int -> 'a ] */ +CAMLprim value caml_array_get(value array, value index) +{ +#ifdef FLAT_FLOAT_ARRAY + if (Tag_val(array) == Double_array_tag) + return caml_array_get_float(array, index); +#else + CAMLassert (Tag_val(array) != Double_array_tag); +#endif + return caml_array_get_addr(array, index); +} + +/* [ floatarray -> int -> float ] */ +CAMLprim value caml_floatarray_get(value array, value index) +{ + intnat idx = Long_val(index); + double d; + value res; + + CAMLassert (Tag_val(array) == Double_array_tag); + if (idx < 0 || idx >= Wosize_val(array) / Double_wosize) + caml_array_bound_error(); + d = Double_flat_field(array, idx); +#define Setup_for_gc +#define Restore_after_gc + Alloc_small(res, Double_wosize, Double_tag); +#undef Setup_for_gc +#undef Restore_after_gc + Store_double_val(res, d); + return res; +} + +/* [ 'a array -> int -> 'a -> unit ] where 'a != float */ +CAMLprim value caml_array_set_addr(value array, value index, value newval) +{ + intnat idx = Long_val(index); + if (idx < 0 || idx >= Wosize_val(array)) caml_array_bound_error(); + Modify(&Field(array, idx), newval); + return Val_unit; +} + +/* [ float array -> int -> float -> unit ] */ +CAMLprim value caml_array_set_float(value array, value index, value newval) +{ + intnat idx = Long_val(index); +#ifdef FLAT_FLOAT_ARRAY + double d = Double_val (newval); + if (idx < 0 || idx >= Wosize_val(array) / Double_wosize) + caml_array_bound_error(); + Store_double_flat_field(array, idx, d); +#else + CAMLassert (Tag_val (array) != Double_array_tag); + if (idx < 0 || idx >= Wosize_val(array)) caml_array_bound_error(); + Modify(&Field(array, idx), newval); +#endif + return Val_unit; +} + +/* [ 'a array -> int -> 'a -> unit ] */ +CAMLprim value caml_array_set(value array, value index, value newval) +{ +#ifdef FLAT_FLOAT_ARRAY + if (Tag_val(array) == Double_array_tag) + return caml_array_set_float(array, index, newval); +#else + CAMLassert (Tag_val(array) != Double_array_tag); +#endif + return caml_array_set_addr(array, index, newval); +} + +/* [ floatarray -> int -> float -> unit ] */ +CAMLprim value caml_floatarray_set(value array, value index, value newval) +{ + intnat idx = Long_val(index); + double d = Double_val (newval); + CAMLassert (Tag_val(array) == Double_array_tag); + if (idx < 0 || idx >= Wosize_val(array) / Double_wosize) + caml_array_bound_error(); + Store_double_flat_field(array, idx, d); + return Val_unit; +} + +/* [ float array -> int -> float ] */ +CAMLprim value caml_array_unsafe_get_float(value array, value index) +{ + intnat idx = Long_val (index); +#ifdef FLAT_FLOAT_ARRAY + double d; + value res; + + d = Double_flat_field(array, idx); +#define Setup_for_gc +#define Restore_after_gc + Alloc_small(res, Double_wosize, Double_tag); +#undef Setup_for_gc +#undef Restore_after_gc + Store_double_val(res, d); + return res; +#else /* FLAT_FLOAT_ARRAY */ + CAMLassert (Tag_val(array) != Double_array_tag); + return Field(array, idx); +#endif /* FLAT_FLOAT_ARRAY */ +} + +/* [ 'a array -> int -> 'a ] */ +CAMLprim value caml_array_unsafe_get(value array, value index) +{ +#ifdef FLAT_FLOAT_ARRAY + if (Tag_val(array) == Double_array_tag) + return caml_array_unsafe_get_float(array, index); +#else + CAMLassert (Tag_val(array) != Double_array_tag); +#endif + return Field(array, Long_val(index)); +} + +/* [ floatarray -> int -> float ] */ +CAMLprim value caml_floatarray_unsafe_get(value array, value index) +{ + intnat idx = Long_val(index); + double d; + value res; + + CAMLassert (Tag_val(array) == Double_array_tag); + d = Double_flat_field(array, idx); +#define Setup_for_gc +#define Restore_after_gc + Alloc_small(res, Double_wosize, Double_tag); +#undef Setup_for_gc +#undef Restore_after_gc + Store_double_val(res, d); + return res; +} + +/* [ 'a array -> int -> 'a -> unit ] where 'a != float */ +CAMLprim value caml_array_unsafe_set_addr(value array, value index,value newval) +{ + intnat idx = Long_val(index); + Modify(&Field(array, idx), newval); + return Val_unit; +} + +/* [ float array -> int -> float -> unit ] */ +CAMLprim value caml_array_unsafe_set_float(value array,value index,value newval) +{ + intnat idx = Long_val(index); +#ifdef FLAT_FLOAT_ARRAY + double d = Double_val (newval); + Store_double_flat_field(array, idx, d); +#else + Modify(&Field(array, idx), newval); +#endif + return Val_unit; +} + +/* [ 'a array -> int -> 'a -> unit ] */ +CAMLprim value caml_array_unsafe_set(value array, value index, value newval) +{ +#ifdef FLAT_FLOAT_ARRAY + if (Tag_val(array) == Double_array_tag) + return caml_array_unsafe_set_float(array, index, newval); +#else + CAMLassert (Tag_val(array) != Double_array_tag); +#endif + return caml_array_unsafe_set_addr(array, index, newval); +} + +/* [ floatarray -> int -> float -> unit ] */ +CAMLprim value caml_floatarray_unsafe_set(value array, value index,value newval) +{ + intnat idx = Long_val(index); + double d = Double_val (newval); + Store_double_flat_field(array, idx, d); + return Val_unit; +} + +/* [len] is a [value] representing number of floats. */ +/* [ int -> floatarray ] */ +CAMLprim value caml_floatarray_create(value len) +{ + mlsize_t wosize = Long_val(len) * Double_wosize; + value result; + if (wosize <= Max_young_wosize){ + if (wosize == 0) + return Atom(0); + else +#define Setup_for_gc +#define Restore_after_gc + Alloc_small (result, wosize, Double_array_tag); +#undef Setup_for_gc +#undef Restore_after_gc + }else if (wosize > Max_wosize) + caml_invalid_argument("Array.Floatarray.create"); + else { + result = caml_alloc_shr (wosize, Double_array_tag); + result = caml_check_urgent_gc (result); + } + return result; +} + +/* [len] is a [value] representing number of floats */ +/* [ int -> float array ] */ +CAMLprim value caml_make_float_vect(value len) +{ +#ifdef FLAT_FLOAT_ARRAY + return caml_floatarray_create (len); +#else + return caml_alloc (Long_val (len), 0); +#endif +} + +/* [len] is a [value] representing number of words or floats */ +/* Spacetime profiling assumes that this function is only called from OCaml. */ +CAMLprim value caml_make_vect(value len, value init) +{ + CAMLparam2 (len, init); + CAMLlocal1 (res); + mlsize_t size, i; + + size = Long_val(len); + if (size == 0) { + res = Atom(0); +#ifdef FLAT_FLOAT_ARRAY + } else if (Is_block(init) + && Is_in_value_area(init) + && Tag_val(init) == Double_tag) { + mlsize_t wsize; + double d; + d = Double_val(init); + wsize = size * Double_wosize; + if (wsize > Max_wosize) caml_invalid_argument("Array.make"); + res = caml_alloc(wsize, Double_array_tag); + for (i = 0; i < size; i++) { + Store_double_flat_field(res, i, d); + } +#endif + } else { + if (size <= Max_young_wosize) { + uintnat profinfo; + Get_my_profinfo_with_cached_backtrace(profinfo, size); + res = caml_alloc_small_with_my_or_given_profinfo(size, 0, profinfo); + for (i = 0; i < size; i++) Field(res, i) = init; + } + else if (size > Max_wosize) caml_invalid_argument("Array.make"); + else if (Is_block(init) && Is_young(init)) { + /* We don't want to create so many major-to-minor references, + so [init] is moved to the major heap by doing a minor GC. */ + CAML_INSTR_INT ("force_minor/make_vect@", 1); + caml_request_minor_gc (); + caml_gc_dispatch (); + res = caml_alloc_shr(size, 0); + for (i = 0; i < size; i++) Field(res, i) = init; + res = caml_check_urgent_gc (res); + } + else { + res = caml_alloc_shr(size, 0); + for (i = 0; i < size; i++) caml_initialize(&Field(res, i), init); + res = caml_check_urgent_gc (res); + } + } + CAMLreturn (res); +} + +/* This primitive is used internally by the compiler to compile + explicit array expressions. + For float arrays when FLAT_FLOAT_ARRAY is true, it takes an array of + boxed floats and returns the corresponding flat-allocated [float array]. + In all other cases, it just returns its argument unchanged. +*/ +CAMLprim value caml_make_array(value init) +{ +#ifdef FLAT_FLOAT_ARRAY + CAMLparam1 (init); + mlsize_t wsize, size, i; + CAMLlocal2 (v, res); + + size = Wosize_val(init); + if (size == 0) { + CAMLreturn (init); + } else { + v = Field(init, 0); + if (Is_long(v) + || ! Is_in_value_area(v) + || Tag_val(v) != Double_tag) { + CAMLreturn (init); + } else { + wsize = size * Double_wosize; + if (wsize <= Max_young_wosize) { + res = caml_alloc_small(wsize, Double_array_tag); + } else { + res = caml_alloc_shr(wsize, Double_array_tag); + res = caml_check_urgent_gc(res); + } + for (i = 0; i < size; i++) { + double d = Double_val(Field(init, i)); + Store_double_flat_field(res, i, d); + } + CAMLreturn (res); + } + } +#else + return init; +#endif +} + +/* Blitting */ + +CAMLprim value caml_array_blit(value a1, value ofs1, value a2, value ofs2, + value n) +{ + value * src, * dst; + intnat count; + +#ifdef FLAT_FLOAT_ARRAY + if (Tag_val(a2) == Double_array_tag) { + /* Arrays of floats. The values being copied are floats, not + pointer, so we can do a direct copy. memmove takes care of + potential overlap between the copied areas. */ + memmove((double *)a2 + Long_val(ofs2), + (double *)a1 + Long_val(ofs1), + Long_val(n) * sizeof(double)); + return Val_unit; + } +#endif + CAMLassert (Tag_val(a2) != Double_array_tag); + if (Is_young(a2)) { + /* Arrays of values, destination is in young generation. + Here too we can do a direct copy since this cannot create + old-to-young pointers, nor mess up with the incremental major GC. + Again, memmove takes care of overlap. */ + memmove(&Field(a2, Long_val(ofs2)), + &Field(a1, Long_val(ofs1)), + Long_val(n) * sizeof(value)); + return Val_unit; + } + /* Array of values, destination is in old generation. + We must use caml_modify. */ + count = Long_val(n); + if (a1 == a2 && Long_val(ofs1) < Long_val(ofs2)) { + /* Copy in descending order */ + for (dst = &Field(a2, Long_val(ofs2) + count - 1), + src = &Field(a1, Long_val(ofs1) + count - 1); + count > 0; + count--, src--, dst--) { + caml_modify(dst, *src); + } + } else { + /* Copy in ascending order */ + for (dst = &Field(a2, Long_val(ofs2)), src = &Field(a1, Long_val(ofs1)); + count > 0; + count--, src++, dst++) { + caml_modify(dst, *src); + } + } + /* Many caml_modify in a row can create a lot of old-to-young refs. + Give the minor GC a chance to run if it needs to. */ + caml_check_urgent_gc(Val_unit); + return Val_unit; +} + +/* A generic function for extraction and concatenation of sub-arrays */ + +static value caml_array_gather(intnat num_arrays, + value arrays[/*num_arrays*/], + intnat offsets[/*num_arrays*/], + intnat lengths[/*num_arrays*/]) +{ + CAMLparamN(arrays, num_arrays); + value res; /* no need to register it as a root */ +#ifdef FLAT_FLOAT_ARRAY + int isfloat = 0; + mlsize_t wsize; +#endif + mlsize_t i, size, count, pos; + value * src; + + /* Determine total size and whether result array is an array of floats */ + size = 0; + for (i = 0; i < num_arrays; i++) { + if (mlsize_t_max - lengths[i] < size) caml_invalid_argument("Array.concat"); + size += lengths[i]; +#ifdef FLAT_FLOAT_ARRAY + if (Tag_val(arrays[i]) == Double_array_tag) isfloat = 1; +#endif + } + if (size == 0) { + /* If total size = 0, just return empty array */ + res = Atom(0); + } +#ifdef FLAT_FLOAT_ARRAY + else if (isfloat) { + /* This is an array of floats. We can use memcpy directly. */ + if (size > Max_wosize/Double_wosize) caml_invalid_argument("Array.concat"); + wsize = size * Double_wosize; + res = caml_alloc(wsize, Double_array_tag); + for (i = 0, pos = 0; i < num_arrays; i++) { + memcpy((double *)res + pos, + (double *)arrays[i] + offsets[i], + lengths[i] * sizeof(double)); + pos += lengths[i]; + } + CAMLassert(pos == size); + } +#endif + else if (size <= Max_young_wosize) { + /* Array of values, small enough to fit in young generation. + We can use memcpy directly. */ + res = caml_alloc_small(size, 0); + for (i = 0, pos = 0; i < num_arrays; i++) { + memcpy(&Field(res, pos), + &Field(arrays[i], offsets[i]), + lengths[i] * sizeof(value)); + pos += lengths[i]; + } + CAMLassert(pos == size); + } + else if (size > Max_wosize) { + /* Array of values, too big. */ + caml_invalid_argument("Array.concat"); + } else { + /* Array of values, must be allocated in old generation and filled + using caml_initialize. */ + res = caml_alloc_shr(size, 0); + for (i = 0, pos = 0; i < num_arrays; i++) { + for (src = &Field(arrays[i], offsets[i]), count = lengths[i]; + count > 0; + count--, src++, pos++) { + caml_initialize(&Field(res, pos), *src); + } + } + CAMLassert(pos == size); + + /* Many caml_initialize in a row can create a lot of old-to-young + refs. Give the minor GC a chance to run if it needs to. */ + res = caml_check_urgent_gc(res); + } + CAMLreturn (res); +} + +CAMLprim value caml_array_sub(value a, value ofs, value len) +{ + value arrays[1] = { a }; + intnat offsets[1] = { Long_val(ofs) }; + intnat lengths[1] = { Long_val(len) }; + return caml_array_gather(1, arrays, offsets, lengths); +} + +CAMLprim value caml_array_append(value a1, value a2) +{ + value arrays[2] = { a1, a2 }; + intnat offsets[2] = { 0, 0 }; + intnat lengths[2] = { caml_array_length(a1), caml_array_length(a2) }; + return caml_array_gather(2, arrays, offsets, lengths); +} + +CAMLprim value caml_array_concat(value al) +{ +#define STATIC_SIZE 16 + value static_arrays[STATIC_SIZE], * arrays; + intnat static_offsets[STATIC_SIZE], * offsets; + intnat static_lengths[STATIC_SIZE], * lengths; + intnat n, i; + value l, res; + + /* Length of list = number of arrays */ + for (n = 0, l = al; l != Val_int(0); l = Field(l, 1)) n++; + /* Allocate extra storage if too many arrays */ + if (n <= STATIC_SIZE) { + arrays = static_arrays; + offsets = static_offsets; + lengths = static_lengths; + } else { + arrays = caml_stat_alloc(n * sizeof(value)); + offsets = caml_stat_alloc_noexc(n * sizeof(intnat)); + if (offsets == NULL) { + caml_stat_free(arrays); + caml_raise_out_of_memory(); + } + lengths = caml_stat_alloc_noexc(n * sizeof(value)); + if (lengths == NULL) { + caml_stat_free(offsets); + caml_stat_free(arrays); + caml_raise_out_of_memory(); + } + } + /* Build the parameters to caml_array_gather */ + for (i = 0, l = al; l != Val_int(0); l = Field(l, 1), i++) { + arrays[i] = Field(l, 0); + offsets[i] = 0; + lengths[i] = caml_array_length(Field(l, 0)); + } + /* Do the concatenation */ + res = caml_array_gather(n, arrays, offsets, lengths); + /* Free the extra storage if needed */ + if (n > STATIC_SIZE) { + caml_stat_free(arrays); + caml_stat_free(offsets); + caml_stat_free(lengths); + } + return res; +} diff --git a/test/monniaux/ocaml/byterun/backtrace.c b/test/monniaux/ocaml/byterun/backtrace.c new file mode 100644 index 00000000..8dfe9b7e --- /dev/null +++ b/test/monniaux/ocaml/byterun/backtrace.c @@ -0,0 +1,349 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Stack backtrace for uncaught exceptions */ + +#include +#include +#include + +#include "caml/alloc.h" +#include "caml/memory.h" +#include "caml/backtrace.h" +#include "caml/backtrace_prim.h" +#include "caml/fail.h" + +/* The table of debug information fragments */ +struct ext_table caml_debug_info; + +CAMLexport int32_t caml_backtrace_active = 0; +CAMLexport int32_t caml_backtrace_pos = 0; +CAMLexport backtrace_slot * caml_backtrace_buffer = NULL; +CAMLexport value caml_backtrace_last_exn = Val_unit; + +void caml_init_backtrace(void) +{ + caml_register_global_root(&caml_backtrace_last_exn); +} + +/* Start or stop the backtrace machinery */ +CAMLprim value caml_record_backtrace(value vflag) +{ + int flag = Int_val(vflag); + + if (flag != caml_backtrace_active) { + caml_backtrace_active = flag; + caml_backtrace_pos = 0; + caml_backtrace_last_exn = Val_unit; + /* Note: We do lazy initialization of caml_backtrace_buffer when + needed in order to simplify the interface with the thread + library (thread creation doesn't need to allocate + caml_backtrace_buffer). So we don't have to allocate it here. + */ + } + return Val_unit; +} + +/* Return the status of the backtrace machinery */ +CAMLprim value caml_backtrace_status(value vunit) +{ + return Val_bool(caml_backtrace_active); +} + +/* Print location information -- same behavior as in Printexc + + note that the test for compiler-inserted raises is slightly redundant: + (!li->loc_valid && li->loc_is_raise) + caml_debuginfo_location guarantees that when li->loc_valid is + 0, then li->loc_is_raise is always 1, so the latter test is + useless. We kept it to keep code identical to the byterun/ + implementation. */ +static void print_location(struct caml_loc_info * li, int index) +{ + char * info; + char * inlined; + + /* Ignore compiler-inserted raise */ + if (!li->loc_valid && li->loc_is_raise) return; + + if (li->loc_is_raise) { + /* Initial raise if index == 0, re-raise otherwise */ + if (index == 0) + info = "Raised at"; + else + info = "Re-raised at"; + } else { + if (index == 0) + info = "Raised by primitive operation at"; + else + info = "Called from"; + } + if (li->loc_is_inlined) { + inlined = " (inlined)"; + } else { + inlined = ""; + } + if (! li->loc_valid) { + fprintf(stderr, "%s unknown location%s\n", info, inlined); + } else { + fprintf (stderr, "%s file \"%s\"%s, line %d, characters %d-%d\n", + info, li->loc_filename, inlined, li->loc_lnum, + li->loc_startchr, li->loc_endchr); + } +} + +/* Print a backtrace */ +CAMLexport void caml_print_exception_backtrace(void) +{ + int i; + struct caml_loc_info li; + debuginfo dbg; + + if (!caml_debug_info_available()) { + fprintf(stderr, "(Cannot print stack backtrace: " + "no debug information available)\n"); + return; + } + + for (i = 0; i < caml_backtrace_pos; i++) { + for (dbg = caml_debuginfo_extract(caml_backtrace_buffer[i]); + dbg != NULL; + dbg = caml_debuginfo_next(dbg)) + { + caml_debuginfo_location(dbg, &li); + print_location(&li, i); + } + } +} + +/* Get a copy of the latest backtrace */ +CAMLprim value caml_get_exception_raw_backtrace(value unit) +{ + CAMLparam0(); + CAMLlocal1(res); + + /* Beware: the allocations below may cause finalizers to be run, and another + backtrace---possibly of a different length---to be stashed (for example + if the finalizer raises then catches an exception). We choose to ignore + any such finalizer backtraces and return the original one. */ + + if (!caml_backtrace_active || + caml_backtrace_buffer == NULL || + caml_backtrace_pos == 0) { + res = caml_alloc(0, 0); + } + else { + backtrace_slot saved_caml_backtrace_buffer[BACKTRACE_BUFFER_SIZE]; + int saved_caml_backtrace_pos; + intnat i; + + saved_caml_backtrace_pos = caml_backtrace_pos; + + if (saved_caml_backtrace_pos > BACKTRACE_BUFFER_SIZE) { + saved_caml_backtrace_pos = BACKTRACE_BUFFER_SIZE; + } + + memcpy(saved_caml_backtrace_buffer, caml_backtrace_buffer, + saved_caml_backtrace_pos * sizeof(backtrace_slot)); + + res = caml_alloc(saved_caml_backtrace_pos, 0); + for (i = 0; i < saved_caml_backtrace_pos; i++) { + Field(res, i) = Val_backtrace_slot(saved_caml_backtrace_buffer[i]); + } + } + + CAMLreturn(res); +} + +/* Copy back a backtrace and exception to the global state. + This function should be used only with Printexc.raw_backtrace */ +/* noalloc (caml value): so no CAMLparam* CAMLreturn* */ +CAMLprim value caml_restore_raw_backtrace(value exn, value backtrace) +{ + intnat i; + mlsize_t bt_size; + + caml_backtrace_last_exn = exn; + + bt_size = Wosize_val(backtrace); + if(bt_size > BACKTRACE_BUFFER_SIZE){ + bt_size = BACKTRACE_BUFFER_SIZE; + } + + /* We don't allocate if the backtrace is empty (no -g or backtrace + not activated) */ + if(bt_size == 0){ + caml_backtrace_pos = 0; + return Val_unit; + } + + /* Allocate if needed and copy the backtrace buffer */ + if (caml_backtrace_buffer == NULL && caml_alloc_backtrace_buffer() == -1){ + return Val_unit; + } + + caml_backtrace_pos = bt_size; + for(i=0; i < caml_backtrace_pos; i++){ + caml_backtrace_buffer[i] = Backtrace_slot_val(Field(backtrace, i)); + } + + return Val_unit; +} + +#define Val_debuginfo(bslot) (Val_long((uintnat)(bslot)>>1)) +#define Debuginfo_val(vslot) ((debuginfo)(Long_val(vslot) << 1)) + +/* Convert the raw backtrace to a data structure usable from OCaml */ +static value caml_convert_debuginfo(debuginfo dbg) +{ + CAMLparam0(); + CAMLlocal2(p, fname); + struct caml_loc_info li; + + caml_debuginfo_location(dbg, &li); + + if (li.loc_valid) { + fname = caml_copy_string(li.loc_filename); + p = caml_alloc_small(6, 0); + Field(p, 0) = Val_bool(li.loc_is_raise); + Field(p, 1) = fname; + Field(p, 2) = Val_int(li.loc_lnum); + Field(p, 3) = Val_int(li.loc_startchr); + Field(p, 4) = Val_int(li.loc_endchr); + Field(p, 5) = Val_bool(li.loc_is_inlined); + } else { + p = caml_alloc_small(1, 1); + Field(p, 0) = Val_bool(li.loc_is_raise); + } + + CAMLreturn(p); +} + +CAMLprim value caml_convert_raw_backtrace_slot(value slot) +{ + if (!caml_debug_info_available()) + caml_failwith("No debug information available"); + + return (caml_convert_debuginfo(Debuginfo_val(slot))); +} + +/* Convert the raw backtrace to a data structure usable from OCaml */ +CAMLprim value caml_convert_raw_backtrace(value bt) +{ + CAMLparam1(bt); + CAMLlocal1(array); + intnat i, index; + + if (!caml_debug_info_available()) + caml_failwith("No debug information available"); + + for (i = 0, index = 0; i < Wosize_val(bt); ++i) + { + debuginfo dbg; + for (dbg = caml_debuginfo_extract(Backtrace_slot_val(Field(bt, i))); + dbg != NULL; + dbg = caml_debuginfo_next(dbg)) + index++; + } + + array = caml_alloc(index, 0); + + for (i = 0, index = 0; i < Wosize_val(bt); ++i) + { + debuginfo dbg; + for (dbg = caml_debuginfo_extract(Backtrace_slot_val(Field(bt, i))); + dbg != NULL; + dbg = caml_debuginfo_next(dbg)) + { + Store_field(array, index, caml_convert_debuginfo(dbg)); + index++; + } + } + + CAMLreturn(array); +} + +CAMLprim value caml_raw_backtrace_length(value bt) +{ + return Val_int(Wosize_val(bt)); +} + +CAMLprim value caml_raw_backtrace_slot(value bt, value index) +{ + uintnat i; + debuginfo dbg; + + i = Long_val(index); + if (i >= Wosize_val(bt)) + caml_invalid_argument("Printexc.get_raw_backtrace_slot: " + "index out of bounds"); + dbg = caml_debuginfo_extract(Backtrace_slot_val(Field(bt, i))); + return Val_debuginfo(dbg); +} + +CAMLprim value caml_raw_backtrace_next_slot(value slot) +{ + debuginfo dbg; + + CAMLparam1(slot); + CAMLlocal1(v); + + dbg = Debuginfo_val(slot); + dbg = caml_debuginfo_next(dbg); + + if (dbg == NULL) + v = Val_int(0); /* None */ + else + { + v = caml_alloc(1, 0); + Field(v, 0) = Val_debuginfo(dbg); + } + + CAMLreturn(v); +} + +/* the function below is deprecated: we previously returned directly + the OCaml-usable representation, instead of the raw backtrace as an + abstract type, but this has a large performance overhead if you + store a lot of backtraces and print only some of them. + + It is not used by the Printexc library anymore, or anywhere else in + the compiler, but we have kept it in case some user still depends + on it as an external. */ +CAMLprim value caml_get_exception_backtrace(value unit) +{ + CAMLparam0(); + CAMLlocal3(arr, res, backtrace); + intnat i; + + if (!caml_debug_info_available()) { + res = Val_int(0); /* None */ + } else { + backtrace = caml_get_exception_raw_backtrace(Val_unit); + + arr = caml_alloc(Wosize_val(backtrace), 0); + for (i = 0; i < Wosize_val(backtrace); i++) { + backtrace_slot slot = Backtrace_slot_val(Field(backtrace, i)); + debuginfo dbg = caml_debuginfo_extract(slot); + Store_field(arr, i, caml_convert_debuginfo(dbg)); + } + + res = caml_alloc_small(1, 0); Field(res, 0) = arr; /* Some */ + } + + CAMLreturn(res); +} diff --git a/test/monniaux/ocaml/byterun/backtrace_prim.c b/test/monniaux/ocaml/byterun/backtrace_prim.c new file mode 100644 index 00000000..e69b2568 --- /dev/null +++ b/test/monniaux/ocaml/byterun/backtrace_prim.c @@ -0,0 +1,457 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Stack backtrace for uncaught exceptions */ + +#include +#include +#include +#include + +#include "caml/config.h" +#ifdef HAS_UNISTD +#include +#endif + +#include "caml/mlvalues.h" +#include "caml/alloc.h" +#include "caml/custom.h" +#include "caml/io.h" +#include "caml/instruct.h" +#include "caml/intext.h" +#include "caml/exec.h" +#include "caml/fix_code.h" +#include "caml/memory.h" +#include "caml/startup.h" +#include "caml/stacks.h" +#include "caml/sys.h" +#include "caml/backtrace.h" +#include "caml/fail.h" +#include "caml/backtrace_prim.h" + +/* The table of debug information fragments */ +struct ext_table caml_debug_info; + +CAMLexport char_os * caml_cds_file = NULL; + +/* Location of fields in the Instruct.debug_event record */ +enum { + EV_POS = 0, + EV_MODULE = 1, + EV_LOC = 2, + EV_KIND = 3 +}; + +/* Location of fields in the Location.t record. */ +enum { + LOC_START = 0, + LOC_END = 1, + LOC_GHOST = 2 +}; + +/* Location of fields in the Lexing.position record. */ +enum { + POS_FNAME = 0, + POS_LNUM = 1, + POS_BOL = 2, + POS_CNUM = 3 +}; + +/* Runtime representation of the debug information, optimized + for quick lookup */ +struct ev_info { + code_t ev_pc; + char *ev_filename; + int ev_lnum; + int ev_startchr; + int ev_endchr; +}; + +struct debug_info { + code_t start; + code_t end; + mlsize_t num_events; + struct ev_info *events; + int already_read; +}; + +static struct debug_info *find_debug_info(code_t pc) +{ + int i; + for (i = 0; i < caml_debug_info.size; i++) { + struct debug_info *di = caml_debug_info.contents[i]; + if (pc >= di->start && pc < di->end) + return di; + } + return NULL; +} + +static int cmp_ev_info(const void *a, const void *b) +{ + code_t pc_a = ((const struct ev_info*)a)->ev_pc; + code_t pc_b = ((const struct ev_info*)b)->ev_pc; + if (pc_a > pc_b) return 1; + if (pc_a < pc_b) return -1; + return 0; +} + +static struct ev_info *process_debug_events(code_t code_start, value events_heap, + mlsize_t *num_events) +{ + CAMLparam1(events_heap); + CAMLlocal3(l, ev, ev_start); + mlsize_t i, j; + struct ev_info *events; + + /* Compute the size of the required event buffer. */ + *num_events = 0; + for (i = 0; i < caml_array_length(events_heap); i++) + for (l = Field(events_heap, i); l != Val_int(0); l = Field(l, 1)) + (*num_events)++; + + if (*num_events == 0) + CAMLreturnT(struct ev_info *, NULL); + + events = caml_stat_alloc_noexc(*num_events * sizeof(struct ev_info)); + if(events == NULL) + caml_fatal_error ("caml_add_debug_info: out of memory"); + + j = 0; + for (i = 0; i < caml_array_length(events_heap); i++) { + for (l = Field(events_heap, i); l != Val_int(0); l = Field(l, 1)) { + ev = Field(l, 0); + + events[j].ev_pc = (code_t)((char*)code_start + + Long_val(Field(ev, EV_POS))); + + ev_start = Field(Field(ev, EV_LOC), LOC_START); + + { + uintnat fnsz = caml_string_length(Field(ev_start, POS_FNAME)) + 1; + events[j].ev_filename = (char*)caml_stat_alloc_noexc(fnsz); + if(events[j].ev_filename == NULL) + caml_fatal_error ("caml_add_debug_info: out of memory"); + memcpy(events[j].ev_filename, + String_val(Field(ev_start, POS_FNAME)), + fnsz); + } + + events[j].ev_lnum = Int_val(Field(ev_start, POS_LNUM)); + events[j].ev_startchr = + Int_val(Field(ev_start, POS_CNUM)) + - Int_val(Field(ev_start, POS_BOL)); + events[j].ev_endchr = + Int_val(Field(Field(Field(ev, EV_LOC), LOC_END), POS_CNUM)) + - Int_val(Field(ev_start, POS_BOL)); + + j++; + } + } + + CAMLassert(j == *num_events); + + qsort(events, *num_events, sizeof(struct ev_info), cmp_ev_info); + + CAMLreturnT(struct ev_info *, events); +} + +/* Processes a (Instruct.debug_event list array) into a form suitable + for quick lookup and registers it for the (code_start,code_size) pc range. */ +CAMLprim value caml_add_debug_info(code_t code_start, value code_size, + value events_heap) +{ + CAMLparam1(events_heap); + struct debug_info *debug_info; + + /* build the OCaml-side debug_info value */ + debug_info = caml_stat_alloc(sizeof(struct debug_info)); + + debug_info->start = code_start; + debug_info->end = (code_t)((char*) code_start + Long_val(code_size)); + if (events_heap == Val_unit) { + debug_info->events = NULL; + debug_info->num_events = 0; + debug_info->already_read = 0; + } else { + debug_info->events = + process_debug_events(code_start, events_heap, &debug_info->num_events); + debug_info->already_read = 1; + } + + caml_ext_table_add(&caml_debug_info, debug_info); + + CAMLreturn(Val_unit); +} + +CAMLprim value caml_remove_debug_info(code_t start) +{ + CAMLparam0(); + CAMLlocal2(dis, prev); + + int i; + for (i = 0; i < caml_debug_info.size; i++) { + struct debug_info *di = caml_debug_info.contents[i]; + if (di->start == start) { + /* note that caml_ext_table_remove calls caml_stat_free on the + removed resource, bracketing the caml_stat_alloc call in + caml_add_debug_info. */ + caml_ext_table_remove(&caml_debug_info, di); + break; + } + } + + CAMLreturn(Val_unit); +} + +int caml_alloc_backtrace_buffer(void){ + CAMLassert(caml_backtrace_pos == 0); + caml_backtrace_buffer = + caml_stat_alloc_noexc(BACKTRACE_BUFFER_SIZE * sizeof(code_t)); + if (caml_backtrace_buffer == NULL) return -1; + return 0; +} + +/* Store the return addresses contained in the given stack fragment + into the backtrace array */ + +void caml_stash_backtrace(value exn, code_t pc, value * sp, int reraise) +{ + if (pc != NULL) pc = pc - 1; + if (exn != caml_backtrace_last_exn || !reraise) { + caml_backtrace_pos = 0; + caml_backtrace_last_exn = exn; + } + + if (caml_backtrace_buffer == NULL && caml_alloc_backtrace_buffer() == -1) + return; + + if (caml_backtrace_pos >= BACKTRACE_BUFFER_SIZE) return; + /* testing the code region is needed: PR#1554 */ + if (find_debug_info(pc) != NULL) + caml_backtrace_buffer[caml_backtrace_pos++] = pc; + + /* Traverse the stack and put all values pointing into bytecode + into the backtrace buffer. */ + for (/*nothing*/; sp < caml_trapsp; sp++) { + code_t p = (code_t) *sp; + if (caml_backtrace_pos >= BACKTRACE_BUFFER_SIZE) break; + if (find_debug_info(p) != NULL) + caml_backtrace_buffer[caml_backtrace_pos++] = p; + } +} + +/* returns the next frame pointer (or NULL if none is available); + updates *sp to point to the following one, and *trsp to the next + trap frame, which we will skip when we reach it */ + +code_t caml_next_frame_pointer(value ** sp, value ** trsp) +{ + while (*sp < caml_stack_high) { + code_t *p = (code_t*) (*sp)++; + if(&Trap_pc(*trsp) == p) { + *trsp = Trap_link(*trsp); + continue; + } + + if (find_debug_info(*p) != NULL) + return *p; + } + return NULL; +} + +/* Stores upto [max_frames_value] frames of the current call stack to + return to the user. This is used not in an exception-raising + context, but only when the user requests to save the trace + (hopefully less often). Instead of using a bounded buffer as + [caml_stash_backtrace], we first traverse the stack to compute the + right size, then allocate space for the trace. */ + +CAMLprim value caml_get_current_callstack(value max_frames_value) +{ + CAMLparam1(max_frames_value); + CAMLlocal1(trace); + + /* we use `intnat` here because, were it only `int`, passing `max_int` + from the OCaml side would overflow on 64bits machines. */ + intnat max_frames = Long_val(max_frames_value); + intnat trace_size; + + /* first compute the size of the trace */ + { + value * sp = caml_extern_sp; + value * trsp = caml_trapsp; + + for (trace_size = 0; trace_size < max_frames; trace_size++) { + code_t p = caml_next_frame_pointer(&sp, &trsp); + if (p == NULL) break; + } + } + + trace = caml_alloc(trace_size, 0); + + /* then collect the trace */ + { + value * sp = caml_extern_sp; + value * trsp = caml_trapsp; + uintnat trace_pos; + + for (trace_pos = 0; trace_pos < trace_size; trace_pos++) { + code_t p = caml_next_frame_pointer(&sp, &trsp); + CAMLassert(p != NULL); + Field(trace, trace_pos) = Val_backtrace_slot(p); + } + } + + CAMLreturn(trace); +} + +/* Read the debugging info contained in the current bytecode executable. */ + +#ifndef O_BINARY +#define O_BINARY 0 +#endif + +static void read_main_debug_info(struct debug_info *di) +{ + CAMLparam0(); + CAMLlocal3(events, evl, l); + char_os *exec_name; + int fd, num_events, orig, i; + struct channel *chan; + struct exec_trailer trail; + + CAMLassert(di->already_read == 0); + di->already_read = 1; + + if (caml_cds_file != NULL) { + exec_name = caml_cds_file; + } else { + exec_name = caml_exe_name; + } + + fd = caml_attempt_open(&exec_name, &trail, 1); + if (fd < 0){ + caml_fatal_error ("executable program file not found"); + CAMLreturn0; + } + + caml_read_section_descriptors(fd, &trail); + if (caml_seek_optional_section(fd, &trail, "DBUG") != -1) { + chan = caml_open_descriptor_in(fd); + + num_events = caml_getword(chan); + events = caml_alloc(num_events, 0); + + for (i = 0; i < num_events; i++) { + orig = caml_getword(chan); + evl = caml_input_val(chan); + caml_input_val(chan); /* Skip the list of absolute directory names */ + /* Relocate events in event list */ + for (l = evl; l != Val_int(0); l = Field(l, 1)) { + value ev = Field(l, 0); + Field(ev, EV_POS) = Val_long(Long_val(Field(ev, EV_POS)) + orig); + } + /* Record event list */ + Store_field(events, i, evl); + } + + caml_close_channel(chan); + + di->events = process_debug_events(caml_start_code, events, &di->num_events); + } + + CAMLreturn0; +} + +CAMLexport void caml_init_debug_info(void) +{ + caml_ext_table_init(&caml_debug_info, 1); + caml_add_debug_info(caml_start_code, Val_long(caml_code_size), Val_unit); +} + +int caml_debug_info_available(void) +{ + return (caml_debug_info.size != 0); +} + +/* Search the event index for the given PC. Return -1 if not found. */ + +static struct ev_info *event_for_location(code_t pc) +{ + uintnat low, high; + struct debug_info *di = find_debug_info(pc); + + if (di == NULL) + return NULL; + + if (!di->already_read) + read_main_debug_info(di); + + if (di->num_events == 0) + return NULL; + + low = 0; + high = di->num_events; + while (low+1 < high) { + uintnat m = (low+high)/2; + if(pc < di->events[m].ev_pc) high = m; + else low = m; + } + if (di->events[low].ev_pc == pc) + return &di->events[low]; + /* ocamlc sometimes moves an event past a following PUSH instruction; + allow mismatch by 1 instruction. */ + if (di->events[low].ev_pc == pc + 1) + return &di->events[low]; + if (low+1 < di->num_events && di->events[low+1].ev_pc == pc + 1) + return &di->events[low+1]; + + return NULL; +} + +/* Extract location information for the given PC */ + +void caml_debuginfo_location(debuginfo dbg, + /*out*/ struct caml_loc_info * li) +{ + code_t pc = dbg; + struct ev_info *event = event_for_location(pc); + li->loc_is_raise = + caml_is_instruction(*pc, RAISE) || + caml_is_instruction(*pc, RERAISE); + if (event == NULL) { + li->loc_valid = 0; + return; + } + li->loc_valid = 1; + li->loc_is_inlined = 0; + li->loc_filename = event->ev_filename; + li->loc_lnum = event->ev_lnum; + li->loc_startchr = event->ev_startchr; + li->loc_endchr = event->ev_endchr; +} + +debuginfo caml_debuginfo_extract(backtrace_slot slot) +{ + return (debuginfo)slot; +} + +debuginfo caml_debuginfo_next(debuginfo dbg) +{ + /* No inlining in bytecode */ + return NULL; +} diff --git a/test/monniaux/ocaml/byterun/bigarray.c b/test/monniaux/ocaml/byterun/bigarray.c new file mode 100644 index 00000000..3e376799 --- /dev/null +++ b/test/monniaux/ocaml/byterun/bigarray.c @@ -0,0 +1,1223 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Manuel Serrano and Xavier Leroy, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include +#include +#include +#include "caml/alloc.h" +#include "caml/bigarray.h" +#include "caml/custom.h" +#include "caml/fail.h" +#include "caml/intext.h" +#include "caml/hash.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/signals.h" + +#define int8 caml_ba_int8 +#define uint8 caml_ba_uint8 +#define int16 caml_ba_int16 +#define uint16 caml_ba_uint16 + +/* Compute the number of elements of a big array */ + +CAMLexport uintnat caml_ba_num_elts(struct caml_ba_array * b) +{ + uintnat num_elts; + int i; + num_elts = 1; + for (i = 0; i < b->num_dims; i++) num_elts = num_elts * b->dim[i]; + return num_elts; +} + +/* Size in bytes of a bigarray element, indexed by bigarray kind */ + +CAMLexport int caml_ba_element_size[] = +{ 4 /*FLOAT32*/, 8 /*FLOAT64*/, + 1 /*SINT8*/, 1 /*UINT8*/, + 2 /*SINT16*/, 2 /*UINT16*/, + 4 /*INT32*/, 8 /*INT64*/, + sizeof(value) /*CAML_INT*/, sizeof(value) /*NATIVE_INT*/, + 8 /*COMPLEX32*/, 16 /*COMPLEX64*/, + 1 /*CHAR*/ +}; + +/* Compute the number of bytes for the elements of a big array */ + +CAMLexport uintnat caml_ba_byte_size(struct caml_ba_array * b) +{ + return caml_ba_num_elts(b) + * caml_ba_element_size[b->flags & CAML_BA_KIND_MASK]; +} + +/* Operation table for bigarrays */ + +CAMLexport struct custom_operations caml_ba_ops = { + "_bigarray", + caml_ba_finalize, + caml_ba_compare, + caml_ba_hash, + caml_ba_serialize, + caml_ba_deserialize, + custom_compare_ext_default +}; + +/* Allocation of a big array */ + +#define CAML_BA_MAX_MEMORY 1024*1024*1024 +/* 1 Gb -- after allocating that much, it's probably worth speeding + up the major GC */ + +/* [caml_ba_alloc] will allocate a new bigarray object in the heap. + If [data] is NULL, the memory for the contents is also allocated + (with [malloc]) by [caml_ba_alloc]. + [data] cannot point into the OCaml heap. + [dim] may point into an object in the OCaml heap. +*/ +CAMLexport value +caml_ba_alloc(int flags, int num_dims, void * data, intnat * dim) +{ + uintnat num_elts, asize, size; + int i; + value res; + struct caml_ba_array * b; + intnat dimcopy[CAML_BA_MAX_NUM_DIMS]; + + CAMLassert(num_dims >= 0 && num_dims <= CAML_BA_MAX_NUM_DIMS); + CAMLassert((flags & CAML_BA_KIND_MASK) <= CAML_BA_CHAR); + for (i = 0; i < num_dims; i++) dimcopy[i] = dim[i]; + size = 0; + if (data == NULL) { + num_elts = 1; + for (i = 0; i < num_dims; i++) { + if (caml_umul_overflow(num_elts, dimcopy[i], &num_elts)) + caml_raise_out_of_memory(); + } + if (caml_umul_overflow(num_elts, + caml_ba_element_size[flags & CAML_BA_KIND_MASK], + &size)) + caml_raise_out_of_memory(); + data = malloc(size); + if (data == NULL && size != 0) caml_raise_out_of_memory(); + flags |= CAML_BA_MANAGED; + } + asize = SIZEOF_BA_ARRAY + num_dims * sizeof(intnat); + res = caml_alloc_custom(&caml_ba_ops, asize, size, CAML_BA_MAX_MEMORY); + b = Caml_ba_array_val(res); + b->data = data; + b->num_dims = num_dims; + b->flags = flags; + b->proxy = NULL; + for (i = 0; i < num_dims; i++) b->dim[i] = dimcopy[i]; + return res; +} + +/* Same as caml_ba_alloc, but dimensions are passed as a list of + arguments */ + +CAMLexport value caml_ba_alloc_dims(int flags, int num_dims, void * data, ...) +{ + va_list ap; + intnat dim[CAML_BA_MAX_NUM_DIMS]; + int i; + value res; + + CAMLassert(num_dims <= CAML_BA_MAX_NUM_DIMS); + va_start(ap, data); + for (i = 0; i < num_dims; i++) dim[i] = va_arg(ap, intnat); + va_end(ap); + res = caml_ba_alloc(flags, num_dims, data, dim); + return res; +} + +/* Finalization of a big array */ + +CAMLexport void caml_ba_finalize(value v) +{ + struct caml_ba_array * b = Caml_ba_array_val(v); + + switch (b->flags & CAML_BA_MANAGED_MASK) { + case CAML_BA_EXTERNAL: + break; + case CAML_BA_MANAGED: + if (b->proxy == NULL) { + free(b->data); + } else { + if (-- b->proxy->refcount == 0) { + free(b->proxy->data); + free(b->proxy); + } + } + break; + case CAML_BA_MAPPED_FILE: + /* Bigarrays for mapped files use a different finalization method */ + default: + CAMLassert(0); + } +} + +/* Comparison of two big arrays */ + +CAMLexport int caml_ba_compare(value v1, value v2) +{ + struct caml_ba_array * b1 = Caml_ba_array_val(v1); + struct caml_ba_array * b2 = Caml_ba_array_val(v2); + uintnat n, num_elts; + intnat flags1, flags2; + int i; + + /* Compare kind & layout in case the arguments are of different types */ + flags1 = b1->flags & (CAML_BA_KIND_MASK | CAML_BA_LAYOUT_MASK); + flags2 = b2->flags & (CAML_BA_KIND_MASK | CAML_BA_LAYOUT_MASK); + if (flags1 != flags2) return flags2 - flags1; + /* Compare number of dimensions */ + if (b1->num_dims != b2->num_dims) return b2->num_dims - b1->num_dims; + /* Same number of dimensions: compare dimensions lexicographically */ + for (i = 0; i < b1->num_dims; i++) { + intnat d1 = b1->dim[i]; + intnat d2 = b2->dim[i]; + if (d1 != d2) return d1 < d2 ? -1 : 1; + } + /* Same dimensions: compare contents lexicographically */ + num_elts = caml_ba_num_elts(b1); + +#define DO_INTEGER_COMPARISON(type) \ + { type * p1 = b1->data; type * p2 = b2->data; \ + for (n = 0; n < num_elts; n++) { \ + type e1 = *p1++; type e2 = *p2++; \ + if (e1 < e2) return -1; \ + if (e1 > e2) return 1; \ + } \ + return 0; \ + } +#define DO_FLOAT_COMPARISON(type) \ + { type * p1 = b1->data; type * p2 = b2->data; \ + for (n = 0; n < num_elts; n++) { \ + type e1 = *p1++; type e2 = *p2++; \ + if (e1 < e2) return -1; \ + if (e1 > e2) return 1; \ + if (e1 != e2) { \ + caml_compare_unordered = 1; \ + if (e1 == e1) return 1; \ + if (e2 == e2) return -1; \ + } \ + } \ + return 0; \ + } + + switch (b1->flags & CAML_BA_KIND_MASK) { + case CAML_BA_COMPLEX32: + num_elts *= 2; /*fallthrough*/ + case CAML_BA_FLOAT32: + DO_FLOAT_COMPARISON(float); + case CAML_BA_COMPLEX64: + num_elts *= 2; /*fallthrough*/ + case CAML_BA_FLOAT64: + DO_FLOAT_COMPARISON(double); + case CAML_BA_CHAR: + DO_INTEGER_COMPARISON(caml_ba_uint8); + case CAML_BA_SINT8: + DO_INTEGER_COMPARISON(caml_ba_int8); + case CAML_BA_UINT8: + DO_INTEGER_COMPARISON(caml_ba_uint8); + case CAML_BA_SINT16: + DO_INTEGER_COMPARISON(caml_ba_int16); + case CAML_BA_UINT16: + DO_INTEGER_COMPARISON(caml_ba_uint16); + case CAML_BA_INT32: + DO_INTEGER_COMPARISON(int32_t); + case CAML_BA_INT64: + DO_INTEGER_COMPARISON(int64_t); + case CAML_BA_CAML_INT: + case CAML_BA_NATIVE_INT: + DO_INTEGER_COMPARISON(intnat); + default: + CAMLassert(0); + return 0; /* should not happen */ + } +#undef DO_INTEGER_COMPARISON +#undef DO_FLOAT_COMPARISON +} + +/* Hashing of a bigarray */ + +CAMLexport intnat caml_ba_hash(value v) +{ + struct caml_ba_array * b = Caml_ba_array_val(v); + intnat num_elts, n; + uint32_t h, w; + int i; + + num_elts = 1; + for (i = 0; i < b->num_dims; i++) num_elts = num_elts * b->dim[i]; + h = 0; + + switch (b->flags & CAML_BA_KIND_MASK) { + case CAML_BA_CHAR: + case CAML_BA_SINT8: + case CAML_BA_UINT8: { + caml_ba_uint8 * p = b->data; + if (num_elts > 256) num_elts = 256; + for (n = 0; n + 4 <= num_elts; n += 4, p += 4) { + w = p[0] | (p[1] << 8) | (p[2] << 16) | (p[3] << 24); + h = caml_hash_mix_uint32(h, w); + } + w = 0; + switch (num_elts & 3) { + case 3: w = p[2] << 16; /* fallthrough */ + case 2: w |= p[1] << 8; /* fallthrough */ + case 1: w |= p[0]; + h = caml_hash_mix_uint32(h, w); + } + break; + } + case CAML_BA_SINT16: + case CAML_BA_UINT16: { + caml_ba_uint16 * p = b->data; + if (num_elts > 128) num_elts = 128; + for (n = 0; n + 2 <= num_elts; n += 2, p += 2) { + w = p[0] | (p[1] << 16); + h = caml_hash_mix_uint32(h, w); + } + if ((num_elts & 1) != 0) + h = caml_hash_mix_uint32(h, p[0]); + break; + } + case CAML_BA_INT32: + { + uint32_t * p = b->data; + if (num_elts > 64) num_elts = 64; + for (n = 0; n < num_elts; n++, p++) h = caml_hash_mix_uint32(h, *p); + break; + } + case CAML_BA_CAML_INT: + case CAML_BA_NATIVE_INT: + { + intnat * p = b->data; + if (num_elts > 64) num_elts = 64; + for (n = 0; n < num_elts; n++, p++) h = caml_hash_mix_intnat(h, *p); + break; + } + case CAML_BA_INT64: + { + int64_t * p = b->data; + if (num_elts > 32) num_elts = 32; + for (n = 0; n < num_elts; n++, p++) h = caml_hash_mix_int64(h, *p); + break; + } + case CAML_BA_COMPLEX32: + num_elts *= 2; /* fallthrough */ + case CAML_BA_FLOAT32: + { + float * p = b->data; + if (num_elts > 64) num_elts = 64; + for (n = 0; n < num_elts; n++, p++) h = caml_hash_mix_float(h, *p); + break; + } + case CAML_BA_COMPLEX64: + num_elts *= 2; /* fallthrough */ + case CAML_BA_FLOAT64: + { + double * p = b->data; + if (num_elts > 32) num_elts = 32; + for (n = 0; n < num_elts; n++, p++) h = caml_hash_mix_double(h, *p); + break; + } + } + return h; +} + +static void caml_ba_serialize_longarray(void * data, + intnat num_elts, + intnat min_val, intnat max_val) +{ +#ifdef ARCH_SIXTYFOUR + int overflow_32 = 0; + intnat * p, n; + for (n = 0, p = data; n < num_elts; n++, p++) { + if (*p < min_val || *p > max_val) { overflow_32 = 1; break; } + } + if (overflow_32) { + caml_serialize_int_1(1); + caml_serialize_block_8(data, num_elts); + } else { + caml_serialize_int_1(0); + for (n = 0, p = data; n < num_elts; n++, p++) + caml_serialize_int_4((int32_t) *p); + } +#else + caml_serialize_int_1(0); + caml_serialize_block_4(data, num_elts); +#endif +} + +CAMLexport void caml_ba_serialize(value v, + uintnat * wsize_32, + uintnat * wsize_64) +{ + struct caml_ba_array * b = Caml_ba_array_val(v); + intnat num_elts; + int i; + + /* Serialize header information */ + caml_serialize_int_4(b->num_dims); + caml_serialize_int_4(b->flags & (CAML_BA_KIND_MASK | CAML_BA_LAYOUT_MASK)); + /* On a 64-bit machine, if any of the dimensions is >= 2^32, + the size of the marshaled data will be >= 2^32 and + extern_value() will fail. So, it is safe to write the dimensions + as 32-bit unsigned integers. */ + for (i = 0; i < b->num_dims; i++) caml_serialize_int_4(b->dim[i]); + /* Compute total number of elements */ + num_elts = 1; + for (i = 0; i < b->num_dims; i++) num_elts = num_elts * b->dim[i]; + /* Serialize elements */ + switch (b->flags & CAML_BA_KIND_MASK) { + case CAML_BA_CHAR: + case CAML_BA_SINT8: + case CAML_BA_UINT8: + caml_serialize_block_1(b->data, num_elts); break; + case CAML_BA_SINT16: + case CAML_BA_UINT16: + caml_serialize_block_2(b->data, num_elts); break; + case CAML_BA_FLOAT32: + case CAML_BA_INT32: + caml_serialize_block_4(b->data, num_elts); break; + case CAML_BA_COMPLEX32: + caml_serialize_block_4(b->data, num_elts * 2); break; + case CAML_BA_FLOAT64: + case CAML_BA_INT64: + caml_serialize_block_8(b->data, num_elts); break; + case CAML_BA_COMPLEX64: + caml_serialize_block_8(b->data, num_elts * 2); break; + case CAML_BA_CAML_INT: + caml_ba_serialize_longarray(b->data, num_elts, -0x40000000, 0x3FFFFFFF); + break; + case CAML_BA_NATIVE_INT: + caml_ba_serialize_longarray(b->data, num_elts, -0x80000000, 0x7FFFFFFF); + break; + } + /* Compute required size in OCaml heap. Assumes struct caml_ba_array + is exactly 4 + num_dims words */ + CAMLassert(SIZEOF_BA_ARRAY == 4 * sizeof(value)); + *wsize_32 = (4 + b->num_dims) * 4; + *wsize_64 = (4 + b->num_dims) * 8; +} + +static void caml_ba_deserialize_longarray(void * dest, intnat num_elts) +{ + int sixty = caml_deserialize_uint_1(); +#ifdef ARCH_SIXTYFOUR + if (sixty) { + caml_deserialize_block_8(dest, num_elts); + } else { + intnat * p, n; + for (n = 0, p = dest; n < num_elts; n++, p++) + *p = caml_deserialize_sint_4(); + } +#else + if (sixty) + caml_deserialize_error("input_value: cannot read bigarray " + "with 64-bit OCaml ints"); + caml_deserialize_block_4(dest, num_elts); +#endif +} + +CAMLexport uintnat caml_ba_deserialize(void * dst) +{ + struct caml_ba_array * b = dst; + int i; + uintnat num_elts, size; + + /* Read back header information */ + b->num_dims = caml_deserialize_uint_4(); + if (b->num_dims < 0 || b->num_dims > CAML_BA_MAX_NUM_DIMS) + caml_deserialize_error("input_value: wrong number of bigarray dimensions"); + b->flags = caml_deserialize_uint_4() | CAML_BA_MANAGED; + b->proxy = NULL; + for (i = 0; i < b->num_dims; i++) b->dim[i] = caml_deserialize_uint_4(); + /* Compute total number of elements. Watch out for overflows (MPR#7765). */ + num_elts = 1; + for (i = 0; i < b->num_dims; i++) { + if (caml_umul_overflow(num_elts, b->dim[i], &num_elts)) + caml_deserialize_error("input_value: size overflow for bigarray"); + } + /* Determine array size in bytes. Watch out for overflows (MPR#7765). */ + if ((b->flags & CAML_BA_KIND_MASK) > CAML_BA_CHAR) + caml_deserialize_error("input_value: bad bigarray kind"); + if (caml_umul_overflow(num_elts, + caml_ba_element_size[b->flags & CAML_BA_KIND_MASK], + &size)) + caml_deserialize_error("input_value: size overflow for bigarray"); + /* Allocate room for data */ + b->data = malloc(size); + if (b->data == NULL) + caml_deserialize_error("input_value: out of memory for bigarray"); + /* Read data */ + switch (b->flags & CAML_BA_KIND_MASK) { + case CAML_BA_CHAR: + case CAML_BA_SINT8: + case CAML_BA_UINT8: + caml_deserialize_block_1(b->data, num_elts); break; + case CAML_BA_SINT16: + case CAML_BA_UINT16: + caml_deserialize_block_2(b->data, num_elts); break; + case CAML_BA_FLOAT32: + case CAML_BA_INT32: + caml_deserialize_block_4(b->data, num_elts); break; + case CAML_BA_COMPLEX32: + caml_deserialize_block_4(b->data, num_elts * 2); break; + case CAML_BA_FLOAT64: + case CAML_BA_INT64: + caml_deserialize_block_8(b->data, num_elts); break; + case CAML_BA_COMPLEX64: + caml_deserialize_block_8(b->data, num_elts * 2); break; + case CAML_BA_CAML_INT: + case CAML_BA_NATIVE_INT: + caml_ba_deserialize_longarray(b->data, num_elts); break; + } + /* PR#5516: use C99's flexible array types if possible */ + return SIZEOF_BA_ARRAY + b->num_dims * sizeof(intnat); +} + +/* Allocate a bigarray from OCaml */ + +CAMLprim value caml_ba_create(value vkind, value vlayout, value vdim) +{ + intnat dim[CAML_BA_MAX_NUM_DIMS]; + mlsize_t num_dims; + int i, flags; + + num_dims = Wosize_val(vdim); + /* here num_dims is unsigned (mlsize_t) so no need to check (num_dims >= 0) */ + if (num_dims > CAML_BA_MAX_NUM_DIMS) + caml_invalid_argument("Bigarray.create: bad number of dimensions"); + for (i = 0; i < num_dims; i++) { + dim[i] = Long_val(Field(vdim, i)); + if (dim[i] < 0) + caml_invalid_argument("Bigarray.create: negative dimension"); + } + flags = Caml_ba_kind_val(vkind) | Caml_ba_layout_val(vlayout); + return caml_ba_alloc(flags, num_dims, NULL, dim); +} + +/* Given a big array and a vector of indices, check that the indices + are within the bounds and return the offset of the corresponding + array element in the data part of the array. */ + +static long caml_ba_offset(struct caml_ba_array * b, intnat * index) +{ + intnat offset; + int i; + + offset = 0; + if ((b->flags & CAML_BA_LAYOUT_MASK) == CAML_BA_C_LAYOUT) { + /* C-style layout: row major, indices start at 0 */ + for (i = 0; i < b->num_dims; i++) { + if ((uintnat) index[i] >= (uintnat) b->dim[i]) + caml_array_bound_error(); + offset = offset * b->dim[i] + index[i]; + } + } else { + /* Fortran-style layout: column major, indices start at 1 */ + for (i = b->num_dims - 1; i >= 0; i--) { + if ((uintnat) (index[i] - 1) >= (uintnat) b->dim[i]) + caml_array_bound_error(); + offset = offset * b->dim[i] + (index[i] - 1); + } + } + return offset; +} + +/* Helper function to allocate a record of two double floats */ + +static value copy_two_doubles(double d0, double d1) +{ + value res = caml_alloc_small(2 * Double_wosize, Double_array_tag); + Store_double_field(res, 0, d0); + Store_double_field(res, 1, d1); + return res; +} + +/* Generic code to read from a big array */ + +value caml_ba_get_N(value vb, value * vind, int nind) +{ + struct caml_ba_array * b = Caml_ba_array_val(vb); + intnat index[CAML_BA_MAX_NUM_DIMS]; + int i; + intnat offset; + + /* Check number of indices = number of dimensions of array + (maybe not necessary if ML typing guarantees this) */ + if (nind != b->num_dims) + caml_invalid_argument("Bigarray.get: wrong number of indices"); + /* Compute offset and check bounds */ + for (i = 0; i < b->num_dims; i++) index[i] = Long_val(vind[i]); + offset = caml_ba_offset(b, index); + /* Perform read */ + switch ((b->flags) & CAML_BA_KIND_MASK) { + default: + CAMLassert(0); + case CAML_BA_FLOAT32: + return caml_copy_double(((float *) b->data)[offset]); + case CAML_BA_FLOAT64: + return caml_copy_double(((double *) b->data)[offset]); + case CAML_BA_SINT8: + return Val_int(((int8 *) b->data)[offset]); + case CAML_BA_UINT8: + return Val_int(((uint8 *) b->data)[offset]); + case CAML_BA_SINT16: + return Val_int(((int16 *) b->data)[offset]); + case CAML_BA_UINT16: + return Val_int(((uint16 *) b->data)[offset]); + case CAML_BA_INT32: + return caml_copy_int32(((int32_t *) b->data)[offset]); + case CAML_BA_INT64: + return caml_copy_int64(((int64_t *) b->data)[offset]); + case CAML_BA_NATIVE_INT: + return caml_copy_nativeint(((intnat *) b->data)[offset]); + case CAML_BA_CAML_INT: + return Val_long(((intnat *) b->data)[offset]); + case CAML_BA_COMPLEX32: + { float * p = ((float *) b->data) + offset * 2; + return copy_two_doubles(p[0], p[1]); } + case CAML_BA_COMPLEX64: + { double * p = ((double *) b->data) + offset * 2; + return copy_two_doubles(p[0], p[1]); } + case CAML_BA_CHAR: + return Val_int(((unsigned char *) b->data)[offset]); + } +} + +CAMLprim value caml_ba_get_1(value vb, value vind1) +{ + return caml_ba_get_N(vb, &vind1, 1); +} + +CAMLprim value caml_ba_get_2(value vb, value vind1, value vind2) +{ + value vind[2]; + vind[0] = vind1; vind[1] = vind2; + return caml_ba_get_N(vb, vind, 2); +} + +CAMLprim value caml_ba_get_3(value vb, value vind1, value vind2, value vind3) +{ + value vind[3]; + vind[0] = vind1; vind[1] = vind2; vind[2] = vind3; + return caml_ba_get_N(vb, vind, 3); +} + +CAMLprim value caml_ba_get_generic(value vb, value vind) +{ + return caml_ba_get_N(vb, &Field(vind, 0), Wosize_val(vind)); +} + + +CAMLprim value caml_ba_uint8_get16(value vb, value vind) +{ + intnat res; + unsigned char b1, b2; + intnat idx = Long_val(vind); + struct caml_ba_array * b = Caml_ba_array_val(vb); + if (idx < 0 || idx >= b->dim[0] - 1) caml_array_bound_error(); + b1 = ((unsigned char*) b->data)[idx]; + b2 = ((unsigned char*) b->data)[idx+1]; +#ifdef ARCH_BIG_ENDIAN + res = b1 << 8 | b2; +#else + res = b2 << 8 | b1; +#endif + return Val_int(res); +} + +CAMLprim value caml_ba_uint8_get32(value vb, value vind) +{ + intnat res; + unsigned char b1, b2, b3, b4; + intnat idx = Long_val(vind); + struct caml_ba_array * b = Caml_ba_array_val(vb); + if (idx < 0 || idx >= b->dim[0] - 3) caml_array_bound_error(); + b1 = ((unsigned char*) b->data)[idx]; + b2 = ((unsigned char*) b->data)[idx+1]; + b3 = ((unsigned char*) b->data)[idx+2]; + b4 = ((unsigned char*) b->data)[idx+3]; +#ifdef ARCH_BIG_ENDIAN + res = b1 << 24 | b2 << 16 | b3 << 8 | b4; +#else + res = b4 << 24 | b3 << 16 | b2 << 8 | b1; +#endif + return caml_copy_int32(res); +} + +CAMLprim value caml_ba_uint8_get64(value vb, value vind) +{ + uint64_t res; + unsigned char b1, b2, b3, b4, b5, b6, b7, b8; + intnat idx = Long_val(vind); + struct caml_ba_array * b = Caml_ba_array_val(vb); + if (idx < 0 || idx >= b->dim[0] - 7) caml_array_bound_error(); + b1 = ((unsigned char*) b->data)[idx]; + b2 = ((unsigned char*) b->data)[idx+1]; + b3 = ((unsigned char*) b->data)[idx+2]; + b4 = ((unsigned char*) b->data)[idx+3]; + b5 = ((unsigned char*) b->data)[idx+4]; + b6 = ((unsigned char*) b->data)[idx+5]; + b7 = ((unsigned char*) b->data)[idx+6]; + b8 = ((unsigned char*) b->data)[idx+7]; +#ifdef ARCH_BIG_ENDIAN + res = (uint64_t) b1 << 56 | (uint64_t) b2 << 48 + | (uint64_t) b3 << 40 | (uint64_t) b4 << 32 + | (uint64_t) b5 << 24 | (uint64_t) b6 << 16 + | (uint64_t) b7 << 8 | (uint64_t) b8; +#else + res = (uint64_t) b8 << 56 | (uint64_t) b7 << 48 + | (uint64_t) b6 << 40 | (uint64_t) b5 << 32 + | (uint64_t) b4 << 24 | (uint64_t) b3 << 16 + | (uint64_t) b2 << 8 | (uint64_t) b1; +#endif + return caml_copy_int64(res); +} + +/* Generic write to a big array */ + +static value caml_ba_set_aux(value vb, value * vind, intnat nind, value newval) +{ + struct caml_ba_array * b = Caml_ba_array_val(vb); + intnat index[CAML_BA_MAX_NUM_DIMS]; + int i; + intnat offset; + + /* Check number of indices = number of dimensions of array + (maybe not necessary if ML typing guarantees this) */ + if (nind != b->num_dims) + caml_invalid_argument("Bigarray.set: wrong number of indices"); + /* Compute offset and check bounds */ + for (i = 0; i < b->num_dims; i++) index[i] = Long_val(vind[i]); + offset = caml_ba_offset(b, index); + /* Perform write */ + switch (b->flags & CAML_BA_KIND_MASK) { + default: + CAMLassert(0); + case CAML_BA_FLOAT32: + ((float *) b->data)[offset] = Double_val(newval); break; + case CAML_BA_FLOAT64: + ((double *) b->data)[offset] = Double_val(newval); break; + case CAML_BA_CHAR: + case CAML_BA_SINT8: + case CAML_BA_UINT8: + ((int8 *) b->data)[offset] = Int_val(newval); break; + case CAML_BA_SINT16: + case CAML_BA_UINT16: + ((int16 *) b->data)[offset] = Int_val(newval); break; + case CAML_BA_INT32: + ((int32_t *) b->data)[offset] = Int32_val(newval); break; + case CAML_BA_INT64: + ((int64_t *) b->data)[offset] = Int64_val(newval); break; + case CAML_BA_NATIVE_INT: + ((intnat *) b->data)[offset] = Nativeint_val(newval); break; + case CAML_BA_CAML_INT: + ((intnat *) b->data)[offset] = Long_val(newval); break; + case CAML_BA_COMPLEX32: + { float * p = ((float *) b->data) + offset * 2; + p[0] = Double_field(newval, 0); + p[1] = Double_field(newval, 1); + break; } + case CAML_BA_COMPLEX64: + { double * p = ((double *) b->data) + offset * 2; + p[0] = Double_field(newval, 0); + p[1] = Double_field(newval, 1); + break; } + } + return Val_unit; +} + +CAMLprim value caml_ba_set_1(value vb, value vind1, value newval) +{ + return caml_ba_set_aux(vb, &vind1, 1, newval); +} + +CAMLprim value caml_ba_set_2(value vb, value vind1, value vind2, value newval) +{ + value vind[2]; + vind[0] = vind1; vind[1] = vind2; + return caml_ba_set_aux(vb, vind, 2, newval); +} + +CAMLprim value caml_ba_set_3(value vb, value vind1, value vind2, value vind3, + value newval) +{ + value vind[3]; + vind[0] = vind1; vind[1] = vind2; vind[2] = vind3; + return caml_ba_set_aux(vb, vind, 3, newval); +} + +value caml_ba_set_N(value vb, value * vind, int nargs) +{ + return caml_ba_set_aux(vb, vind, nargs - 1, vind[nargs - 1]); +} + +CAMLprim value caml_ba_set_generic(value vb, value vind, value newval) +{ + return caml_ba_set_aux(vb, &Field(vind, 0), Wosize_val(vind), newval); +} + +CAMLprim value caml_ba_uint8_set16(value vb, value vind, value newval) +{ + unsigned char b1, b2; + intnat val; + intnat idx = Long_val(vind); + struct caml_ba_array * b = Caml_ba_array_val(vb); + if (idx < 0 || idx >= b->dim[0] - 1) caml_array_bound_error(); + val = Long_val(newval); +#ifdef ARCH_BIG_ENDIAN + b1 = 0xFF & val >> 8; + b2 = 0xFF & val; +#else + b2 = 0xFF & val >> 8; + b1 = 0xFF & val; +#endif + ((unsigned char*) b->data)[idx] = b1; + ((unsigned char*) b->data)[idx+1] = b2; + return Val_unit; +} + +CAMLprim value caml_ba_uint8_set32(value vb, value vind, value newval) +{ + unsigned char b1, b2, b3, b4; + intnat idx = Long_val(vind); + intnat val; + struct caml_ba_array * b = Caml_ba_array_val(vb); + if (idx < 0 || idx >= b->dim[0] - 3) caml_array_bound_error(); + val = Int32_val(newval); +#ifdef ARCH_BIG_ENDIAN + b1 = 0xFF & val >> 24; + b2 = 0xFF & val >> 16; + b3 = 0xFF & val >> 8; + b4 = 0xFF & val; +#else + b4 = 0xFF & val >> 24; + b3 = 0xFF & val >> 16; + b2 = 0xFF & val >> 8; + b1 = 0xFF & val; +#endif + ((unsigned char*) b->data)[idx] = b1; + ((unsigned char*) b->data)[idx+1] = b2; + ((unsigned char*) b->data)[idx+2] = b3; + ((unsigned char*) b->data)[idx+3] = b4; + return Val_unit; +} + +CAMLprim value caml_ba_uint8_set64(value vb, value vind, value newval) +{ + unsigned char b1, b2, b3, b4, b5, b6, b7, b8; + intnat idx = Long_val(vind); + int64_t val; + struct caml_ba_array * b = Caml_ba_array_val(vb); + if (idx < 0 || idx >= b->dim[0] - 7) caml_array_bound_error(); + val = Int64_val(newval); +#ifdef ARCH_BIG_ENDIAN + b1 = 0xFF & val >> 56; + b2 = 0xFF & val >> 48; + b3 = 0xFF & val >> 40; + b4 = 0xFF & val >> 32; + b5 = 0xFF & val >> 24; + b6 = 0xFF & val >> 16; + b7 = 0xFF & val >> 8; + b8 = 0xFF & val; +#else + b8 = 0xFF & val >> 56; + b7 = 0xFF & val >> 48; + b6 = 0xFF & val >> 40; + b5 = 0xFF & val >> 32; + b4 = 0xFF & val >> 24; + b3 = 0xFF & val >> 16; + b2 = 0xFF & val >> 8; + b1 = 0xFF & val; +#endif + ((unsigned char*) b->data)[idx] = b1; + ((unsigned char*) b->data)[idx+1] = b2; + ((unsigned char*) b->data)[idx+2] = b3; + ((unsigned char*) b->data)[idx+3] = b4; + ((unsigned char*) b->data)[idx+4] = b5; + ((unsigned char*) b->data)[idx+5] = b6; + ((unsigned char*) b->data)[idx+6] = b7; + ((unsigned char*) b->data)[idx+7] = b8; + return Val_unit; +} + +/* Return the number of dimensions of a big array */ + +CAMLprim value caml_ba_num_dims(value vb) +{ + struct caml_ba_array * b = Caml_ba_array_val(vb); + return Val_long(b->num_dims); +} + +/* Return the n-th dimension of a big array */ + +CAMLprim value caml_ba_dim(value vb, value vn) +{ + struct caml_ba_array * b = Caml_ba_array_val(vb); + intnat n = Long_val(vn); + if (n < 0 || n >= b->num_dims) caml_invalid_argument("Bigarray.dim"); + return Val_long(b->dim[n]); +} + +CAMLprim value caml_ba_dim_1(value vb) +{ + return caml_ba_dim(vb, Val_int(0)); +} + +CAMLprim value caml_ba_dim_2(value vb) +{ + return caml_ba_dim(vb, Val_int(1)); +} + +CAMLprim value caml_ba_dim_3(value vb) +{ + return caml_ba_dim(vb, Val_int(2)); +} + +/* Return the kind of a big array */ + +CAMLprim value caml_ba_kind(value vb) +{ + return Val_caml_ba_kind(Caml_ba_array_val(vb)->flags & CAML_BA_KIND_MASK); +} + +/* Return the layout of a big array */ + +CAMLprim value caml_ba_layout(value vb) +{ + int layout = Caml_ba_array_val(vb)->flags & CAML_BA_LAYOUT_MASK; + return Val_caml_ba_layout(layout); +} + +/* Create / update proxy to indicate that b2 is a sub-array of b1 */ + +static void caml_ba_update_proxy(struct caml_ba_array * b1, + struct caml_ba_array * b2) +{ + struct caml_ba_proxy * proxy; + /* Nothing to do for un-managed arrays */ + if ((b1->flags & CAML_BA_MANAGED_MASK) == CAML_BA_EXTERNAL) return; + if (b1->proxy != NULL) { + /* If b1 is already a proxy for a larger array, increment refcount of + proxy */ + b2->proxy = b1->proxy; + ++ b1->proxy->refcount; + } else { + /* Otherwise, create proxy and attach it to both b1 and b2 */ + proxy = malloc(sizeof(struct caml_ba_proxy)); + if (proxy == NULL) caml_raise_out_of_memory(); + proxy->refcount = 2; /* original array + sub array */ + proxy->data = b1->data; + proxy->size = + b1->flags & CAML_BA_MAPPED_FILE ? caml_ba_byte_size(b1) : 0; + b1->proxy = proxy; + b2->proxy = proxy; + } +} + +/* Slicing */ + +CAMLprim value caml_ba_slice(value vb, value vind) +{ + CAMLparam2 (vb, vind); + #define b ((struct caml_ba_array *) Caml_ba_array_val(vb)) + CAMLlocal1 (res); + intnat index[CAML_BA_MAX_NUM_DIMS]; + int num_inds, i; + intnat offset; + intnat * sub_dims; + char * sub_data; + + /* Check number of indices <= number of dimensions of array */ + num_inds = Wosize_val(vind); + if (num_inds > b->num_dims) + caml_invalid_argument("Bigarray.slice: too many indices"); + /* Compute offset and check bounds */ + if ((b->flags & CAML_BA_LAYOUT_MASK) == CAML_BA_C_LAYOUT) { + /* We slice from the left */ + for (i = 0; i < num_inds; i++) index[i] = Long_val(Field(vind, i)); + for (/*nothing*/; i < b->num_dims; i++) index[i] = 0; + offset = caml_ba_offset(b, index); + sub_dims = b->dim + num_inds; + } else { + /* We slice from the right */ + for (i = 0; i < num_inds; i++) + index[b->num_dims - num_inds + i] = Long_val(Field(vind, i)); + for (i = 0; i < b->num_dims - num_inds; i++) index[i] = 1; + offset = caml_ba_offset(b, index); + sub_dims = b->dim; + } + sub_data = + (char *) b->data + + offset * caml_ba_element_size[b->flags & CAML_BA_KIND_MASK]; + /* Allocate an OCaml bigarray to hold the result */ + res = caml_ba_alloc(b->flags, b->num_dims - num_inds, sub_data, sub_dims); + /* Create or update proxy in case of managed bigarray */ + caml_ba_update_proxy(b, Caml_ba_array_val(res)); + /* Return result */ + CAMLreturn (res); + + #undef b +} + +/* Changing the layout of an array (memory is shared) */ + +CAMLprim value caml_ba_change_layout(value vb, value vlayout) +{ + CAMLparam2 (vb, vlayout); + CAMLlocal1 (res); + #define b ((struct caml_ba_array *) Caml_ba_array_val(vb)) + /* if the layout is different, change the flags and reverse the dimensions */ + if (Caml_ba_layout_val(vlayout) != (b->flags & CAML_BA_LAYOUT_MASK)) { + /* change the flags to reflect the new layout */ + int flags = (b->flags & (CAML_BA_KIND_MASK | CAML_BA_MANAGED_MASK)) + | Caml_ba_layout_val(vlayout); + /* reverse the dimensions */ + intnat new_dim[CAML_BA_MAX_NUM_DIMS]; + unsigned int i; + for(i = 0; i < b->num_dims; i++) new_dim[i] = b->dim[b->num_dims - i - 1]; + res = caml_ba_alloc(flags, b->num_dims, b->data, new_dim); + caml_ba_update_proxy(b, Caml_ba_array_val(res)); + CAMLreturn(res); + } else { + /* otherwise, do nothing */ + CAMLreturn(vb); + } + #undef b +} + + +/* Extracting a sub-array of same number of dimensions */ + +CAMLprim value caml_ba_sub(value vb, value vofs, value vlen) +{ + CAMLparam3 (vb, vofs, vlen); + CAMLlocal1 (res); + #define b ((struct caml_ba_array *) Caml_ba_array_val(vb)) + intnat ofs = Long_val(vofs); + intnat len = Long_val(vlen); + int i, changed_dim; + intnat mul; + char * sub_data; + + /* Compute offset and check bounds */ + if ((b->flags & CAML_BA_LAYOUT_MASK) == CAML_BA_C_LAYOUT) { + /* We reduce the first dimension */ + mul = 1; + for (i = 1; i < b->num_dims; i++) mul *= b->dim[i]; + changed_dim = 0; + } else { + /* We reduce the last dimension */ + mul = 1; + for (i = 0; i < b->num_dims - 1; i++) mul *= b->dim[i]; + changed_dim = b->num_dims - 1; + ofs--; /* Fortran arrays start at 1 */ + } + if (ofs < 0 || len < 0 || ofs + len > b->dim[changed_dim]) + caml_invalid_argument("Bigarray.sub: bad sub-array"); + sub_data = + (char *) b->data + + ofs * mul * caml_ba_element_size[b->flags & CAML_BA_KIND_MASK]; + /* Allocate an OCaml bigarray to hold the result */ + res = caml_ba_alloc(b->flags, b->num_dims, sub_data, b->dim); + /* Doctor the changed dimension */ + Caml_ba_array_val(res)->dim[changed_dim] = len; + /* Create or update proxy in case of managed bigarray */ + caml_ba_update_proxy(b, Caml_ba_array_val(res)); + /* Return result */ + CAMLreturn (res); + + #undef b +} + +/* Copying a big array into another one */ + +#define LEAVE_RUNTIME_OP_CUTOFF 4096 +#define is_mmapped(ba) ((ba)->flags & CAML_BA_MAPPED_FILE) + +CAMLprim value caml_ba_blit(value vsrc, value vdst) +{ + CAMLparam2(vsrc, vdst); + struct caml_ba_array * src = Caml_ba_array_val(vsrc); + struct caml_ba_array * dst = Caml_ba_array_val(vdst); + void *src_data = src->data; + void *dst_data = dst->data; + int i; + intnat num_bytes; + int leave_runtime; + + /* Check same numbers of dimensions and same dimensions */ + if (src->num_dims != dst->num_dims) goto blit_error; + for (i = 0; i < src->num_dims; i++) + if (src->dim[i] != dst->dim[i]) goto blit_error; + /* Compute number of bytes in array data */ + num_bytes = + caml_ba_num_elts(src) + * caml_ba_element_size[src->flags & CAML_BA_KIND_MASK]; + leave_runtime = + ( + (num_bytes >= LEAVE_RUNTIME_OP_CUTOFF*sizeof(long)) + || is_mmapped(src) + || is_mmapped(dst) + ); + /* Do the copying */ + if (leave_runtime) caml_enter_blocking_section(); + memmove (dst_data, src_data, num_bytes); + if (leave_runtime) caml_leave_blocking_section(); + CAMLreturn (Val_unit); + blit_error: + caml_invalid_argument("Bigarray.blit: dimension mismatch"); + CAMLreturn (Val_unit); /* not reached */ +} + +/* Filling a big array with a given value */ + +#define FILL_GEN_LOOP(n_ops, loop) do{ \ + int leave_runtime = ((n_ops >= LEAVE_RUNTIME_OP_CUTOFF) || is_mmapped(b)); \ + if (leave_runtime) caml_enter_blocking_section(); \ + loop; \ + if (leave_runtime) caml_leave_blocking_section(); \ +}while(0) + +#define FILL_SCALAR_LOOP \ + FILL_GEN_LOOP(num_elts, \ + for (p = data; num_elts > 0; p++, num_elts--) *p = init) + +#define FILL_COMPLEX_LOOP \ + FILL_GEN_LOOP(num_elts + num_elts, \ + for (p = data; num_elts > 0; num_elts--) { *p++ = init0; *p++ = init1; }) + +CAMLprim value caml_ba_fill(value vb, value vinit) +{ + CAMLparam1(vb); + struct caml_ba_array * b = Caml_ba_array_val(vb); + void *data = b->data; + intnat num_elts = caml_ba_num_elts(b); + + switch (b->flags & CAML_BA_KIND_MASK) { + default: + CAMLassert(0); + case CAML_BA_FLOAT32: { + float init = Double_val(vinit); + float * p; + FILL_SCALAR_LOOP; + break; + } + case CAML_BA_FLOAT64: { + double init = Double_val(vinit); + double * p; + FILL_SCALAR_LOOP; + break; + } + case CAML_BA_CHAR: + case CAML_BA_SINT8: + case CAML_BA_UINT8: { + int init = Int_val(vinit); + unsigned char * p; + FILL_SCALAR_LOOP; + break; + } + case CAML_BA_SINT16: + case CAML_BA_UINT16: { + int init = Int_val(vinit); + int16 * p; + FILL_SCALAR_LOOP; + break; + } + case CAML_BA_INT32: { + int32_t init = Int32_val(vinit); + int32_t * p; + FILL_SCALAR_LOOP; + break; + } + case CAML_BA_INT64: { + int64_t init = Int64_val(vinit); + int64_t * p; + FILL_SCALAR_LOOP; + break; + } + case CAML_BA_NATIVE_INT: { + intnat init = Nativeint_val(vinit); + intnat * p; + FILL_SCALAR_LOOP; + break; + } + case CAML_BA_CAML_INT: { + intnat init = Long_val(vinit); + intnat * p; + FILL_SCALAR_LOOP; + break; + } + case CAML_BA_COMPLEX32: { + float init0 = Double_field(vinit, 0); + float init1 = Double_field(vinit, 1); + float * p; + FILL_COMPLEX_LOOP; + break; + } + case CAML_BA_COMPLEX64: { + double init0 = Double_field(vinit, 0); + double init1 = Double_field(vinit, 1); + double * p; + FILL_COMPLEX_LOOP; + break; + } + } + CAMLreturn (Val_unit); +} + +/* Reshape an array: change dimensions and number of dimensions, preserving + array contents */ + +CAMLprim value caml_ba_reshape(value vb, value vdim) +{ + CAMLparam2 (vb, vdim); + CAMLlocal1 (res); +#define b ((struct caml_ba_array *) Caml_ba_array_val(vb)) + intnat dim[CAML_BA_MAX_NUM_DIMS]; + mlsize_t num_dims; + uintnat num_elts; + int i; + + num_dims = Wosize_val(vdim); + /* here num_dims is unsigned (mlsize_t) so no need to check (num_dims >= 0) */ + if (num_dims > CAML_BA_MAX_NUM_DIMS) + caml_invalid_argument("Bigarray.reshape: bad number of dimensions"); + num_elts = 1; + for (i = 0; i < num_dims; i++) { + dim[i] = Long_val(Field(vdim, i)); + if (dim[i] < 0) + caml_invalid_argument("Bigarray.reshape: negative dimension"); + num_elts *= dim[i]; + } + /* Check that sizes agree */ + if (num_elts != caml_ba_num_elts(b)) + caml_invalid_argument("Bigarray.reshape: size mismatch"); + /* Create bigarray with same data and new dimensions */ + res = caml_ba_alloc(b->flags, num_dims, b->data, dim); + /* Create or update proxy in case of managed bigarray */ + caml_ba_update_proxy(b, Caml_ba_array_val(res)); + /* Return result */ + CAMLreturn (res); + +#undef b +} diff --git a/test/monniaux/ocaml/byterun/callback.c b/test/monniaux/ocaml/byterun/callback.c new file mode 100644 index 00000000..9c479b30 --- /dev/null +++ b/test/monniaux/ocaml/byterun/callback.c @@ -0,0 +1,262 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Callbacks from C to OCaml */ + +#include +#include "caml/callback.h" +#include "caml/fail.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" + +#ifndef NATIVE_CODE + +/* Bytecode callbacks */ + +#include "caml/interp.h" +#include "caml/instruct.h" +#include "caml/fix_code.h" +#include "caml/stacks.h" + +CAMLexport int caml_callback_depth = 0; + +#ifndef LOCAL_CALLBACK_BYTECODE +static opcode_t callback_code[] = { ACC, 0, APPLY, 0, POP, 1, STOP }; +#endif + + +#ifdef THREADED_CODE + +static int callback_code_threaded = 0; + +static void thread_callback(void) +{ + caml_thread_code(callback_code, sizeof(callback_code)); + callback_code_threaded = 1; +} + +#define Init_callback() if (!callback_code_threaded) thread_callback() + +#else + +#define Init_callback() + +#endif + +CAMLexport value caml_callbackN_exn(value closure, int narg, value args[]) +{ + int i; + value res; + + /* some alternate bytecode implementations (e.g. a JIT translator) + might require that the bytecode is kept in a local variable on + the C stack */ +#ifdef LOCAL_CALLBACK_BYTECODE + opcode_t local_callback_code[7]; +#endif + + CAMLassert(narg + 4 <= 256); + + caml_extern_sp -= narg + 4; + for (i = 0; i < narg; i++) caml_extern_sp[i] = args[i]; /* arguments */ +#ifndef LOCAL_CALLBACK_BYTECODE + caml_extern_sp[narg] = (value) (callback_code + 4); /* return address */ + caml_extern_sp[narg + 1] = Val_unit; /* environment */ + caml_extern_sp[narg + 2] = Val_long(0); /* extra args */ + caml_extern_sp[narg + 3] = closure; + Init_callback(); + callback_code[1] = narg + 3; + callback_code[3] = narg; + res = caml_interprete(callback_code, sizeof(callback_code)); +#else /*have LOCAL_CALLBACK_BYTECODE*/ + caml_extern_sp[narg] = (value) (local_callback_code + 4); /* return address */ + caml_extern_sp[narg + 1] = Val_unit; /* environment */ + caml_extern_sp[narg + 2] = Val_long(0); /* extra args */ + caml_extern_sp[narg + 3] = closure; + local_callback_code[0] = ACC; + local_callback_code[1] = narg + 3; + local_callback_code[2] = APPLY; + local_callback_code[3] = narg; + local_callback_code[4] = POP; + local_callback_code[5] = 1; + local_callback_code[6] = STOP; +#ifdef THREADED_CODE + caml_thread_code(local_callback_code, sizeof(local_callback_code)); +#endif /*THREADED_CODE*/ + res = caml_interprete(local_callback_code, sizeof(local_callback_code)); + caml_release_bytecode(local_callback_code, sizeof(local_callback_code)); +#endif /*LOCAL_CALLBACK_BYTECODE*/ + if (Is_exception_result(res)) caml_extern_sp += narg + 4; /* PR#1228 */ + return res; +} + +CAMLexport value caml_callback_exn(value closure, value arg1) +{ + value arg[1]; + arg[0] = arg1; + return caml_callbackN_exn(closure, 1, arg); +} + +CAMLexport value caml_callback2_exn(value closure, value arg1, value arg2) +{ + value arg[2]; + arg[0] = arg1; + arg[1] = arg2; + return caml_callbackN_exn(closure, 2, arg); +} + +CAMLexport value caml_callback3_exn(value closure, + value arg1, value arg2, value arg3) +{ + value arg[3]; + arg[0] = arg1; + arg[1] = arg2; + arg[2] = arg3; + return caml_callbackN_exn(closure, 3, arg); +} + +#else + +/* Native-code callbacks. caml_callback[123]_exn are implemented in asm. */ + +CAMLexport value caml_callbackN_exn(value closure, int narg, value args[]) +{ + CAMLparam1 (closure); + CAMLxparamN (args, narg); + CAMLlocal1 (res); + int i; + + res = closure; + for (i = 0; i < narg; /*nothing*/) { + /* Pass as many arguments as possible */ + switch (narg - i) { + case 1: + res = caml_callback_exn(res, args[i]); + if (Is_exception_result(res)) CAMLreturn (res); + i += 1; + break; + case 2: + res = caml_callback2_exn(res, args[i], args[i + 1]); + if (Is_exception_result(res)) CAMLreturn (res); + i += 2; + break; + default: + res = caml_callback3_exn(res, args[i], args[i + 1], args[i + 2]); + if (Is_exception_result(res)) CAMLreturn (res); + i += 3; + break; + } + } + CAMLreturn (res); +} + +#endif + +/* Exception-propagating variants of the above */ + +CAMLexport value caml_callback (value closure, value arg) +{ + value res = caml_callback_exn(closure, arg); + if (Is_exception_result(res)) caml_raise(Extract_exception(res)); + return res; +} + +CAMLexport value caml_callback2 (value closure, value arg1, value arg2) +{ + value res = caml_callback2_exn(closure, arg1, arg2); + if (Is_exception_result(res)) caml_raise(Extract_exception(res)); + return res; +} + +CAMLexport value caml_callback3 (value closure, value arg1, value arg2, + value arg3) +{ + value res = caml_callback3_exn(closure, arg1, arg2, arg3); + if (Is_exception_result(res)) caml_raise(Extract_exception(res)); + return res; +} + +CAMLexport value caml_callbackN (value closure, int narg, value args[]) +{ + value res = caml_callbackN_exn(closure, narg, args); + if (Is_exception_result(res)) caml_raise(Extract_exception(res)); + return res; +} + +/* Naming of OCaml values */ + +struct named_value { + value val; + struct named_value * next; + char name[1]; +}; + +#define Named_value_size 13 + +static struct named_value * named_value_table[Named_value_size] = { NULL, }; + +static unsigned int hash_value_name(char const *name) +{ + unsigned int h; + for (h = 0; *name != 0; name++) h = h * 19 + *name; + return h % Named_value_size; +} + +CAMLprim value caml_register_named_value(value vname, value val) +{ + struct named_value * nv; + const char * name = String_val(vname); + size_t namelen = strlen(name); + unsigned int h = hash_value_name(name); + + for (nv = named_value_table[h]; nv != NULL; nv = nv->next) { + if (strcmp(name, nv->name) == 0) { + nv->val = val; + return Val_unit; + } + } + nv = (struct named_value *) + caml_stat_alloc(sizeof(struct named_value) + namelen); + memcpy(nv->name, name, namelen + 1); + nv->val = val; + nv->next = named_value_table[h]; + named_value_table[h] = nv; + caml_register_global_root(&nv->val); + return Val_unit; +} + +CAMLexport value * caml_named_value(char const *name) +{ + struct named_value * nv; + for (nv = named_value_table[hash_value_name(name)]; + nv != NULL; + nv = nv->next) { + if (strcmp(name, nv->name) == 0) return &nv->val; + } + return NULL; +} + +CAMLexport void caml_iterate_named_values(caml_named_action f) +{ + int i; + for(i = 0; i < Named_value_size; i++){ + struct named_value * nv; + for (nv = named_value_table[i]; nv != NULL; nv = nv->next) { + f( &nv->val, nv->name ); + } + } +} diff --git a/test/monniaux/ocaml/byterun/caml/address_class.h b/test/monniaux/ocaml/byterun/caml/address_class.h new file mode 100644 index 00000000..85e22d32 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/address_class.h @@ -0,0 +1,85 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Classification of addresses for GC and runtime purposes. */ + +#ifndef CAML_ADDRESS_CLASS_H +#define CAML_ADDRESS_CLASS_H + +#include "config.h" +#include "misc.h" +#include "mlvalues.h" + +/* Use the following macros to test an address for the different classes + it might belong to. */ + +#define Is_young(val) \ + (CAMLassert (Is_block (val)), \ + (addr)(val) < (addr)caml_young_end && (addr)(val) > (addr)caml_young_start) + +#define Is_in_heap(a) (Classify_addr(a) & In_heap) + +#define Is_in_heap_or_young(a) (Classify_addr(a) & (In_heap | In_young)) + +#define Is_in_value_area(a) \ + (Classify_addr(a) & (In_heap | In_young | In_static_data)) + +#define Is_in_code_area(pc) \ + ( ((char *)(pc) >= caml_code_area_start && \ + (char *)(pc) <= caml_code_area_end) \ + || (Classify_addr(pc) & In_code_area) ) + +#define Is_in_static_data(a) (Classify_addr(a) & In_static_data) + +/***********************************************************************/ +/* The rest of this file is private and may change without notice. */ + +extern value *caml_young_start, *caml_young_end; +extern char * caml_code_area_start, * caml_code_area_end; + +#define Not_in_heap 0 +#define In_heap 1 +#define In_young 2 +#define In_static_data 4 +#define In_code_area 8 + +#ifdef ARCH_SIXTYFOUR + +/* 64 bits: Represent page table as a sparse hash table */ +int caml_page_table_lookup(void * addr); +#define Classify_addr(a) (caml_page_table_lookup((void *)(a))) + +#else + +/* 32 bits: Represent page table as a 2-level array */ +#define Pagetable2_log 11 +#define Pagetable2_size (1 << Pagetable2_log) +#define Pagetable1_log (Page_log + Pagetable2_log) +#define Pagetable1_size (1 << (32 - Pagetable1_log)) +CAMLextern unsigned char * caml_page_table[Pagetable1_size]; + +#define Pagetable_index1(a) (((uintnat)(a)) >> Pagetable1_log) +#define Pagetable_index2(a) \ + ((((uintnat)(a)) >> Page_log) & (Pagetable2_size - 1)) +#define Classify_addr(a) \ + caml_page_table[Pagetable_index1(a)][Pagetable_index2(a)] + +#endif + +int caml_page_table_add(int kind, void * start, void * end); +int caml_page_table_remove(int kind, void * start, void * end); +int caml_page_table_initialize(mlsize_t bytesize); + +#endif /* CAML_ADDRESS_CLASS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/alloc.h b/test/monniaux/ocaml/byterun/caml/alloc.h new file mode 100644 index 00000000..81fff858 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/alloc.h @@ -0,0 +1,82 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_ALLOC_H +#define CAML_ALLOC_H + + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "misc.h" +#include "mlvalues.h" + +#ifdef __cplusplus +extern "C" { +#endif + +CAMLextern value caml_alloc (mlsize_t wosize, tag_t); +CAMLextern value caml_alloc_small (mlsize_t wosize, tag_t); +CAMLextern value caml_alloc_tuple (mlsize_t wosize); +CAMLextern value caml_alloc_float_array (mlsize_t len); +CAMLextern value caml_alloc_string (mlsize_t len); /* len in bytes (chars) */ +CAMLextern value caml_alloc_initialized_string (mlsize_t len, const char *); +CAMLextern value caml_copy_string (char const *); +CAMLextern value caml_copy_string_array (char const **); +CAMLextern value caml_copy_double (double); +CAMLextern value caml_copy_int32 (int32_t); /* defined in [ints.c] */ +CAMLextern value caml_copy_int64 (int64_t); /* defined in [ints.c] */ +CAMLextern value caml_copy_nativeint (intnat); /* defined in [ints.c] */ +CAMLextern value caml_alloc_array (value (*funct) (char const *), + char const ** array); +CAMLextern value caml_alloc_sprintf(const char * format, ...) +#ifdef __GNUC__ + __attribute__ ((format (printf, 1, 2))) +#endif +; + +CAMLextern value caml_alloc_with_profinfo (mlsize_t, tag_t, intnat); +CAMLextern value caml_alloc_small_with_my_or_given_profinfo ( + mlsize_t, tag_t, uintnat); +CAMLextern value caml_alloc_small_with_profinfo (mlsize_t, tag_t, intnat); + +typedef void (*final_fun)(value); +CAMLextern value caml_alloc_final (mlsize_t wosize, + final_fun, /*finalization function*/ + mlsize_t, /*resources consumed*/ + mlsize_t /*max resources*/); + +CAMLextern int caml_convert_flag_list (value, int *); + +/* Convenience functions to deal with unboxable types. */ +static inline value caml_alloc_unboxed (value arg) { return arg; } +static inline value caml_alloc_boxed (value arg) { + value result = caml_alloc_small (1, 0); + Field (result, 0) = arg; + return result; +} +static inline value caml_field_unboxed (value arg) { return arg; } +static inline value caml_field_boxed (value arg) { return Field (arg, 0); } + +/* Unannotated unboxable types are boxed by default. (may change in the + future) */ +#define caml_alloc_unboxable caml_alloc_boxed +#define caml_field_unboxable caml_field_boxed + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_ALLOC_H */ diff --git a/test/monniaux/ocaml/byterun/caml/backtrace.h b/test/monniaux/ocaml/byterun/caml/backtrace.h new file mode 100644 index 00000000..fc0baf2d --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/backtrace.h @@ -0,0 +1,136 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2001 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_BACKTRACE_H +#define CAML_BACKTRACE_H + +#ifdef CAML_INTERNALS + +#include "mlvalues.h" +#include "exec.h" + +/* Runtime support for backtrace generation. + * + * It has two kind of users: + * - high-level API to capture and decode backtraces; + * - low-level runtime routines, to introspect machine state and determine + * whether a backtrace should be generated when using "raise". + * + * Backtrace generation is split in multiple steps. + * The lowest-level one, done by [backtrace_prim.c] just fills the + * [caml_backtrace_buffer] variable each time a frame is unwinded. + * At that point, we don't know whether the backtrace will be useful or not so + * this code should be as fast as possible. + * + * If the backtrace happens to be useful, later passes will read + * [caml_backtrace_buffer] and turn it into a [raw_backtrace] and then a + * [backtrace]. + * This is done in [backtrace.c] and [stdlib/printexc.ml]. + * + * Content of buffers + * ------------------ + * + * [caml_backtrace_buffer] (really cheap) + * Backend and process image dependent, abstracted by C-type backtrace_slot. + * [raw_backtrace] (cheap) + * OCaml values of abstract type [Printexc.raw_backtrace_slot], + * still backend and process image dependent (unsafe to marshal). + * [backtrace] (more expensive) + * OCaml values of algebraic data-type [Printexc.backtrace_slot] + */ + +/* Non zero iff backtraces are recorded. + * One should use to change this variable [caml_record_backtrace]. + */ +CAMLextern int caml_backtrace_active; + +/* The [backtrace_slot] type represents values stored in the + * [caml_backtrace_buffer]. In bytecode, it is the same as a + * [code_t], in native code it as a [frame_descr *]. The difference + * doesn't matter for code outside [backtrace_prim.c], so it is just + * exposed as a [backtrace_slot]. + */ +typedef void * backtrace_slot; + +/* The [caml_backtrace_buffer] and [caml_backtrace_last_exn] + * variables are valid only if [caml_backtrace_active != 0]. + * + * They are part of the state specific to each thread, and threading libraries + * are responsible for copying them on context switch. + * See [otherlibs/systhreads/st_stubs.c] and [otherlibs/threads/scheduler.c]. + */ + +/* [caml_backtrace_buffer] is filled by runtime when unwinding stack. + * It is an array ranging from [0] to [caml_backtrace_pos - 1]. + * [caml_backtrace_pos] is always zero if [!caml_backtrace_active]. + * + * Its maximum size is determined by [BACKTRACE_BUFFER_SIZE] from + * [backtrace_prim.h], but this shouldn't affect users. + */ +CAMLextern backtrace_slot * caml_backtrace_buffer; +CAMLextern int caml_backtrace_pos; + +/* [caml_backtrace_last_exn] stores the last exception value that was raised, + * iff [caml_backtrace_active != 0]. + * It is tested for equality to determine whether a raise is a re-raise of the + * same exception. + * + * FIXME: this shouldn't matter anymore. Since OCaml 4.02, non-parameterized + * exceptions are constant, so physical equality is no longer appropriate. + * raise and re-raise are distinguished by: + * - passing reraise = 1 to [caml_stash_backtrace] (see below) in the bytecode + * interpreter; + * - directly resetting [caml_backtrace_pos] to 0 in native runtimes for raise. + */ +CAMLextern value caml_backtrace_last_exn; + +/* [caml_record_backtrace] toggle backtrace recording on and off. + * This function can be called at runtime by user-code, or during + * initialization if backtraces were requested. + * + * It might be called before GC initialization, so it shouldn't do OCaml + * allocation. + */ +CAMLprim value caml_record_backtrace(value vflag); + + +#ifndef NATIVE_CODE + +/* Path to the file containing debug information, if any, or NULL. */ +CAMLextern char_os * caml_cds_file; + +/* Primitive called _only_ by runtime to record unwinded frames to + * backtrace. A similar primitive exists for native code, but with a + * different prototype. */ +extern void caml_stash_backtrace(value exn, code_t pc, value * sp, int reraise); + +#endif + + +/* Default (C-level) printer for backtraces. It is called if an + * exception causes a termination of the program or of a thread. + * + * [Printexc] provide a higher-level printer mimicking its output but making + * use of registered exception printers, and is used when possible in place of + * this function after [Printexc] initialization. + */ +CAMLextern void caml_print_exception_backtrace(void); + +void caml_init_backtrace(void); +CAMLexport void caml_init_debug_info(void); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_BACKTRACE_H */ diff --git a/test/monniaux/ocaml/byterun/caml/backtrace_prim.h b/test/monniaux/ocaml/byterun/caml/backtrace_prim.h new file mode 100644 index 00000000..2484b294 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/backtrace_prim.h @@ -0,0 +1,91 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2001 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_BACKTRACE_PRIM_H +#define CAML_BACKTRACE_PRIM_H + +#ifdef CAML_INTERNALS + +#include "backtrace.h" + +/* Backtrace generation is split in [backtrace.c] and [backtrace_prim.c]. + * + * [backtrace_prim.c] contains all backend-specific code, and has two different + * implementations in [byterun/backtrace_prim.c] and [asmrun/backtrace_prim.c]. + * + * [backtrace.c] has a unique implementation, and expose a uniform + * higher level API above [backtrace_prim.c]. + */ + +/* Extract location information for the given raw_backtrace_slot */ + +struct caml_loc_info { + int loc_valid; + int loc_is_raise; + char * loc_filename; + int loc_lnum; + int loc_startchr; + int loc_endchr; + int loc_is_inlined; +}; + +/* When compiling with -g, backtrace slots have debug info associated. + * When a call is inlined in native mode, debuginfos form a linked list. + */ +typedef void * debuginfo; + +/* Check availability of debug information before extracting a trace. + * Relevant for bytecode, always true for native code. */ +int caml_debug_info_available(void); + +/* Return debuginfo associated to a slot or NULL. */ +debuginfo caml_debuginfo_extract(backtrace_slot slot); + +/* In case of an inlined call return next debuginfo or NULL otherwise. */ +debuginfo caml_debuginfo_next(debuginfo dbg); + +/* Extract locations from backtrace_slot */ +void caml_debuginfo_location(debuginfo dbg, /*out*/ struct caml_loc_info * li); + +/* In order to prevent the GC from walking through the debug + information (which have no headers), we transform slots to 31/63 bits + ocaml integers by shifting them by 1 to the right. We do not lose + information as slots are aligned. + + In particular, we do not need to use [caml_modify] when setting + an array element with such a value. + */ +#define Val_backtrace_slot(bslot) (Val_long(((uintnat)(bslot))>>1)) +#define Backtrace_slot_val(vslot) ((backtrace_slot)(Long_val(vslot) << 1)) + +/* Allocate the caml_backtrace_buffer. Returns 0 on success, -1 otherwise */ +int caml_alloc_backtrace_buffer(void); + +#define BACKTRACE_BUFFER_SIZE 1024 + +/* Besides decoding backtrace info, [backtrace_prim] has two other + * responsibilities: + * + * It defines the [caml_stash_backtrace] function, which is called to quickly + * fill the backtrace buffer by walking the stack when an exception is raised. + * + * It also defines the [caml_get_current_callstack] OCaml primitive, which also + * walks the stack but directly turns it into a [raw_backtrace] and is called + * explicitly. + */ + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_BACKTRACE_PRIM_H */ diff --git a/test/monniaux/ocaml/byterun/caml/bigarray.h b/test/monniaux/ocaml/byterun/caml/bigarray.h new file mode 100644 index 00000000..fc1fb145 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/bigarray.h @@ -0,0 +1,134 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Manuel Serrano and Xavier Leroy, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_BIGARRAY_H +#define CAML_BIGARRAY_H + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "config.h" +#include "mlvalues.h" + +typedef signed char caml_ba_int8; +typedef unsigned char caml_ba_uint8; +#if defined(HAS_STDINT_H) +typedef int16_t caml_ba_int16; +typedef uint16_t caml_ba_uint16; +#elif SIZEOF_SHORT == 2 +typedef short caml_ba_int16; +typedef unsigned short caml_ba_uint16; +#else +#error "No 16-bit integer type available" +#endif + +#define CAML_BA_MAX_NUM_DIMS 16 + +enum caml_ba_kind { + CAML_BA_FLOAT32, /* Single-precision floats */ + CAML_BA_FLOAT64, /* Double-precision floats */ + CAML_BA_SINT8, /* Signed 8-bit integers */ + CAML_BA_UINT8, /* Unsigned 8-bit integers */ + CAML_BA_SINT16, /* Signed 16-bit integers */ + CAML_BA_UINT16, /* Unsigned 16-bit integers */ + CAML_BA_INT32, /* Signed 32-bit integers */ + CAML_BA_INT64, /* Signed 64-bit integers */ + CAML_BA_CAML_INT, /* OCaml-style integers (signed 31 or 63 bits) */ + CAML_BA_NATIVE_INT, /* Platform-native long integers (32 or 64 bits) */ + CAML_BA_COMPLEX32, /* Single-precision complex */ + CAML_BA_COMPLEX64, /* Double-precision complex */ + CAML_BA_CHAR, /* Characters */ + CAML_BA_KIND_MASK = 0xFF /* Mask for kind in flags field */ +}; + +#define Caml_ba_kind_val(v) Int_val(v) + +#define Val_caml_ba_kind(k) Val_int(k) + +enum caml_ba_layout { + CAML_BA_C_LAYOUT = 0, /* Row major, indices start at 0 */ + CAML_BA_FORTRAN_LAYOUT = 0x100, /* Column major, indices start at 1 */ + CAML_BA_LAYOUT_MASK = 0x100, /* Mask for layout in flags field */ + CAML_BA_LAYOUT_SHIFT = 8 /* Bit offset of layout flag */ +}; + +#define Caml_ba_layout_val(v) (Int_val(v) << CAML_BA_LAYOUT_SHIFT) + +#define Val_caml_ba_layout(l) Val_int(l >> CAML_BA_LAYOUT_SHIFT) + +enum caml_ba_managed { + CAML_BA_EXTERNAL = 0, /* Data is not allocated by OCaml */ + CAML_BA_MANAGED = 0x200, /* Data is allocated by OCaml */ + CAML_BA_MAPPED_FILE = 0x400, /* Data is a memory mapped file */ + CAML_BA_MANAGED_MASK = 0x600 /* Mask for "managed" bits in flags field */ +}; + +struct caml_ba_proxy { + intnat refcount; /* Reference count */ + void * data; /* Pointer to base of actual data */ + uintnat size; /* Size of data in bytes (if mapped file) */ +}; + +struct caml_ba_array { + void * data; /* Pointer to raw data */ + intnat num_dims; /* Number of dimensions */ + intnat flags; /* Kind of element array + memory layout + allocation status */ + struct caml_ba_proxy * proxy; /* The proxy for sub-arrays, or NULL */ + /* PR#5516: use C99's flexible array types if possible */ +#if (__STDC_VERSION__ >= 199901L) + intnat dim[] /*[num_dims]*/; /* Size in each dimension */ +#else + intnat dim[1] /*[num_dims]*/; /* Size in each dimension */ +#endif +}; + +/* Size of struct caml_ba_array, in bytes, without dummy first dimension */ +#if (__STDC_VERSION__ >= 199901L) +#define SIZEOF_BA_ARRAY sizeof(struct caml_ba_array) +#else +#define SIZEOF_BA_ARRAY (sizeof(struct caml_ba_array) - sizeof(intnat)) +#endif + +#define Caml_ba_array_val(v) ((struct caml_ba_array *) Data_custom_val(v)) + +#define Caml_ba_data_val(v) (Caml_ba_array_val(v)->data) + +#ifdef __cplusplus +extern "C" { +#endif + +CAMLextern value + caml_ba_alloc(int flags, int num_dims, void * data, intnat * dim); +CAMLextern value caml_ba_alloc_dims(int flags, int num_dims, void * data, + ... /*dimensions, with type intnat */); +CAMLextern uintnat caml_ba_byte_size(struct caml_ba_array * b); +CAMLextern uintnat caml_ba_num_elts(struct caml_ba_array * b); + +#ifdef __cplusplus +} +#endif + +#ifdef CAML_INTERNALS + +CAMLextern int caml_ba_element_size[]; +CAMLextern void caml_ba_finalize(value v); +CAMLextern int caml_ba_compare(value v1, value v2); +CAMLextern intnat caml_ba_hash(value v); +CAMLextern void caml_ba_serialize(value, uintnat *, uintnat *); +CAMLextern uintnat caml_ba_deserialize(void * dst); + +#endif + +#endif /* CAML_BIGARRAY_H */ diff --git a/test/monniaux/ocaml/byterun/caml/callback.h b/test/monniaux/ocaml/byterun/caml/callback.h new file mode 100644 index 00000000..93208b7a --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/callback.h @@ -0,0 +1,63 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Callbacks from C to OCaml */ + +#ifndef CAML_CALLBACK_H +#define CAML_CALLBACK_H + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "mlvalues.h" + +#ifdef __cplusplus +extern "C" { +#endif + +CAMLextern value caml_callback (value closure, value arg); +CAMLextern value caml_callback2 (value closure, value arg1, value arg2); +CAMLextern value caml_callback3 (value closure, value arg1, value arg2, + value arg3); +CAMLextern value caml_callbackN (value closure, int narg, value args[]); + +CAMLextern value caml_callback_exn (value closure, value arg); +CAMLextern value caml_callback2_exn (value closure, value arg1, value arg2); +CAMLextern value caml_callback3_exn (value closure, + value arg1, value arg2, value arg3); +CAMLextern value caml_callbackN_exn (value closure, int narg, value args[]); + +#define Make_exception_result(v) ((v) | 2) +#define Is_exception_result(v) (((v) & 3) == 2) +#define Extract_exception(v) ((v) & ~3) + +CAMLextern value * caml_named_value (char const * name); +typedef void (*caml_named_action) (value*, char *); +CAMLextern void caml_iterate_named_values(caml_named_action f); + +CAMLextern void caml_main (char_os ** argv); +CAMLextern void caml_startup (char_os ** argv); +CAMLextern value caml_startup_exn (char_os ** argv); +CAMLextern void caml_startup_pooled (char_os ** argv); +CAMLextern value caml_startup_pooled_exn (char_os ** argv); +CAMLextern void caml_shutdown (void); + +CAMLextern int caml_callback_depth; + +#ifdef __cplusplus +} +#endif + +#endif diff --git a/test/monniaux/ocaml/byterun/caml/compact.h b/test/monniaux/ocaml/byterun/caml/compact.h new file mode 100644 index 00000000..e29d0e86 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/compact.h @@ -0,0 +1,31 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_COMPACT_H +#define CAML_COMPACT_H + +#ifdef CAML_INTERNALS + +#include "config.h" +#include "misc.h" +#include "mlvalues.h" + +void caml_compact_heap (void); +void caml_compact_heap_maybe (void); +void caml_invert_root (value v, value *p); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_COMPACT_H */ diff --git a/test/monniaux/ocaml/byterun/caml/compare.h b/test/monniaux/ocaml/byterun/caml/compare.h new file mode 100644 index 00000000..54b71581 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/compare.h @@ -0,0 +1,25 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, Projet Moscova, INRIA Rocquencourt */ +/* */ +/* Copyright 2003 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_COMPARE_H +#define CAML_COMPARE_H + +#ifdef CAML_INTERNALS + +CAMLextern int caml_compare_unordered; + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_COMPARE_H */ diff --git a/test/monniaux/ocaml/byterun/caml/compatibility.h b/test/monniaux/ocaml/byterun/caml/compatibility.h new file mode 100644 index 00000000..082943c6 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/compatibility.h @@ -0,0 +1,375 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Moscova, INRIA Rocquencourt */ +/* */ +/* Copyright 2003 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* definitions for compatibility with old identifiers */ + +#ifndef CAML_COMPATIBILITY_H +#define CAML_COMPATIBILITY_H + +/* internal global variables renamed between 4.02.1 and 4.03.0 */ +#define caml_stat_top_heap_size Bsize_wsize(caml_stat_top_heap_wsz) +#define caml_stat_heap_size Bsize_wsize(caml_stat_heap_wsz) + +#ifndef CAML_NAME_SPACE + +/* + #define --> CAMLextern (defined with CAMLexport or CAMLprim) + (rien) --> CAMLprim + g --> global C identifier + x --> special case + + SP* signals the special cases: + - when the identifier was not simply prefixed with [caml_] + - when the [caml_] version was already used for something else, and + was renamed out of the way (watch out for [caml_alloc] and + [caml_array_bound_error] in *.s) +*/ + +/* a faire: + - ui_* (reverifier que win32.c n'en depend pas) +*/ + + +/* **** alloc.c */ +#define alloc caml_alloc /*SP*/ +#define alloc_small caml_alloc_small +#define alloc_tuple caml_alloc_tuple +#define alloc_string caml_alloc_string +#define alloc_final caml_alloc_final +#define copy_string caml_copy_string +#define alloc_array caml_alloc_array +#define copy_string_array caml_copy_string_array +#define convert_flag_list caml_convert_flag_list + +/* **** array.c */ + +/* **** backtrace.c */ +#define backtrace_active caml_backtrace_active +#define backtrace_pos caml_backtrace_pos +#define backtrace_buffer caml_backtrace_buffer +#define backtrace_last_exn caml_backtrace_last_exn +#define print_exception_backtrace caml_print_exception_backtrace + +/* **** callback.c */ +#define callback_depth caml_callback_depth +#define callbackN_exn caml_callbackN_exn +#define callback_exn caml_callback_exn +#define callback2_exn caml_callback2_exn +#define callback3_exn caml_callback3_exn +#define callback caml_callback +#define callback2 caml_callback2 +#define callback3 caml_callback3 +#define callbackN caml_callbackN + +/* **** compact.c */ + +/* **** compare.c */ +#define compare_unordered caml_compare_unordered + +/* **** custom.c */ +#define alloc_custom caml_alloc_custom +#define register_custom_operations caml_register_custom_operations + +/* **** debugger.c */ + +/* **** dynlink.c */ + +/* **** extern.c */ +#define output_val caml_output_val +#define output_value_to_malloc caml_output_value_to_malloc +#define output_value_to_block caml_output_value_to_block +#define serialize_int_1 caml_serialize_int_1 +#define serialize_int_2 caml_serialize_int_2 +#define serialize_int_4 caml_serialize_int_4 +#define serialize_int_8 caml_serialize_int_8 +#define serialize_float_4 caml_serialize_float_4 +#define serialize_float_8 caml_serialize_float_8 +#define serialize_block_1 caml_serialize_block_1 +#define serialize_block_2 caml_serialize_block_2 +#define serialize_block_4 caml_serialize_block_4 +#define serialize_block_8 caml_serialize_block_8 +#define serialize_block_float_8 caml_serialize_block_float_8 + +/* **** fail.c */ +#define external_raise caml_external_raise +#define mlraise caml_raise /*SP*/ +#define raise_constant caml_raise_constant +#define raise_with_arg caml_raise_with_arg +#define raise_with_string caml_raise_with_string +#define failwith caml_failwith +#define invalid_argument caml_invalid_argument +#define array_bound_error caml_array_bound_error /*SP*/ +#define raise_out_of_memory caml_raise_out_of_memory +#define raise_stack_overflow caml_raise_stack_overflow +#define raise_sys_error caml_raise_sys_error +#define raise_end_of_file caml_raise_end_of_file +#define raise_zero_divide caml_raise_zero_divide +#define raise_not_found caml_raise_not_found +#define raise_sys_blocked_io caml_raise_sys_blocked_io +/* **** asmrun/fail.c */ +/* **** asmrun/.s */ + +/* **** finalise.c */ + +/* **** fix_code.c */ + +/* **** floats.c */ +/*#define Double_val caml_Double_val done in mlvalues.h as needed */ +/*#define Store_double_val caml_Store_double_val done in mlvalues.h as needed */ +#define copy_double caml_copy_double + +/* **** freelist.c */ + +/* **** gc_ctrl.c */ + +/* **** globroots.c */ +#define register_global_root caml_register_global_root +#define remove_global_root caml_remove_global_root + +/* **** hash.c */ +#define hash_variant caml_hash_variant + +/* **** instrtrace.c */ + +/* **** intern.c */ +#define input_val caml_input_val +#define input_val_from_string caml_input_val_from_string +#define input_value_from_malloc caml_input_value_from_malloc +#define input_value_from_block caml_input_value_from_block +#define deserialize_uint_1 caml_deserialize_uint_1 +#define deserialize_sint_1 caml_deserialize_sint_1 +#define deserialize_uint_2 caml_deserialize_uint_2 +#define deserialize_sint_2 caml_deserialize_sint_2 +#define deserialize_uint_4 caml_deserialize_uint_4 +#define deserialize_sint_4 caml_deserialize_sint_4 +#define deserialize_uint_8 caml_deserialize_uint_8 +#define deserialize_sint_8 caml_deserialize_sint_8 +#define deserialize_float_4 caml_deserialize_float_4 +#define deserialize_float_8 caml_deserialize_float_8 +#define deserialize_block_1 caml_deserialize_block_1 +#define deserialize_block_2 caml_deserialize_block_2 +#define deserialize_block_4 caml_deserialize_block_4 +#define deserialize_block_8 caml_deserialize_block_8 +#define deserialize_block_float_8 caml_deserialize_block_float_8 +#define deserialize_error caml_deserialize_error + +/* **** interp.c */ + +/* **** ints.c */ +#define int32_ops caml_int32_ops +#define copy_int32 caml_copy_int32 +/*#define Int64_val caml_Int64_val *** done in mlvalues.h as needed */ +#define int64_ops caml_int64_ops +#define copy_int64 caml_copy_int64 +#define nativeint_ops caml_nativeint_ops +#define copy_nativeint caml_copy_nativeint + +/* **** io.c */ +#define channel_mutex_free caml_channel_mutex_free +#define channel_mutex_lock caml_channel_mutex_lock +#define channel_mutex_unlock caml_channel_mutex_unlock +#define channel_mutex_unlock_exn caml_channel_mutex_unlock_exn +#define all_opened_channels caml_all_opened_channels +#define open_descriptor_in caml_open_descriptor_in /*SP*/ +#define open_descriptor_out caml_open_descriptor_out /*SP*/ +#define close_channel caml_close_channel /*SP*/ +#define channel_size caml_channel_size /*SP*/ +#define channel_binary_mode caml_channel_binary_mode +#define flush_partial caml_flush_partial /*SP*/ +#define flush caml_flush /*SP*/ +#define putword caml_putword +#define putblock caml_putblock +#define really_putblock caml_really_putblock +#define seek_out caml_seek_out /*SP*/ +#define pos_out caml_pos_out /*SP*/ +#define do_read caml_do_read +#define refill caml_refill +#define getword caml_getword +#define getblock caml_getblock +#define really_getblock caml_really_getblock +#define seek_in caml_seek_in /*SP*/ +#define pos_in caml_pos_in /*SP*/ +#define input_scan_line caml_input_scan_line /*SP*/ +#define finalize_channel caml_finalize_channel +#define alloc_channel caml_alloc_channel +/*#define Val_file_offset caml_Val_file_offset *** done in io.h as needed */ +/*#define File_offset_val caml_File_offset_val *** done in io.h as needed */ + +/* **** lexing.c */ + +/* **** main.c */ +/* *** no change */ + +/* **** major_gc.c */ +#define heap_start caml_heap_start +#define page_table caml_page_table + +/* **** md5.c */ +#define md5_string caml_md5_string +#define md5_chan caml_md5_chan +#define MD5Init caml_MD5Init +#define MD5Update caml_MD5Update +#define MD5Final caml_MD5Final +#define MD5Transform caml_MD5Transform + +/* **** memory.c */ +#define alloc_shr caml_alloc_shr +#define initialize caml_initialize +#define modify caml_modify +#define stat_alloc caml_stat_alloc +#define stat_free caml_stat_free +#define stat_resize caml_stat_resize + +/* **** meta.c */ + +/* **** minor_gc.c */ +#define young_start caml_young_start +#define young_end caml_young_end +#define young_ptr caml_young_ptr +#define young_limit caml_young_limit +#define ref_table caml_ref_table +#define minor_collection caml_minor_collection +#define check_urgent_gc caml_check_urgent_gc + +/* **** misc.c */ + +/* **** obj.c */ + +/* **** parsing.c */ + +/* **** prims.c */ + +/* **** printexc.c */ +#define format_caml_exception caml_format_exception /*SP*/ + +/* **** roots.c */ +#define local_roots caml_local_roots +#define scan_roots_hook caml_scan_roots_hook +#define do_local_roots caml_do_local_roots + +/* **** signals.c */ +#define pending_signals caml_pending_signals +#define something_to_do caml_something_to_do +#define enter_blocking_section_hook caml_enter_blocking_section_hook +#define leave_blocking_section_hook caml_leave_blocking_section_hook +#define try_leave_blocking_section_hook caml_try_leave_blocking_section_hook +#define async_action_hook caml_async_action_hook +#define enter_blocking_section caml_enter_blocking_section +#define leave_blocking_section caml_leave_blocking_section +#define convert_signal_number caml_convert_signal_number +/* **** asmrun/signals.c */ +#define garbage_collection caml_garbage_collection + +/* **** stacks.c */ +#define stack_low caml_stack_low +#define stack_high caml_stack_high +#define stack_threshold caml_stack_threshold +#define extern_sp caml_extern_sp +#define trapsp caml_trapsp +#define trap_barrier caml_trap_barrier + +/* **** startup.c */ +#define atom_table caml_atom_table +/* **** asmrun/startup.c */ +#define static_data_start caml_static_data_start +#define static_data_end caml_static_data_end + +/* **** str.c */ +#define string_length caml_string_length + +/* **** sys.c */ +#define sys_error caml_sys_error +#define sys_exit caml_sys_exit + +/* **** terminfo.c */ + +/* **** unix.c & win32.c */ +#define search_exe_in_path caml_search_exe_in_path + +/* **** weak.c */ + +/* **** asmcomp/asmlink.ml */ + +/* **** asmcomp/cmmgen.ml */ + +/* **** asmcomp/asmlink.ml, asmcomp/cmmgen.ml, asmcomp/compilenv.ml */ + +/* ************************************************************* */ + +/* **** otherlibs/bigarray */ +#define int8 caml_ba_int8 +#define uint8 caml_ba_uint8 +#define int16 caml_ba_int16 +#define uint16 caml_ba_uint16 +#define MAX_NUM_DIMS CAML_BA_MAX_NUM_DIMS +#define caml_bigarray_kind caml_ba_kind +#define BIGARRAY_FLOAT32 CAML_BA_FLOAT32 +#define BIGARRAY_FLOAT64 CAML_BA_FLOAT64 +#define BIGARRAY_SINT8 CAML_BA_SINT8 +#define BIGARRAY_UINT8 CAML_BA_UINT8 +#define BIGARRAY_SINT16 CAML_BA_SINT16 +#define BIGARRAY_UINT16 CAML_BA_UINT16 +#define BIGARRAY_INT32 CAML_BA_INT32 +#define BIGARRAY_INT64 CAML_BA_INT64 +#define BIGARRAY_CAML_INT CAML_BA_CAML_INT +#define BIGARRAY_NATIVE_INT CAML_BA_NATIVE_INT +#define BIGARRAY_COMPLEX32 CAML_BA_COMPLEX32 +#define BIGARRAY_COMPLEX64 CAML_BA_COMPLEX64 +#define BIGARRAY_KIND_MASK CAML_BA_KIND_MASK +#define caml_bigarray_layout caml_ba_layout +#define BIGARRAY_C_LAYOUT CAML_BA_C_LAYOUT +#define BIGARRAY_FORTRAN_LAYOUT CAML_BA_FORTRAN_LAYOUT +#define BIGARRAY_LAYOUT_MASK CAML_BA_LAYOUT_MASK +#define caml_bigarray_managed caml_ba_managed +#define BIGARRAY_EXTERNAL CAML_BA_EXTERNAL +#define BIGARRAY_MANAGED CAML_BA_MANAGED +#define BIGARRAY_MAPPED_FILE CAML_BA_MAPPED_FILE +#define BIGARRAY_MANAGED_MASK CAML_BA_MANAGED_MASK +#define caml_bigarray_proxy caml_ba_proxy +#define caml_bigarray caml_ba_array +#define Bigarray_val Caml_ba_array_val +#define Data_bigarray_val Caml_ba_data_val +#define alloc_bigarray caml_ba_alloc +#define alloc_bigarray_dims caml_ba_alloc_dims +#define bigarray_map_file caml_ba_map_file +#define bigarray_unmap_file caml_ba_unmap_file +#define bigarray_element_size caml_ba_element_size +#define bigarray_byte_size caml_ba_byte_size +#define bigarray_deserialize caml_ba_deserialize +#define MAX_BIGARRAY_MEMORY CAML_BA_MAX_MEMORY +#define bigarray_create caml_ba_create +#define bigarray_get_N caml_ba_get_N +#define bigarray_get_1 caml_ba_get_1 +#define bigarray_get_2 caml_ba_get_2 +#define bigarray_get_3 caml_ba_get_3 +#define bigarray_get_generic caml_ba_get_generic +#define bigarray_set_1 caml_ba_set_1 +#define bigarray_set_2 caml_ba_set_2 +#define bigarray_set_3 caml_ba_set_3 +#define bigarray_set_N caml_ba_set_N +#define bigarray_set_generic caml_ba_set_generic +#define bigarray_num_dims caml_ba_num_dims +#define bigarray_dim caml_ba_dim +#define bigarray_kind caml_ba_kind +#define bigarray_layout caml_ba_layout +#define bigarray_slice caml_ba_slice +#define bigarray_sub caml_ba_sub +#define bigarray_blit caml_ba_blit +#define bigarray_fill caml_ba_fill +#define bigarray_reshape caml_ba_reshape +#define bigarray_init caml_ba_init + +#endif /* CAML_NAME_SPACE */ +#endif /* CAML_COMPATIBILITY_H */ diff --git a/test/monniaux/ocaml/byterun/caml/config.h b/test/monniaux/ocaml/byterun/caml/config.h new file mode 100644 index 00000000..06d8d368 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/config.h @@ -0,0 +1,208 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_CONFIG_H +#define CAML_CONFIG_H + +/* */ +/* */ +/* */ +#include "m.h" +#include "s.h" +#ifdef BOOTSTRAPPING_FLEXLINK +#undef SUPPORT_DYNAMIC_LINKING +#endif + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif + +#include + +#ifdef HAS_STDINT_H +#include +#endif + +#ifndef ARCH_SIZET_PRINTF_FORMAT +#define ARCH_SIZET_PRINTF_FORMAT "z" +#endif + +/* Types for 32-bit integers, 64-bit integers, and + native integers (as wide as a pointer type) */ + +#ifndef ARCH_INT32_TYPE +#if SIZEOF_INT == 4 +#define ARCH_INT32_TYPE int +#define ARCH_UINT32_TYPE unsigned int +#define ARCH_INT32_PRINTF_FORMAT "" +#elif SIZEOF_LONG == 4 +#define ARCH_INT32_TYPE long +#define ARCH_UINT32_TYPE unsigned long +#define ARCH_INT32_PRINTF_FORMAT "l" +#elif SIZEOF_SHORT == 4 +#define ARCH_INT32_TYPE short +#define ARCH_UINT32_TYPE unsigned short +#define ARCH_INT32_PRINTF_FORMAT "" +#else +#error "No 32-bit integer type available" +#endif +#endif + +#ifndef ARCH_INT64_TYPE +#if SIZEOF_LONG == 8 +#define ARCH_INT64_TYPE long +#define ARCH_UINT64_TYPE unsigned long +#define ARCH_INT64_PRINTF_FORMAT "l" +#elif SIZEOF_LONGLONG == 8 +#define ARCH_INT64_TYPE long long +#define ARCH_UINT64_TYPE unsigned long long +#define ARCH_INT64_PRINTF_FORMAT "ll" +#else +#error "No 64-bit integer type available" +#endif +#endif + +#ifndef HAS_STDINT_H +/* Not a C99 compiler, typically MSVC. Define the C99 types we use. */ +typedef ARCH_INT32_TYPE int32_t; +typedef ARCH_UINT32_TYPE uint32_t; +typedef ARCH_INT64_TYPE int64_t; +typedef ARCH_UINT64_TYPE uint64_t; +#if SIZEOF_SHORT == 2 +typedef short int16_t; +typedef unsigned short uint16_t; +#else +#error "No 16-bit integer type available" +#endif +#endif + +#if SIZEOF_PTR == SIZEOF_LONG +/* Standard models: ILP32 or I32LP64 */ +typedef long intnat; +typedef unsigned long uintnat; +#define ARCH_INTNAT_PRINTF_FORMAT "l" +#elif SIZEOF_PTR == SIZEOF_INT +/* Hypothetical IP32L64 model */ +typedef int intnat; +typedef unsigned int uintnat; +#define ARCH_INTNAT_PRINTF_FORMAT "" +#elif SIZEOF_PTR == 8 +/* Win64 model: IL32P64 */ +typedef int64_t intnat; +typedef uint64_t uintnat; +#define ARCH_INTNAT_PRINTF_FORMAT ARCH_INT64_PRINTF_FORMAT +#else +#error "No integer type available to represent pointers" +#endif + +/* Endianness of floats */ + +/* ARCH_FLOAT_ENDIANNESS encodes the byte order of doubles as follows: + the value [0xabcdefgh] means that the least significant byte of the + float is at byte offset [a], the next lsb at [b], ..., and the + most significant byte at [h]. */ + +#if defined(__arm__) && !defined(__ARM_EABI__) +#define ARCH_FLOAT_ENDIANNESS 0x45670123 +#elif defined(ARCH_BIG_ENDIAN) +#define ARCH_FLOAT_ENDIANNESS 0x76543210 +#else +#define ARCH_FLOAT_ENDIANNESS 0x01234567 +#endif + + +/* We use threaded code interpretation if the compiler provides labels + as first-class values (GCC 2.x). */ + +#if defined(__GNUC__) && __GNUC__ >= 2 && !defined(DEBUG) \ + && !defined (SHRINKED_GNUC) && !defined(CAML_JIT) +#define THREADED_CODE +#endif + + +/* Memory model parameters */ + +/* The size of a page for memory management (in bytes) is [1 << Page_log]. + [Page_size] must be a multiple of [sizeof (value)]. + [Page_log] must be be >= 8 and <= 20. + Do not change the definition of [Page_size]. */ +#define Page_log 12 /* A page is 4 kilobytes. */ +#define Page_size (1 << Page_log) + +/* Initial size of stack (bytes). */ +#define Stack_size (4096 * sizeof(value)) + +/* Minimum free size of stack (bytes); below that, it is reallocated. */ +#define Stack_threshold (256 * sizeof(value)) + +/* Default maximum size of the stack (words). */ +#define Max_stack_def (1024 * 1024) + + +/* Maximum size of a block allocated in the young generation (words). */ +/* Must be > 4 */ +#define Max_young_wosize 256 +#define Max_young_whsize (Whsize_wosize (Max_young_wosize)) + + +/* Minimum size of the minor zone (words). + This must be at least [2 * Max_young_whsize]. */ +#define Minor_heap_min 4096 + +/* Maximum size of the minor zone (words). + Must be greater than or equal to [Minor_heap_min]. +*/ +#define Minor_heap_max (1 << 28) + +/* Default size of the minor zone. (words) */ +#define Minor_heap_def 16384 + + +/* Minimum size increment when growing the heap (words). + Must be a multiple of [Page_size / sizeof (value)]. */ +#define Heap_chunk_min (15 * Page_size) + +/* Default size increment when growing the heap. + If this is <= 1000, it's a percentage of the current heap size. + If it is > 1000, it's a number of words. */ +#define Heap_chunk_def 15 + +/* Default initial size of the major heap (words); + Must be a multiple of [Page_size / sizeof (value)]. */ +#define Init_heap_def (4 * Page_size) +/* (about 512 kB for a 32-bit platform, 1 MB for a 64-bit platform.) */ + + +/* Default speed setting for the major GC. The heap will grow until + the dead objects and the free list represent this percentage of the + total size of live objects. */ +#define Percent_free_def 80 + +/* Default setting for the compacter: 500% + (i.e. trigger the compacter when 5/6 of the heap is free or garbage) + This can be set quite high because the overhead is over-estimated + when fragmentation occurs. + */ +#define Max_percent_free_def 500 + +/* Default setting for the major GC slice smoothing window: 1 + (i.e. no smoothing) +*/ +#define Major_window_def 1 + +/* Maximum size of the major GC slice smoothing window. */ +#define Max_major_window 50 + +#endif /* CAML_CONFIG_H */ diff --git a/test/monniaux/ocaml/byterun/caml/custom.h b/test/monniaux/ocaml/byterun/caml/custom.h new file mode 100644 index 00000000..6bc3aa95 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/custom.h @@ -0,0 +1,73 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Manuel Serrano and Xavier Leroy, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_CUSTOM_H +#define CAML_CUSTOM_H + + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "mlvalues.h" + +struct custom_operations { + char *identifier; + void (*finalize)(value v); + int (*compare)(value v1, value v2); + intnat (*hash)(value v); + void (*serialize)(value v, + /*out*/ uintnat * bsize_32 /*size in bytes*/, + /*out*/ uintnat * bsize_64 /*size in bytes*/); + uintnat (*deserialize)(void * dst); + int (*compare_ext)(value v1, value v2); +}; + +#define custom_finalize_default NULL +#define custom_compare_default NULL +#define custom_hash_default NULL +#define custom_serialize_default NULL +#define custom_deserialize_default NULL +#define custom_compare_ext_default NULL + +#define Custom_ops_val(v) (*((struct custom_operations **) (v))) + +#ifdef __cplusplus +extern "C" { +#endif + + +CAMLextern value caml_alloc_custom(struct custom_operations * ops, + uintnat size, /*size in bytes*/ + mlsize_t mem, /*resources consumed*/ + mlsize_t max /*max resources*/); + +CAMLextern void caml_register_custom_operations(struct custom_operations * ops); + +CAMLextern int caml_compare_unordered; + /* Used by custom comparison to report unordered NaN-like cases. */ + +#ifdef CAML_INTERNALS +extern struct custom_operations * caml_find_custom_operations(char * ident); +extern struct custom_operations * + caml_final_custom_operations(void (*fn)(value)); + +extern void caml_init_custom_operations(void); +#endif /* CAML_INTERNALS */ + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_CUSTOM_H */ diff --git a/test/monniaux/ocaml/byterun/caml/debugger.h b/test/monniaux/ocaml/byterun/caml/debugger.h new file mode 100644 index 00000000..c98f35a8 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/debugger.h @@ -0,0 +1,117 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Interface with the debugger */ + +#ifndef CAML_DEBUGGER_H +#define CAML_DEBUGGER_H + +#ifdef CAML_INTERNALS + +#include "misc.h" +#include "mlvalues.h" + +CAMLextern int caml_debugger_in_use; +CAMLextern int caml_debugger_fork_mode; /* non-zero for parent */ +extern uintnat caml_event_count; + +enum event_kind { + EVENT_COUNT, BREAKPOINT, PROGRAM_START, PROGRAM_EXIT, + TRAP_BARRIER, UNCAUGHT_EXC +}; + +void caml_debugger_init (void); +void caml_debugger (enum event_kind event); +void caml_debugger_cleanup_fork (void); + +/* Communication protocol */ + +/* Requests from the debugger to the runtime system */ + +enum debugger_request { + REQ_SET_EVENT = 'e', /* uint32_t pos */ + /* Set an event on the instruction at position pos */ + REQ_SET_BREAKPOINT = 'B', /* uint32_t pos, (char k) */ + /* Set a breakpoint at position pos */ + /* In profiling mode, the breakpoint kind is set to k */ + REQ_RESET_INSTR = 'i', /* uint32_t pos */ + /* Clear an event or breapoint at position pos, restores initial instr. */ + REQ_CHECKPOINT = 'c', /* no args */ + /* Checkpoint the runtime system by forking a child process. + Reply is pid of child process or -1 if checkpoint failed. */ + REQ_GO = 'g', /* uint32_t n */ + /* Run the program for n events. + Reply is one of debugger_reply described below. */ + REQ_STOP = 's', /* no args */ + /* Terminate the runtime system */ + REQ_WAIT = 'w', /* no args */ + /* Reap one dead child (a discarded checkpoint). */ + REQ_INITIAL_FRAME = '0', /* no args */ + /* Set current frame to bottom frame (the one currently executing). + Reply is stack offset and current pc. */ + REQ_GET_FRAME = 'f', /* no args */ + /* Return current frame location (stack offset + current pc). */ + REQ_SET_FRAME = 'S', /* uint32_t stack_offset */ + /* Set current frame to given stack offset. No reply. */ + REQ_UP_FRAME = 'U', /* uint32_t n */ + /* Move one frame up. Argument n is size of current frame (in words). + Reply is stack offset and current pc, or -1 if top of stack reached. */ + REQ_SET_TRAP_BARRIER = 'b', /* uint32_t offset */ + /* Set the trap barrier at the given offset. */ + REQ_GET_LOCAL = 'L', /* uint32_t slot_number */ + /* Return the local variable at the given slot in the current frame. + Reply is one value. */ + REQ_GET_ENVIRONMENT = 'E', /* uint32_t slot_number */ + /* Return the local variable at the given slot in the heap environment + of the current frame. Reply is one value. */ + REQ_GET_GLOBAL = 'G', /* uint32_t global_number */ + /* Return the specified global variable. Reply is one value. */ + REQ_GET_ACCU = 'A', /* no args */ + /* Return the current contents of the accumulator. Reply is one value. */ + REQ_GET_HEADER = 'H', /* mlvalue v */ + /* As REQ_GET_OBJ, but sends only the header. */ + REQ_GET_FIELD = 'F', /* mlvalue v, uint32_t fieldnum */ + /* As REQ_GET_OBJ, but sends only one field. */ + REQ_MARSHAL_OBJ = 'M', /* mlvalue v */ + /* Send a copy of the data structure rooted at v, using the same + format as [caml_output_value]. */ + REQ_GET_CLOSURE_CODE = 'C', /* mlvalue v */ + /* Send the code address of the given closure. + Reply is one uint32_t. */ + REQ_SET_FORK_MODE = 'K' /* uint32_t m */ + /* Set whether to follow the child (m=0) or the parent on fork. */ +}; + +/* Replies to a REQ_GO request. All replies are followed by three uint32_t: + - the value of the event counter + - the position of the stack + - the current pc. */ + +enum debugger_reply { + REP_EVENT = 'e', + /* Event counter reached 0. */ + REP_BREAKPOINT = 'b', + /* Breakpoint hit. */ + REP_EXITED = 'x', + /* Program exited by calling exit or reaching the end of the source. */ + REP_TRAP = 's', + /* Trap barrier crossed. */ + REP_UNCAUGHT_EXC = 'u' + /* Program exited due to a stray exception. */ +}; + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_DEBUGGER_H */ diff --git a/test/monniaux/ocaml/byterun/caml/dynlink.h b/test/monniaux/ocaml/byterun/caml/dynlink.h new file mode 100644 index 00000000..92f4e235 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/dynlink.h @@ -0,0 +1,46 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Dynamic loading of C primitives. */ + +#ifndef CAML_DYNLINK_H +#define CAML_DYNLINK_H + +#ifdef CAML_INTERNALS + +#include "misc.h" + +/* Build the table of primitives, given a search path, a list + of shared libraries, and a list of primitive names + (all three 0-separated in char arrays). + Abort the runtime system on error. + Calling this frees caml_shared_libs_path (not touching its contents). */ +extern void caml_build_primitive_table(char_os * lib_path, + char_os * libs, + char * req_prims); + +/* The search path for shared libraries */ +extern struct ext_table caml_shared_libs_path; + +/* Build the table of primitives as a copy of the builtin primitive table. + Used for executables generated by ocamlc -output-obj. */ +extern void caml_build_primitive_table_builtin(void); + +/* Unload all the previously loaded shared libraries */ +extern void caml_free_shared_libs(void); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_DYNLINK_H */ diff --git a/test/monniaux/ocaml/byterun/caml/exec.h b/test/monniaux/ocaml/byterun/caml/exec.h new file mode 100644 index 00000000..38ab7ae8 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/exec.h @@ -0,0 +1,65 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* exec.h : format of executable bytecode files */ + +#ifndef CAML_EXEC_H +#define CAML_EXEC_H + +#ifdef CAML_INTERNALS + +/* Executable bytecode files are composed of a number of sections, + identified by 4-character names. A table of contents at the + end of the file lists the section names along with their sizes, + in the order in which they appear in the file: + + offset 0 ---> initial junk + data for section 1 + data for section 2 + ... + data for section N + table of contents: + descriptor for section 1 + ... + descriptor for section N + trailer + end of file ---> +*/ + +/* Structure of t.o.c. entries + Numerical quantities are 32-bit unsigned integers, big endian */ + +struct section_descriptor { + char name[4]; /* Section name */ + uint32_t len; /* Length of data in bytes */ +}; + +/* Structure of the trailer. */ + +struct exec_trailer { + uint32_t num_sections; /* Number of sections */ + char magic[12]; /* The magic number */ + struct section_descriptor * section; /* Not part of file */ +}; + +#define TRAILER_SIZE (4+12) + +/* Magic number for this release */ + +#define EXEC_MAGIC "Caml1999X023" + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_EXEC_H */ diff --git a/test/monniaux/ocaml/byterun/caml/fail.h b/test/monniaux/ocaml/byterun/caml/fail.h new file mode 100644 index 00000000..54907e42 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/fail.h @@ -0,0 +1,144 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_FAIL_H +#define CAML_FAIL_H + +#ifdef CAML_INTERNALS +#include +#endif /* CAML_INTERNALS */ + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "misc.h" +#include "mlvalues.h" + +#ifdef CAML_INTERNALS +#define OUT_OF_MEMORY_EXN 0 /* "Out_of_memory" */ +#define SYS_ERROR_EXN 1 /* "Sys_error" */ +#define FAILURE_EXN 2 /* "Failure" */ +#define INVALID_EXN 3 /* "Invalid_argument" */ +#define END_OF_FILE_EXN 4 /* "End_of_file" */ +#define ZERO_DIVIDE_EXN 5 /* "Division_by_zero" */ +#define NOT_FOUND_EXN 6 /* "Not_found" */ +#define MATCH_FAILURE_EXN 7 /* "Match_failure" */ +#define STACK_OVERFLOW_EXN 8 /* "Stack_overflow" */ +#define SYS_BLOCKED_IO 9 /* "Sys_blocked_io" */ +#define ASSERT_FAILURE_EXN 10 /* "Assert_failure" */ +#define UNDEFINED_RECURSIVE_MODULE_EXN 11 /* "Undefined_recursive_module" */ + +#ifdef POSIX_SIGNALS +struct longjmp_buffer { + sigjmp_buf buf; +}; +#elif defined(__MINGW64__) && defined(__GNUC__) && __GNUC__ >= 4 +/* MPR#7638: issues with setjmp/longjmp in Mingw64, use GCC builtins instead */ +struct longjmp_buffer { + intptr_t buf[5]; +}; +#define sigsetjmp(buf,save) __builtin_setjmp(buf) +#define siglongjmp(buf,val) __builtin_longjmp(buf,val) +#else +struct longjmp_buffer { + jmp_buf buf; +}; +#define sigsetjmp(buf,save) setjmp(buf) +#define siglongjmp(buf,val) longjmp(buf,val) +#endif + +CAMLextern struct longjmp_buffer * caml_external_raise; +extern value caml_exn_bucket; +int caml_is_special_exception(value exn); + +#endif /* CAML_INTERNALS */ + +#ifdef __cplusplus +extern "C" { +#endif + +CAMLnoreturn_start +CAMLextern void caml_raise (value bucket) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_constant (value tag) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_with_arg (value tag, value arg) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_with_args (value tag, int nargs, value arg[]) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_with_string (value tag, char const * msg) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_failwith (char const *msg) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_failwith_value (value msg) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_invalid_argument (char const *msg) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_invalid_argument_value (value msg) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_out_of_memory (void) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_stack_overflow (void) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_sys_error (value) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_end_of_file (void) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_zero_divide (void) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_not_found (void) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_array_bound_error (void) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_raise_sys_blocked_io (void) +CAMLnoreturn_end; + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_FAIL_H */ diff --git a/test/monniaux/ocaml/byterun/caml/finalise.h b/test/monniaux/ocaml/byterun/caml/finalise.h new file mode 100644 index 00000000..5315ac21 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/finalise.h @@ -0,0 +1,36 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Moscova, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_FINALISE_H +#define CAML_FINALISE_H + +#ifdef CAML_INTERNALS + +#include "roots.h" + +void caml_final_update_mark_phase (void); +void caml_final_update_clean_phase (void); +void caml_final_do_calls (void); +void caml_final_do_roots (scanning_action f); +void caml_final_invert_finalisable_values (); +void caml_final_oldify_young_roots (); +void caml_final_empty_young (void); +void caml_final_update_minor_roots(void); +value caml_final_register (value f, value v); +void caml_final_invariant_check(void); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_FINALISE_H */ diff --git a/test/monniaux/ocaml/byterun/caml/fix_code.h b/test/monniaux/ocaml/byterun/caml/fix_code.h new file mode 100644 index 00000000..7e5633d6 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/fix_code.h @@ -0,0 +1,45 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Handling of blocks of bytecode (endianness switch, threading). */ + +#ifndef CAML_FIX_CODE_H +#define CAML_FIX_CODE_H + +#ifdef CAML_INTERNALS + +#include "config.h" +#include "misc.h" +#include "mlvalues.h" + +extern code_t caml_start_code; +extern asize_t caml_code_size; +extern unsigned char * caml_saved_code; + +void caml_init_code_fragments(void); +void caml_load_code (int fd, asize_t len); +void caml_fixup_endianness (code_t code, asize_t len); +void caml_set_instruction (code_t pos, opcode_t instr); +int caml_is_instruction (opcode_t instr1, opcode_t instr2); + +#ifdef THREADED_CODE +extern char ** caml_instr_table; +extern char * caml_instr_base; +void caml_thread_code (code_t code, asize_t len); +#endif + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_FIX_CODE_H */ diff --git a/test/monniaux/ocaml/byterun/caml/freelist.h b/test/monniaux/ocaml/byterun/caml/freelist.h new file mode 100644 index 00000000..54e0e822 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/freelist.h @@ -0,0 +1,38 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Free lists of heap blocks. */ + +#ifndef CAML_FREELIST_H +#define CAML_FREELIST_H + +#ifdef CAML_INTERNALS + +#include "misc.h" +#include "mlvalues.h" + +extern asize_t caml_fl_cur_wsz; + +header_t *caml_fl_allocate (mlsize_t wo_sz); +void caml_fl_init_merge (void); +void caml_fl_reset (void); +header_t *caml_fl_merge_block (value); +void caml_fl_add_blocks (value); +void caml_make_free_blocks (value *, mlsize_t wsz, int, int); +void caml_set_allocation_policy (uintnat); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_FREELIST_H */ diff --git a/test/monniaux/ocaml/byterun/caml/gc.h b/test/monniaux/ocaml/byterun/caml/gc.h new file mode 100644 index 00000000..a646a756 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/gc.h @@ -0,0 +1,79 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_GC_H +#define CAML_GC_H + + +#include "mlvalues.h" + +#define Caml_white (0 << 8) +#define Caml_gray (1 << 8) +#define Caml_blue (2 << 8) +#define Caml_black (3 << 8) + +#define Color_hd(hd) ((color_t) ((hd) & Caml_black)) +#define Color_hp(hp) (Color_hd (Hd_hp (hp))) +#define Color_val(val) (Color_hd (Hd_val (val))) + +#define Is_white_hd(hd) (Color_hd (hd) == Caml_white) +#define Is_gray_hd(hd) (Color_hd (hd) == Caml_gray) +#define Is_blue_hd(hd) (Color_hd (hd) == Caml_blue) +#define Is_black_hd(hd) (Color_hd (hd) == Caml_black) + +#define Whitehd_hd(hd) (((hd) & ~Caml_black)/*| Caml_white*/) +#define Grayhd_hd(hd) (((hd) & ~Caml_black) | Caml_gray) +#define Blackhd_hd(hd) (((hd)/*& ~Caml_black*/)| Caml_black) +#define Bluehd_hd(hd) (((hd) & ~Caml_black) | Caml_blue) + +/* This depends on the layout of the header. See [mlvalues.h]. */ +#define Make_header(wosize, tag, color) \ + (/*CAMLassert ((wosize) <= Max_wosize),*/ \ + ((header_t) (((header_t) (wosize) << 10) \ + + (color) \ + + (tag_t) (tag))) \ + ) + +#ifdef WITH_PROFINFO +#define Make_header_with_profinfo(wosize, tag, color, profinfo) \ + (Make_header(wosize, tag, color) \ + | ((((intnat) profinfo) & PROFINFO_MASK) << PROFINFO_SHIFT) \ + ) +#else +#define Make_header_with_profinfo(wosize, tag, color, profinfo) \ + Make_header(wosize, tag, color) +#endif + +#ifdef WITH_SPACETIME +struct ext_table; +extern uintnat caml_spacetime_my_profinfo(struct ext_table**, uintnat); +#define Make_header_allocated_here(wosize, tag, color) \ + (Make_header_with_profinfo(wosize, tag, color, \ + caml_spacetime_my_profinfo(NULL, wosize)) \ + ) +#else +#define Make_header_allocated_here Make_header +#endif + +#define Is_white_val(val) (Color_val(val) == Caml_white) +#define Is_gray_val(val) (Color_val(val) == Caml_gray) +#define Is_blue_val(val) (Color_val(val) == Caml_blue) +#define Is_black_val(val) (Color_val(val) == Caml_black) + +/* For extern.c */ +#define Colornum_hd(hd) ((color_t) (((hd) >> 8) & 3)) +#define Coloredhd_hd(hd,colnum) (((hd) & ~Caml_black) | ((colnum) << 8)) + +#endif /* CAML_GC_H */ diff --git a/test/monniaux/ocaml/byterun/caml/gc_ctrl.h b/test/monniaux/ocaml/byterun/caml/gc_ctrl.h new file mode 100644 index 00000000..ebf1a40b --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/gc_ctrl.h @@ -0,0 +1,58 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_GC_CTRL_H +#define CAML_GC_CTRL_H + +#ifdef CAML_INTERNALS + +#include "misc.h" + +extern double + caml_stat_minor_words, + caml_stat_promoted_words, + caml_stat_major_words; + +extern intnat + caml_stat_minor_collections, + caml_stat_major_collections, + caml_stat_heap_wsz, + caml_stat_top_heap_wsz, + caml_stat_compactions, + caml_stat_heap_chunks; + +uintnat caml_normalize_heap_increment (uintnat); + +/* + minor_size: cf. minor_heap_size in gc.mli + major_size: Size in words of the initial major heap + major_incr: cf. major_heap_increment in gc.mli + percent_fr: cf. space_overhead in gc.mli + percent_m : cf. max_overhead in gc.mli + window : cf. window_size in gc.mli +*/ +void caml_init_gc (uintnat minor_size, uintnat major_size, uintnat major_incr, + uintnat percent_fr, uintnat percent_m, uintnat window); + + +CAMLextern value caml_gc_stat(value v); + +#ifdef DEBUG +void caml_heap_check (void); +#endif + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_GC_CTRL_H */ diff --git a/test/monniaux/ocaml/byterun/caml/globroots.h b/test/monniaux/ocaml/byterun/caml/globroots.h new file mode 100644 index 00000000..10fe66f5 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/globroots.h @@ -0,0 +1,31 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2001 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Registration of global memory roots */ + +#ifndef CAML_GLOBROOTS_H +#define CAML_GLOBROOTS_H + +#ifdef CAML_INTERNALS + +#include "mlvalues.h" +#include "roots.h" + +void caml_scan_global_roots(scanning_action f); +void caml_scan_global_young_roots(scanning_action f); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_GLOBROOTS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/hash.h b/test/monniaux/ocaml/byterun/caml/hash.h new file mode 100644 index 00000000..fcc7589f --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/hash.h @@ -0,0 +1,39 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Gallium, INRIA Rocquencourt */ +/* */ +/* Copyright 2011 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Auxiliary functions for custom hash functions */ + +#ifndef CAML_HASH_H +#define CAML_HASH_H + +#include "mlvalues.h" + +#ifdef __cplusplus +extern "C" { +#endif + +CAMLextern uint32_t caml_hash_mix_uint32(uint32_t h, uint32_t d); +CAMLextern uint32_t caml_hash_mix_intnat(uint32_t h, intnat d); +CAMLextern uint32_t caml_hash_mix_int64(uint32_t h, int64_t d); +CAMLextern uint32_t caml_hash_mix_double(uint32_t h, double d); +CAMLextern uint32_t caml_hash_mix_float(uint32_t h, float d); +CAMLextern uint32_t caml_hash_mix_string(uint32_t h, value s); + +#ifdef __cplusplus +} +#endif + + +#endif /* CAML_HASH_H */ diff --git a/test/monniaux/ocaml/byterun/caml/hooks.h b/test/monniaux/ocaml/byterun/caml/hooks.h new file mode 100644 index 00000000..5877dd25 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/hooks.h @@ -0,0 +1,42 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Fabrice Le Fessant, INRIA de Paris */ +/* */ +/* Copyright 2016 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_HOOKS_H +#define CAML_HOOKS_H + +#include "misc.h" +#include "memory.h" + +#ifdef __cplusplus +extern "C" { +#endif + +#ifdef CAML_INTERNALS + +#ifdef NATIVE_CODE + +/* executed just before calling the entry point of a dynamically + loaded native code module. */ +CAMLextern void (*caml_natdynlink_hook)(void* handle, const char* unit); + +#endif /* NATIVE_CODE */ + +#endif /* CAML_INTERNALS */ + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_HOOKS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/instrtrace.h b/test/monniaux/ocaml/byterun/caml/instrtrace.h new file mode 100644 index 00000000..2e42a80a --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/instrtrace.h @@ -0,0 +1,35 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Trace the instructions executed */ + +#ifndef _instrtrace_ +#define _instrtrace_ + +#ifdef CAML_INTERNALS + +#include "mlvalues.h" +#include "misc.h" + +extern intnat caml_icount; +void caml_stop_here (void); +void caml_disasm_instr (code_t pc); +void caml_trace_value_file (value v, code_t prog, int proglen, FILE * f); +void caml_trace_accu_sp_file(value accu, value * sp, code_t prog, int proglen, + FILE * f); + +#endif /* CAML_INTERNALS */ + +#endif diff --git a/test/monniaux/ocaml/byterun/caml/instruct.h b/test/monniaux/ocaml/byterun/caml/instruct.h new file mode 100644 index 00000000..5c10df4f --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/instruct.h @@ -0,0 +1,68 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* The instruction set. */ + +#ifndef CAML_INSTRUCT_H +#define CAML_INSTRUCT_H + +#ifdef CAML_INTERNALS + +enum instructions { + ACC0, ACC1, ACC2, ACC3, ACC4, ACC5, ACC6, ACC7, + ACC, PUSH, + PUSHACC0, PUSHACC1, PUSHACC2, PUSHACC3, + PUSHACC4, PUSHACC5, PUSHACC6, PUSHACC7, + PUSHACC, POP, ASSIGN, + ENVACC1, ENVACC2, ENVACC3, ENVACC4, ENVACC, + PUSHENVACC1, PUSHENVACC2, PUSHENVACC3, PUSHENVACC4, PUSHENVACC, + PUSH_RETADDR, APPLY, APPLY1, APPLY2, APPLY3, + APPTERM, APPTERM1, APPTERM2, APPTERM3, + RETURN, RESTART, GRAB, + CLOSURE, CLOSUREREC, + OFFSETCLOSUREM2, OFFSETCLOSURE0, OFFSETCLOSURE2, OFFSETCLOSURE, + PUSHOFFSETCLOSUREM2, PUSHOFFSETCLOSURE0, + PUSHOFFSETCLOSURE2, PUSHOFFSETCLOSURE, + GETGLOBAL, PUSHGETGLOBAL, GETGLOBALFIELD, PUSHGETGLOBALFIELD, SETGLOBAL, + ATOM0, ATOM, PUSHATOM0, PUSHATOM, + MAKEBLOCK, MAKEBLOCK1, MAKEBLOCK2, MAKEBLOCK3, MAKEFLOATBLOCK, + GETFIELD0, GETFIELD1, GETFIELD2, GETFIELD3, GETFIELD, GETFLOATFIELD, + SETFIELD0, SETFIELD1, SETFIELD2, SETFIELD3, SETFIELD, SETFLOATFIELD, + VECTLENGTH, GETVECTITEM, SETVECTITEM, + GETBYTESCHAR, SETBYTESCHAR, + BRANCH, BRANCHIF, BRANCHIFNOT, SWITCH, BOOLNOT, + PUSHTRAP, POPTRAP, RAISE, + CHECK_SIGNALS, + C_CALL1, C_CALL2, C_CALL3, C_CALL4, C_CALL5, C_CALLN, + CONST0, CONST1, CONST2, CONST3, CONSTINT, + PUSHCONST0, PUSHCONST1, PUSHCONST2, PUSHCONST3, PUSHCONSTINT, + NEGINT, ADDINT, SUBINT, MULINT, DIVINT, MODINT, + ANDINT, ORINT, XORINT, LSLINT, LSRINT, ASRINT, + EQ, NEQ, LTINT, LEINT, GTINT, GEINT, + OFFSETINT, OFFSETREF, ISINT, + GETMETHOD, + BEQ, BNEQ, BLTINT, BLEINT, BGTINT, BGEINT, + ULTINT, UGEINT, + BULTINT, BUGEINT, + GETPUBMET, GETDYNMET, + STOP, + EVENT, BREAK, + RERAISE, RAISE_NOTRACE, + GETSTRINGCHAR, +FIRST_UNIMPLEMENTED_OP}; + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_INSTRUCT_H */ diff --git a/test/monniaux/ocaml/byterun/caml/int64_emul.h b/test/monniaux/ocaml/byterun/caml/int64_emul.h new file mode 100644 index 00000000..c1cddcc0 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/int64_emul.h @@ -0,0 +1,293 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2002 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Software emulation of 64-bit integer arithmetic, for C compilers + that do not support it. */ + +#ifndef CAML_INT64_EMUL_H +#define CAML_INT64_EMUL_H + +#ifdef CAML_INTERNALS + +#include + +#ifdef ARCH_BIG_ENDIAN +#define I64_literal(hi,lo) { hi, lo } +#else +#define I64_literal(hi,lo) { lo, hi } +#endif + +#define I64_split(x,hi,lo) (hi = (x).h, lo = (x).l) + +/* Unsigned comparison */ +static int I64_ucompare(uint64_t x, uint64_t y) +{ + if (x.h > y.h) return 1; + if (x.h < y.h) return -1; + if (x.l > y.l) return 1; + if (x.l < y.l) return -1; + return 0; +} + +#define I64_ult(x, y) (I64_ucompare(x, y) < 0) + +/* Signed comparison */ +static int I64_compare(int64_t x, int64_t y) +{ + if ((int32_t)x.h > (int32_t)y.h) return 1; + if ((int32_t)x.h < (int32_t)y.h) return -1; + if (x.l > y.l) return 1; + if (x.l < y.l) return -1; + return 0; +} + +/* Negation */ +static int64_t I64_neg(int64_t x) +{ + int64_t res; + res.l = -x.l; + res.h = ~x.h; + if (res.l == 0) res.h++; + return res; +} + +/* Addition */ +static int64_t I64_add(int64_t x, int64_t y) +{ + int64_t res; + res.l = x.l + y.l; + res.h = x.h + y.h; + if (res.l < x.l) res.h++; + return res; +} + +/* Subtraction */ +static int64_t I64_sub(int64_t x, int64_t y) +{ + int64_t res; + res.l = x.l - y.l; + res.h = x.h - y.h; + if (x.l < y.l) res.h--; + return res; +} + +/* Multiplication */ +static int64_t I64_mul(int64_t x, int64_t y) +{ + int64_t res; + uint32_t prod00 = (x.l & 0xFFFF) * (y.l & 0xFFFF); + uint32_t prod10 = (x.l >> 16) * (y.l & 0xFFFF); + uint32_t prod01 = (x.l & 0xFFFF) * (y.l >> 16); + uint32_t prod11 = (x.l >> 16) * (y.l >> 16); + res.l = prod00; + res.h = prod11 + (prod01 >> 16) + (prod10 >> 16); + prod01 = prod01 << 16; res.l += prod01; if (res.l < prod01) res.h++; + prod10 = prod10 << 16; res.l += prod10; if (res.l < prod10) res.h++; + res.h += x.l * y.h + x.h * y.l; + return res; +} + +#define I64_is_zero(x) (((x).l | (x).h) == 0) +#define I64_is_negative(x) ((int32_t) (x).h < 0) +#define I64_is_min_int(x) ((x).l == 0 && (x).h == 0x80000000U) +#define I64_is_minus_one(x) (((x).l & (x).h) == 0xFFFFFFFFU) + +/* Bitwise operations */ +static int64_t I64_and(int64_t x, int64_t y) +{ + int64_t res; + res.l = x.l & y.l; + res.h = x.h & y.h; + return res; +} + +static int64_t I64_or(int64_t x, int64_t y) +{ + int64_t res; + res.l = x.l | y.l; + res.h = x.h | y.h; + return res; +} + +static int64_t I64_xor(int64_t x, int64_t y) +{ + int64_t res; + res.l = x.l ^ y.l; + res.h = x.h ^ y.h; + return res; +} + +/* Shifts */ +static int64_t I64_lsl(int64_t x, int s) +{ + int64_t res; + s = s & 63; + if (s == 0) return x; + if (s < 32) { + res.l = x.l << s; + res.h = (x.h << s) | (x.l >> (32 - s)); + } else { + res.l = 0; + res.h = x.l << (s - 32); + } + return res; +} + +static int64_t I64_lsr(int64_t x, int s) +{ + int64_t res; + s = s & 63; + if (s == 0) return x; + if (s < 32) { + res.l = (x.l >> s) | (x.h << (32 - s)); + res.h = x.h >> s; + } else { + res.l = x.h >> (s - 32); + res.h = 0; + } + return res; +} + +static int64_t I64_asr(int64_t x, int s) +{ + int64_t res; + s = s & 63; + if (s == 0) return x; + if (s < 32) { + res.l = (x.l >> s) | (x.h << (32 - s)); + res.h = (int32_t) x.h >> s; + } else { + res.l = (int32_t) x.h >> (s - 32); + res.h = (int32_t) x.h >> 31; + } + return res; +} + +/* Division and modulus */ + +#define I64_SHL1(x) x.h = (x.h << 1) | (x.l >> 31); x.l <<= 1 +#define I64_SHR1(x) x.l = (x.l >> 1) | (x.h << 31); x.h >>= 1 + +static void I64_udivmod(uint64_t modulus, uint64_t divisor, + uint64_t * quo, uint64_t * mod) +{ + int64_t quotient, mask; + int cmp; + + quotient.h = 0; quotient.l = 0; + mask.h = 0; mask.l = 1; + while ((int32_t) divisor.h >= 0) { + cmp = I64_ucompare(divisor, modulus); + I64_SHL1(divisor); + I64_SHL1(mask); + if (cmp >= 0) break; + } + while (mask.l | mask.h) { + if (I64_ucompare(modulus, divisor) >= 0) { + quotient.h |= mask.h; quotient.l |= mask.l; + modulus = I64_sub(modulus, divisor); + } + I64_SHR1(mask); + I64_SHR1(divisor); + } + *quo = quotient; + *mod = modulus; +} + +static int64_t I64_div(int64_t x, int64_t y) +{ + int64_t q, r; + int32_t sign; + + sign = x.h ^ y.h; + if ((int32_t) x.h < 0) x = I64_neg(x); + if ((int32_t) y.h < 0) y = I64_neg(y); + I64_udivmod(x, y, &q, &r); + if (sign < 0) q = I64_neg(q); + return q; +} + +static int64_t I64_mod(int64_t x, int64_t y) +{ + int64_t q, r; + int32_t sign; + + sign = x.h; + if ((int32_t) x.h < 0) x = I64_neg(x); + if ((int32_t) y.h < 0) y = I64_neg(y); + I64_udivmod(x, y, &q, &r); + if (sign < 0) r = I64_neg(r); + return r; +} + +/* Coercions */ + +static int64_t I64_of_int32(int32_t x) +{ + int64_t res; + res.l = x; + res.h = x >> 31; + return res; +} + +#define I64_to_int32(x) ((int32_t) (x).l) + +/* Note: we assume sizeof(intnat) = 4 here, which is true otherwise + autoconfiguration would have selected native 64-bit integers */ +#define I64_of_intnat I64_of_int32 +#define I64_to_intnat I64_to_int32 + +static double I64_to_double(int64_t x) +{ + double res; + int32_t sign = x.h; + if (sign < 0) x = I64_neg(x); + res = ldexp((double) x.h, 32) + x.l; + if (sign < 0) res = -res; + return res; +} + +static int64_t I64_of_double(double f) +{ + int64_t res; + double frac, integ; + int neg; + + neg = (f < 0); + f = fabs(f); + frac = modf(ldexp(f, -32), &integ); + res.h = (uint32_t) integ; + res.l = (uint32_t) ldexp(frac, 32); + if (neg) res = I64_neg(res); + return res; +} + +static int64_t I64_bswap(int64_t x) +{ + int64_t res; + res.h = (((x.l & 0x000000FF) << 24) | + ((x.l & 0x0000FF00) << 8) | + ((x.l & 0x00FF0000) >> 8) | + ((x.l & 0xFF000000) >> 24)); + res.l = (((x.h & 0x000000FF) << 24) | + ((x.h & 0x0000FF00) << 8) | + ((x.h & 0x00FF0000) >> 8) | + ((x.h & 0xFF000000) >> 24)); + return res; +} + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_INT64_EMUL_H */ diff --git a/test/monniaux/ocaml/byterun/caml/int64_format.h b/test/monniaux/ocaml/byterun/caml/int64_format.h new file mode 100644 index 00000000..40250ed9 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/int64_format.h @@ -0,0 +1,111 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2002 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* printf-like formatting of 64-bit integers, in case the C library + printf() function does not support them. */ + +#ifndef CAML_INT64_FORMAT_H +#define CAML_INT64_FORMAT_H + +#ifdef CAML_INTERNALS + +static void I64_format(char * buffer, char * fmt, int64_t x) +{ + static char conv_lower[] = "0123456789abcdef"; + static char conv_upper[] = "0123456789ABCDEF"; + char rawbuffer[24]; + char justify, signstyle, filler, alternate, signedconv; + int base, width, sign, i, rawlen; + char * cvtbl; + char * p, * r; + int64_t wbase, digit; + + /* Parsing of format */ + justify = '+'; + signstyle = '-'; + filler = ' '; + alternate = 0; + base = 0; + signedconv = 0; + width = 0; + cvtbl = conv_lower; + for (p = fmt; *p != 0; p++) { + switch (*p) { + case '-': + justify = '-'; break; + case '+': case ' ': + signstyle = *p; break; + case '0': + filler = '0'; break; + case '#': + alternate = 1; break; + case '1': case '2': case '3': case '4': case '5': + case '6': case '7': case '8': case '9': + width = atoi(p); + while (p[1] >= '0' && p[1] <= '9') p++; + break; + case 'd': case 'i': + signedconv = 1; /* fallthrough */ + case 'u': + base = 10; break; + case 'x': + base = 16; break; + case 'X': + base = 16; cvtbl = conv_upper; break; + case 'o': + base = 8; break; + } + } + if (base == 0) { buffer[0] = 0; return; } + /* Do the conversion */ + sign = 1; + if (signedconv && I64_is_negative(x)) { sign = -1; x = I64_neg(x); } + r = rawbuffer + sizeof(rawbuffer); + wbase = I64_of_int32(base); + do { + I64_udivmod(x, wbase, &x, &digit); + *--r = cvtbl[I64_to_int32(digit)]; + } while (! I64_is_zero(x)); + rawlen = rawbuffer + sizeof(rawbuffer) - r; + /* Adjust rawlen to reflect additional chars (sign, etc) */ + if (signedconv && (sign < 0 || signstyle != '-')) rawlen++; + if (alternate) { + if (base == 8) rawlen += 1; + if (base == 16) rawlen += 2; + } + /* Do the formatting */ + p = buffer; + if (justify == '+' && filler == ' ') { + for (i = rawlen; i < width; i++) *p++ = ' '; + } + if (signedconv) { + if (sign < 0) *p++ = '-'; + else if (signstyle != '-') *p++ = signstyle; + } + if (alternate && base == 8) *p++ = '0'; + if (alternate && base == 16) { *p++ = '0'; *p++ = 'x'; } + if (justify == '+' && filler == '0') { + for (i = rawlen; i < width; i++) *p++ = '0'; + } + while (r < rawbuffer + sizeof(rawbuffer)) *p++ = *r++; + if (justify == '-') { + for (i = rawlen; i < width; i++) *p++ = ' '; + } + *p = 0; +} + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_INT64_FORMAT_H */ diff --git a/test/monniaux/ocaml/byterun/caml/int64_native.h b/test/monniaux/ocaml/byterun/caml/int64_native.h new file mode 100644 index 00000000..7df66511 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/int64_native.h @@ -0,0 +1,67 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2002 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Wrapper macros around native 64-bit integer arithmetic, + so that it has the same interface as the software emulation + provided in int64_emul.h */ + +#ifndef CAML_INT64_NATIVE_H +#define CAML_INT64_NATIVE_H + +#ifdef CAML_INTERNALS + +#define I64_literal(hi,lo) ((int64_t)(hi) << 32 | (lo)) +#define I64_split(x,hi,lo) (hi = (uint32_t)((x)>>32), lo = (uint32_t)(x)) +#define I64_compare(x,y) (((x) > (y)) - ((x) < (y))) +#define I64_ult(x,y) ((uint64_t)(x) < (uint64_t)(y)) +#define I64_neg(x) (-(x)) +#define I64_add(x,y) ((x) + (y)) +#define I64_sub(x,y) ((x) - (y)) +#define I64_mul(x,y) ((x) * (y)) +#define I64_is_zero(x) ((x) == 0) +#define I64_is_negative(x) ((x) < 0) +#define I64_is_min_int(x) ((x) == ((int64_t)1 << 63)) +#define I64_is_minus_one(x) ((x) == -1) + +#define I64_div(x,y) ((x) / (y)) +#define I64_mod(x,y) ((x) % (y)) +#define I64_udivmod(x,y,quo,rem) \ + (*(rem) = (uint64_t)(x) % (uint64_t)(y), \ + *(quo) = (uint64_t)(x) / (uint64_t)(y)) +#define I64_and(x,y) ((x) & (y)) +#define I64_or(x,y) ((x) | (y)) +#define I64_xor(x,y) ((x) ^ (y)) +#define I64_lsl(x,y) ((x) << (y)) +#define I64_asr(x,y) ((x) >> (y)) +#define I64_lsr(x,y) ((uint64_t)(x) >> (y)) +#define I64_to_intnat(x) ((intnat) (x)) +#define I64_of_intnat(x) ((intnat) (x)) +#define I64_to_int32(x) ((int32_t) (x)) +#define I64_of_int32(x) ((int64_t) (x)) +#define I64_to_double(x) ((double)(x)) +#define I64_of_double(x) ((int64_t)(x)) + +#define I64_bswap(x) ((((x) & 0x00000000000000FFULL) << 56) | \ + (((x) & 0x000000000000FF00ULL) << 40) | \ + (((x) & 0x0000000000FF0000ULL) << 24) | \ + (((x) & 0x00000000FF000000ULL) << 8) | \ + (((x) & 0x000000FF00000000ULL) >> 8) | \ + (((x) & 0x0000FF0000000000ULL) >> 24) | \ + (((x) & 0x00FF000000000000ULL) >> 40) | \ + (((x) & 0xFF00000000000000ULL) >> 56)) + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_INT64_NATIVE_H */ diff --git a/test/monniaux/ocaml/byterun/caml/interp.h b/test/monniaux/ocaml/byterun/caml/interp.h new file mode 100644 index 00000000..d1ebdc01 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/interp.h @@ -0,0 +1,37 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* The bytecode interpreter */ + +#ifndef CAML_INTERP_H +#define CAML_INTERP_H + +#ifdef CAML_INTERNALS + +#include "misc.h" +#include "mlvalues.h" + +/* interpret a bytecode */ +value caml_interprete (code_t prog, asize_t prog_size); + +/* tell the runtime that a bytecode program might be needed */ +void caml_prepare_bytecode(code_t prog, asize_t prog_size); + +/* tell the runtime that a bytecode program is no more needed */ +void caml_release_bytecode(code_t prog, asize_t prog_size); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_INTERP_H */ diff --git a/test/monniaux/ocaml/byterun/caml/intext.h b/test/monniaux/ocaml/byterun/caml/intext.h new file mode 100644 index 00000000..f67c98b5 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/intext.h @@ -0,0 +1,207 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Structured input/output */ + +#ifndef CAML_INTEXT_H +#define CAML_INTEXT_H + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "misc.h" +#include "mlvalues.h" + +#ifdef CAML_INTERNALS +#include "io.h" + +/* Magic number */ + +#define Intext_magic_number_small 0x8495A6BE +#define Intext_magic_number_big 0x8495A6BF + +/* Header format for the "small" model: 20 bytes + 0 "small" magic number + 4 length of marshaled data, in bytes + 8 number of shared blocks + 12 size in words when read on a 32-bit platform + 16 size in words when read on a 64-bit platform + The 4 numbers are 32 bits each, in big endian. + + Header format for the "big" model: 32 bytes + 0 "big" magic number + 4 four reserved bytes, currently set to 0 + 8 length of marshaled data, in bytes + 16 number of shared blocks + 24 size in words when read on a 64-bit platform + The 3 numbers are 64 bits each, in big endian. +*/ + +/* Codes for the compact format */ + +#define PREFIX_SMALL_BLOCK 0x80 +#define PREFIX_SMALL_INT 0x40 +#define PREFIX_SMALL_STRING 0x20 +#define CODE_INT8 0x0 +#define CODE_INT16 0x1 +#define CODE_INT32 0x2 +#define CODE_INT64 0x3 +#define CODE_SHARED8 0x4 +#define CODE_SHARED16 0x5 +#define CODE_SHARED32 0x6 +#define CODE_SHARED64 0x14 +#define CODE_BLOCK32 0x8 +#define CODE_BLOCK64 0x13 +#define CODE_STRING8 0x9 +#define CODE_STRING32 0xA +#define CODE_STRING64 0x15 +#define CODE_DOUBLE_BIG 0xB +#define CODE_DOUBLE_LITTLE 0xC +#define CODE_DOUBLE_ARRAY8_BIG 0xD +#define CODE_DOUBLE_ARRAY8_LITTLE 0xE +#define CODE_DOUBLE_ARRAY32_BIG 0xF +#define CODE_DOUBLE_ARRAY32_LITTLE 0x7 +#define CODE_DOUBLE_ARRAY64_BIG 0x16 +#define CODE_DOUBLE_ARRAY64_LITTLE 0x17 +#define CODE_CODEPOINTER 0x10 +#define CODE_INFIXPOINTER 0x11 +#define CODE_CUSTOM 0x12 + +#if ARCH_FLOAT_ENDIANNESS == 0x76543210 +#define CODE_DOUBLE_NATIVE CODE_DOUBLE_BIG +#define CODE_DOUBLE_ARRAY8_NATIVE CODE_DOUBLE_ARRAY8_BIG +#define CODE_DOUBLE_ARRAY32_NATIVE CODE_DOUBLE_ARRAY32_BIG +#define CODE_DOUBLE_ARRAY64_NATIVE CODE_DOUBLE_ARRAY64_BIG +#else +#define CODE_DOUBLE_NATIVE CODE_DOUBLE_LITTLE +#define CODE_DOUBLE_ARRAY8_NATIVE CODE_DOUBLE_ARRAY8_LITTLE +#define CODE_DOUBLE_ARRAY32_NATIVE CODE_DOUBLE_ARRAY32_LITTLE +#define CODE_DOUBLE_ARRAY64_NATIVE CODE_DOUBLE_ARRAY64_LITTLE +#endif + +/* Size-ing data structures for extern. Chosen so that + sizeof(struct trail_block) and sizeof(struct output_block) + are slightly below 8Kb. */ + +#define ENTRIES_PER_TRAIL_BLOCK 1025 +#define SIZE_EXTERN_OUTPUT_BLOCK 8100 + +/* The entry points */ + +void caml_output_val (struct channel * chan, value v, value flags); + /* Output [v] with flags [flags] on the channel [chan]. */ + +#endif /* CAML_INTERNALS */ + +#ifdef __cplusplus +extern "C" { +#endif + +CAMLextern void caml_output_value_to_malloc(value v, value flags, + /*out*/ char ** buf, + /*out*/ intnat * len); + /* Output [v] with flags [flags] to a memory buffer allocated with + malloc. On return, [*buf] points to the buffer and [*len] + contains the number of bytes in buffer. */ +CAMLextern intnat caml_output_value_to_block(value v, value flags, + char * data, intnat len); + /* Output [v] with flags [flags] to a user-provided memory buffer. + [data] points to the start of this buffer, and [len] is its size + in bytes. Return the number of bytes actually written in buffer. + Raise [Failure] if buffer is too short. */ + +#ifdef CAML_INTERNALS +value caml_input_val (struct channel * chan); + /* Read a structured value from the channel [chan]. */ + +extern value caml_input_value_to_outside_heap (value channel); + /* As for [caml_input_value], but the value is unmarshalled into + malloc blocks that are not added to the heap. Not for the + casual user. */ + +extern int caml_extern_allow_out_of_heap; + /* Permit the marshaller to traverse structures that look like OCaml + values but do not live in the OCaml heap. */ + +extern value caml_output_value(value vchan, value v, value flags); +#endif /* CAML_INTERNALS */ + +CAMLextern value caml_input_val_from_string (value str, intnat ofs); + /* Read a structured value from the OCaml string [str], starting + at offset [ofs]. */ +CAMLextern value caml_input_value_from_malloc(char * data, intnat ofs); + /* Read a structured value from a malloced buffer. [data] points + to the beginning of the buffer, and [ofs] is the offset of the + beginning of the externed data in this buffer. The buffer is + deallocated with [free] on return, or if an exception is raised. */ +CAMLextern value caml_input_value_from_block(char * data, intnat len); + /* Read a structured value from a user-provided buffer. [data] points + to the beginning of the externed data in this buffer, + and [len] is the length in bytes of valid data in this buffer. + The buffer is never deallocated by this routine. */ + +/* Functions for writing user-defined marshallers */ + +CAMLextern void caml_serialize_int_1(int i); +CAMLextern void caml_serialize_int_2(int i); +CAMLextern void caml_serialize_int_4(int32_t i); +CAMLextern void caml_serialize_int_8(int64_t i); +CAMLextern void caml_serialize_float_4(float f); +CAMLextern void caml_serialize_float_8(double f); +CAMLextern void caml_serialize_block_1(void * data, intnat len); +CAMLextern void caml_serialize_block_2(void * data, intnat len); +CAMLextern void caml_serialize_block_4(void * data, intnat len); +CAMLextern void caml_serialize_block_8(void * data, intnat len); +CAMLextern void caml_serialize_block_float_8(void * data, intnat len); + +CAMLextern int caml_deserialize_uint_1(void); +CAMLextern int caml_deserialize_sint_1(void); +CAMLextern int caml_deserialize_uint_2(void); +CAMLextern int caml_deserialize_sint_2(void); +CAMLextern uint32_t caml_deserialize_uint_4(void); +CAMLextern int32_t caml_deserialize_sint_4(void); +CAMLextern uint64_t caml_deserialize_uint_8(void); +CAMLextern int64_t caml_deserialize_sint_8(void); +CAMLextern float caml_deserialize_float_4(void); +CAMLextern double caml_deserialize_float_8(void); +CAMLextern void caml_deserialize_block_1(void * data, intnat len); +CAMLextern void caml_deserialize_block_2(void * data, intnat len); +CAMLextern void caml_deserialize_block_4(void * data, intnat len); +CAMLextern void caml_deserialize_block_8(void * data, intnat len); +CAMLextern void caml_deserialize_block_float_8(void * data, intnat len); +CAMLextern void caml_deserialize_error(char * msg); + +#ifdef CAML_INTERNALS + +/* Auxiliary stuff for sending code pointers */ + +struct code_fragment { + char * code_start; + char * code_end; + unsigned char digest[16]; + char digest_computed; +}; + +CAMLextern struct code_fragment * caml_extern_find_code(char *addr); + +extern struct ext_table caml_code_fragments_table; + +#endif /* CAML_INTERNALS */ + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_INTEXT_H */ diff --git a/test/monniaux/ocaml/byterun/caml/io.h b/test/monniaux/ocaml/byterun/caml/io.h new file mode 100644 index 00000000..87de679e --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/io.h @@ -0,0 +1,126 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Buffered input/output */ + +#ifndef CAML_IO_H +#define CAML_IO_H + +#ifdef CAML_INTERNALS + +#include "misc.h" +#include "mlvalues.h" + +#ifndef IO_BUFFER_SIZE +#define IO_BUFFER_SIZE 65536 +#endif + +#if defined(_WIN32) +typedef __int64 file_offset; +#elif defined(HAS_OFF_T) +#include +typedef off_t file_offset; +#else +typedef long file_offset; +#endif + +struct channel { + int fd; /* Unix file descriptor */ + file_offset offset; /* Absolute position of fd in the file */ + char * end; /* Physical end of the buffer */ + char * curr; /* Current position in the buffer */ + char * max; /* Logical end of the buffer (for input) */ + void * mutex; /* Placeholder for mutex (for systhreads) */ + struct channel * next, * prev;/* Double chaining of channels (flush_all) */ + int revealed; /* For Cash only */ + int old_revealed; /* For Cash only */ + int refcount; /* For flush_all and for Cash */ + int flags; /* Bitfield */ + char buff[IO_BUFFER_SIZE]; /* The buffer itself */ + char * name; /* Optional name (to report fd leaks) */ +}; + +enum { + CHANNEL_FLAG_FROM_SOCKET = 1, /* For Windows */ +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + CHANNEL_FLAG_BLOCKING_WRITE = 2, /* Don't release master lock when writing */ +#endif + CHANNEL_FLAG_MANAGED_BY_GC = 4, /* Free and close using GC finalization */ +}; + +/* For an output channel: + [offset] is the absolute position of the beginning of the buffer [buff]. + For an input channel: + [offset] is the absolute position of the logical end of the buffer, [max]. +*/ + +/* Functions and macros that can be called from C. Take arguments of + type struct channel *. No locking is performed. */ + +#define caml_putch(channel, ch) do{ \ + if ((channel)->curr >= (channel)->end) caml_flush_partial(channel); \ + *((channel)->curr)++ = (ch); \ +}while(0) + +#define caml_getch(channel) \ + ((channel)->curr >= (channel)->max \ + ? caml_refill(channel) \ + : (unsigned char) *((channel)->curr)++) + +CAMLextern struct channel * caml_open_descriptor_in (int); +CAMLextern struct channel * caml_open_descriptor_out (int); +CAMLextern void caml_close_channel (struct channel *); +CAMLextern int caml_channel_binary_mode (struct channel *); +CAMLextern value caml_alloc_channel(struct channel *chan); + +CAMLextern int caml_flush_partial (struct channel *); +CAMLextern void caml_flush (struct channel *); +CAMLextern void caml_putword (struct channel *, uint32_t); +CAMLextern int caml_putblock (struct channel *, char *, intnat); +CAMLextern void caml_really_putblock (struct channel *, char *, intnat); + +CAMLextern unsigned char caml_refill (struct channel *); +CAMLextern uint32_t caml_getword (struct channel *); +CAMLextern int caml_getblock (struct channel *, char *, intnat); +CAMLextern intnat caml_really_getblock (struct channel *, char *, intnat); + +/* Extract a struct channel * from the heap object representing it */ + +#define Channel(v) (*((struct channel **) (Data_custom_val(v)))) + +/* The locking machinery */ + +CAMLextern void (*caml_channel_mutex_free) (struct channel *); +CAMLextern void (*caml_channel_mutex_lock) (struct channel *); +CAMLextern void (*caml_channel_mutex_unlock) (struct channel *); +CAMLextern void (*caml_channel_mutex_unlock_exn) (void); + +CAMLextern struct channel * caml_all_opened_channels; + +#define Lock(channel) \ + if (caml_channel_mutex_lock != NULL) (*caml_channel_mutex_lock)(channel) +#define Unlock(channel) \ + if (caml_channel_mutex_unlock != NULL) (*caml_channel_mutex_unlock)(channel) +#define Unlock_exn() \ + if (caml_channel_mutex_unlock_exn != NULL) (*caml_channel_mutex_unlock_exn)() + +/* Conversion between file_offset and int64_t */ + +#define Val_file_offset(fofs) caml_copy_int64(fofs) +#define File_offset_val(v) ((file_offset) Int64_val(v)) + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_IO_H */ diff --git a/test/monniaux/ocaml/byterun/caml/m.h b/test/monniaux/ocaml/byterun/caml/m.h new file mode 100644 index 00000000..c54b4f0d --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/m.h @@ -0,0 +1,86 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Processor dependencies */ + +#define ARCH_SIXTYFOUR + +/* Define ARCH_SIXTYFOUR if the processor has a natural word size of 64 bits. + That is, sizeof(char *) = 8. + Otherwise, leave ARCH_SIXTYFOUR undefined. + This assumes sizeof(char *) = 4. */ + +/* #define ARCH_BIG_ENDIAN */ + +/* Define ARCH_BIG_ENDIAN if the processor is big endian (the most + significant byte of an integer stored in memory comes first). + Leave ARCH_BIG_ENDIAN undefined if the processor is little-endian + (the least significant byte comes first). +*/ + +#define ARCH_ALIGN_DOUBLE + +/* Define ARCH_ALIGN_DOUBLE if the processor requires doubles to be + doubleword-aligned. Leave ARCH_ALIGN_DOUBLE undefined if the processor + supports word-aligned doubles. */ + +#undef ARCH_CODE32 + +/* Define ARCH_CODE32 if, on a 64-bit machine, code pointers fit in 32 bits, + i.e. the code segment resides in the low 4G of the addressing space. + ARCH_CODE32 is ignored on 32-bit machines. */ + +#define SIZEOF_INT 4 +#define SIZEOF_LONG 8 +#define SIZEOF_PTR 8 +#define SIZEOF_SHORT 2 + +/* Define SIZEOF_INT, SIZEOF_LONG, SIZEOF_PTR and SIZEOF_SHORT + to the sizes in bytes of the C types "int", "long", "char *" and "short", + respectively. */ + +#define ARCH_INT64_TYPE long long +#define ARCH_UINT64_TYPE unsigned long long + +/* Define ARCH_INT64_TYPE and ARCH_UINT64_TYPE to 64-bit integer types, + typically "long long" and "unsigned long long" on 32-bit platforms, + and "long" and "unsigned long" on 64-bit platforms. + If the C compiler doesn't support any 64-bit integer type, + leave both ARCH_INT64_TYPE and ARCH_UINT64_TYPE undefined. */ + +#define ARCH_INT64_PRINTF_FORMAT "ll" + +/* Define ARCH_INT64_PRINTF_FORMAT to the printf format used for formatting + values of type ARCH_INT64_TYPE. This is usually "ll" on 32-bit + platforms and "l" on 64-bit platforms. + Leave undefined if ARCH_INT64_TYPE is undefined. */ + +#define ARCH_ALIGN_INT64 + +/* Define ARCH_ALIGN_INT64 if the processor requires 64-bit integers to be + doubleword-aligned. Leave ARCH_ALIGN_INT64 undefined if the processor + supports word-aligned 64-bit integers. Leave undefined if + 64-bit integers are not supported. */ + +#undef NONSTANDARD_DIV_MOD + +/* Leave NONSTANDARD_DIV_MOD undefined if the C operators / and % implement + round-towards-zero semantics, as specified by ISO C 9x and implemented + by most contemporary processors. Otherwise, or if you don't know, + define NONSTANDARD_DIV_MOD: this will select a slower but correct + software emulation of division and modulus. */ + +#define PROFINFO_WIDTH 0 +#define INT64_LITERAL(s) s ## LL diff --git a/test/monniaux/ocaml/byterun/caml/major_gc.h b/test/monniaux/ocaml/byterun/caml/major_gc.h new file mode 100644 index 00000000..813f8a78 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/major_gc.h @@ -0,0 +1,95 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_MAJOR_GC_H +#define CAML_MAJOR_GC_H + +#ifdef CAML_INTERNALS + +#include "freelist.h" +#include "misc.h" + +typedef struct { + void *block; /* address of the malloced block this chunk lives in */ + asize_t alloc; /* in bytes, used for compaction */ + asize_t size; /* in bytes */ + char *next; +} heap_chunk_head; + +#define Chunk_size(c) (((heap_chunk_head *) (c)) [-1]).size +#define Chunk_alloc(c) (((heap_chunk_head *) (c)) [-1]).alloc +#define Chunk_next(c) (((heap_chunk_head *) (c)) [-1]).next +#define Chunk_block(c) (((heap_chunk_head *) (c)) [-1]).block + +extern int caml_gc_phase; +extern int caml_gc_subphase; +extern uintnat caml_allocated_words; +extern double caml_extra_heap_resources; +extern uintnat caml_dependent_size, caml_dependent_allocated; +extern uintnat caml_fl_wsz_at_phase_change; + +#define Phase_mark 0 +#define Phase_clean 1 +#define Phase_sweep 2 +#define Phase_idle 3 + +/* Subphase of mark */ +#define Subphase_mark_roots 10 +/* Subphase_mark_roots: At the end of this subphase all the global + roots are marked. */ +#define Subphase_mark_main 11 +/* Subphase_mark_main: At the end of this subphase all the value alive at + the start of this subphase and created during it are marked. */ +#define Subphase_mark_final 12 +/* Subphase_mark_final: At the start of this subphase register which + value with an ocaml finalizer are not marked, the associated + finalizer will be run later. So we mark now these values as alive, + since they must be available for their finalizer. + */ + +CAMLextern char *caml_heap_start; +extern uintnat total_heap_size; +extern char *caml_gc_sweep_hp; + +extern int caml_major_window; +extern double caml_major_ring[Max_major_window]; +extern int caml_major_ring_index; +extern double caml_major_work_credit; +extern double caml_gc_clock; + +/* [caml_major_gc_hook] is called just between the end of the mark + phase and the beginning of the sweep phase of the major GC */ +CAMLextern void (*caml_major_gc_hook)(void); + +void caml_init_major_heap (asize_t); /* size in bytes */ +asize_t caml_clip_heap_chunk_wsz (asize_t wsz); +void caml_darken (value, value *); +void caml_major_collection_slice (intnat); +void major_collection (void); +void caml_finish_major_cycle (void); +void caml_set_major_window (int); + +/* Forces finalisation of all heap-allocated values, + disregarding both local and global roots. + + Warning: finalisation is performed by means of forced sweeping, which may + result in pointers referencing nonexistent values; therefore the function + should only be used on runtime shutdown. +*/ +void caml_finalise_heap (void); + +#endif /* CAML_INTERNALiS */ + +#endif /* CAML_MAJOR_GC_H */ diff --git a/test/monniaux/ocaml/byterun/caml/md5.h b/test/monniaux/ocaml/byterun/caml/md5.h new file mode 100644 index 00000000..e83c16cd --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/md5.h @@ -0,0 +1,47 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1999 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* MD5 message digest */ + +#ifndef CAML_MD5_H +#define CAML_MD5_H + +#ifdef CAML_INTERNALS + +#include "mlvalues.h" +#include "io.h" + +CAMLextern value caml_md5_string (value str, value ofs, value len); +CAMLextern value caml_md5_chan (value vchan, value len); +CAMLextern void caml_md5_block(unsigned char digest[16], + void * data, uintnat len); + +CAMLextern value caml_md5_channel(struct channel *chan, intnat toread); + +struct MD5Context { + uint32_t buf[4]; + uint32_t bits[2]; + unsigned char in[64]; +}; + +CAMLextern void caml_MD5Init (struct MD5Context *context); +CAMLextern void caml_MD5Update (struct MD5Context *context, unsigned char *buf, + uintnat len); +CAMLextern void caml_MD5Final (unsigned char *digest, struct MD5Context *ctx); +CAMLextern void caml_MD5Transform (uint32_t *buf, uint32_t *in); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_MD5_H */ diff --git a/test/monniaux/ocaml/byterun/caml/memory.h b/test/monniaux/ocaml/byterun/caml/memory.h new file mode 100644 index 00000000..1ba489b2 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/memory.h @@ -0,0 +1,595 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Allocation macros and functions */ + +#ifndef CAML_MEMORY_H +#define CAML_MEMORY_H + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "config.h" +#ifdef CAML_INTERNALS +#include "gc.h" +#include "major_gc.h" +#include "minor_gc.h" +#endif /* CAML_INTERNALS */ +#include "misc.h" +#include "mlvalues.h" + +#ifdef __cplusplus +extern "C" { +#endif + + +CAMLextern value caml_alloc_shr (mlsize_t wosize, tag_t); +#ifdef WITH_PROFINFO +CAMLextern value caml_alloc_shr_with_profinfo (mlsize_t, tag_t, intnat); +CAMLextern value caml_alloc_shr_preserving_profinfo (mlsize_t, tag_t, + header_t); +#else +#define caml_alloc_shr_with_profinfo(size, tag, profinfo) \ + caml_alloc_shr(size, tag) +#define caml_alloc_shr_preserving_profinfo(size, tag, header) \ + caml_alloc_shr(size, tag) +#endif /* WITH_PROFINFO */ +CAMLextern value caml_alloc_shr_no_raise (mlsize_t wosize, tag_t); +CAMLextern void caml_adjust_gc_speed (mlsize_t, mlsize_t); +CAMLextern void caml_alloc_dependent_memory (mlsize_t bsz); +CAMLextern void caml_free_dependent_memory (mlsize_t bsz); +CAMLextern void caml_modify (value *, value); +CAMLextern void caml_initialize (value *, value); +CAMLextern value caml_check_urgent_gc (value); +CAMLextern int caml_init_alloc_for_heap (void); +CAMLextern char *caml_alloc_for_heap (asize_t request); /* Size in bytes. */ +CAMLextern void caml_free_for_heap (char *mem); +CAMLextern void caml_disown_for_heap (char *mem); +CAMLextern int caml_add_to_heap (char *mem); +CAMLextern color_t caml_allocation_color (void *hp); + +CAMLextern int caml_huge_fallback_count; + + +/* [caml_stat_*] functions below provide an interface to the static memory + manager built into the runtime, which can be used for managing static + (that is, non-moving) blocks of heap memory. + + Function arguments that have type [caml_stat_block] must always be pointers + to blocks returned by the [caml_stat_*] functions below. Attempting to use + these functions on memory blocks allocated by a different memory manager + (e.g. the one from the C runtime) will cause undefined behaviour. +*/ +typedef void* caml_stat_block; + +#ifdef CAML_INTERNALS + +/* The pool must be initialized with a call to [caml_stat_create_pool] + before it is possible to use any of the [caml_stat_*] functions below. + + If the pool is not initialized, [caml_stat_*] functions will still work in + backward compatibility mode, becoming thin wrappers around [malloc] family + of functions. In this case, calling [caml_stat_destroy_pool] will not free + the claimed heap memory, resulting in leaks. +*/ +CAMLextern void caml_stat_create_pool(void); + +/* [caml_stat_destroy_pool] frees all the heap memory claimed by the pool. + + Once the pool is destroyed, [caml_stat_*] functions will continue to work + in backward compatibility mode, becoming thin wrappers around [malloc] + family of functions. +*/ +CAMLextern void caml_stat_destroy_pool(void); + +#endif /* CAML_INTERNALS */ + +/* [caml_stat_alloc(size)] allocates a memory block of the requested [size] + (in bytes) and returns a pointer to it. It throws an OCaml exception in case + the request fails, and so requires the runtime lock to be held. +*/ +CAMLextern caml_stat_block caml_stat_alloc(asize_t); + +/* [caml_stat_alloc_noexc(size)] allocates a memory block of the requested [size] + (in bytes) and returns a pointer to it, or NULL in case the request fails. +*/ +CAMLextern caml_stat_block caml_stat_alloc_noexc(asize_t); + +/* [caml_stat_alloc_aligned(size, modulo, block*)] allocates a memory block of + the requested [size] (in bytes), the starting address of which is aligned to + the provided [modulo] value. The function returns the aligned address, as + well as the unaligned [block] (as an output parameter). It throws an OCaml + exception in case the request fails, and so requires the runtime lock. +*/ +CAMLextern void* caml_stat_alloc_aligned(asize_t, int modulo, caml_stat_block*); + +/* [caml_stat_alloc_aligned_noexc] is a variant of [caml_stat_alloc_aligned] + that returns NULL in case the request fails, and doesn't require the runtime + lock to be held. +*/ +CAMLextern void* caml_stat_alloc_aligned_noexc(asize_t, int modulo, + caml_stat_block*); + +/* [caml_stat_calloc_noexc(num, size)] allocates a block of memory for an array + of [num] elements, each of them [size] bytes long, and initializes all its + bits to zero, effectively allocating a zero-initialized memory block of + [num * size] bytes. It returns NULL in case the request fails. +*/ +CAMLextern caml_stat_block caml_stat_calloc_noexc(asize_t, asize_t); + +/* [caml_stat_free(block)] deallocates the provided [block]. */ +CAMLextern void caml_stat_free(caml_stat_block); + +/* [caml_stat_resize(block, size)] changes the size of the provided [block] to + [size] bytes. The function may move the memory block to a new location (whose + address is returned by the function). The content of the [block] is preserved + up to the smaller of the new and old sizes, even if the block is moved to a + new location. If the new size is larger, the value of the newly allocated + portion is indeterminate. The function throws an OCaml exception in case the + request fails, and so requires the runtime lock to be held. +*/ +CAMLextern caml_stat_block caml_stat_resize(caml_stat_block, asize_t); + +/* [caml_stat_resize_noexc] is a variant of [caml_stat_resize] that returns NULL + in case the request fails, and doesn't require the runtime lock. +*/ +CAMLextern caml_stat_block caml_stat_resize_noexc(caml_stat_block, asize_t); + + +/* A [caml_stat_block] containing a NULL-terminated string */ +typedef char* caml_stat_string; + +/* [caml_stat_strdup(s)] returns a pointer to a heap-allocated string which is a + copy of the NULL-terminated string [s]. It throws an OCaml exception in case + the request fails, and so requires the runtime lock to be held. +*/ +CAMLextern caml_stat_string caml_stat_strdup(const char *s); +#ifdef _WIN32 +CAMLextern wchar_t* caml_stat_wcsdup(const wchar_t *s); +#endif + +/* [caml_stat_strdup_noexc] is a variant of [caml_stat_strdup] that returns NULL + in case the request fails, and doesn't require the runtime lock. +*/ +CAMLextern caml_stat_string caml_stat_strdup_noexc(const char *s); + +/* [caml_stat_strconcat(nargs, strings)] concatenates NULL-terminated [strings] + (an array of [char*] of size [nargs]) into a new string, dropping all NULLs, + except for the very last one. It throws an OCaml exception in case the + request fails, and so requires the runtime lock to be held. +*/ +CAMLextern caml_stat_string caml_stat_strconcat(int n, ...); +#ifdef _WIN32 +CAMLextern wchar_t* caml_stat_wcsconcat(int n, ...); +#endif + + +/* void caml_shrink_heap (char *); Only used in compact.c */ + +#ifdef CAML_INTERNALS + +extern uintnat caml_use_huge_pages; + +#ifdef HAS_HUGE_PAGES +#include +#define Heap_page_size HUGE_PAGE_SIZE +#define Round_mmap_size(x) \ + (((x) + (Heap_page_size - 1)) & ~ (Heap_page_size - 1)) +#endif + + +int caml_page_table_add(int kind, void * start, void * end); +int caml_page_table_remove(int kind, void * start, void * end); +int caml_page_table_initialize(mlsize_t bytesize); + +#ifdef DEBUG +#define DEBUG_clear(result, wosize) do{ \ + uintnat caml__DEBUG_i; \ + for (caml__DEBUG_i = 0; caml__DEBUG_i < (wosize); ++ caml__DEBUG_i){ \ + Field ((result), caml__DEBUG_i) = Debug_uninit_minor; \ + } \ +}while(0) +#else +#define DEBUG_clear(result, wosize) +#endif + +#define Alloc_small_with_profinfo(result, wosize, tag, profinfo) do { \ + CAMLassert ((wosize) >= 1); \ + CAMLassert ((tag_t) (tag) < 256); \ + CAMLassert ((wosize) <= Max_young_wosize); \ + caml_young_ptr -= Whsize_wosize (wosize); \ + if (caml_young_ptr < caml_young_trigger){ \ + caml_young_ptr += Whsize_wosize (wosize); \ + CAML_INSTR_INT ("force_minor/alloc_small@", 1); \ + Setup_for_gc; \ + caml_gc_dispatch (); \ + Restore_after_gc; \ + caml_young_ptr -= Whsize_wosize (wosize); \ + } \ + Hd_hp (caml_young_ptr) = \ + Make_header_with_profinfo ((wosize), (tag), Caml_black, profinfo); \ + (result) = Val_hp (caml_young_ptr); \ + DEBUG_clear ((result), (wosize)); \ +}while(0) + +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) +extern uintnat caml_spacetime_my_profinfo(struct ext_table**, uintnat); +#define Alloc_small(result, wosize, tag) \ + Alloc_small_with_profinfo(result, wosize, tag, \ + caml_spacetime_my_profinfo(NULL, wosize)) +#else +#define Alloc_small(result, wosize, tag) \ + Alloc_small_with_profinfo(result, wosize, tag, (uintnat) 0) +#endif + +/* Deprecated alias for [caml_modify] */ + +#define Modify(fp,val) caml_modify((fp), (val)) + +#endif /* CAML_INTERNALS */ + +struct caml__roots_block { + struct caml__roots_block *next; + intnat ntables; + intnat nitems; + value *tables [5]; +}; + +CAMLextern struct caml__roots_block *caml_local_roots; /* defined in roots.c */ + +/* The following macros are used to declare C local variables and + function parameters of type [value]. + + The function body must start with one of the [CAMLparam] macros. + If the function has no parameter of type [value], use [CAMLparam0]. + If the function has 1 to 5 [value] parameters, use the corresponding + [CAMLparam] with the parameters as arguments. + If the function has more than 5 [value] parameters, use [CAMLparam5] + for the first 5 parameters, and one or more calls to the [CAMLxparam] + macros for the others. + If the function takes an array of [value]s as argument, use + [CAMLparamN] to declare it (or [CAMLxparamN] if you already have a + call to [CAMLparam] for some other arguments). + + If you need local variables of type [value], declare them with one + or more calls to the [CAMLlocal] macros at the beginning of the + function, after the call to CAMLparam. Use [CAMLlocalN] (at the + beginning of the function) to declare an array of [value]s. + + Your function may raise an exception or return a [value] with the + [CAMLreturn] macro. Its argument is simply the [value] returned by + your function. Do NOT directly return a [value] with the [return] + keyword. If your function returns void, use [CAMLreturn0]. If you + un-register the local roots (i.e. undo the effects of the [CAMLparam*] + and [CAMLlocal] macros) without returning immediately, use [CAMLdrop]. + + All the identifiers beginning with "caml__" are reserved by OCaml. + Do not use them for anything (local or global variables, struct or + union tags, macros, etc.) +*/ + +#define CAMLparam0() \ + struct caml__roots_block *caml__frame = caml_local_roots + +#define CAMLparam1(x) \ + CAMLparam0 (); \ + CAMLxparam1 (x) + +#define CAMLparam2(x, y) \ + CAMLparam0 (); \ + CAMLxparam2 (x, y) + +#define CAMLparam3(x, y, z) \ + CAMLparam0 (); \ + CAMLxparam3 (x, y, z) + +#define CAMLparam4(x, y, z, t) \ + CAMLparam0 (); \ + CAMLxparam4 (x, y, z, t) + +#define CAMLparam5(x, y, z, t, u) \ + CAMLparam0 (); \ + CAMLxparam5 (x, y, z, t, u) + +#define CAMLparamN(x, size) \ + CAMLparam0 (); \ + CAMLxparamN (x, (size)) + +/* CAMLunused is preserved for compatibility reasons. + Instead of the legacy GCC/Clang-only + CAMLunused foo; + you should prefer + CAMLunused_start foo CAMLunused_end; + which supports both GCC/Clang and MSVC. +*/ +#if defined(__GNUC__) && (__GNUC__ > 2 || (__GNUC__ == 2 && __GNUC_MINOR__ > 7)) + #define CAMLunused_start __attribute__ ((unused)) + #define CAMLunused_end + #define CAMLunused __attribute__ ((unused)) +#elif _MSC_VER >= 1500 + #define CAMLunused_start __pragma( warning (push) ) \ + __pragma( warning (disable:4189 ) ) + #define CAMLunused_end __pragma( warning (pop)) + #define CAMLunused +#else + #define CAMLunused_start + #define CAMLunused_end + #define CAMLunused +#endif + +#define CAMLxparam1(x) \ + struct caml__roots_block caml__roots_##x; \ + CAMLunused_start int caml__dummy_##x = ( \ + (void) caml__frame, \ + (caml__roots_##x.next = caml_local_roots), \ + (caml_local_roots = &caml__roots_##x), \ + (caml__roots_##x.nitems = 1), \ + (caml__roots_##x.ntables = 1), \ + (caml__roots_##x.tables [0] = &x), \ + 0) \ + CAMLunused_end + +#define CAMLxparam2(x, y) \ + struct caml__roots_block caml__roots_##x; \ + CAMLunused_start int caml__dummy_##x = ( \ + (void) caml__frame, \ + (caml__roots_##x.next = caml_local_roots), \ + (caml_local_roots = &caml__roots_##x), \ + (caml__roots_##x.nitems = 1), \ + (caml__roots_##x.ntables = 2), \ + (caml__roots_##x.tables [0] = &x), \ + (caml__roots_##x.tables [1] = &y), \ + 0) \ + CAMLunused_end + +#define CAMLxparam3(x, y, z) \ + struct caml__roots_block caml__roots_##x; \ + CAMLunused_start int caml__dummy_##x = ( \ + (void) caml__frame, \ + (caml__roots_##x.next = caml_local_roots), \ + (caml_local_roots = &caml__roots_##x), \ + (caml__roots_##x.nitems = 1), \ + (caml__roots_##x.ntables = 3), \ + (caml__roots_##x.tables [0] = &x), \ + (caml__roots_##x.tables [1] = &y), \ + (caml__roots_##x.tables [2] = &z), \ + 0) \ + CAMLunused_end + +#define CAMLxparam4(x, y, z, t) \ + struct caml__roots_block caml__roots_##x; \ + CAMLunused_start int caml__dummy_##x = ( \ + (void) caml__frame, \ + (caml__roots_##x.next = caml_local_roots), \ + (caml_local_roots = &caml__roots_##x), \ + (caml__roots_##x.nitems = 1), \ + (caml__roots_##x.ntables = 4), \ + (caml__roots_##x.tables [0] = &x), \ + (caml__roots_##x.tables [1] = &y), \ + (caml__roots_##x.tables [2] = &z), \ + (caml__roots_##x.tables [3] = &t), \ + 0) \ + CAMLunused_end + +#define CAMLxparam5(x, y, z, t, u) \ + struct caml__roots_block caml__roots_##x; \ + CAMLunused_start int caml__dummy_##x = ( \ + (void) caml__frame, \ + (caml__roots_##x.next = caml_local_roots), \ + (caml_local_roots = &caml__roots_##x), \ + (caml__roots_##x.nitems = 1), \ + (caml__roots_##x.ntables = 5), \ + (caml__roots_##x.tables [0] = &x), \ + (caml__roots_##x.tables [1] = &y), \ + (caml__roots_##x.tables [2] = &z), \ + (caml__roots_##x.tables [3] = &t), \ + (caml__roots_##x.tables [4] = &u), \ + 0) \ + CAMLunused_end + +#define CAMLxparamN(x, size) \ + struct caml__roots_block caml__roots_##x; \ + CAMLunused_start int caml__dummy_##x = ( \ + (void) caml__frame, \ + (caml__roots_##x.next = caml_local_roots), \ + (caml_local_roots = &caml__roots_##x), \ + (caml__roots_##x.nitems = (size)), \ + (caml__roots_##x.ntables = 1), \ + (caml__roots_##x.tables[0] = &(x[0])), \ + 0) \ + CAMLunused_end + +#define CAMLlocal1(x) \ + value x = Val_unit; \ + CAMLxparam1 (x) + +#define CAMLlocal2(x, y) \ + value x = Val_unit, y = Val_unit; \ + CAMLxparam2 (x, y) + +#define CAMLlocal3(x, y, z) \ + value x = Val_unit, y = Val_unit, z = Val_unit; \ + CAMLxparam3 (x, y, z) + +#define CAMLlocal4(x, y, z, t) \ + value x = Val_unit, y = Val_unit, z = Val_unit, t = Val_unit; \ + CAMLxparam4 (x, y, z, t) + +#define CAMLlocal5(x, y, z, t, u) \ + value x = Val_unit, y = Val_unit, z = Val_unit, t = Val_unit, u = Val_unit; \ + CAMLxparam5 (x, y, z, t, u) + +#define CAMLlocalN(x, size) \ + value x [(size)]; \ + int caml__i_##x; \ + for (caml__i_##x = 0; caml__i_##x < size; caml__i_##x ++) { \ + x[caml__i_##x] = Val_unit; \ + } \ + CAMLxparamN (x, (size)) + + +#define CAMLdrop caml_local_roots = caml__frame + +#define CAMLreturn0 do{ \ + CAMLdrop; \ + return; \ +}while (0) + +#define CAMLreturnT(type, result) do{ \ + type caml__temp_result = (result); \ + CAMLdrop; \ + return caml__temp_result; \ +}while(0) + +#define CAMLreturn(result) CAMLreturnT(value, result) + +#define CAMLnoreturn ((void) caml__frame) + + +/* convenience macro */ +#define Store_field(block, offset, val) do{ \ + mlsize_t caml__temp_offset = (offset); \ + value caml__temp_val = (val); \ + caml_modify (&Field ((block), caml__temp_offset), caml__temp_val); \ +}while(0) + +/* + NOTE: [Begin_roots] and [End_roots] are superseded by [CAMLparam]*, + [CAMLxparam]*, [CAMLlocal]*, [CAMLreturn]. + + [Begin_roots] and [End_roots] are used for C variables that are GC roots. + It must contain all values in C local variables and function parameters + at the time the minor GC is called. + Usage: + After initialising your local variables to legal OCaml values, but before + calling allocation functions, insert [Begin_roots_n(v1, ... vn)], where + v1 ... vn are your variables of type [value] that you want to be updated + across allocations. + At the end, insert [End_roots()]. + + Note that [Begin_roots] opens a new block, and [End_roots] closes it. + Thus they must occur in matching pairs at the same brace nesting level. + + You can use [Val_unit] as a dummy initial value for your variables. +*/ + +#define Begin_root Begin_roots1 + +#define Begin_roots1(r0) { \ + struct caml__roots_block caml__roots_block; \ + caml__roots_block.next = caml_local_roots; \ + caml_local_roots = &caml__roots_block; \ + caml__roots_block.nitems = 1; \ + caml__roots_block.ntables = 1; \ + caml__roots_block.tables[0] = &(r0); + +#define Begin_roots2(r0, r1) { \ + struct caml__roots_block caml__roots_block; \ + caml__roots_block.next = caml_local_roots; \ + caml_local_roots = &caml__roots_block; \ + caml__roots_block.nitems = 1; \ + caml__roots_block.ntables = 2; \ + caml__roots_block.tables[0] = &(r0); \ + caml__roots_block.tables[1] = &(r1); + +#define Begin_roots3(r0, r1, r2) { \ + struct caml__roots_block caml__roots_block; \ + caml__roots_block.next = caml_local_roots; \ + caml_local_roots = &caml__roots_block; \ + caml__roots_block.nitems = 1; \ + caml__roots_block.ntables = 3; \ + caml__roots_block.tables[0] = &(r0); \ + caml__roots_block.tables[1] = &(r1); \ + caml__roots_block.tables[2] = &(r2); + +#define Begin_roots4(r0, r1, r2, r3) { \ + struct caml__roots_block caml__roots_block; \ + caml__roots_block.next = caml_local_roots; \ + caml_local_roots = &caml__roots_block; \ + caml__roots_block.nitems = 1; \ + caml__roots_block.ntables = 4; \ + caml__roots_block.tables[0] = &(r0); \ + caml__roots_block.tables[1] = &(r1); \ + caml__roots_block.tables[2] = &(r2); \ + caml__roots_block.tables[3] = &(r3); + +#define Begin_roots5(r0, r1, r2, r3, r4) { \ + struct caml__roots_block caml__roots_block; \ + caml__roots_block.next = caml_local_roots; \ + caml_local_roots = &caml__roots_block; \ + caml__roots_block.nitems = 1; \ + caml__roots_block.ntables = 5; \ + caml__roots_block.tables[0] = &(r0); \ + caml__roots_block.tables[1] = &(r1); \ + caml__roots_block.tables[2] = &(r2); \ + caml__roots_block.tables[3] = &(r3); \ + caml__roots_block.tables[4] = &(r4); + +#define Begin_roots_block(table, size) { \ + struct caml__roots_block caml__roots_block; \ + caml__roots_block.next = caml_local_roots; \ + caml_local_roots = &caml__roots_block; \ + caml__roots_block.nitems = (size); \ + caml__roots_block.ntables = 1; \ + caml__roots_block.tables[0] = (table); + +#define End_roots() caml_local_roots = caml__roots_block.next; } + + +/* [caml_register_global_root] registers a global C variable as a memory root + for the duration of the program, or until [caml_remove_global_root] is + called. */ + +CAMLextern void caml_register_global_root (value *); + +/* [caml_remove_global_root] removes a memory root registered on a global C + variable with [caml_register_global_root]. */ + +CAMLextern void caml_remove_global_root (value *); + +/* [caml_register_generational_global_root] registers a global C + variable as a memory root for the duration of the program, or until + [caml_remove_generational_global_root] is called. + The program guarantees that the value contained in this variable + will not be assigned directly. If the program needs to change + the value of this variable, it must do so by calling + [caml_modify_generational_global_root]. The [value *] pointer + passed to [caml_register_generational_global_root] must contain + a valid OCaml value before the call. + In return for these constraints, scanning of memory roots during + minor collection is made more efficient. */ + +CAMLextern void caml_register_generational_global_root (value *); + +/* [caml_remove_generational_global_root] removes a memory root + registered on a global C variable with + [caml_register_generational_global_root]. */ + +CAMLextern void caml_remove_generational_global_root (value *); + +/* [caml_modify_generational_global_root(r, newval)] + modifies the value contained in [r], storing [newval] inside. + In other words, the assignment [*r = newval] is performed, + but in a way that is compatible with the optimized scanning of + generational global roots. [r] must be a global memory root + previously registered with [caml_register_generational_global_root]. */ + +CAMLextern void caml_modify_generational_global_root(value *r, value newval); + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_MEMORY_H */ diff --git a/test/monniaux/ocaml/byterun/caml/minor_gc.h b/test/monniaux/ocaml/byterun/caml/minor_gc.h new file mode 100644 index 00000000..6c48c761 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/minor_gc.h @@ -0,0 +1,120 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_MINOR_GC_H +#define CAML_MINOR_GC_H + + +#include "address_class.h" +#include "config.h" + +CAMLextern value *caml_young_start, *caml_young_end; +CAMLextern value *caml_young_alloc_start, *caml_young_alloc_end; +CAMLextern value *caml_young_ptr, *caml_young_limit; +CAMLextern value *caml_young_trigger; +extern asize_t caml_minor_heap_wsz; +extern int caml_in_minor_collection; +extern double caml_extra_heap_resources_minor; + +#define CAML_TABLE_STRUCT(t) { \ + t *base; \ + t *end; \ + t *threshold; \ + t *ptr; \ + t *limit; \ + asize_t size; \ + asize_t reserve; \ +} + +struct caml_ref_table CAML_TABLE_STRUCT(value *); +CAMLextern struct caml_ref_table caml_ref_table; + +struct caml_ephe_ref_elt { + value ephe; /* an ephemeron in major heap */ + mlsize_t offset; /* the offset that points in the minor heap */ +}; + +struct caml_ephe_ref_table CAML_TABLE_STRUCT(struct caml_ephe_ref_elt); +CAMLextern struct caml_ephe_ref_table caml_ephe_ref_table; + +struct caml_custom_elt { + value block; /* The finalized block in the minor heap. */ + mlsize_t mem; /* The parameters for adjusting GC speed. */ + mlsize_t max; +}; + +struct caml_custom_table CAML_TABLE_STRUCT(struct caml_custom_elt); +CAMLextern struct caml_custom_table caml_custom_table; + +extern void caml_set_minor_heap_size (asize_t); /* size in bytes */ +extern void caml_empty_minor_heap (void); +CAMLextern void caml_gc_dispatch (void); +CAMLextern void garbage_collection (void); /* def in asmrun/signals_asm.c */ +extern void caml_realloc_ref_table (struct caml_ref_table *); +extern void caml_alloc_table (struct caml_ref_table *, asize_t, asize_t); +extern void caml_realloc_ephe_ref_table (struct caml_ephe_ref_table *); +extern void caml_alloc_ephe_table (struct caml_ephe_ref_table *, + asize_t, asize_t); +extern void caml_realloc_custom_table (struct caml_custom_table *); +extern void caml_alloc_custom_table (struct caml_custom_table *, + asize_t, asize_t); +extern void caml_oldify_one (value, value *); +extern void caml_oldify_mopup (void); + +#define Oldify(p) do{ \ + value __oldify__v__ = *p; \ + if (Is_block (__oldify__v__) && Is_young (__oldify__v__)){ \ + caml_oldify_one (__oldify__v__, (p)); \ + } \ + }while(0) + +static inline void add_to_ref_table (struct caml_ref_table *tbl, value *p) +{ + if (tbl->ptr >= tbl->limit){ + CAMLassert (tbl->ptr == tbl->limit); + caml_realloc_ref_table (tbl); + } + *tbl->ptr++ = p; +} + +static inline void add_to_ephe_ref_table (struct caml_ephe_ref_table *tbl, + value ar, mlsize_t offset) +{ + struct caml_ephe_ref_elt *ephe_ref; + if (tbl->ptr >= tbl->limit){ + CAMLassert (tbl->ptr == tbl->limit); + caml_realloc_ephe_ref_table (tbl); + } + ephe_ref = tbl->ptr++; + ephe_ref->ephe = ar; + ephe_ref->offset = offset; + CAMLassert(ephe_ref->offset < Wosize_val(ephe_ref->ephe)); +} + +static inline void add_to_custom_table (struct caml_custom_table *tbl, value v, + mlsize_t mem, mlsize_t max) +{ + struct caml_custom_elt *elt; + if (tbl->ptr >= tbl->limit){ + CAMLassert (tbl->ptr == tbl->limit); + caml_realloc_custom_table (tbl); + } + elt = tbl->ptr++; + elt->block = v; + elt->mem = mem; + elt->max = max; +} + +#endif /* CAML_MINOR_GC_H */ diff --git a/test/monniaux/ocaml/byterun/caml/misc.h b/test/monniaux/ocaml/byterun/caml/misc.h new file mode 100644 index 00000000..5d668e56 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/misc.h @@ -0,0 +1,525 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Miscellaneous macros and variables. */ + +#ifndef CAML_MISC_H +#define CAML_MISC_H + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "config.h" + +/* Standard definitions */ + +#include +#include + +/* Basic types and constants */ + +typedef size_t asize_t; + +#ifndef NULL +#define NULL 0 +#endif + +#ifdef CAML_INTERNALS +typedef char * addr; +#endif /* CAML_INTERNALS */ + +/* Noreturn is preserved for compatibility reasons. + Instead of the legacy GCC/Clang-only + foo Noreturn; + you should prefer + CAMLnoreturn_start foo CAMLnoreturn_end; + which supports both GCC/Clang and MSVC. + + Note: CAMLnoreturn is a different macro defined in memory.h, + to be used in function bodies rather than as a prototype attribute. +*/ +#ifdef __GNUC__ + /* Works only in GCC 2.5 and later */ + #define CAMLnoreturn_start + #define CAMLnoreturn_end __attribute__ ((noreturn)) + #define Noreturn __attribute__ ((noreturn)) +#elif _MSC_VER >= 1500 + #define CAMLnoreturn_start __declspec(noreturn) + #define CAMLnoreturn_end + #define Noreturn +#else + #define CAMLnoreturn_start + #define CAMLnoreturn_end + #define Noreturn +#endif + + + +/* Export control (to mark primitives and to handle Windows DLL) */ + +#define CAMLexport +#define CAMLprim +#define CAMLextern extern + +/* Weak function definitions that can be overridden by external libs */ +/* Conservatively restricted to ELF and MacOSX platforms */ +#if defined(__GNUC__) && (defined (__ELF__) || defined(__APPLE__)) +#define CAMLweakdef __attribute__((weak)) +#else +#define CAMLweakdef +#endif + +#ifdef __cplusplus +extern "C" { +#endif + +/* GC timing hooks. These can be assigned by the user. + [caml_minor_gc_begin_hook] must not allocate nor change any heap value. + The others can allocate and even call back to OCaml code. +*/ +typedef void (*caml_timing_hook) (void); +extern caml_timing_hook caml_major_slice_begin_hook, caml_major_slice_end_hook; +extern caml_timing_hook caml_minor_gc_begin_hook, caml_minor_gc_end_hook; +extern caml_timing_hook caml_finalise_begin_hook, caml_finalise_end_hook; + +/* Assertions */ + +#ifdef DEBUG +#define CAMLassert(x) \ + ((x) ? (void) 0 : (void) caml_failed_assert ( #x , __FILE__, __LINE__)) +CAMLnoreturn_start +CAMLextern int caml_failed_assert (char *, char *, int) +CAMLnoreturn_end; +#else +#define CAMLassert(x) ((void) 0) +#endif + +CAMLnoreturn_start +CAMLextern void caml_fatal_error (char *msg) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_fatal_error_arg (char *fmt, char *arg) +CAMLnoreturn_end; + +CAMLnoreturn_start +CAMLextern void caml_fatal_error_arg2 (char *fmt1, char *arg1, + char *fmt2, char *arg2) +CAMLnoreturn_end; + +/* Detection of available C built-in functions, the Clang way. */ + +#ifdef __has_builtin +#define Caml_has_builtin(x) __has_builtin(x) +#else +#define Caml_has_builtin(x) 0 +#endif + +/* Integer arithmetic with overflow detection. + The functions return 0 if no overflow, 1 if overflow. + The result of the operation is always stored at [*res]. + If no overflow is reported, this is the exact result. + If overflow is reported, this is the exact result modulo 2 to the word size. +*/ + +static inline int caml_uadd_overflow(uintnat a, uintnat b, uintnat * res) +{ +#if __GNUC__ >= 5 || Caml_has_builtin(__builtin_add_overflow) + return __builtin_add_overflow(a, b, res); +#else + uintnat c = a + b; + *res = c; + return c < a; +#endif +} + +static inline int caml_usub_overflow(uintnat a, uintnat b, uintnat * res) +{ +#if __GNUC__ >= 5 || Caml_has_builtin(__builtin_sub_overflow) + return __builtin_sub_overflow(a, b, res); +#else + uintnat c = a - b; + *res = c; + return a < b; +#endif +} + +#if __GNUC__ >= 5 || Caml_has_builtin(__builtin_mul_overflow) +static inline int caml_umul_overflow(uintnat a, uintnat b, uintnat * res) +{ + return __builtin_mul_overflow(a, b, res); +} +#else +extern int caml_umul_overflow(uintnat a, uintnat b, uintnat * res); +#endif + +/* Windows Unicode support */ + +#ifdef _WIN32 + +typedef wchar_t char_os; + +#define _T(x) L ## x + +#define access_os _waccess +#define open_os _wopen +#define stat_os _wstati64 +#define unlink_os _wunlink +#define rename_os caml_win32_rename +#define chdir_os _wchdir +#define getcwd_os _wgetcwd +#define system_os _wsystem +#define rmdir_os _wrmdir +#define putenv_os _wputenv +#define chmod_os _wchmod +#define execv_os _wexecv +#define execve_os _wexecve +#define execvp_os _wexecvp +#define execvpe_os _wexecvpe +#define strcmp_os wcscmp +#define strlen_os wcslen +#define sscanf_os swscanf + +#define caml_stat_strdup_os caml_stat_wcsdup +#define caml_stat_strconcat_os caml_stat_wcsconcat + +#define caml_stat_strdup_to_os caml_stat_strdup_to_utf16 +#define caml_stat_strdup_of_os caml_stat_strdup_of_utf16 +#define caml_copy_string_of_os caml_copy_string_of_utf16 + +#else /* _WIN32 */ + +typedef char char_os; + +#define _T(x) x + +#define access_os access +#define open_os open +#define stat_os stat +#define unlink_os unlink +#define rename_os rename +#define chdir_os chdir +#define getcwd_os getcwd +#define system_os system +#define rmdir_os rmdir +#define putenv_os putenv +#define chmod_os chmod +#define execv_os execv +#define execve_os execve +#define execvp_os execvp +#define execvpe_os execvpe +#define strcmp_os strcmp +#define strlen_os strlen +#define sscanf_os sscanf + +#define caml_stat_strdup_os caml_stat_strdup +#define caml_stat_strconcat_os caml_stat_strconcat + +#define caml_stat_strdup_to_os caml_stat_strdup +#define caml_stat_strdup_of_os caml_stat_strdup +#define caml_copy_string_of_os caml_copy_string + +#endif /* _WIN32 */ + + +/* Use macros for some system calls being called from OCaml itself. + These calls can be either traced for security reasons, or changed to + virtualize the program. */ + + +#ifndef CAML_WITH_CPLUGINS + +#define CAML_SYS_EXIT(retcode) exit(retcode) +#define CAML_SYS_OPEN(filename,flags,perm) open_os(filename,flags,perm) +#define CAML_SYS_CLOSE(fd) close(fd) +#define CAML_SYS_STAT(filename,st) stat_os(filename,st) +#define CAML_SYS_UNLINK(filename) unlink_os(filename) +#define CAML_SYS_RENAME(old_name,new_name) rename_os(old_name, new_name) +#define CAML_SYS_CHDIR(dirname) chdir_os(dirname) +#define CAML_SYS_GETENV(varname) getenv(varname) +#define CAML_SYS_SYSTEM(command) system_os(command) +#define CAML_SYS_READ_DIRECTORY(dirname,tbl) caml_read_directory(dirname,tbl) + +#else + + +#define CAML_CPLUGINS_EXIT 0 +#define CAML_CPLUGINS_OPEN 1 +#define CAML_CPLUGINS_CLOSE 2 +#define CAML_CPLUGINS_STAT 3 +#define CAML_CPLUGINS_UNLINK 4 +#define CAML_CPLUGINS_RENAME 5 +#define CAML_CPLUGINS_CHDIR 6 +#define CAML_CPLUGINS_GETENV 7 +#define CAML_CPLUGINS_SYSTEM 8 +#define CAML_CPLUGINS_READ_DIRECTORY 9 +#define CAML_CPLUGINS_PRIMS_MAX 9 + +#define CAML_CPLUGINS_PRIMS_BITMAP ((1 << CAML_CPLUGINS_PRIMS_MAX)-1) + +extern intnat (*caml_cplugins_prim)(int,intnat,intnat,intnat); + +#define CAML_SYS_PRIM_1(code,prim,arg1) \ + (caml_cplugins_prim == NULL) ? prim(arg1) : \ + caml_cplugins_prim(code,(intnat) (arg1),0,0) +#define CAML_SYS_STRING_PRIM_1(code,prim,arg1) \ + (caml_cplugins_prim == NULL) ? prim(arg1) : \ + (char_os*)caml_cplugins_prim(code,(intnat) (arg1),0,0) +#define CAML_SYS_VOID_PRIM_1(code,prim,arg1) \ + (caml_cplugins_prim == NULL) ? prim(arg1) : \ + (void)caml_cplugins_prim(code,(intnat) (arg1),0,0) +#define CAML_SYS_PRIM_2(code,prim,arg1,arg2) \ + (caml_cplugins_prim == NULL) ? prim(arg1,arg2) : \ + caml_cplugins_prim(code,(intnat) (arg1), (intnat) (arg2),0) +#define CAML_SYS_PRIM_3(code,prim,arg1,arg2,arg3) \ + (caml_cplugins_prim == NULL) ? prim(arg1,arg2,arg3) : \ + caml_cplugins_prim(code,(intnat) (arg1), (intnat) (arg2),(intnat) (arg3)) + +#define CAML_SYS_EXIT(retcode) \ + CAML_SYS_VOID_PRIM_1(CAML_CPLUGINS_EXIT,exit,retcode) +#define CAML_SYS_OPEN(filename,flags,perm) \ + CAML_SYS_PRIM_3(CAML_CPLUGINS_OPEN,open_os,filename,flags,perm) +#define CAML_SYS_CLOSE(fd) \ + CAML_SYS_PRIM_1(CAML_CPLUGINS_CLOSE,close,fd) +#define CAML_SYS_STAT(filename,st) \ + CAML_SYS_PRIM_2(CAML_CPLUGINS_STAT,stat_os,filename,st) +#define CAML_SYS_UNLINK(filename) \ + CAML_SYS_PRIM_1(CAML_CPLUGINS_UNLINK,unlink_os,filename) +#define CAML_SYS_RENAME(old_name,new_name) \ + CAML_SYS_PRIM_2(CAML_CPLUGINS_RENAME,rename_os,old_name,new_name) +#define CAML_SYS_CHDIR(dirname) \ + CAML_SYS_PRIM_1(CAML_CPLUGINS_CHDIR,chdir_os,dirname) +#define CAML_SYS_GETENV(varname) \ + CAML_SYS_STRING_PRIM_1(CAML_CPLUGINS_GETENV,getenv,varname) +#define CAML_SYS_SYSTEM(command) \ + CAML_SYS_PRIM_1(CAML_CPLUGINS_SYSTEM,system_os,command) +#define CAML_SYS_READ_DIRECTORY(dirname,tbl) \ + CAML_SYS_PRIM_2(CAML_CPLUGINS_READ_DIRECTORY,caml_read_directory, \ + dirname,tbl) + +#define CAML_CPLUGIN_CONTEXT_API 0 + +struct cplugin_context { + int api_version; + int prims_bitmap; + char_os *exe_name; + char_os** argv; + char_os *plugin; /* absolute filename of plugin, do a copy if you need it ! */ + char *ocaml_version; +/* end of CAML_CPLUGIN_CONTEXT_API version 0 */ +}; + +extern void caml_cplugins_init(char_os * exe_name, char_os **argv); + +/* A plugin MUST define a symbol "caml_cplugin_init" with the prototype: + +void caml_cplugin_init(struct cplugin_context *ctx) +*/ + +/* to write plugins for CAML_SYS_READ_DIRECTORY, we will need the + definition of struct ext_table to be public. */ + +#endif /* CAML_WITH_CPLUGINS */ + +/* Data structures */ + +struct ext_table { + int size; + int capacity; + void ** contents; +}; + +extern void caml_ext_table_init(struct ext_table * tbl, int init_capa); +extern int caml_ext_table_add(struct ext_table * tbl, void * data); +extern void caml_ext_table_remove(struct ext_table * tbl, void * data); +extern void caml_ext_table_free(struct ext_table * tbl, int free_entries); +extern void caml_ext_table_clear(struct ext_table * tbl, int free_entries); + +CAMLextern int caml_read_directory(char_os * dirname, struct ext_table * contents); + +/* Deprecated aliases */ +#define caml_aligned_malloc caml_stat_alloc_aligned_noexc +#define caml_strdup caml_stat_strdup +#define caml_strconcat caml_stat_strconcat + +#ifdef CAML_INTERNALS + +/* GC flags and messages */ + +extern uintnat caml_verb_gc; +void caml_gc_message (int, char *, ...) +#ifdef __GNUC__ + __attribute__ ((format (printf, 2, 3))) +#endif +; + +/* Runtime warnings */ +extern uintnat caml_runtime_warnings; +int caml_runtime_warnings_active(void); + +#ifdef DEBUG +#ifdef ARCH_SIXTYFOUR +#define Debug_tag(x) (0xD700D7D7D700D6D7ul \ + | ((uintnat) (x) << 16) \ + | ((uintnat) (x) << 48)) +#else +#define Debug_tag(x) (0xD700D6D7ul | ((uintnat) (x) << 16)) +#endif /* ARCH_SIXTYFOUR */ + +/* + 00 -> free words in minor heap + 01 -> fields of free list blocks in major heap + 03 -> heap chunks deallocated by heap shrinking + 04 -> fields deallocated by [caml_obj_truncate] + 10 -> uninitialised fields of minor objects + 11 -> uninitialised fields of major objects + 15 -> uninitialised words of [caml_stat_alloc_aligned] blocks + 85 -> filler bytes of [caml_stat_alloc_aligned] + 99 -> the magic prefix of a memory block allocated by [caml_stat_alloc] + + special case (byte by byte): + D7 -> uninitialised words of [caml_stat_alloc] blocks +*/ +#define Debug_free_minor Debug_tag (0x00) +#define Debug_free_major Debug_tag (0x01) +#define Debug_free_shrink Debug_tag (0x03) +#define Debug_free_truncate Debug_tag (0x04) +#define Debug_uninit_minor Debug_tag (0x10) +#define Debug_uninit_major Debug_tag (0x11) +#define Debug_uninit_align Debug_tag (0x15) +#define Debug_filler_align Debug_tag (0x85) +#define Debug_pool_magic Debug_tag (0x99) + +#define Debug_uninit_stat 0xD7 + +/* Note: the first argument is in fact a [value] but we don't have this + type available yet because we can't include [mlvalues.h] in this file. +*/ +extern void caml_set_fields (intnat v, unsigned long, unsigned long); +#endif /* DEBUG */ + + +/* snprintf emulation for Win32 */ + +#if defined(_WIN32) && !defined(_UCRT) +extern int caml_snprintf(char * buf, size_t size, const char * format, ...); +#define snprintf caml_snprintf +#endif + +#ifdef CAML_INSTR +/* Timers and counters for GC latency profiling (Linux-only) */ + +#include +#include + +extern intnat caml_stat_minor_collections; +extern intnat CAML_INSTR_STARTTIME, CAML_INSTR_STOPTIME; + +struct CAML_INSTR_BLOCK { + struct timespec ts[10]; + char *tag[10]; + int index; + struct CAML_INSTR_BLOCK *next; +}; + +extern struct CAML_INSTR_BLOCK *CAML_INSTR_LOG; + +/* Declare a timer/counter name. [t] must be a new variable name. */ +#define CAML_INSTR_DECLARE(t) \ + struct CAML_INSTR_BLOCK *t = NULL + +/* Allocate the data block for a given name. + [t] must have been declared with [CAML_INSTR_DECLARE]. */ +#define CAML_INSTR_ALLOC(t) do{ \ + if (caml_stat_minor_collections >= CAML_INSTR_STARTTIME \ + && caml_stat_minor_collections < CAML_INSTR_STOPTIME){ \ + t = caml_stat_alloc_noexc (sizeof (struct CAML_INSTR_BLOCK)); \ + t->index = 0; \ + t->tag[0] = ""; \ + t->next = CAML_INSTR_LOG; \ + CAML_INSTR_LOG = t; \ + } \ + }while(0) + +/* Allocate the data block and start the timer. + [t] must have been declared with [CAML_INSTR_DECLARE] + and allocated with [CAML_INSTR_ALLOC]. */ +#define CAML_INSTR_START(t, msg) do{ \ + if (t != NULL){ \ + t->tag[0] = msg; \ + clock_gettime (CLOCK_REALTIME, &(t->ts[0])); \ + } \ + }while(0) + +/* Declare a timer, allocate its data, and start it. + [t] must be a new variable name. */ +#define CAML_INSTR_SETUP(t, msg) \ + CAML_INSTR_DECLARE (t); \ + CAML_INSTR_ALLOC (t); \ + CAML_INSTR_START (t, msg) + +/* Record an intermediate time within a given timer. + [t] must have been declared, allocated, and started. */ +#define CAML_INSTR_TIME(t, msg) do{ \ + if (t != NULL){ \ + ++ t->index; \ + t->tag[t->index] = (msg); \ + clock_gettime (CLOCK_REALTIME, &(t->ts[t->index])); \ + } \ + }while(0) + +/* Record an integer data point. + If [msg] ends with # it will be interpreted as an integer-valued event. + If it ends with @ it will be interpreted as an event counter. +*/ +#define CAML_INSTR_INT(msg, data) do{ \ + CAML_INSTR_SETUP (__caml_tmp, ""); \ + if (__caml_tmp != NULL){ \ + __caml_tmp->index = 1; \ + __caml_tmp->tag[1] = msg; \ + __caml_tmp->ts[1].tv_sec = 0; \ + __caml_tmp->ts[1].tv_nsec = (data); \ + } \ + }while(0) + +/* This function is called at the start of the program to set up + the data for the above macros. +*/ +extern void CAML_INSTR_INIT (void); + +/* This function is automatically called by the runtime to output + the collected data to the dump file. */ +extern void CAML_INSTR_ATEXIT (void); + +#else /* CAML_INSTR */ + +#define CAML_INSTR_DECLARE(t) /**/ +#define CAML_INSTR_ALLOC(t) /**/ +#define CAML_INSTR_START(t, name) /**/ +#define CAML_INSTR_SETUP(t, name) /**/ +#define CAML_INSTR_TIME(t, msg) /**/ +#define CAML_INSTR_INT(msg, c) /**/ +#define CAML_INSTR_INIT() /**/ +#define CAML_INSTR_ATEXIT() /**/ + +#endif /* CAML_INSTR */ + +#endif /* CAML_INTERNALS */ + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_MISC_H */ diff --git a/test/monniaux/ocaml/byterun/caml/mlvalues.h b/test/monniaux/ocaml/byterun/caml/mlvalues.h new file mode 100644 index 00000000..ec30b20a --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/mlvalues.h @@ -0,0 +1,377 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_MLVALUES_H +#define CAML_MLVALUES_H + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "config.h" +#include "misc.h" + +#ifdef __cplusplus +extern "C" { +#endif + +/* Definitions + + word: Four bytes on 32 and 16 bit architectures, + eight bytes on 64 bit architectures. + long: A C integer having the same number of bytes as a word. + val: The ML representation of something. A long or a block or a pointer + outside the heap. If it is a block, it is the (encoded) address + of an object. If it is a long, it is encoded as well. + block: Something allocated. It always has a header and some + fields or some number of bytes (a multiple of the word size). + field: A word-sized val which is part of a block. + bp: Pointer to the first byte of a block. (a char *) + op: Pointer to the first field of a block. (a value *) + hp: Pointer to the header of a block. (a char *) + int32_t: Four bytes on all architectures. + int64_t: Eight bytes on all architectures. + + Remark: A block size is always a multiple of the word size, and at least + one word plus the header. + + bosize: Size (in bytes) of the "bytes" part. + wosize: Size (in words) of the "fields" part. + bhsize: Size (in bytes) of the block with its header. + whsize: Size (in words) of the block with its header. + + hd: A header. + tag: The value of the tag field of the header. + color: The value of the color field of the header. + This is for use only by the GC. +*/ + +typedef intnat value; +typedef uintnat header_t; +typedef uintnat mlsize_t; +typedef unsigned int tag_t; /* Actually, an unsigned char */ +typedef uintnat color_t; +typedef uintnat mark_t; + +/* Longs vs blocks. */ +#define Is_long(x) (((x) & 1) != 0) +#define Is_block(x) (((x) & 1) == 0) + +/* Conversion macro names are always of the form "to_from". */ +/* Example: Val_long as in "Val from long" or "Val of long". */ +#define Val_long(x) ((intnat) (((uintnat)(x) << 1)) + 1) +#define Long_val(x) ((x) >> 1) +#define Max_long (((intnat)1 << (8 * sizeof(value) - 2)) - 1) +#define Min_long (-((intnat)1 << (8 * sizeof(value) - 2))) +#define Val_int(x) Val_long(x) +#define Int_val(x) ((int) Long_val(x)) +#define Unsigned_long_val(x) ((uintnat)(x) >> 1) +#define Unsigned_int_val(x) ((int) Unsigned_long_val(x)) + +/* Structure of the header: + +For 16-bit and 32-bit architectures: + +--------+-------+-----+ + | wosize | color | tag | + +--------+-------+-----+ +bits 31 10 9 8 7 0 + +For 64-bit architectures: + + +--------+-------+-----+ + | wosize | color | tag | + +--------+-------+-----+ +bits 63 10 9 8 7 0 + +For x86-64 with Spacetime profiling: + P = PROFINFO_WIDTH (as set by "configure", currently 26 bits, giving a + maximum block size of just under 4Gb) + +----------------+----------------+-------------+ + | profiling info | wosize | color | tag | + +----------------+----------------+-------------+ +bits 63 (64-P) (63-P) 10 9 8 7 0 + +*/ + +#define Tag_hd(hd) ((tag_t) ((hd) & 0xFF)) + +#define Gen_profinfo_shift(width) (64 - (width)) +#define Gen_profinfo_mask(width) ((1ull << (width)) - 1ull) +#define Gen_profinfo_hd(width, hd) \ + (((mlsize_t) ((hd) >> (Gen_profinfo_shift(width)))) \ + & (Gen_profinfo_mask(width))) + +#ifdef WITH_PROFINFO +#define PROFINFO_SHIFT (Gen_profinfo_shift(PROFINFO_WIDTH)) +#define PROFINFO_MASK (Gen_profinfo_mask(PROFINFO_WIDTH)) +#define Hd_no_profinfo(hd) ((hd) & ~(PROFINFO_MASK << PROFINFO_SHIFT)) +#define Wosize_hd(hd) ((mlsize_t) ((Hd_no_profinfo(hd)) >> 10)) +#define Profinfo_hd(hd) (Gen_profinfo_hd(PROFINFO_WIDTH, hd)) +#else +#define Wosize_hd(hd) ((mlsize_t) ((hd) >> 10)) +#endif /* WITH_PROFINFO */ + +#define Hd_val(val) (((header_t *) (val)) [-1]) /* Also an l-value. */ +#define Hd_op(op) (Hd_val (op)) /* Also an l-value. */ +#define Hd_bp(bp) (Hd_val (bp)) /* Also an l-value. */ +#define Hd_hp(hp) (* ((header_t *) (hp))) /* Also an l-value. */ +#define Hp_val(val) (((header_t *) (val)) - 1) +#define Hp_op(op) (Hp_val (op)) +#define Hp_bp(bp) (Hp_val (bp)) +#define Val_op(op) ((value) (op)) +#define Val_hp(hp) ((value) (((header_t *) (hp)) + 1)) +#define Op_hp(hp) ((value *) Val_hp (hp)) +#define Bp_hp(hp) ((char *) Val_hp (hp)) + +#define Num_tags (1 << 8) +#ifdef ARCH_SIXTYFOUR +#define Max_wosize (((intnat)1 << (54-PROFINFO_WIDTH)) - 1) +#else +#define Max_wosize ((1 << 22) - 1) +#endif /* ARCH_SIXTYFOUR */ + +#define Wosize_val(val) (Wosize_hd (Hd_val (val))) +#define Wosize_op(op) (Wosize_val (op)) +#define Wosize_bp(bp) (Wosize_val (bp)) +#define Wosize_hp(hp) (Wosize_hd (Hd_hp (hp))) +#define Whsize_wosize(sz) ((sz) + 1) +#define Wosize_whsize(sz) ((sz) - 1) +#define Wosize_bhsize(sz) ((sz) / sizeof (value) - 1) +#define Bsize_wsize(sz) ((sz) * sizeof (value)) +#define Wsize_bsize(sz) ((sz) / sizeof (value)) +#define Bhsize_wosize(sz) (Bsize_wsize (Whsize_wosize (sz))) +#define Bhsize_bosize(sz) ((sz) + sizeof (header_t)) +#define Bosize_val(val) (Bsize_wsize (Wosize_val (val))) +#define Bosize_op(op) (Bosize_val (Val_op (op))) +#define Bosize_bp(bp) (Bosize_val (Val_bp (bp))) +#define Bosize_hd(hd) (Bsize_wsize (Wosize_hd (hd))) +#define Whsize_hp(hp) (Whsize_wosize (Wosize_hp (hp))) +#define Whsize_val(val) (Whsize_hp (Hp_val (val))) +#define Whsize_bp(bp) (Whsize_val (Val_bp (bp))) +#define Whsize_hd(hd) (Whsize_wosize (Wosize_hd (hd))) +#define Bhsize_hp(hp) (Bsize_wsize (Whsize_hp (hp))) +#define Bhsize_hd(hd) (Bsize_wsize (Whsize_hd (hd))) + +#define Profinfo_val(val) (Profinfo_hd (Hd_val (val))) + +#ifdef ARCH_BIG_ENDIAN +#define Tag_val(val) (((unsigned char *) (val)) [-1]) + /* Also an l-value. */ +#define Tag_hp(hp) (((unsigned char *) (hp)) [sizeof(value)-1]) + /* Also an l-value. */ +#else +#define Tag_val(val) (((unsigned char *) (val)) [-sizeof(value)]) + /* Also an l-value. */ +#define Tag_hp(hp) (((unsigned char *) (hp)) [0]) + /* Also an l-value. */ +#endif + +/* The lowest tag for blocks containing no value. */ +#define No_scan_tag 251 + + +/* 1- If tag < No_scan_tag : a tuple of fields. */ + +/* Pointer to the first field. */ +#define Op_val(x) ((value *) (x)) +/* Fields are numbered from 0. */ +#define Field(x, i) (((value *)(x)) [i]) /* Also an l-value. */ + +typedef int32_t opcode_t; +typedef opcode_t * code_t; + +/* NOTE: [Forward_tag] and [Infix_tag] must be just under + [No_scan_tag], with [Infix_tag] the lower one. + See [caml_oldify_one] in minor_gc.c for more details. + + NOTE: Update stdlib/obj.ml whenever you change the tags. + */ + +/* Forward_tag: forwarding pointer that the GC may silently shortcut. + See stdlib/lazy.ml. */ +#define Forward_tag 250 +#define Forward_val(v) Field(v, 0) + +/* If tag == Infix_tag : an infix header inside a closure */ +/* Infix_tag must be odd so that the infix header is scanned as an integer */ +/* Infix_tag must be 1 modulo 4 and infix headers can only occur in blocks + with tag Closure_tag (see compact.c). */ + +#define Infix_tag 249 +#define Infix_offset_hd(hd) (Bosize_hd(hd)) +#define Infix_offset_val(v) Infix_offset_hd(Hd_val(v)) + +/* Another special case: objects */ +#define Object_tag 248 +#define Class_val(val) Field((val), 0) +#define Oid_val(val) Long_val(Field((val), 1)) +CAMLextern value caml_get_public_method (value obj, value tag); +/* Called as: + caml_callback(caml_get_public_method(obj, caml_hash_variant(name)), obj) */ +/* caml_get_public_method returns 0 if tag not in the table. + Note however that tags being hashed, same tag does not necessarily mean + same method name. */ + +/* Special case of tuples of fields: closures */ +#define Closure_tag 247 +#define Code_val(val) (((code_t *) (val)) [0]) /* Also an l-value. */ + +/* This tag is used (with Forward_tag) to implement lazy values. + See major_gc.c and stdlib/lazy.ml. */ +#define Lazy_tag 246 + +/* Another special case: variants */ +CAMLextern value caml_hash_variant(char const * tag); + +/* 2- If tag >= No_scan_tag : a sequence of bytes. */ + +/* Pointer to the first byte */ +#define Bp_val(v) ((char *) (v)) +#define Val_bp(p) ((value) (p)) +/* Bytes are numbered from 0. */ +#define Byte(x, i) (((char *) (x)) [i]) /* Also an l-value. */ +#define Byte_u(x, i) (((unsigned char *) (x)) [i]) /* Also an l-value. */ + +/* Abstract things. Their contents is not traced by the GC; therefore they + must not contain any [value]. Must have odd number so that headers with + this tag cannot be mistaken for pointers (see caml_obj_truncate). +*/ +#define Abstract_tag 251 +#define Data_abstract_val(v) ((void*) Op_val(v)) + +/* Strings. */ +#define String_tag 252 +#ifdef CAML_SAFE_STRING +#define String_val(x) ((const char *) Bp_val(x)) +#else +#define String_val(x) ((char *) Bp_val(x)) +#endif +#define Bytes_val(x) ((unsigned char *) Bp_val(x)) +CAMLextern mlsize_t caml_string_length (value); /* size in bytes */ +CAMLextern int caml_string_is_c_safe (value); + /* true if string contains no '\0' null characters */ + +/* Floating-point numbers. */ +#define Double_tag 253 +#define Double_wosize ((sizeof(double) / sizeof(value))) +#ifndef ARCH_ALIGN_DOUBLE +#define Double_val(v) (* (double *)(v)) +#define Store_double_val(v,d) (* (double *)(v) = (d)) +#else +CAMLextern double caml_Double_val (value); +CAMLextern void caml_Store_double_val (value,double); +#define Double_val(v) caml_Double_val(v) +#define Store_double_val(v,d) caml_Store_double_val(v,d) +#endif + +/* Arrays of floating-point numbers. */ +#define Double_array_tag 254 + +/* The [_flat_field] macros are for [floatarray] values and float-only records. +*/ +#define Double_flat_field(v,i) Double_val((value)((double *)(v) + (i))) +#define Store_double_flat_field(v,i,d) do{ \ + mlsize_t caml__temp_i = (i); \ + double caml__temp_d = (d); \ + Store_double_val((value)((double *) (v) + caml__temp_i), caml__temp_d); \ +}while(0) + +/* The [_array_field] macros are for [float array]. */ +#ifdef FLAT_FLOAT_ARRAY + #define Double_array_field(v,i) Double_flat_field(v,i) + #define Store_double_array_field(v,i,d) Store_double_flat_field(v,i,d) +#else + #define Double_array_field(v,i) Double_val (Field(v,i)) + CAMLextern void caml_Store_double_array_field (value, mlsize_t, double); + #define Store_double_array_field(v,i,d) caml_Store_double_array_field (v,i,d) +#endif + +/* The old [_field] macros are for backward compatibility only. + They work with [floatarray], float-only records, and [float array]. */ +#ifdef FLAT_FLOAT_ARRAY + #define Double_field(v,i) Double_flat_field(v,i) + #define Store_double_field(v,i,d) Store_double_flat_field(v,i,d) +#else + static inline double Double_field (value v, mlsize_t i) { + if (Tag_val (v) == Double_array_tag){ + return Double_flat_field (v, i); + }else{ + return Double_array_field (v, i); + } + } + static inline void Store_double_field (value v, mlsize_t i, double d) { + if (Tag_val (v) == Double_array_tag){ + Store_double_flat_field (v, i, d); + }else{ + Store_double_array_field (v, i, d); + } + } +#endif /* FLAT_FLOAT_ARRAY */ + +CAMLextern mlsize_t caml_array_length (value); /* size in items */ +CAMLextern int caml_is_double_array (value); /* 0 is false, 1 is true */ + + +/* Custom blocks. They contain a pointer to a "method suite" + of functions (for finalization, comparison, hashing, etc) + followed by raw data. The contents of custom blocks is not traced by + the GC; therefore, they must not contain any [value]. + See [custom.h] for operations on method suites. */ +#define Custom_tag 255 +#define Data_custom_val(v) ((void *) &Field((v), 1)) +struct custom_operations; /* defined in [custom.h] */ + +/* Int32.t, Int64.t and Nativeint.t are represented as custom blocks. */ + +#define Int32_val(v) (*((int32_t *) Data_custom_val(v))) +#define Nativeint_val(v) (*((intnat *) Data_custom_val(v))) +#ifndef ARCH_ALIGN_INT64 +#define Int64_val(v) (*((int64_t *) Data_custom_val(v))) +#else +CAMLextern int64_t caml_Int64_val(value v); +#define Int64_val(v) caml_Int64_val(v) +#endif + +/* 3- Atoms are 0-tuples. They are statically allocated once and for all. */ + +CAMLextern header_t caml_atom_table[]; +#define Atom(tag) (Val_hp (&(caml_atom_table [(tag)]))) + +/* Booleans are integers 0 or 1 */ + +#define Val_bool(x) Val_int((x) != 0) +#define Bool_val(x) Int_val(x) +#define Val_false Val_int(0) +#define Val_true Val_int(1) +#define Val_not(x) (Val_false + Val_true - (x)) + +/* The unit value is 0 (tagged) */ + +#define Val_unit Val_int(0) + +/* List constructors */ +#define Val_emptylist Val_int(0) +#define Tag_cons 0 + +/* The table of global identifiers */ + +extern value caml_global_data; + +CAMLextern value caml_set_oo_id(value obj); + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_MLVALUES_H */ diff --git a/test/monniaux/ocaml/byterun/caml/osdeps.h b/test/monniaux/ocaml/byterun/caml/osdeps.h new file mode 100644 index 00000000..b65503d8 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/osdeps.h @@ -0,0 +1,152 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2001 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Operating system - specific stuff */ + +#ifndef CAML_OSDEPS_H +#define CAML_OSDEPS_H + +#ifdef _WIN32 +extern unsigned short caml_win32_major; +extern unsigned short caml_win32_minor; +extern unsigned short caml_win32_build; +extern unsigned short caml_win32_revision; +#endif + +#ifdef CAML_INTERNALS + +#include "misc.h" +#include "memory.h" + +/* Read at most [n] bytes from file descriptor [fd] into buffer [buf]. + [flags] indicates whether [fd] is a socket + (bit [CHANNEL_FLAG_FROM_SOCKET] is set in this case, see [io.h]). + (This distinction matters for Win32, but not for Unix.) + Return number of bytes read. + In case of error, raises [Sys_error] or [Sys_blocked_io]. */ +extern int caml_read_fd(int fd, int flags, void * buf, int n); + +/* Write at most [n] bytes from buffer [buf] onto file descriptor [fd]. + [flags] indicates whether [fd] is a socket + (bit [CHANNEL_FLAG_FROM_SOCKET] is set in this case, see [io.h]). + (This distinction matters for Win32, but not for Unix.) + Return number of bytes written. + In case of error, raises [Sys_error] or [Sys_blocked_io]. */ +extern int caml_write_fd(int fd, int flags, void * buf, int n); + +/* Decompose the given path into a list of directories, and add them + to the given table. */ +extern char_os * caml_decompose_path(struct ext_table * tbl, char_os * path); + +/* Search the given file in the given list of directories. + If not found, return a copy of [name]. */ +extern char_os * caml_search_in_path(struct ext_table * path, const char_os * name); + +/* Same, but search an executable name in the system path for executables. */ +CAMLextern char_os * caml_search_exe_in_path(const char_os * name); + +/* Same, but search a shared library in the given path. */ +extern char_os * caml_search_dll_in_path(struct ext_table * path, const char_os * name); + +/* Open a shared library and return a handle on it. + If [for_execution] is true, perform full symbol resolution and + execute initialization code so that functions from the shared library + can be called. If [for_execution] is false, functions from this + shared library will not be called, but just checked for presence, + so symbol resolution can be skipped. + If [global] is true, symbols from the shared library can be used + to resolve for other libraries to be opened later on. + Return [NULL] on error. */ +extern void * caml_dlopen(char_os * libname, int for_execution, int global); + +/* Close a shared library handle */ +extern void caml_dlclose(void * handle); + +/* Look up the given symbol in the given shared library. + Return [NULL] if not found, or symbol value if found. */ +extern void * caml_dlsym(void * handle, const char * name); + +extern void * caml_globalsym(const char * name); + +/* Return an error message describing the most recent dynlink failure. */ +extern char * caml_dlerror(void); + +/* Add to [contents] the (short) names of the files contained in + the directory named [dirname]. No entries are added for [.] and [..]. + Return 0 on success, -1 on error; set errno in the case of error. */ +extern int caml_read_directory(char_os * dirname, struct ext_table * contents); + +/* Recover executable name if possible (/proc/sef/exe under Linux, + GetModuleFileName under Windows). Return NULL on error, + string allocated with [caml_stat_alloc] on success. */ +extern char_os * caml_executable_name(void); + +/* Secure version of [getenv]: returns NULL if the process has special + privileges (setuid bit, setgid bit, capabilities). +*/ +extern char_os *caml_secure_getenv(char_os const *var); + +/* If [fd] refers to a terminal or console, return the number of rows + (lines) that it displays. Otherwise, or if the number of rows + cannot be determined, return -1. */ +extern int caml_num_rows_fd(int fd); + +#ifdef _WIN32 + +extern int caml_win32_rename(const wchar_t *, const wchar_t *); + +extern void caml_probe_win32_version(void); +extern void caml_setup_win32_terminal(void); +extern void caml_restore_win32_terminal(void); + +extern wchar_t *caml_win32_getenv(wchar_t const *); + +/* Windows Unicode support */ + +extern int win_multi_byte_to_wide_char(const char* s, int slen, wchar_t *out, int outlen); +extern int win_wide_char_to_multi_byte(const wchar_t* s, int slen, char *out, int outlen); + +/* [caml_stat_strdup_to_utf16(s)] returns a NULL-terminated copy of [s], + re-encoded in UTF-16. The encoding of [s] is assumed to be UTF-8 if + [caml_windows_unicode_runtime_enabled] is non-zero **and** [s] is valid + UTF-8, or the current Windows code page otherwise. + + The returned string is allocated with [caml_stat_alloc], so it should be free + using [caml_stat_free]. +*/ +extern wchar_t* caml_stat_strdup_to_utf16(const char *s); + +/* [caml_stat_strdup_of_utf16(s)] returns a NULL-terminated copy of [s], + re-encoded in UTF-8 if [caml_windows_unicode_runtime_enabled] is non-zero or + the current Windows code page otherwise. + + The returned string is allocated with [caml_stat_alloc], so it should be free + using [caml_stat_free]. +*/ +extern char* caml_stat_strdup_of_utf16(const wchar_t *s); + +/* [caml_copy_string_of_utf16(s)] returns an OCaml string containing a copy of + [s] re-encoded in UTF-8 if [caml_windows_unicode_runtime_enabled] is non-zero + or in the current code page otherwise. +*/ +extern value caml_copy_string_of_utf16(const wchar_t *s); + +extern int caml_win32_isatty(int fd); + +#endif /* _WIN32 */ + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_OSDEPS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/prims.h b/test/monniaux/ocaml/byterun/caml/prims.h new file mode 100644 index 00000000..147cd98a --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/prims.h @@ -0,0 +1,40 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Interface with C primitives. */ + +#ifndef CAML_PRIMS_H +#define CAML_PRIMS_H + +#ifdef CAML_INTERNALS + +typedef value (*c_primitive)(); + +extern c_primitive caml_builtin_cprim[]; +extern char * caml_names_of_builtin_cprim[]; + +extern struct ext_table caml_prim_table; +#ifdef DEBUG +extern struct ext_table caml_prim_name_table; +#endif + +#define Primitive(n) ((c_primitive)(caml_prim_table.contents[n])) + +extern char * caml_section_table; +extern asize_t caml_section_table_size; + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_PRIMS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/printexc.h b/test/monniaux/ocaml/byterun/caml/printexc.h new file mode 100644 index 00000000..92c5af53 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/printexc.h @@ -0,0 +1,35 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2001 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_PRINTEXC_H +#define CAML_PRINTEXC_H + + +#include "misc.h" +#include "mlvalues.h" + +#ifdef __cplusplus +extern "C" { +#endif + + +CAMLextern char * caml_format_exception (value); +CAMLnoreturn_start void caml_fatal_uncaught_exception (value) CAMLnoreturn_end; + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_PRINTEXC_H */ diff --git a/test/monniaux/ocaml/byterun/caml/reverse.h b/test/monniaux/ocaml/byterun/caml/reverse.h new file mode 100644 index 00000000..a186078e --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/reverse.h @@ -0,0 +1,92 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Swap byte-order in 16, 32, and 64-bit integers or floats */ + +#ifndef CAML_REVERSE_H +#define CAML_REVERSE_H + +#ifdef CAML_INTERNALS + +#define Reverse_16(dst,src) { \ + char * _p, * _q; \ + char _a; \ + _p = (char *) (src); \ + _q = (char *) (dst); \ + _a = _p[0]; \ + _q[0] = _p[1]; \ + _q[1] = _a; \ +} + +#define Reverse_32(dst,src) { \ + char * _p, * _q; \ + char _a, _b; \ + _p = (char *) (src); \ + _q = (char *) (dst); \ + _a = _p[0]; \ + _b = _p[1]; \ + _q[0] = _p[3]; \ + _q[1] = _p[2]; \ + _q[3] = _a; \ + _q[2] = _b; \ +} + +#define Reverse_64(dst,src) { \ + char * _p, * _q; \ + char _a, _b; \ + _p = (char *) (src); \ + _q = (char *) (dst); \ + _a = _p[0]; \ + _b = _p[1]; \ + _q[0] = _p[7]; \ + _q[1] = _p[6]; \ + _q[7] = _a; \ + _q[6] = _b; \ + _a = _p[2]; \ + _b = _p[3]; \ + _q[2] = _p[5]; \ + _q[3] = _p[4]; \ + _q[5] = _a; \ + _q[4] = _b; \ +} + +#define Perm_index(perm,i) ((perm >> (i * 4)) & 0xF) + +#define Permute_64(dst,perm_dst,src,perm_src) { \ + char * _p; \ + char _a, _b, _c, _d, _e, _f, _g, _h; \ + _p = (char *) (src); \ + _a = _p[Perm_index(perm_src, 0)]; \ + _b = _p[Perm_index(perm_src, 1)]; \ + _c = _p[Perm_index(perm_src, 2)]; \ + _d = _p[Perm_index(perm_src, 3)]; \ + _e = _p[Perm_index(perm_src, 4)]; \ + _f = _p[Perm_index(perm_src, 5)]; \ + _g = _p[Perm_index(perm_src, 6)]; \ + _h = _p[Perm_index(perm_src, 7)]; \ + _p = (char *) (dst); \ + _p[Perm_index(perm_dst, 0)] = _a; \ + _p[Perm_index(perm_dst, 1)] = _b; \ + _p[Perm_index(perm_dst, 2)] = _c; \ + _p[Perm_index(perm_dst, 3)] = _d; \ + _p[Perm_index(perm_dst, 4)] = _e; \ + _p[Perm_index(perm_dst, 5)] = _f; \ + _p[Perm_index(perm_dst, 6)] = _g; \ + _p[Perm_index(perm_dst, 7)] = _h; \ +} + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_REVERSE_H */ diff --git a/test/monniaux/ocaml/byterun/caml/roots.h b/test/monniaux/ocaml/byterun/caml/roots.h new file mode 100644 index 00000000..fed345d3 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/roots.h @@ -0,0 +1,44 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_ROOTS_H +#define CAML_ROOTS_H + +#ifdef CAML_INTERNALS + +#include "misc.h" +#include "memory.h" + +typedef void (*scanning_action) (value, value *); + +void caml_oldify_local_roots (void); +void caml_darken_all_roots_start (void); +intnat caml_darken_all_roots_slice (intnat); +void caml_do_roots (scanning_action, int); +extern uintnat caml_incremental_roots_count; +#ifndef NATIVE_CODE +CAMLextern void caml_do_local_roots (scanning_action, value *, value *, + struct caml__roots_block *); +#else +CAMLextern void caml_do_local_roots(scanning_action f, char * bottom_of_stack, + uintnat last_retaddr, value * gc_regs, + struct caml__roots_block * local_roots); +#endif + +CAMLextern void (*caml_scan_roots_hook) (scanning_action); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_ROOTS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/s.h b/test/monniaux/ocaml/byterun/caml/s.h new file mode 100644 index 00000000..add8c1cc --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/s.h @@ -0,0 +1,217 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Operating system and standard library dependencies. */ + +/* 0. Operating system type string. */ + +#define OCAML_OS_TYPE "Unix" +/* #define OCAML_OS_TYPE "Win32" */ +/* #define OCAML_OS_TYPE "MacOS" */ +#define HAS_STDINT_H 1 + +/* 1. For the runtime system. */ + +/* #define POSIX_SIGNALS */ + +/* Define POSIX_SIGNALS if signal handling is POSIX-compliant. + In particular, sigaction(), sigprocmask() and the operations on + sigset_t are provided. */ + +/* #define BSD_SIGNALS */ + +/* Define BSD_SIGNALS if signal handlers have the BSD semantics: the handler + remains attached to the signal when the signal is received. Leave it + undefined if signal handlers have the System V semantics: the signal + resets the behavior to default. */ + +/* #define HAS_SIGSETMASK */ + +/* Define HAS_SIGSETMASK if you have sigsetmask(), as in BSD. */ + +/* #define SUPPORT_DYNAMIC_LINKING */ + +/* Define SUPPORT_DYNAMIC_LINKING if dynamic loading of C stub code + via dlopen() is available. */ + +#define HAS_C99_FLOAT_OPS + +/* Define HAS_C99_FLOAT_OPS if conforms to ISO C99. + In particular, it should provide expm1(), log1p(), hypot(), copysign(). */ + +/* 2. For the Unix library. */ + +/* #define HAS_SOCKETS */ + +/* Define HAS_SOCKETS if you have BSD sockets. */ + +/* #define HAS_SOCKLEN_T */ + +/* Define HAS_SOCKLEN_T if the type socklen_t is defined in + /usr/include/sys/socket.h. */ + +#define HAS_UNISTD + +/* Define HAS_UNISTD if you have /usr/include/unistd.h. */ + +#define HAS_DIRENT + +/* Define HAS_DIRENT if you have /usr/include/dirent.h and the result of + readdir() is of type struct dirent *. + Otherwise, we'll load /usr/include/sys/dir.h, and readdir() is expected to + return a struct direct *. */ + +/* #define HAS_REWINDDIR */ + +/* Define HAS_REWINDDIR if you have rewinddir(). */ + +/* #define HAS_LOCKF */ + +/* Define HAS_LOCKF if the library provides the lockf() function. */ + +/* #define HAS_MKFIFO */ + +/* Define HAS_MKFIFO if the library provides the mkfifo() function. */ + +/* #define HAS_GETCWD */ + +/* Define HAS_GETCWD if the library provides the getcwd() function. */ + +#define HAS_GETPRIORITY + +/* Define HAS_GETPRIORITY if the library provides getpriority() and + setpriority(). Otherwise, we'll use nice(). */ + +#define HAS_UTIME +#define HAS_UTIMES + +/* Define HAS_UTIME if you have /usr/include/utime.h and the library + provides utime(). Define HAS_UTIMES if the library provides utimes(). */ + +#define HAS_DUP2 + +/* Define HAS_DUP2 if you have dup2(). */ + +#define HAS_FCHMOD + +/* Define HAS_FCHMOD if you have fchmod() and fchown(). */ + +#define HAS_TRUNCATE + +/* Define HAS_TRUNCATE if you have truncate() and + ftruncate(). */ + +#define HAS_SELECT + +/* Define HAS_SELECT if you have select(). */ + +#define HAS_SYS_SELECT_H + +/* Define HAS_SYS_SELECT_H if /usr/include/sys/select.h exists + and should be included before using select(). */ + +#define HAS_NANOSLEEP +/* Define HAS_NANOSLEEP if you have nanosleep(). */ + +#define HAS_SYMLINK + +/* Define HAS_SYMLINK if you have symlink() and readlink() and lstat(). */ + +#define HAS_WAIT4 +#define HAS_WAITPID + +/* Define HAS_WAIT4 if you have wait4(). + Define HAS_WAITPID if you have waitpid(). */ + +#define HAS_GETGROUPS + +/* Define HAS_GETGROUPS if you have getgroups(). */ + +#define HAS_SETGROUPS + +/* Define HAS_SETGROUPS if you have setgroups(). */ + +#define HAS_INITGROUPS + +/* Define HAS_INITGROUPS if you have initgroups(). */ + +#define HAS_TERMIOS + +/* Define HAS_TERMIOS if you have /usr/include/termios.h and it is + Posix-compliant. */ + +#define HAS_ASYNC_IO + +/* Define HAS_ASYNC_IO if BSD-style asynchronous I/O are supported + (the process can request to be sent a SIGIO signal when a descriptor + is ready for reading). */ + +#define HAS_SETITIMER + +/* Define HAS_SETITIMER if you have setitimer(). */ + +#define HAS_GETHOSTNAME + +/* Define HAS_GETHOSTNAME if you have gethostname(). */ + +#define HAS_UNAME + +/* Define HAS_UNAME if you have uname(). */ + +#define HAS_GETTIMEOFDAY + +/* Define HAS_GETTIMEOFDAY if you have gettimeofday(). */ + +#define HAS_MKTIME + +/* Define HAS_MKTIME if you have mktime(). */ + +#define HAS_SETSID + +/* Define HAS_SETSID if you have setsid(). */ + +#define HAS_PUTENV + +/* Define HAS_PUTENV if you have putenv(). */ + +#define HAS_LOCALE + +/* Define HAS_LOCALE if you have the include file and the + setlocale() function. */ + +#define HAS_MMAP + +/* Define HAS_MMAP if you have the include file and the + functions mmap() and munmap(). */ + +#define HAS_GETHOSTBYNAME_R 6 + +/* Define HAS_GETHOSTBYNAME_R if gethostbyname_r() is available. + The value of this symbol is the number of arguments of + gethostbyname_r(): either 5 or 6 depending on prototype. + (5 is the Solaris version, 6 is the Linux version). */ + +#define HAS_GETHOSTBYADDR_R 8 + +/* Define HAS_GETHOSTBYADDR_R if gethostbyname_r() is available. + The value of this symbol is the number of arguments of + gethostbyaddr_r(): either 7 or 8 depending on prototype. + (7 is the Solaris version, 8 is the Linux version). */ + +#define HAS_NICE + +/* Define HAS_NICE if you have nice(). */ + +#define OCAML_STDLIB_DIR "/opt/ccomp/ocaml/4.07.1/lib/ocaml" diff --git a/test/monniaux/ocaml/byterun/caml/signals.h b/test/monniaux/ocaml/byterun/caml/signals.h new file mode 100644 index 00000000..99924e4f --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/signals.h @@ -0,0 +1,59 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_SIGNALS_H +#define CAML_SIGNALS_H + +#ifndef CAML_NAME_SPACE +#include "compatibility.h" +#endif +#include "misc.h" +#include "mlvalues.h" + +#ifdef __cplusplus +extern "C" { +#endif + +#ifdef CAML_INTERNALS +CAMLextern intnat volatile caml_signals_are_pending; +CAMLextern intnat volatile caml_pending_signals[]; +CAMLextern int volatile caml_something_to_do; +extern int volatile caml_requested_major_slice; +extern int volatile caml_requested_minor_gc; + +void caml_request_major_slice (void); +void caml_request_minor_gc (void); +CAMLextern int caml_convert_signal_number (int); +CAMLextern int caml_rev_convert_signal_number (int); +void caml_execute_signal(int signal_number, int in_signal_handler); +void caml_record_signal(int signal_number); +void caml_process_pending_signals(void); +void caml_process_event(void); +int caml_set_signal_action(int signo, int action); + +CAMLextern void (*caml_enter_blocking_section_hook)(void); +CAMLextern void (*caml_leave_blocking_section_hook)(void); +CAMLextern int (*caml_try_leave_blocking_section_hook)(void); +CAMLextern void (* volatile caml_async_action_hook)(void); +#endif /* CAML_INTERNALS */ + +CAMLextern void caml_enter_blocking_section (void); +CAMLextern void caml_leave_blocking_section (void); + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_SIGNALS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/signals_machdep.h b/test/monniaux/ocaml/byterun/caml/signals_machdep.h new file mode 100644 index 00000000..ae2145ba --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/signals_machdep.h @@ -0,0 +1,74 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Processor-specific operation: atomic "read and clear" */ + +#ifndef CAML_SIGNALS_MACHDEP_H +#define CAML_SIGNALS_MACHDEP_H + +#ifdef CAML_INTERNALS + +#if defined(__GNUC__) && defined(__ATOMIC_SEQ_CST) \ + && defined(__GCC_ATOMIC_LONG_LOCK_FREE) && 0 + +/* Use the "atomic" builtins of GCC and Clang */ +#define Read_and_clear(dst,src) \ + ((dst) = __atomic_exchange_n(&(src), 0, __ATOMIC_SEQ_CST)) + +#elif defined(__GNUC__) && (defined(__i386__) || (defined(__x86_64__) \ + && defined(__ILP32__))) + +#define Read_and_clear(dst,src) \ + asm("xorl %0, %0; xchgl %0, %1" \ + : "=r" (dst), "=m" (src) \ + : "m" (src)) + +#elif defined(__GNUC__) && defined(__x86_64__) + +#define Read_and_clear(dst,src) \ + asm("xorq %0, %0; xchgq %0, %1" \ + : "=r" (dst), "=m" (src) \ + : "m" (src)) + +#elif defined(__GNUC__) && defined(__ppc__) + +#define Read_and_clear(dst,src) \ + asm("0: lwarx %0, 0, %1\n\t" \ + "stwcx. %2, 0, %1\n\t" \ + "bne- 0b" \ + : "=&r" (dst) \ + : "r" (&(src)), "r" (0) \ + : "cr0", "memory") + +#elif defined(__GNUC__) && defined(__ppc64__) + +#define Read_and_clear(dst,src) \ + asm("0: ldarx %0, 0, %1\n\t" \ + "stdcx. %2, 0, %1\n\t" \ + "bne- 0b" \ + : "=&r" (dst) \ + : "r" (&(src)), "r" (0) \ + : "cr0", "memory") + +#else + +/* Default, non-atomic implementation */ +#define Read_and_clear(dst,src) ((dst) = (src), (src) = 0) + +#endif + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_SIGNALS_MACHDEP_H */ diff --git a/test/monniaux/ocaml/byterun/caml/spacetime.h b/test/monniaux/ocaml/byterun/caml/spacetime.h new file mode 100644 index 00000000..5bcc9232 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/spacetime.h @@ -0,0 +1,203 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Mark Shinwell and Leo White, Jane Street Europe */ +/* */ +/* Copyright 2013--2016, Jane Street Group, LLC */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_SPACETIME_H +#define CAML_SPACETIME_H + +#include "io.h" +#include "misc.h" +#include "stack.h" + +/* Runtime support for Spacetime profiling. + * This header file is not intended for the casual user. + * + * The implementation is split into three files: + * 1. spacetime.c: core management of the instrumentation; + * 2. spacetime_snapshot.c: the taking of heap snapshots; + * 3. spacetime_offline.c: functions that are also used when examining + * saved profiling data. + */ + +typedef enum { + CALL, + ALLOCATION +} c_node_type; + +/* All pointers between nodes point at the word immediately after the + GC headers, and everything is traversable using the normal OCaml rules. + + On entry to an OCaml function: + If the node hole pointer register has the bottom bit set, then the function + is being tail called or called from a self-recursive call site: + - If the node hole is empty, the callee must create a new node and link + it into the tail chain. The node hole pointer will point at the tail + chain. + - Otherwise the node should be used as normal. + Otherwise (not a tail call): + - If the node hole is empty, the callee must create a new node, but the + tail chain is untouched. + - Otherwise the node should be used as normal. +*/ + +/* Classification of nodes (OCaml or C) with corresponding GC tags. */ +#define OCaml_node_tag 0 +#define C_node_tag 1 +#define Is_ocaml_node(node) (Is_block(node) && Tag_val(node) == OCaml_node_tag) +#define Is_c_node(node) (Is_block(node) && Tag_val(node) == C_node_tag) + +/* The header words are: + 1. The node program counter. + 2. The tail link. */ +#define Node_num_header_words 2 + +/* The "node program counter" at the start of an OCaml node. */ +#define Node_pc(node) (Field(node, 0)) +#define Encode_node_pc(pc) (((value) pc) | 1) +#define Decode_node_pc(encoded_pc) ((void*) (encoded_pc & ~1)) + +/* The circular linked list of tail-called functions within OCaml nodes. */ +#define Tail_link(node) (Field(node, 1)) + +/* The convention for pointers from OCaml nodes to other nodes. There are + two special cases: + 1. [Val_unit] means "uninitialized", and further, that this is not a + tail call point. (Tail call points are pre-initialized, as in case 2.) + 2. If the bottom bit is set, and the value is not [Val_unit], this is a + tail call point. */ +#define Encode_tail_caller_node(node) ((node) | 1) +#define Decode_tail_caller_node(node) ((node) & ~1) +#define Is_tail_caller_node_encoded(node) (((node) & 1) == 1) + +/* Allocation points within OCaml nodes. + The "profinfo" value looks exactly like a black Infix_tag header. + This enables us to point just after it and return such pointer as a valid + OCaml value. (Used for the list of all allocation points. We could do + without this and instead just encode the list pointers as integers, but + this would mean that the structure was destroyed on marshalling. This + might not be a great problem since it is intended that the total counts + be obtained via snapshots, but it seems neater and easier to use + Infix_tag. + The "count" is just an OCaml integer giving the total number of words + (including headers) allocated at the point. + The "pointer to next allocation point" points to the "count" word of the + next allocation point in the linked list of all allocation points. + There is no special encoding needed by virtue of the [Infix_tag] trick. */ +#define Alloc_point_profinfo(node, offset) (Field(node, offset)) +#define Alloc_point_count(node, offset) (Field(node, offset + 1)) +#define Alloc_point_next_ptr(node, offset) (Field(node, offset + 2)) + +/* Direct call points (tail or non-tail) within OCaml nodes. + They hold a pointer to the child node and (if the compiler was so + configured) a call count. + The call site and callee are both recorded in the shape. */ +#define Direct_callee_node(node,offset) (Field(node, offset)) +#define Direct_call_count(node,offset) (Field(node, offset + 1)) +#define Encode_call_point_pc(pc) (((value) pc) | 1) +#define Decode_call_point_pc(pc) ((void*) (((value) pc) & ~((uintnat) 1))) + +/* Indirect call points (tail or non-tail) within OCaml nodes. + They hold a linked list of (PC upon entry to the callee, pointer to + child node) pairs. The linked list is encoded using C nodes and should + be thought of as part of the OCaml node itself. */ +#define Indirect_num_fields 1 +#define Indirect_pc_linked_list(node,offset) (Field(node, offset)) + +/* Encodings of the program counter value within a C node. */ +#define Encode_c_node_pc_for_call(pc) ((((value) pc) << 2) | 3) +#define Encode_c_node_pc_for_alloc_point(pc) ((((value) pc) << 2) | 1) +#define Decode_c_node_pc(pc) ((void*) (((uintnat) (pc)) >> 2)) + +typedef struct { + /* The layout and encoding of this structure must match that of the + allocation points within OCaml nodes, so that the linked list + traversal across all allocation points works correctly. */ + value profinfo; /* encoded using [Infix_tag] (see above) */ + value count; + /* [next] is [Val_unit] for the end of the list. + Otherwise it points at the second word of this [allocation_point] + structure. */ + value next; +} allocation_point; + +typedef struct { + value callee_node; + value call_count; +} call_point; + +typedef struct { + /* CR-soon mshinwell: delete [gc_header], all the offset arithmetic will + then go away */ + uintnat gc_header; + uintnat pc; /* see above for encodings */ + union { + call_point call; /* for CALL */ + allocation_point allocation; /* for ALLOCATION */ + } data; + value next; /* [Val_unit] for the end of the list */ +} c_node; /* CR-soon mshinwell: rename to dynamic_node */ + +typedef struct shape_table { + uint64_t* table; + struct shape_table* next; +} shape_table; + +extern uint64_t** caml_spacetime_static_shape_tables; +extern shape_table* caml_spacetime_dynamic_shape_tables; + +typedef struct ext_table* spacetime_unwind_info_cache; + +extern value caml_spacetime_trie_root; +extern value* caml_spacetime_trie_node_ptr; +extern value* caml_spacetime_finaliser_trie_root; + +extern allocation_point* caml_all_allocation_points; + +extern void caml_spacetime_initialize(void); +extern uintnat caml_spacetime_my_profinfo( + spacetime_unwind_info_cache*, uintnat); +extern c_node_type caml_spacetime_classify_c_node(c_node* node); +extern c_node* caml_spacetime_c_node_of_stored_pointer(value); +extern c_node* caml_spacetime_c_node_of_stored_pointer_not_null(value); +extern value caml_spacetime_stored_pointer_of_c_node(c_node* node); +extern void caml_spacetime_register_thread(value*, value*); +extern void caml_spacetime_register_shapes(void*); +extern value caml_spacetime_frame_table(void); +extern value caml_spacetime_shape_table(void); +extern void caml_spacetime_save_snapshot (struct channel *chan, + double time_override, + int use_time_override); +extern value caml_spacetime_timestamp(double time_override, + int use_time_override); +extern void caml_spacetime_automatic_snapshot (void); + +/* For use in runtime functions that are executed from OCaml + code, to save the overhead of using libunwind every time. */ +#ifdef WITH_SPACETIME +#define Get_my_profinfo_with_cached_backtrace(profinfo, size) \ + do { \ + static spacetime_unwind_info_cache spacetime_unwind_info = NULL; \ + profinfo = caml_spacetime_my_profinfo(&spacetime_unwind_info, size); \ + } \ + while (0); +#else +#define Get_my_profinfo_with_cached_backtrace(profinfo, size) \ + profinfo = (uintnat) 0; +#endif + +#else + +#define Get_my_profinfo_with_cached_backtrace(profinfo, size) \ + profinfo = (uintnat) 0; + +#endif diff --git a/test/monniaux/ocaml/byterun/caml/stack.h b/test/monniaux/ocaml/byterun/caml/stack.h new file mode 100644 index 00000000..26686398 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/stack.h @@ -0,0 +1,124 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Machine-dependent interface with the asm code */ + +#ifndef CAML_STACK_H +#define CAML_STACK_H + +#ifdef CAML_INTERNALS + +/* Macros to access the stack frame */ + +#ifdef TARGET_i386 +#define Saved_return_address(sp) *((intnat *)((sp) - 4)) +#ifndef SYS_win32 +#define Callback_link(sp) ((struct caml_context *)((sp) + 16)) +#else +#define Callback_link(sp) ((struct caml_context *)((sp) + 8)) +#endif +#endif + +#ifdef TARGET_power +#if defined(MODEL_ppc) +#define Saved_return_address(sp) *((intnat *)((sp) - 4)) +#define Callback_link(sp) ((struct caml_context *)((sp) + 16)) +#elif defined(MODEL_ppc64) +#define Saved_return_address(sp) *((intnat *)((sp) + 16)) +#define Callback_link(sp) ((struct caml_context *)((sp) + (48 + 32))) +#elif defined(MODEL_ppc64le) +#define Saved_return_address(sp) *((intnat *)((sp) + 16)) +#define Callback_link(sp) ((struct caml_context *)((sp) + (32 + 32))) +#else +#error "TARGET_power: wrong MODEL" +#endif +#define Already_scanned(sp, retaddr) ((retaddr) & 1) +#define Mask_already_scanned(retaddr) ((retaddr) & ~1) +#define Mark_scanned(sp, retaddr) Saved_return_address(sp) = (retaddr) | 1 +#endif + +#ifdef TARGET_s390x +#define Saved_return_address(sp) *((intnat *)((sp) - SIZEOF_PTR)) +#define Trap_frame_size 16 +#define Callback_link(sp) ((struct caml_context *)((sp) + Trap_frame_size)) +#endif + +#ifdef TARGET_arm +#define Saved_return_address(sp) *((intnat *)((sp) - 4)) +#define Callback_link(sp) ((struct caml_context *)((sp) + 8)) +#endif + +#ifdef TARGET_amd64 +#define Saved_return_address(sp) *((intnat *)((sp) - 8)) +#define Callback_link(sp) ((struct caml_context *)((sp) + 16)) +#endif + +#ifdef TARGET_arm64 +#define Saved_return_address(sp) *((intnat *)((sp) - 8)) +#define Callback_link(sp) ((struct caml_context *)((sp) + 16)) +#endif + +/* Structure of OCaml callback contexts */ + +struct caml_context { + char * bottom_of_stack; /* beginning of OCaml stack chunk */ + uintnat last_retaddr; /* last return address in OCaml code */ + value * gc_regs; /* pointer to register block */ +#ifdef WITH_SPACETIME + void* trie_node; +#endif +}; + +/* Structure of frame descriptors */ + +typedef struct { + uintnat retaddr; + unsigned short frame_size; + unsigned short num_live; + unsigned short live_ofs[1]; +} frame_descr; + +/* Hash table of frame descriptors */ + +extern frame_descr ** caml_frame_descriptors; +extern int caml_frame_descriptors_mask; + +#define Hash_retaddr(addr) \ + (((uintnat)(addr) >> 3) & caml_frame_descriptors_mask) + +extern void caml_init_frame_descriptors(void); +extern void caml_register_frametable(intnat *); +extern void caml_unregister_frametable(intnat *); +extern void caml_register_dyn_global(void *); + +extern uintnat caml_stack_usage (void); +extern uintnat (*caml_stack_usage_hook)(void); + +/* Declaration of variables used in the asm code */ +extern char * caml_top_of_stack; +extern char * caml_bottom_of_stack; +extern uintnat caml_last_return_address; +extern value * caml_gc_regs; +extern char * caml_exception_pointer; +extern value * caml_globals[]; +extern char caml_globals_map[]; +extern intnat caml_globals_inited; +extern intnat * caml_frametable[]; + +CAMLextern frame_descr * caml_next_frame_descriptor(uintnat * pc, char ** sp); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_STACK_H */ diff --git a/test/monniaux/ocaml/byterun/caml/stacks.h b/test/monniaux/ocaml/byterun/caml/stacks.h new file mode 100644 index 00000000..18ec0ac3 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/stacks.h @@ -0,0 +1,46 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* structure of the stacks */ + +#ifndef CAML_STACKS_H +#define CAML_STACKS_H + +#ifdef CAML_INTERNALS + +#include "misc.h" +#include "mlvalues.h" +#include "memory.h" + +CAMLextern value * caml_stack_low; +CAMLextern value * caml_stack_high; +CAMLextern value * caml_stack_threshold; +CAMLextern value * caml_extern_sp; +CAMLextern value * caml_trapsp; +CAMLextern value * caml_trap_barrier; + +#define Trap_pc(tp) (((code_t *)(tp))[0]) +#define Trap_link(tp) (((value **)(tp))[1]) + +void caml_init_stack (uintnat init_max_size); +void caml_realloc_stack (asize_t required_size); +void caml_change_max_stack_size (uintnat new_max_size); +uintnat caml_stack_usage (void); + +CAMLextern uintnat (*caml_stack_usage_hook)(void); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_STACKS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/startup.h b/test/monniaux/ocaml/byterun/caml/startup.h new file mode 100644 index 00000000..2b26e916 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/startup.h @@ -0,0 +1,52 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2001 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_STARTUP_H +#define CAML_STARTUP_H + +#ifdef CAML_INTERNALS + +#include "mlvalues.h" +#include "exec.h" + +CAMLextern void caml_main(char_os **argv); + +CAMLextern void caml_startup_code( + code_t code, asize_t code_size, + char *data, asize_t data_size, + char *section_table, asize_t section_table_size, + int pooling, + char_os **argv); + +CAMLextern value caml_startup_code_exn( + code_t code, asize_t code_size, + char *data, asize_t data_size, + char *section_table, asize_t section_table_size, + int pooling, + char_os **argv); + +enum { FILE_NOT_FOUND = -1, BAD_BYTECODE = -2 }; + +extern int caml_attempt_open(char_os **name, struct exec_trailer *trail, + int do_open_script); +extern void caml_read_section_descriptors(int fd, struct exec_trailer *trail); +extern int32_t caml_seek_optional_section(int fd, struct exec_trailer *trail, + char *name); +extern int32_t caml_seek_section(int fd, struct exec_trailer *trail, + char *name); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_STARTUP_H */ diff --git a/test/monniaux/ocaml/byterun/caml/startup_aux.h b/test/monniaux/ocaml/byterun/caml/startup_aux.h new file mode 100644 index 00000000..7286bc98 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/startup_aux.h @@ -0,0 +1,44 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, Jane Street Group, LLC */ +/* */ +/* Copyright 2015 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_STARTUP_AUX_H +#define CAML_STARTUP_AUX_H + +#ifdef CAML_INTERNALS + +#include "config.h" + +extern void caml_init_atom_table (void); + +extern uintnat caml_init_percent_free; +extern uintnat caml_init_max_percent_free; +extern uintnat caml_init_minor_heap_wsz; +extern uintnat caml_init_heap_chunk_sz; +extern uintnat caml_init_heap_wsz; +extern uintnat caml_init_max_stack_wsz; +extern uintnat caml_init_major_window; +extern uintnat caml_trace_level; +extern uintnat caml_cleanup_on_exit; + +extern void caml_parse_ocamlrunparam (void); + +/* Common entry point to caml_startup. + Returns 0 if the runtime is already initialized. + If [pooling] is 0, [caml_stat_*] functions will not be backed by a pool. */ +extern int caml_startup_aux (int pooling); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_STARTUP_AUX_H */ diff --git a/test/monniaux/ocaml/byterun/caml/sys.h b/test/monniaux/ocaml/byterun/caml/sys.h new file mode 100644 index 00000000..0f372771 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/sys.h @@ -0,0 +1,45 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#ifndef CAML_SYS_H +#define CAML_SYS_H + +#ifdef CAML_INTERNALS + +#include "misc.h" + +#ifdef __cplusplus +extern "C" { +#endif + +#define NO_ARG Val_int(0) + +CAMLextern void caml_sys_error (value); +CAMLextern void caml_sys_io_error (value); +CAMLextern double caml_sys_time_unboxed(value); +CAMLextern void caml_sys_init (char_os * exe_name, char_os ** argv); +CAMLextern value caml_sys_exit (value); +extern double caml_sys_time_unboxed(value); +CAMLextern value caml_sys_get_argv(value unit); + +extern char_os * caml_exe_name; + +#ifdef __cplusplus +} +#endif + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_SYS_H */ diff --git a/test/monniaux/ocaml/byterun/caml/ui.h b/test/monniaux/ocaml/byterun/caml/ui.h new file mode 100644 index 00000000..3047ba7f --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/ui.h @@ -0,0 +1,32 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Function declarations for non-Unix user interfaces */ + +#ifndef CAML_UI_H +#define CAML_UI_H + +#ifdef CAML_INTERNALS + +#include "config.h" + +void ui_exit (int return_code); +int ui_read (int file_desc, char *buf, unsigned int length); +int ui_write (int file_desc, char *buf, unsigned int length); +void ui_print_stderr (char *format, void *arg); + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_UI_H */ diff --git a/test/monniaux/ocaml/byterun/caml/weak.h b/test/monniaux/ocaml/byterun/caml/weak.h new file mode 100644 index 00000000..a8563867 --- /dev/null +++ b/test/monniaux/ocaml/byterun/caml/weak.h @@ -0,0 +1,93 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1997 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Operations on weak arrays */ + +#ifndef CAML_WEAK_H +#define CAML_WEAK_H + +#ifdef CAML_INTERNALS + +#include "mlvalues.h" + +extern value caml_ephe_list_head; +extern value caml_ephe_none; + + +/** The first field 0: weak list; + second field 1: data; + others 2..: keys; + + A weak pointer is an ephemeron with the data at caml_ephe_none + If fields are added, don't forget to update weak.ml [additional_values]. + */ + +#define CAML_EPHE_LINK_OFFSET 0 +#define CAML_EPHE_DATA_OFFSET 1 +#define CAML_EPHE_FIRST_KEY 2 + + +/* In the header, in order to let major_gc.c + and weak.c see the body of the function */ +static inline void caml_ephe_clean (value v){ + value child; + int release_data = 0; + mlsize_t size, i; + header_t hd; + CAMLassert(caml_gc_phase == Phase_clean); + + hd = Hd_val (v); + size = Wosize_hd (hd); + for (i = 2; i < size; i++){ + child = Field (v, i); + ephemeron_again: + if (child != caml_ephe_none + && Is_block (child) && Is_in_heap_or_young (child)){ + if (Tag_val (child) == Forward_tag){ + value f = Forward_val (child); + if (Is_block (f)) { + if (!Is_in_value_area(f) || Tag_val (f) == Forward_tag + || Tag_val (f) == Lazy_tag || Tag_val (f) == Double_tag){ + /* Do not short-circuit the pointer. */ + }else{ + Field (v, i) = child = f; + if (Is_block (f) && Is_young (f)) + add_to_ephe_ref_table(&caml_ephe_ref_table, v, i); + goto ephemeron_again; + } + } + } + if (Is_white_val (child) && !Is_young (child)){ + release_data = 1; + Field (v, i) = caml_ephe_none; + } + } + } + + child = Field (v, 1); + if(child != caml_ephe_none){ + if (release_data){ + Field (v, 1) = caml_ephe_none; + } else { + /* The mark phase must have marked it */ + CAMLassert( !(Is_block (child) && Is_in_heap (child) + && Is_white_val (child)) ); + } + } +} + +#endif /* CAML_INTERNALS */ + +#endif /* CAML_WEAK_H */ diff --git a/test/monniaux/ocaml/byterun/compact.c b/test/monniaux/ocaml/byterun/compact.c new file mode 100644 index 00000000..7b7188ab --- /dev/null +++ b/test/monniaux/ocaml/byterun/compact.c @@ -0,0 +1,562 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include + +#include "caml/address_class.h" +#include "caml/config.h" +#include "caml/finalise.h" +#include "caml/freelist.h" +#include "caml/gc.h" +#include "caml/gc_ctrl.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/roots.h" +#include "caml/weak.h" +#include "caml/compact.h" + +extern uintnat caml_percent_free; /* major_gc.c */ +extern void caml_shrink_heap (char *); /* memory.c */ + +/* Encoded headers: the color is stored in the 2 least significant bits. + (For pointer inversion, we need to distinguish headers from pointers.) + s is a Wosize, t is a tag, and c is a color (a two-bit number) + + For the purpose of compaction, "colors" are: + 0: pointers (direct or inverted) + 1: integer or (unencoded) infix header + 2: inverted pointer for infix header + 3: integer or encoded (noninfix) header + + XXX Should be fixed: + XXX The above assumes that all roots are aligned on a 4-byte boundary, + XXX which is not always guaranteed by C. + XXX (see [caml_register_global_roots]) + XXX Should be able to fix it to only assume 2-byte alignment. +*/ +#ifdef WITH_PROFINFO +#define Make_ehd(s,t,c,p) \ + (((s) << 10) | (t) << 2 | (c) | ((p) << PROFINFO_SHIFT)) +#else +#define Make_ehd(s,t,c,p) (((s) << 10) | (t) << 2 | (c)) +#endif +#define Whsize_ehd(h) Whsize_hd (h) +#define Wosize_ehd(h) Wosize_hd (h) +#define Tag_ehd(h) (((h) >> 2) & 0xFF) +#define Profinfo_ehd(hd) Profinfo_hd(hd) +#define Ecolor(w) ((w) & 3) + +typedef uintnat word; + +static void invert_pointer_at (word *p) +{ + word q = *p; + CAMLassert (Ecolor ((intnat) p) == 0); + + /* Use Ecolor (q) == 0 instead of Is_block (q) because q could be an + inverted pointer for an infix header (with Ecolor == 2). */ + if (Ecolor (q) == 0 && Is_in_heap (q)){ + switch (Ecolor (Hd_val (q))){ + case 0: + case 3: /* Pointer or header: insert in inverted list. */ + *p = Hd_val (q); + Hd_val (q) = (header_t) p; + break; + case 1: /* Infix header: make inverted infix list. */ + /* Double inversion: the last of the inverted infix list points to + the next infix header in this block. The last of the last list + contains the original block header. */ + { + /* This block as a value. */ + value val = (value) q - Infix_offset_val (q); + /* Get the block header. */ + word *hp = (word *) Hp_val (val); + + while (Ecolor (*hp) == 0) hp = (word *) *hp; + CAMLassert (Ecolor (*hp) == 3); + if (Tag_ehd (*hp) == Closure_tag){ + /* This is the first infix found in this block. */ + /* Save original header. */ + *p = *hp; + /* Link inverted infix list. */ + Hd_val (q) = (header_t) ((word) p | 2); + /* Change block header's tag to Infix_tag, and change its size + to point to the infix list. */ + *hp = Make_ehd (Wosize_bhsize (q - val), Infix_tag, 3, (uintnat) 0); + }else{ + CAMLassert (Tag_ehd (*hp) == Infix_tag); + /* Point the last of this infix list to the current first infix + list of the block. */ + *p = (word) &Field (val, Wosize_ehd (*hp)) | 1; + /* Point the head of this infix list to the above. */ + Hd_val (q) = (header_t) ((word) p | 2); + /* Change block header's size to point to this infix list. */ + *hp = Make_ehd (Wosize_bhsize (q - val), Infix_tag, 3, (uintnat) 0); + } + } + break; + case 2: /* Inverted infix list: insert. */ + *p = Hd_val (q); + Hd_val (q) = (header_t) ((word) p | 2); + break; + } + } +} + +void caml_invert_root (value v, value *p) +{ + invert_pointer_at ((word *) p); +} + +static char *compact_fl; + +static void init_compact_allocate (void) +{ + char *ch = caml_heap_start; + while (ch != NULL){ + Chunk_alloc (ch) = 0; + ch = Chunk_next (ch); + } + compact_fl = caml_heap_start; +} + +/* [size] is a number of bytes and includes the header size */ +static char *compact_allocate (mlsize_t size) +{ + char *chunk, *adr; + + while (Chunk_size (compact_fl) - Chunk_alloc (compact_fl) <= Bhsize_wosize (3) + && Chunk_size (Chunk_next (compact_fl)) + - Chunk_alloc (Chunk_next (compact_fl)) + <= Bhsize_wosize (3)){ + compact_fl = Chunk_next (compact_fl); + } + chunk = compact_fl; + while (Chunk_size (chunk) - Chunk_alloc (chunk) < size){ + chunk = Chunk_next (chunk); + CAMLassert (chunk != NULL); + } + adr = chunk + Chunk_alloc (chunk); + Chunk_alloc (chunk) += size; + return adr; +} + +static void do_compaction (void) +{ + char *ch, *chend; + CAMLassert (caml_gc_phase == Phase_idle); + caml_gc_message (0x10, "Compacting heap...\n"); + +#ifdef DEBUG + caml_heap_check (); +#endif + + /* First pass: encode all noninfix headers. */ + { + ch = caml_heap_start; + while (ch != NULL){ + header_t *p = (header_t *) ch; + + chend = ch + Chunk_size (ch); + while ((char *) p < chend){ + header_t hd = Hd_hp (p); + mlsize_t sz = Wosize_hd (hd); + + if (Is_blue_hd (hd)){ + /* Free object. Give it a string tag. */ + Hd_hp (p) = Make_ehd (sz, String_tag, 3, (uintnat) 0); + }else{ + CAMLassert (Is_white_hd (hd)); + /* Live object. Keep its tag. */ + Hd_hp (p) = Make_ehd (sz, Tag_hd (hd), 3, Profinfo_hd (hd)); + } + p += Whsize_wosize (sz); + } + ch = Chunk_next (ch); + } + } + + + /* Second pass: invert pointers. + Link infix headers in each block in an inverted list of inverted lists. + Don't forget roots and weak pointers. */ + { + /* Invert roots first because the threads library needs some heap + data structures to find its roots. Fortunately, it doesn't need + the headers (see above). */ + caml_do_roots (caml_invert_root, 1); + /* The values to be finalised are not roots but should still be inverted */ + caml_final_invert_finalisable_values (); + + ch = caml_heap_start; + while (ch != NULL){ + word *p = (word *) ch; + chend = ch + Chunk_size (ch); + + while ((char *) p < chend){ + word q = *p; + size_t sz, i; + tag_t t; + word *infixes; + + while (Ecolor (q) == 0) q = * (word *) q; + sz = Whsize_ehd (q); + t = Tag_ehd (q); + + if (t == Infix_tag){ + /* Get the original header of this block. */ + infixes = p + sz; + q = *infixes; + while (Ecolor (q) != 3) q = * (word *) (q & ~(uintnat)3); + sz = Whsize_ehd (q); + t = Tag_ehd (q); + } + + if (t < No_scan_tag){ + for (i = 1; i < sz; i++) invert_pointer_at (&(p[i])); + } + p += sz; + } + ch = Chunk_next (ch); + } + /* Invert weak pointers. */ + { + value *pp = &caml_ephe_list_head; + value p; + word q; + size_t sz, i; + + while (1){ + p = *pp; + if (p == (value) NULL) break; + q = Hd_val (p); + while (Ecolor (q) == 0) q = * (word *) q; + sz = Wosize_ehd (q); + for (i = 1; i < sz; i++){ + if (Field (p,i) != caml_ephe_none){ + invert_pointer_at ((word *) &(Field (p,i))); + } + } + invert_pointer_at ((word *) pp); + pp = &Field (p, 0); + } + } + } + + + /* Third pass: reallocate virtually; revert pointers; decode headers. + Rebuild infix headers. */ + { + init_compact_allocate (); + ch = caml_heap_start; + while (ch != NULL){ + word *p = (word *) ch; + + chend = ch + Chunk_size (ch); + while ((char *) p < chend){ + word q = *p; + + if (Ecolor (q) == 0 || Tag_ehd (q) == Infix_tag){ + /* There were (normal or infix) pointers to this block. */ + size_t sz; + tag_t t; + char *newadr; +#ifdef WITH_PROFINFO + uintnat profinfo; +#endif + word *infixes = NULL; + + while (Ecolor (q) == 0) q = * (word *) q; + sz = Whsize_ehd (q); + t = Tag_ehd (q); +#ifdef WITH_PROFINFO + profinfo = Profinfo_ehd (q); +#endif + if (t == Infix_tag){ + /* Get the original header of this block. */ + infixes = p + sz; + q = *infixes; + CAMLassert (Ecolor (q) == 2); + while (Ecolor (q) != 3) q = * (word *) (q & ~(uintnat)3); + sz = Whsize_ehd (q); + t = Tag_ehd (q); + } + + newadr = compact_allocate (Bsize_wsize (sz)); + q = *p; + while (Ecolor (q) == 0){ + word next = * (word *) q; + * (word *) q = (word) Val_hp (newadr); + q = next; + } + *p = Make_header_with_profinfo (Wosize_whsize (sz), t, Caml_white, + profinfo); + + if (infixes != NULL){ + /* Rebuild the infix headers and revert the infix pointers. */ + while (Ecolor ((word) infixes) != 3){ + infixes = (word *) ((word) infixes & ~(uintnat) 3); + q = *infixes; + while (Ecolor (q) == 2){ + word next; + q = (word) q & ~(uintnat) 3; + next = * (word *) q; + * (word *) q = (word) Val_hp ((word *) newadr + (infixes - p)); + q = next; + } + CAMLassert (Ecolor (q) == 1 || Ecolor (q) == 3); + /* No need to preserve any profinfo value on the [Infix_tag] + headers; the Spacetime profiling heap snapshot code doesn't + look at them. */ + *infixes = Make_header (infixes - p, Infix_tag, Caml_white); + infixes = (word *) q; + } + } + p += sz; + }else{ + CAMLassert (Ecolor (q) == 3); + /* This is guaranteed only if caml_compact_heap was called after a + nonincremental major GC: CAMLassert (Tag_ehd (q) == String_tag); + */ + /* No pointers to the header and no infix header: + the object was free. */ + *p = Make_header (Wosize_ehd (q), Tag_ehd (q), Caml_blue); + p += Whsize_ehd (q); + } + } + ch = Chunk_next (ch); + } + } + + + /* Fourth pass: reallocate and move objects. + Use the exact same allocation algorithm as pass 3. */ + { + init_compact_allocate (); + ch = caml_heap_start; + while (ch != NULL){ + word *p = (word *) ch; + + chend = ch + Chunk_size (ch); + while ((char *) p < chend){ + word q = *p; + if (Color_hd (q) == Caml_white){ + size_t sz = Bhsize_hd (q); + char *newadr = compact_allocate (sz); + memmove (newadr, p, sz); + p += Wsize_bsize (sz); + }else{ + CAMLassert (Color_hd (q) == Caml_blue); + p += Whsize_hd (q); + } + } + ch = Chunk_next (ch); + } + } + + /* Shrink the heap if needed. */ + { + /* Find the amount of live data and the unshrinkable free space. */ + asize_t live = 0; + asize_t free = 0; + asize_t wanted; + + ch = caml_heap_start; + while (ch != NULL){ + if (Chunk_alloc (ch) != 0){ + live += Wsize_bsize (Chunk_alloc (ch)); + free += Wsize_bsize (Chunk_size (ch) - Chunk_alloc (ch)); + } + ch = Chunk_next (ch); + } + + /* Add up the empty chunks until there are enough, then remove the + other empty chunks. */ + wanted = caml_percent_free * (live / 100 + 1); + ch = caml_heap_start; + while (ch != NULL){ + char *next_chunk = Chunk_next (ch); /* Chunk_next (ch) will be erased */ + + if (Chunk_alloc (ch) == 0){ + if (free < wanted){ + free += Wsize_bsize (Chunk_size (ch)); + }else{ + caml_shrink_heap (ch); + } + } + ch = next_chunk; + } + } + + /* Rebuild the free list. */ + { + ch = caml_heap_start; + caml_fl_reset (); + while (ch != NULL){ + if (Chunk_size (ch) > Chunk_alloc (ch)){ + caml_make_free_blocks ((value *) (ch + Chunk_alloc (ch)), + Wsize_bsize (Chunk_size(ch)-Chunk_alloc(ch)), 1, + Caml_white); + } + ch = Chunk_next (ch); + } + } + ++ caml_stat_compactions; + caml_gc_message (0x10, "done.\n"); +} + +uintnat caml_percent_max; /* used in gc_ctrl.c and memory.c */ + +void caml_compact_heap (void) +{ + uintnat target_wsz, live; + CAML_INSTR_SETUP(tmr, "compact"); + + CAMLassert (caml_young_ptr == caml_young_alloc_end); + CAMLassert (caml_ref_table.ptr == caml_ref_table.base); + CAMLassert (caml_ephe_ref_table.ptr == caml_ephe_ref_table.base); + CAMLassert (caml_custom_table.ptr == caml_custom_table.base); + + do_compaction (); + CAML_INSTR_TIME (tmr, "compact/main"); + /* Compaction may fail to shrink the heap to a reasonable size + because it deals in complete chunks: if a very large chunk + is at the beginning of the heap, everything gets moved to + it and it is not freed. + + In that case, we allocate a new chunk of the desired heap + size, chain it at the beginning of the heap (thus pretending + its address is smaller), and launch a second compaction. + This will move all data to this new chunk and free the + very large chunk. + + See PR#5389 + */ + /* We compute: + freewords = caml_fl_cur_wsz (exact) + heapwords = Wsize_bsize (caml_heap_size) (exact) + live = heapwords - freewords + wanted = caml_percent_free * (live / 100 + 1) (same as in do_compaction) + target_wsz = live + wanted + We add one page to make sure a small difference in counting sizes + won't make [do_compaction] keep the second block (and break all sorts + of invariants). + + We recompact if target_wsz < heap_size / 2 + */ + live = caml_stat_heap_wsz - caml_fl_cur_wsz; + target_wsz = live + caml_percent_free * (live / 100 + 1) + + Wsize_bsize (Page_size); + target_wsz = caml_clip_heap_chunk_wsz (target_wsz); + +#ifdef HAS_HUGE_PAGES + if (caml_use_huge_pages + && Bsize_wsize (caml_stat_heap_wsz) <= HUGE_PAGE_SIZE) + return; +#endif + + if (target_wsz < caml_stat_heap_wsz / 2){ + /* Recompact. */ + char *chunk; + + caml_gc_message (0x10, "Recompacting heap (target=%" + ARCH_INTNAT_PRINTF_FORMAT "uk words)\n", + target_wsz / 1024); + + chunk = caml_alloc_for_heap (Bsize_wsize (target_wsz)); + if (chunk == NULL) return; + /* PR#5757: we need to make the new blocks blue, or they won't be + recognized as free by the recompaction. */ + caml_make_free_blocks ((value *) chunk, + Wsize_bsize (Chunk_size (chunk)), 0, Caml_blue); + if (caml_page_table_add (In_heap, chunk, chunk + Chunk_size (chunk)) != 0){ + caml_free_for_heap (chunk); + return; + } + Chunk_next (chunk) = caml_heap_start; + caml_heap_start = chunk; + ++ caml_stat_heap_chunks; + caml_stat_heap_wsz += Wsize_bsize (Chunk_size (chunk)); + if (caml_stat_heap_wsz > caml_stat_top_heap_wsz){ + caml_stat_top_heap_wsz = caml_stat_heap_wsz; + } + do_compaction (); + CAMLassert (caml_stat_heap_chunks == 1); + CAMLassert (Chunk_next (caml_heap_start) == NULL); + CAMLassert (caml_stat_heap_wsz == Wsize_bsize (Chunk_size (chunk))); + CAML_INSTR_TIME (tmr, "compact/recompact"); + } +} + +void caml_compact_heap_maybe (void) +{ + /* Estimated free+garbage words in the heap: + FW = fl_size_at_phase_change + 3 * (caml_fl_cur_wsz + - caml_fl_wsz_at_phase_change) + FW = 3 * caml_fl_cur_wsz - 2 * caml_fl_wsz_at_phase_change + Estimated live words: LW = caml_stat_heap_wsz - FW + Estimated free percentage: FP = 100 * FW / LW + We compact the heap if FP > caml_percent_max + */ + float fw, fp; + CAMLassert (caml_gc_phase == Phase_idle); + if (caml_percent_max >= 1000000) return; + if (caml_stat_major_collections < 3) return; + if (caml_stat_heap_wsz <= 2 * caml_clip_heap_chunk_wsz (0)) return; + +#ifdef HAS_HUGE_PAGES + if (caml_use_huge_pages + && Bsize_wsize (caml_stat_heap_wsz) <= HUGE_PAGE_SIZE) + return; +#endif + + fw = 3.0 * caml_fl_cur_wsz - 2.0 * caml_fl_wsz_at_phase_change; + if (fw < 0) fw = caml_fl_cur_wsz; + + if (fw >= caml_stat_heap_wsz){ + fp = 1000000.0; + }else{ + fp = 100.0 * fw / (caml_stat_heap_wsz - fw); + if (fp > 1000000.0) fp = 1000000.0; + } + caml_gc_message (0x200, "FL size at phase change = %" + ARCH_INTNAT_PRINTF_FORMAT "u words\n", + (uintnat) caml_fl_wsz_at_phase_change); + caml_gc_message (0x200, "FL current size = %" + ARCH_INTNAT_PRINTF_FORMAT "u words\n", + (uintnat) caml_fl_cur_wsz); + caml_gc_message (0x200, "Estimated overhead = %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", + (uintnat) fp); + if (fp >= caml_percent_max){ + caml_gc_message (0x200, "Automatic compaction triggered.\n"); + caml_empty_minor_heap (); /* minor heap must be empty for compaction */ + caml_finish_major_cycle (); + + fw = caml_fl_cur_wsz; + fp = 100.0 * fw / (caml_stat_heap_wsz - fw); + caml_gc_message (0x200, "Measured overhead: %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", + (uintnat) fp); + if (fp >= caml_percent_max) + caml_compact_heap (); + else + caml_gc_message (0x200, "Automatic compaction aborted.\n"); + + } +} diff --git a/test/monniaux/ocaml/byterun/compare.c b/test/monniaux/ocaml/byterun/compare.c new file mode 100644 index 00000000..382c9dff --- /dev/null +++ b/test/monniaux/ocaml/byterun/compare.c @@ -0,0 +1,363 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include +#include +#include "caml/custom.h" +#include "caml/fail.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" + +#if defined(LACKS_SANE_NAN) && !defined(isnan) +#define isnan _isnan +#endif + +/* Structural comparison on trees. */ + +struct compare_item { value * v1, * v2; mlsize_t count; }; + +#define COMPARE_STACK_INIT_SIZE 8 +#define COMPARE_STACK_MIN_ALLOC_SIZE 32 +#define COMPARE_STACK_MAX_SIZE (1024*1024) +CAMLexport int caml_compare_unordered; + +struct compare_stack { + struct compare_item init_stack[COMPARE_STACK_INIT_SIZE]; + struct compare_item* stack; + struct compare_item* limit; +}; + +/* Free the compare stack if needed */ +static void compare_free_stack(struct compare_stack* stk) +{ + if (stk->stack != stk->init_stack) { + caml_stat_free(stk->stack); + stk->stack = NULL; + } +} + +/* Same, then raise Out_of_memory */ +static void compare_stack_overflow(struct compare_stack* stk) +{ + caml_gc_message (0x04, "Stack overflow in structural comparison\n"); + compare_free_stack(stk); + caml_raise_out_of_memory(); +} + +/* Grow the compare stack */ +static struct compare_item * compare_resize_stack(struct compare_stack* stk, + struct compare_item * sp) +{ + asize_t newsize; + asize_t sp_offset = sp - stk->stack; + struct compare_item * newstack; + + if (stk->stack == stk->init_stack) { + newsize = COMPARE_STACK_MIN_ALLOC_SIZE; + newstack = caml_stat_alloc_noexc(sizeof(struct compare_item) * newsize); + if (newstack == NULL) compare_stack_overflow(stk); + memcpy(newstack, stk->init_stack, + sizeof(struct compare_item) * COMPARE_STACK_INIT_SIZE); + } else { + newsize = 2 * (stk->limit - stk->stack); + if (newsize >= COMPARE_STACK_MAX_SIZE) compare_stack_overflow(stk); + newstack = caml_stat_resize_noexc(stk->stack, + sizeof(struct compare_item) * newsize); + if (newstack == NULL) compare_stack_overflow(stk); + } + stk->stack = newstack; + stk->limit = newstack + newsize; + return newstack + sp_offset; +} + + +static intnat do_compare_val(struct compare_stack* stk, + value v1, value v2, int total); + +static intnat compare_val(value v1, value v2, int total) +{ + struct compare_stack stk; + intnat res; + stk.stack = stk.init_stack; + stk.limit = stk.stack + COMPARE_STACK_INIT_SIZE; + res = do_compare_val(&stk, v1, v2, total); + compare_free_stack(&stk); + return res; +} + +/* Structural comparison */ + + +#define LESS -1 +#define EQUAL 0 +#define GREATER 1 +#define UNORDERED ((intnat)1 << (8 * sizeof(value) - 1)) + +/* The return value of compare_val is as follows: + > 0 v1 is greater than v2 + 0 v1 is equal to v2 + < 0 and > UNORDERED v1 is less than v2 + UNORDERED v1 and v2 cannot be compared */ + +static intnat do_compare_val(struct compare_stack* stk, + value v1, value v2, int total) +{ + struct compare_item * sp; + tag_t t1, t2; + + sp = stk->stack; + while (1) { + if (v1 == v2 && total) goto next_item; + if (Is_long(v1)) { + if (v1 == v2) goto next_item; + if (Is_long(v2)) + return Long_val(v1) - Long_val(v2); + /* Subtraction above cannot overflow and cannot result in UNORDERED */ + if (Is_in_value_area(v2)) { + switch (Tag_val(v2)) { + case Forward_tag: + v2 = Forward_val(v2); + continue; + case Custom_tag: { + int res; + int (*compare)(value v1, value v2) = Custom_ops_val(v2)->compare_ext; + if (compare == NULL) break; /* for backward compatibility */ + caml_compare_unordered = 0; + res = compare(v1, v2); + if (caml_compare_unordered && !total) return UNORDERED; + if (res != 0) return res; + goto next_item; + } + default: /*fallthrough*/; + } + } + return LESS; /* v1 long < v2 block */ + } + if (Is_long(v2)) { + if (Is_in_value_area(v1)) { + switch (Tag_val(v1)) { + case Forward_tag: + v1 = Forward_val(v1); + continue; + case Custom_tag: { + int res; + int (*compare)(value v1, value v2) = Custom_ops_val(v1)->compare_ext; + if (compare == NULL) break; /* for backward compatibility */ + caml_compare_unordered = 0; + res = compare(v1, v2); + if (caml_compare_unordered && !total) return UNORDERED; + if (res != 0) return res; + goto next_item; + } + default: /*fallthrough*/; + } + } + return GREATER; /* v1 block > v2 long */ + } + /* If one of the objects is outside the heap (but is not an atom), + use address comparison. Since both addresses are 2-aligned, + shift lsb off to avoid overflow in subtraction. */ + if (! Is_in_value_area(v1) || ! Is_in_value_area(v2)) { + if (v1 == v2) goto next_item; + return (v1 >> 1) - (v2 >> 1); + /* Subtraction above cannot result in UNORDERED */ + } + t1 = Tag_val(v1); + t2 = Tag_val(v2); + if (t1 == Forward_tag) { v1 = Forward_val (v1); continue; } + if (t2 == Forward_tag) { v2 = Forward_val (v2); continue; } + if (t1 != t2) return (intnat)t1 - (intnat)t2; + switch(t1) { + case String_tag: { + mlsize_t len1, len2; + int res; + if (v1 == v2) break; + len1 = caml_string_length(v1); + len2 = caml_string_length(v2); + res = memcmp(String_val(v1), String_val(v2), len1 <= len2 ? len1 : len2); + if (res < 0) return LESS; + if (res > 0) return GREATER; + if (len1 != len2) return len1 - len2; + break; + } + case Double_tag: { + double d1 = Double_val(v1); + double d2 = Double_val(v2); +#ifdef LACKS_SANE_NAN + if (isnan(d2)) { + if (! total) return UNORDERED; + if (isnan(d1)) break; + return GREATER; + } else if (isnan(d1)) { + if (! total) return UNORDERED; + return LESS; + } +#endif + if (d1 < d2) return LESS; + if (d1 > d2) return GREATER; +#ifndef LACKS_SANE_NAN + if (d1 != d2) { + if (! total) return UNORDERED; + /* One or both of d1 and d2 is NaN. Order according to the + convention NaN = NaN and NaN < f for all other floats f. */ + if (d1 == d1) return GREATER; /* d1 is not NaN, d2 is NaN */ + if (d2 == d2) return LESS; /* d2 is not NaN, d1 is NaN */ + /* d1 and d2 are both NaN, thus equal: continue comparison */ + } +#endif + break; + } + case Double_array_tag: { + mlsize_t sz1 = Wosize_val(v1) / Double_wosize; + mlsize_t sz2 = Wosize_val(v2) / Double_wosize; + mlsize_t i; + if (sz1 != sz2) return sz1 - sz2; + for (i = 0; i < sz1; i++) { + double d1 = Double_flat_field(v1, i); + double d2 = Double_flat_field(v2, i); + #ifdef LACKS_SANE_NAN + if (isnan(d2)) { + if (! total) return UNORDERED; + if (isnan(d1)) break; + return GREATER; + } else if (isnan(d1)) { + if (! total) return UNORDERED; + return LESS; + } + #endif + if (d1 < d2) return LESS; + if (d1 > d2) return GREATER; + #ifndef LACKS_SANE_NAN + if (d1 != d2) { + if (! total) return UNORDERED; + /* See comment for Double_tag case */ + if (d1 == d1) return GREATER; + if (d2 == d2) return LESS; + } + #endif + } + break; + } + case Abstract_tag: + compare_free_stack(stk); + caml_invalid_argument("compare: abstract value"); + case Closure_tag: + case Infix_tag: + compare_free_stack(stk); + caml_invalid_argument("compare: functional value"); + case Object_tag: { + intnat oid1 = Oid_val(v1); + intnat oid2 = Oid_val(v2); + if (oid1 != oid2) return oid1 - oid2; + break; + } + case Custom_tag: { + int res; + int (*compare)(value v1, value v2) = Custom_ops_val(v1)->compare; + /* Hardening against comparisons between different types */ + if (compare != Custom_ops_val(v2)->compare) { + return strcmp(Custom_ops_val(v1)->identifier, + Custom_ops_val(v2)->identifier) < 0 + ? LESS : GREATER; + } + if (compare == NULL) { + compare_free_stack(stk); + caml_invalid_argument("compare: abstract value"); + } + caml_compare_unordered = 0; + res = compare(v1, v2); + if (caml_compare_unordered && !total) return UNORDERED; + if (res != 0) return res; + break; + } + default: { + mlsize_t sz1 = Wosize_val(v1); + mlsize_t sz2 = Wosize_val(v2); + /* Compare sizes first for speed */ + if (sz1 != sz2) return sz1 - sz2; + if (sz1 == 0) break; + /* Remember that we still have to compare fields 1 ... sz - 1 */ + if (sz1 > 1) { + sp++; + if (sp >= stk->limit) sp = compare_resize_stack(stk, sp); + sp->v1 = &Field(v1, 1); + sp->v2 = &Field(v2, 1); + sp->count = sz1 - 1; + } + /* Continue comparison with first field */ + v1 = Field(v1, 0); + v2 = Field(v2, 0); + continue; + } + } + next_item: + /* Pop one more item to compare, if any */ + if (sp == stk->stack) return EQUAL; /* we're done */ + v1 = *((sp->v1)++); + v2 = *((sp->v2)++); + if (--(sp->count) == 0) sp--; + } +} + +CAMLprim value caml_compare(value v1, value v2) +{ + intnat res = compare_val(v1, v2, 1); + /* Free stack if needed */ + if (res < 0) + return Val_int(LESS); + else if (res > 0) + return Val_int(GREATER); + else + return Val_int(EQUAL); +} + +CAMLprim value caml_equal(value v1, value v2) +{ + intnat res = compare_val(v1, v2, 0); + return Val_int(res == 0); +} + +CAMLprim value caml_notequal(value v1, value v2) +{ + intnat res = compare_val(v1, v2, 0); + return Val_int(res != 0); +} + +CAMLprim value caml_lessthan(value v1, value v2) +{ + intnat res = compare_val(v1, v2, 0); + return Val_int(res < 0 && res != UNORDERED); +} + +CAMLprim value caml_lessequal(value v1, value v2) +{ + intnat res = compare_val(v1, v2, 0); + return Val_int(res <= 0 && res != UNORDERED); +} + +CAMLprim value caml_greaterthan(value v1, value v2) +{ + intnat res = compare_val(v1, v2, 0); + return Val_int(res > 0); +} + +CAMLprim value caml_greaterequal(value v1, value v2) +{ + intnat res = compare_val(v1, v2, 0); + return Val_int(res >= 0); +} diff --git a/test/monniaux/ocaml/byterun/custom.c b/test/monniaux/ocaml/byterun/custom.c new file mode 100644 index 00000000..b6a5c4e3 --- /dev/null +++ b/test/monniaux/ocaml/byterun/custom.c @@ -0,0 +1,124 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Manuel Serrano and Xavier Leroy, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include + +#include "caml/alloc.h" +#include "caml/custom.h" +#include "caml/fail.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/signals.h" + +/* [size] is a number of bytes */ +CAMLexport value caml_alloc_custom(struct custom_operations * ops, + uintnat size, + mlsize_t mem, + mlsize_t max) +{ + mlsize_t wosize; + CAMLparam0(); + CAMLlocal1(result); + + wosize = 1 + (size + sizeof(value) - 1) / sizeof(value); + if (wosize <= Max_young_wosize) { + result = caml_alloc_small(wosize, Custom_tag); + Custom_ops_val(result) = ops; + if (ops->finalize != NULL || mem != 0) { + /* Remember that the block needs processing after minor GC. */ + add_to_custom_table (&caml_custom_table, result, mem, max); + /* Keep track of extra resources held by custom block in + minor heap. */ + if (mem != 0) { + if (max == 0) max = 1; + caml_extra_heap_resources_minor += (double) mem / (double) max; + if (caml_extra_heap_resources_minor > 1.0) { + caml_request_minor_gc (); + caml_gc_dispatch (); + } + } + } + } else { + result = caml_alloc_shr(wosize, Custom_tag); + Custom_ops_val(result) = ops; + caml_adjust_gc_speed(mem, max); + result = caml_check_urgent_gc(result); + } + CAMLreturn(result); +} + +struct custom_operations_list { + struct custom_operations * ops; + struct custom_operations_list * next; +}; + +static struct custom_operations_list * custom_ops_table = NULL; + +CAMLexport void caml_register_custom_operations(struct custom_operations * ops) +{ + struct custom_operations_list * l = + caml_stat_alloc(sizeof(struct custom_operations_list)); + CAMLassert(ops->identifier != NULL); + CAMLassert(ops->deserialize != NULL); + l->ops = ops; + l->next = custom_ops_table; + custom_ops_table = l; +} + +struct custom_operations * caml_find_custom_operations(char * ident) +{ + struct custom_operations_list * l; + for (l = custom_ops_table; l != NULL; l = l->next) + if (strcmp(l->ops->identifier, ident) == 0) return l->ops; + return NULL; +} + +static struct custom_operations_list * custom_ops_final_table = NULL; + +struct custom_operations * caml_final_custom_operations(final_fun fn) +{ + struct custom_operations_list * l; + struct custom_operations * ops; + for (l = custom_ops_final_table; l != NULL; l = l->next) + if (l->ops->finalize == fn) return l->ops; + ops = caml_stat_alloc(sizeof(struct custom_operations)); + ops->identifier = "_final"; + ops->finalize = fn; + ops->compare = custom_compare_default; + ops->hash = custom_hash_default; + ops->serialize = custom_serialize_default; + ops->deserialize = custom_deserialize_default; + ops->compare_ext = custom_compare_ext_default; + l = caml_stat_alloc(sizeof(struct custom_operations_list)); + l->ops = ops; + l->next = custom_ops_final_table; + custom_ops_final_table = l; + return ops; +} + +extern struct custom_operations caml_int32_ops, + caml_nativeint_ops, + caml_int64_ops, + caml_ba_ops; + +void caml_init_custom_operations(void) +{ + caml_register_custom_operations(&caml_int32_ops); + caml_register_custom_operations(&caml_nativeint_ops); + caml_register_custom_operations(&caml_int64_ops); + caml_register_custom_operations(&caml_ba_ops); +} diff --git a/test/monniaux/ocaml/byterun/debugger.c b/test/monniaux/ocaml/byterun/debugger.c new file mode 100644 index 00000000..1c416cd6 --- /dev/null +++ b/test/monniaux/ocaml/byterun/debugger.c @@ -0,0 +1,454 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Interface with the byte-code debugger */ + +#ifdef _WIN32 +#include +#endif /* _WIN32 */ + +#include + +#include "caml/alloc.h" +#include "caml/config.h" +#include "caml/debugger.h" +#include "caml/misc.h" +#include "caml/osdeps.h" + +int caml_debugger_in_use = 0; +uintnat caml_event_count; +int caml_debugger_fork_mode = 1; /* parent by default */ + +#if !defined(HAS_SOCKETS) || defined(NATIVE_CODE) + +void caml_debugger_init(void) +{ +} + +void caml_debugger(enum event_kind event) +{ +} + +void caml_debugger_cleanup_fork(void) +{ +} + +#else + +#ifdef HAS_UNISTD +#include +#endif +#include +#include +#ifndef _WIN32 +#include +#include +#include +#include +#include +#include +#else +#define ATOM ATOM_WS +#include +#undef ATOM +#include +#endif + +#include "caml/fail.h" +#include "caml/fix_code.h" +#include "caml/instruct.h" +#include "caml/intext.h" +#include "caml/io.h" +#include "caml/mlvalues.h" +#include "caml/stacks.h" +#include "caml/sys.h" + +static value marshal_flags = Val_emptylist; + +static int sock_domain; /* Socket domain for the debugger */ +static union { /* Socket address for the debugger */ + struct sockaddr s_gen; +#ifndef _WIN32 + struct sockaddr_un s_unix; +#endif + struct sockaddr_in s_inet; +} sock_addr; +static int sock_addr_len; /* Length of sock_addr */ + +static int dbg_socket = -1; /* The socket connected to the debugger */ +static struct channel * dbg_in; /* Input channel on the socket */ +static struct channel * dbg_out;/* Output channel on the socket */ + +static char *dbg_addr = NULL; + +static void open_connection(void) +{ +#ifdef _WIN32 + /* Set socket to synchronous mode so that file descriptor-oriented + functions (read()/write() etc.) can be used */ + + int oldvalue, oldvaluelen, newvalue, retcode; + oldvaluelen = sizeof(oldvalue); + retcode = getsockopt(INVALID_SOCKET, SOL_SOCKET, SO_OPENTYPE, + (char *) &oldvalue, &oldvaluelen); + if (retcode == 0) { + newvalue = SO_SYNCHRONOUS_NONALERT; + setsockopt(INVALID_SOCKET, SOL_SOCKET, SO_OPENTYPE, + (char *) &newvalue, sizeof(newvalue)); + } +#endif + dbg_socket = socket(sock_domain, SOCK_STREAM, 0); +#ifdef _WIN32 + if (retcode == 0) { + /* Restore initial mode */ + setsockopt(INVALID_SOCKET, SOL_SOCKET, SO_OPENTYPE, + (char *) &oldvalue, oldvaluelen); + } +#endif + if (dbg_socket == -1 || + connect(dbg_socket, &sock_addr.s_gen, sock_addr_len) == -1){ + caml_fatal_error_arg2 ("cannot connect to debugger at %s\n", (dbg_addr ? dbg_addr : "(none)"), + "error: %s\n", strerror (errno)); + } +#ifdef _WIN32 + dbg_socket = _open_osfhandle(dbg_socket, 0); + if (dbg_socket == -1) + caml_fatal_error("_open_osfhandle failed"); +#endif + dbg_in = caml_open_descriptor_in(dbg_socket); + dbg_out = caml_open_descriptor_out(dbg_socket); + if (!caml_debugger_in_use) caml_putword(dbg_out, -1); /* first connection */ +#ifdef _WIN32 + caml_putword(dbg_out, _getpid()); +#else + caml_putword(dbg_out, getpid()); +#endif + caml_flush(dbg_out); +} + +static void close_connection(void) +{ + caml_close_channel(dbg_in); + caml_close_channel(dbg_out); + dbg_socket = -1; /* was closed by caml_close_channel */ +} + +#ifdef _WIN32 +static void winsock_startup(void) +{ + WSADATA wsaData; + int err = WSAStartup(MAKEWORD(2, 0), &wsaData); + if (err) caml_fatal_error("WSAStartup failed"); +} + +static void winsock_cleanup(void) +{ + WSACleanup(); +} +#endif + +void caml_debugger_init(void) +{ + char * address; + char_os * a; + size_t a_len; + char * port, * p; + struct hostent * host; + int n; + + caml_register_global_root(&marshal_flags); + marshal_flags = caml_alloc(2, Tag_cons); + Store_field(marshal_flags, 0, Val_int(1)); /* Marshal.Closures */ + Store_field(marshal_flags, 1, Val_emptylist); + + a = caml_secure_getenv(_T("CAML_DEBUG_SOCKET")); + address = a ? caml_stat_strdup_of_os(a) : NULL; + if (address == NULL) return; + if (dbg_addr != NULL) caml_stat_free(dbg_addr); + dbg_addr = address; + +#ifdef _WIN32 + winsock_startup(); + (void)atexit(winsock_cleanup); +#endif + /* Parse the address */ + port = NULL; + for (p = address; *p != 0; p++) { + if (*p == ':') { *p = 0; port = p+1; break; } + } + if (port == NULL) { +#ifndef _WIN32 + /* Unix domain */ + sock_domain = PF_UNIX; + sock_addr.s_unix.sun_family = AF_UNIX; + a_len = strlen(address); + if (a_len >= sizeof(sock_addr.s_unix.sun_path)) { + caml_fatal_error("Debug socket path length exceeds maximum permitted length"); + } + strncpy(sock_addr.s_unix.sun_path, address, + sizeof(sock_addr.s_unix.sun_path) - 1); + sock_addr.s_unix.sun_path[sizeof(sock_addr.s_unix.sun_path) - 1] = '\0'; + sock_addr_len = + ((char *)&(sock_addr.s_unix.sun_path) - (char *)&(sock_addr.s_unix)) + + a_len; +#else + caml_fatal_error("Unix sockets not supported"); +#endif + } else { + /* Internet domain */ + sock_domain = PF_INET; + for (p = (char *) &sock_addr.s_inet, n = sizeof(sock_addr.s_inet); + n > 0; n--) *p++ = 0; + sock_addr.s_inet.sin_family = AF_INET; + sock_addr.s_inet.sin_addr.s_addr = inet_addr(address); + if (sock_addr.s_inet.sin_addr.s_addr == -1) { + host = gethostbyname(address); + if (host == NULL) + caml_fatal_error_arg("Unknown debugging host %s\n", address); + memmove(&sock_addr.s_inet.sin_addr, host->h_addr, host->h_length); + } + sock_addr.s_inet.sin_port = htons(atoi(port)); + sock_addr_len = sizeof(sock_addr.s_inet); + } + open_connection(); + caml_debugger_in_use = 1; + caml_trap_barrier = caml_stack_high; +} + +static value getval(struct channel *chan) +{ + value res; + if (caml_really_getblock(chan, (char *) &res, sizeof(res)) < sizeof(res)) + caml_raise_end_of_file(); /* Bad, but consistent with caml_getword */ + return res; +} + +static void putval(struct channel *chan, value val) +{ + caml_really_putblock(chan, (char *) &val, sizeof(val)); +} + +static void safe_output_value(struct channel *chan, value val) +{ + struct longjmp_buffer raise_buf, * saved_external_raise; + + /* Catch exceptions raised by [caml_output_val] */ + saved_external_raise = caml_external_raise; + if (sigsetjmp(raise_buf.buf, 0) == 0) { + caml_external_raise = &raise_buf; + caml_output_val(chan, val, marshal_flags); + } else { + /* Send wrong magic number, will cause [caml_input_value] to fail */ + caml_really_putblock(chan, "\000\000\000\000", 4); + } + caml_external_raise = saved_external_raise; +} + +#define Pc(sp) ((code_t)((sp)[0])) +#define Env(sp) ((sp)[1]) +#define Extra_args(sp) (Long_val(((sp)[2]))) +#define Locals(sp) ((sp) + 3) + +void caml_debugger(enum event_kind event) +{ + value * frame; + intnat i, pos; + value val; + + if (dbg_socket == -1) return; /* Not connected to a debugger. */ + + /* Reset current frame */ + frame = caml_extern_sp + 1; + + /* Report the event to the debugger */ + switch(event) { + case PROGRAM_START: /* Nothing to report */ + goto command_loop; + case EVENT_COUNT: + caml_putch(dbg_out, REP_EVENT); + break; + case BREAKPOINT: + caml_putch(dbg_out, REP_BREAKPOINT); + break; + case PROGRAM_EXIT: + caml_putch(dbg_out, REP_EXITED); + break; + case TRAP_BARRIER: + caml_putch(dbg_out, REP_TRAP); + break; + case UNCAUGHT_EXC: + caml_putch(dbg_out, REP_UNCAUGHT_EXC); + break; + } + caml_putword(dbg_out, caml_event_count); + if (event == EVENT_COUNT || event == BREAKPOINT) { + caml_putword(dbg_out, caml_stack_high - frame); + caml_putword(dbg_out, (Pc(frame) - caml_start_code) * sizeof(opcode_t)); + } else { + /* No PC and no stack frame associated with other events */ + caml_putword(dbg_out, 0); + caml_putword(dbg_out, 0); + } + caml_flush(dbg_out); + + command_loop: + + /* Read and execute the commands sent by the debugger */ + while(1) { + switch(caml_getch(dbg_in)) { + case REQ_SET_EVENT: + pos = caml_getword(dbg_in); + CAMLassert (pos >= 0); + CAMLassert (pos < caml_code_size); + caml_set_instruction(caml_start_code + pos / sizeof(opcode_t), EVENT); + break; + case REQ_SET_BREAKPOINT: + pos = caml_getword(dbg_in); + CAMLassert (pos >= 0); + CAMLassert (pos < caml_code_size); + caml_set_instruction(caml_start_code + pos / sizeof(opcode_t), BREAK); + break; + case REQ_RESET_INSTR: + pos = caml_getword(dbg_in); + CAMLassert (pos >= 0); + CAMLassert (pos < caml_code_size); + pos = pos / sizeof(opcode_t); + caml_set_instruction(caml_start_code + pos, caml_saved_code[pos]); + break; + case REQ_CHECKPOINT: +#ifndef _WIN32 + i = fork(); + if (i == 0) { + close_connection(); /* Close parent connection. */ + open_connection(); /* Open new connection with debugger */ + } else { + caml_putword(dbg_out, i); + caml_flush(dbg_out); + } +#else + caml_fatal_error("error: REQ_CHECKPOINT command"); + exit(-1); +#endif + break; + case REQ_GO: + caml_event_count = caml_getword(dbg_in); + return; + case REQ_STOP: + exit(0); + break; + case REQ_WAIT: +#ifndef _WIN32 + wait(NULL); +#else + caml_fatal_error("Fatal error: REQ_WAIT command"); + exit(-1); +#endif + break; + case REQ_INITIAL_FRAME: + frame = caml_extern_sp + 1; + /* Fall through */ + case REQ_GET_FRAME: + caml_putword(dbg_out, caml_stack_high - frame); + if (frame < caml_stack_high){ + caml_putword(dbg_out, (Pc(frame) - caml_start_code) * sizeof(opcode_t)); + }else{ + caml_putword (dbg_out, 0); + } + caml_flush(dbg_out); + break; + case REQ_SET_FRAME: + i = caml_getword(dbg_in); + frame = caml_stack_high - i; + break; + case REQ_UP_FRAME: + i = caml_getword(dbg_in); + if (frame + Extra_args(frame) + i + 3 >= caml_stack_high) { + caml_putword(dbg_out, -1); + } else { + frame += Extra_args(frame) + i + 3; + caml_putword(dbg_out, caml_stack_high - frame); + caml_putword(dbg_out, (Pc(frame) - caml_start_code) * sizeof(opcode_t)); + } + caml_flush(dbg_out); + break; + case REQ_SET_TRAP_BARRIER: + i = caml_getword(dbg_in); + caml_trap_barrier = caml_stack_high - i; + break; + case REQ_GET_LOCAL: + i = caml_getword(dbg_in); + putval(dbg_out, Locals(frame)[i]); + caml_flush(dbg_out); + break; + case REQ_GET_ENVIRONMENT: + i = caml_getword(dbg_in); + putval(dbg_out, Field(Env(frame), i)); + caml_flush(dbg_out); + break; + case REQ_GET_GLOBAL: + i = caml_getword(dbg_in); + putval(dbg_out, Field(caml_global_data, i)); + caml_flush(dbg_out); + break; + case REQ_GET_ACCU: + putval(dbg_out, *caml_extern_sp); + caml_flush(dbg_out); + break; + case REQ_GET_HEADER: + val = getval(dbg_in); + caml_putword(dbg_out, Hd_val(val)); + caml_flush(dbg_out); + break; + case REQ_GET_FIELD: + val = getval(dbg_in); + i = caml_getword(dbg_in); + if (Tag_val(val) != Double_array_tag) { + caml_putch(dbg_out, 0); + putval(dbg_out, Field(val, i)); + } else { + double d = Double_flat_field(val, i); + caml_putch(dbg_out, 1); + caml_really_putblock(dbg_out, (char *) &d, 8); + } + caml_flush(dbg_out); + break; + case REQ_MARSHAL_OBJ: + val = getval(dbg_in); + safe_output_value(dbg_out, val); + caml_flush(dbg_out); + break; + case REQ_GET_CLOSURE_CODE: + val = getval(dbg_in); + caml_putword(dbg_out, (Code_val(val)-caml_start_code) * sizeof(opcode_t)); + caml_flush(dbg_out); + break; + case REQ_SET_FORK_MODE: + caml_debugger_fork_mode = caml_getword(dbg_in); + break; + } + } +} + +void caml_debugger_cleanup_fork(void) +{ + /* We could remove all of the breakpoints, but closing the connection + * means that they'll just be skipped anyway. */ + close_connection(); + caml_debugger_in_use = 0; +} + +#endif diff --git a/test/monniaux/ocaml/byterun/dynlink.c b/test/monniaux/ocaml/byterun/dynlink.c new file mode 100644 index 00000000..7c339bf5 --- /dev/null +++ b/test/monniaux/ocaml/byterun/dynlink.c @@ -0,0 +1,300 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Dynamic loading of C primitives. */ + +#include +#include +#include +#include +#include +#include "caml/config.h" +#ifdef HAS_UNISTD +#include +#endif +#include "caml/alloc.h" +#include "caml/dynlink.h" +#include "caml/fail.h" +#include "caml/mlvalues.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/osdeps.h" +#include "caml/prims.h" +#include "caml/signals.h" + +#ifndef NATIVE_CODE + +/* The table of primitives */ +struct ext_table caml_prim_table; + +#ifdef DEBUG +/* The names of primitives (for instrtrace.c) */ +struct ext_table caml_prim_name_table; +#endif + +/* The table of shared libraries currently opened */ +static struct ext_table shared_libs; + +/* The search path for shared libraries */ +struct ext_table caml_shared_libs_path; + +/* Look up the given primitive name in the built-in primitive table, + then in the opened shared libraries (shared_libs) */ +static c_primitive lookup_primitive(char * name) +{ + int i; + void * res; + + for (i = 0; caml_names_of_builtin_cprim[i] != NULL; i++) { + if (strcmp(name, caml_names_of_builtin_cprim[i]) == 0) + return caml_builtin_cprim[i]; + } + for (i = 0; i < shared_libs.size; i++) { + res = caml_dlsym(shared_libs.contents[i], name); + if (res != NULL) return (c_primitive) res; + } + return NULL; +} + +/* Parse the OCAML_STDLIB_DIR/ld.conf file and add the directories + listed there to the search path */ + +#define LD_CONF_NAME _T("ld.conf") + +static char_os * parse_ld_conf(void) +{ + char_os * stdlib, * ldconfname, * wconfig, * p, * q; + char * config; +#ifdef _WIN32 + struct _stati64 st; +#else + struct stat st; +#endif + int ldconf, nread; + + stdlib = caml_secure_getenv(_T("OCAMLLIB")); + if (stdlib == NULL) stdlib = caml_secure_getenv(_T("CAMLLIB")); + if (stdlib == NULL) stdlib = OCAML_STDLIB_DIR; + ldconfname = caml_stat_strconcat_os(3, stdlib, _T("/"), LD_CONF_NAME); + if (stat_os(ldconfname, &st) == -1) { + caml_stat_free(ldconfname); + return NULL; + } + ldconf = open_os(ldconfname, O_RDONLY, 0); + if (ldconf == -1) + caml_fatal_error_arg("Fatal error: cannot read loader config file %s\n", + caml_stat_strdup_of_os(ldconfname)); + config = caml_stat_alloc(st.st_size + 1); + nread = read(ldconf, config, st.st_size); + if (nread == -1) + caml_fatal_error_arg + ("Fatal error: error while reading loader config file %s\n", + caml_stat_strdup_of_os(ldconfname)); + config[nread] = 0; + wconfig = caml_stat_strdup_to_os(config); + caml_stat_free(config); + q = wconfig; + for (p = wconfig; *p != 0; p++) { + if (*p == _T('\n')) { + *p = 0; + caml_ext_table_add(&caml_shared_libs_path, q); + q = p + 1; + } + } + if (q < p) caml_ext_table_add(&caml_shared_libs_path, q); + close(ldconf); + caml_stat_free(ldconfname); + return wconfig; +} + +/* Open the given shared library and add it to shared_libs. + Abort on error. */ +static void open_shared_lib(char_os * name) +{ + char_os * realname; + char * u8; + void * handle; + + realname = caml_search_dll_in_path(&caml_shared_libs_path, name); + u8 = caml_stat_strdup_of_os(realname); + caml_gc_message(0x100, "Loading shared library %s\n", u8); + caml_stat_free(u8); + caml_enter_blocking_section(); + handle = caml_dlopen(realname, 1, 1); + caml_leave_blocking_section(); + if (handle == NULL) + caml_fatal_error_arg2("Fatal error: cannot load shared library %s\n", + caml_stat_strdup_of_os(name), + "Reason: %s\n", caml_dlerror()); + caml_ext_table_add(&shared_libs, handle); + caml_stat_free(realname); +} + +/* Build the table of primitives, given a search path and a list + of shared libraries (both 0-separated in a char array). + Abort the runtime system on error. */ +void caml_build_primitive_table(char_os * lib_path, + char_os * libs, + char * req_prims) +{ + char_os * tofree1, * tofree2; + char_os * p; + char * q; + + /* Initialize the search path for dynamic libraries: + - directories specified on the command line with the -I option + - directories specified in the CAML_LD_LIBRARY_PATH + - directories specified in the executable + - directories specified in the file /ld.conf */ + tofree1 = caml_decompose_path(&caml_shared_libs_path, + caml_secure_getenv(_T("CAML_LD_LIBRARY_PATH"))); + if (lib_path != NULL) + for (p = lib_path; *p != 0; p += strlen_os(p) + 1) + caml_ext_table_add(&caml_shared_libs_path, p); + tofree2 = parse_ld_conf(); + /* Open the shared libraries */ + caml_ext_table_init(&shared_libs, 8); + if (libs != NULL) + for (p = libs; *p != 0; p += strlen_os(p) + 1) + open_shared_lib(p); + /* Build the primitive table */ + caml_ext_table_init(&caml_prim_table, 0x180); +#ifdef DEBUG + caml_ext_table_init(&caml_prim_name_table, 0x180); +#endif + for (q = req_prims; *q != 0; q += strlen(q) + 1) { + c_primitive prim = lookup_primitive(q); + if (prim == NULL) + caml_fatal_error_arg("Fatal error: unknown C primitive `%s'\n", q); + caml_ext_table_add(&caml_prim_table, (void *) prim); +#ifdef DEBUG + caml_ext_table_add(&caml_prim_name_table, caml_stat_strdup(q)); +#endif + } + /* Clean up */ + caml_stat_free(tofree1); + caml_stat_free(tofree2); + caml_ext_table_free(&caml_shared_libs_path, 0); +} + +/* Build the table of primitives as a copy of the builtin primitive table. + Used for executables generated by ocamlc -output-obj. */ + +void caml_build_primitive_table_builtin(void) +{ + int i; + caml_ext_table_init(&caml_prim_table, 0x180); +#ifdef DEBUG + caml_ext_table_init(&caml_prim_name_table, 0x180); +#endif + for (i = 0; caml_builtin_cprim[i] != 0; i++) { + caml_ext_table_add(&caml_prim_table, (void *) caml_builtin_cprim[i]); +#ifdef DEBUG + caml_ext_table_add(&caml_prim_name_table, + caml_stat_strdup(caml_names_of_builtin_cprim[i])); +#endif + } +} + +void caml_free_shared_libs(void) +{ + while (shared_libs.size > 0) + caml_dlclose(shared_libs.contents[--shared_libs.size]); +} + +#endif /* NATIVE_CODE */ + +/** dlopen interface for the bytecode linker **/ + +#define Handle_val(v) (*((void **) (v))) + +CAMLprim value caml_dynlink_open_lib(value mode, value filename) +{ + void * handle; + value result; + char_os * p; + + caml_gc_message(0x100, "Opening shared library %s\n", + String_val(filename)); + p = caml_stat_strdup_to_os(String_val(filename)); + caml_enter_blocking_section(); + handle = caml_dlopen(p, Int_val(mode), 1); + caml_leave_blocking_section(); + caml_stat_free(p); + if (handle == NULL) caml_failwith(caml_dlerror()); + result = caml_alloc_small(1, Abstract_tag); + Handle_val(result) = handle; + return result; +} + +CAMLprim value caml_dynlink_close_lib(value handle) +{ + caml_dlclose(Handle_val(handle)); + return Val_unit; +} + +/*#include */ +CAMLprim value caml_dynlink_lookup_symbol(value handle, value symbolname) +{ + void * symb; + value result; + symb = caml_dlsym(Handle_val(handle), String_val(symbolname)); + /* printf("%s = 0x%lx\n", String_val(symbolname), symb); + fflush(stdout); */ + if (symb == NULL) return Val_unit /*caml_failwith(caml_dlerror())*/; + result = caml_alloc_small(1, Abstract_tag); + Handle_val(result) = symb; + return result; +} + +#ifndef NATIVE_CODE + +CAMLprim value caml_dynlink_add_primitive(value handle) +{ + return Val_int(caml_ext_table_add(&caml_prim_table, Handle_val(handle))); +} + +CAMLprim value caml_dynlink_get_current_libs(value unit) +{ + CAMLparam0(); + CAMLlocal1(res); + int i; + + res = caml_alloc_tuple(shared_libs.size); + for (i = 0; i < shared_libs.size; i++) { + value v = caml_alloc_small(1, Abstract_tag); + Handle_val(v) = shared_libs.contents[i]; + Store_field(res, i, v); + } + CAMLreturn(res); +} + +#else + +value caml_dynlink_add_primitive(value handle) +{ + caml_invalid_argument("dynlink_add_primitive"); + return Val_unit; /* not reached */ +} + +value caml_dynlink_get_current_libs(value unit) +{ + caml_invalid_argument("dynlink_get_current_libs"); + return Val_unit; /* not reached */ +} + +#endif /* NATIVE_CODE */ diff --git a/test/monniaux/ocaml/byterun/extern.c b/test/monniaux/ocaml/byterun/extern.c new file mode 100644 index 00000000..db7163cd --- /dev/null +++ b/test/monniaux/ocaml/byterun/extern.c @@ -0,0 +1,925 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Structured output */ + +/* The interface of this file is "caml/intext.h" */ + +#include +#include "caml/alloc.h" +#include "caml/config.h" +#include "caml/custom.h" +#include "caml/fail.h" +#include "caml/gc.h" +#include "caml/intext.h" +#include "caml/io.h" +#include "caml/md5.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/reverse.h" + +static uintnat obj_counter; /* Number of objects emitted so far */ +static uintnat size_32; /* Size in words of 32-bit block for struct. */ +static uintnat size_64; /* Size in words of 64-bit block for struct. */ + +/* Flags affecting marshaling */ + +enum { + NO_SHARING = 1, /* Flag to ignore sharing */ + CLOSURES = 2, /* Flag to allow marshaling code pointers */ + COMPAT_32 = 4 /* Flag to ensure that output can safely + be read back on a 32-bit platform */ +}; + +static int extern_flags; /* logical or of some of the flags above */ + +/* Trail mechanism to undo forwarding pointers put inside objects */ + +struct trail_entry { + value obj; /* address of object + initial color in low 2 bits */ + value field0; /* initial contents of field 0 */ +}; + +struct trail_block { + struct trail_block * previous; + struct trail_entry entries[ENTRIES_PER_TRAIL_BLOCK]; +}; + +static struct trail_block extern_trail_first; +static struct trail_block * extern_trail_block; +static struct trail_entry * extern_trail_cur, * extern_trail_limit; + + +/* Stack for pending values to marshal */ + +struct extern_item { value * v; mlsize_t count; }; + +#define EXTERN_STACK_INIT_SIZE 256 +#define EXTERN_STACK_MAX_SIZE (1024*1024*100) + +static struct extern_item extern_stack_init[EXTERN_STACK_INIT_SIZE]; + +static struct extern_item * extern_stack = extern_stack_init; +static struct extern_item * extern_stack_limit = extern_stack_init + + EXTERN_STACK_INIT_SIZE; + +/* Forward declarations */ + +CAMLnoreturn_start +static void extern_out_of_memory(void) +CAMLnoreturn_end; + +CAMLnoreturn_start +static void extern_invalid_argument(char *msg) +CAMLnoreturn_end; + +CAMLnoreturn_start +static void extern_failwith(char *msg) +CAMLnoreturn_end; + +CAMLnoreturn_start +static void extern_stack_overflow(void) +CAMLnoreturn_end; + +static void extern_replay_trail(void); +static void free_extern_output(void); + +/* Free the extern stack if needed */ +static void extern_free_stack(void) +{ + if (extern_stack != extern_stack_init) { + caml_stat_free(extern_stack); + /* Reinitialize the globals for next time around */ + extern_stack = extern_stack_init; + extern_stack_limit = extern_stack + EXTERN_STACK_INIT_SIZE; + } +} + +static struct extern_item * extern_resize_stack(struct extern_item * sp) +{ + asize_t newsize = 2 * (extern_stack_limit - extern_stack); + asize_t sp_offset = sp - extern_stack; + struct extern_item * newstack; + + if (newsize >= EXTERN_STACK_MAX_SIZE) extern_stack_overflow(); + if (extern_stack == extern_stack_init) { + newstack = caml_stat_alloc_noexc(sizeof(struct extern_item) * newsize); + if (newstack == NULL) extern_stack_overflow(); + memcpy(newstack, extern_stack_init, + sizeof(struct extern_item) * EXTERN_STACK_INIT_SIZE); + } else { + newstack = caml_stat_resize_noexc(extern_stack, + sizeof(struct extern_item) * newsize); + if (newstack == NULL) extern_stack_overflow(); + } + extern_stack = newstack; + extern_stack_limit = newstack + newsize; + return newstack + sp_offset; +} + +/* Initialize the trail */ + +static void init_extern_trail(void) +{ + extern_trail_block = &extern_trail_first; + extern_trail_cur = extern_trail_block->entries; + extern_trail_limit = extern_trail_block->entries + ENTRIES_PER_TRAIL_BLOCK; +} + +/* Replay the trail, undoing the in-place modifications + performed on objects */ + +static void extern_replay_trail(void) +{ + struct trail_block * blk, * prevblk; + struct trail_entry * ent, * lim; + + blk = extern_trail_block; + lim = extern_trail_cur; + while (1) { + for (ent = &(blk->entries[0]); ent < lim; ent++) { + value obj = ent->obj; + color_t colornum = obj & 3; + obj = obj & ~3; + Hd_val(obj) = Coloredhd_hd(Hd_val(obj), colornum); + Field(obj, 0) = ent->field0; + } + if (blk == &extern_trail_first) break; + prevblk = blk->previous; + caml_stat_free(blk); + blk = prevblk; + lim = &(blk->entries[ENTRIES_PER_TRAIL_BLOCK]); + } + /* Protect against a second call to extern_replay_trail */ + extern_trail_block = &extern_trail_first; + extern_trail_cur = extern_trail_block->entries; +} + +/* Set forwarding pointer on an object and add corresponding entry + to the trail. */ + +static void extern_record_location(value obj) +{ + header_t hdr; + + if (extern_flags & NO_SHARING) return; + if (extern_trail_cur == extern_trail_limit) { + struct trail_block * new_block = caml_stat_alloc_noexc(sizeof(struct trail_block)); + if (new_block == NULL) extern_out_of_memory(); + new_block->previous = extern_trail_block; + extern_trail_block = new_block; + extern_trail_cur = extern_trail_block->entries; + extern_trail_limit = extern_trail_block->entries + ENTRIES_PER_TRAIL_BLOCK; + } + hdr = Hd_val(obj); + extern_trail_cur->obj = obj | Colornum_hd(hdr); + extern_trail_cur->field0 = Field(obj, 0); + extern_trail_cur++; + Hd_val(obj) = Bluehd_hd(hdr); + Field(obj, 0) = (value) obj_counter; + obj_counter++; +} + +/* To buffer the output */ + +static char * extern_userprovided_output; +static char * extern_ptr, * extern_limit; + +struct output_block { + struct output_block * next; + char * end; + char data[SIZE_EXTERN_OUTPUT_BLOCK]; +}; + +static struct output_block * extern_output_first, * extern_output_block; + +static void init_extern_output(void) +{ + extern_userprovided_output = NULL; + extern_output_first = caml_stat_alloc_noexc(sizeof(struct output_block)); + if (extern_output_first == NULL) caml_raise_out_of_memory(); + extern_output_block = extern_output_first; + extern_output_block->next = NULL; + extern_ptr = extern_output_block->data; + extern_limit = extern_output_block->data + SIZE_EXTERN_OUTPUT_BLOCK; +} + +static void close_extern_output(void) +{ + if (extern_userprovided_output == NULL){ + extern_output_block->end = extern_ptr; + } +} + +static void free_extern_output(void) +{ + struct output_block * blk, * nextblk; + + if (extern_userprovided_output != NULL) return; + for (blk = extern_output_first; blk != NULL; blk = nextblk) { + nextblk = blk->next; + caml_stat_free(blk); + } + extern_output_first = NULL; + extern_free_stack(); +} + +static void grow_extern_output(intnat required) +{ + struct output_block * blk; + intnat extra; + + if (extern_userprovided_output != NULL) { + extern_failwith("Marshal.to_buffer: buffer overflow"); + } + extern_output_block->end = extern_ptr; + if (required <= SIZE_EXTERN_OUTPUT_BLOCK / 2) + extra = 0; + else + extra = required; + blk = caml_stat_alloc_noexc(sizeof(struct output_block) + extra); + if (blk == NULL) extern_out_of_memory(); + extern_output_block->next = blk; + extern_output_block = blk; + extern_output_block->next = NULL; + extern_ptr = extern_output_block->data; + extern_limit = extern_output_block->data + SIZE_EXTERN_OUTPUT_BLOCK + extra; +} + +static intnat extern_output_length(void) +{ + struct output_block * blk; + intnat len; + + if (extern_userprovided_output != NULL) { + return extern_ptr - extern_userprovided_output; + } else { + for (len = 0, blk = extern_output_first; blk != NULL; blk = blk->next) + len += blk->end - blk->data; + return len; + } +} + +/* Exception raising, with cleanup */ + +static void extern_out_of_memory(void) +{ + extern_replay_trail(); + free_extern_output(); + caml_raise_out_of_memory(); +} + +static void extern_invalid_argument(char *msg) +{ + extern_replay_trail(); + free_extern_output(); + caml_invalid_argument(msg); +} + +static void extern_failwith(char *msg) +{ + extern_replay_trail(); + free_extern_output(); + caml_failwith(msg); +} + +static void extern_stack_overflow(void) +{ + caml_gc_message (0x04, "Stack overflow in marshaling value\n"); + extern_replay_trail(); + free_extern_output(); + caml_raise_out_of_memory(); +} + +/* Conversion to big-endian */ + +static inline void store16(char * dst, int n) +{ + dst[0] = n >> 8; dst[1] = n; +} + +static inline void store32(char * dst, intnat n) +{ + dst[0] = n >> 24; dst[1] = n >> 16; dst[2] = n >> 8; dst[3] = n; +} + +static inline void store64(char * dst, int64_t n) +{ + dst[0] = n >> 56; dst[1] = n >> 48; dst[2] = n >> 40; dst[3] = n >> 32; + dst[4] = n >> 24; dst[5] = n >> 16; dst[6] = n >> 8; dst[7] = n; +} + +/* Write characters, integers, and blocks in the output buffer */ + +static inline void write(int c) +{ + if (extern_ptr >= extern_limit) grow_extern_output(1); + *extern_ptr++ = c; +} + +static void writeblock(const char * data, intnat len) +{ + if (extern_ptr + len > extern_limit) grow_extern_output(len); + memcpy(extern_ptr, data, len); + extern_ptr += len; +} + +static inline void writeblock_float8(const double * data, intnat ndoubles) +{ +#if ARCH_FLOAT_ENDIANNESS == 0x01234567 || ARCH_FLOAT_ENDIANNESS == 0x76543210 + writeblock((const char *) data, ndoubles * 8); +#else + caml_serialize_block_float_8(data, ndoubles); +#endif +} + +static void writecode8(int code, intnat val) +{ + if (extern_ptr + 2 > extern_limit) grow_extern_output(2); + extern_ptr[0] = code; + extern_ptr[1] = val; + extern_ptr += 2; +} + +static void writecode16(int code, intnat val) +{ + if (extern_ptr + 3 > extern_limit) grow_extern_output(3); + extern_ptr[0] = code; + store16(extern_ptr + 1, val); + extern_ptr += 3; +} + +static void writecode32(int code, intnat val) +{ + if (extern_ptr + 5 > extern_limit) grow_extern_output(5); + extern_ptr[0] = code; + store32(extern_ptr + 1, val); + extern_ptr += 5; +} + +#ifdef ARCH_SIXTYFOUR +static void writecode64(int code, intnat val) +{ + if (extern_ptr + 9 > extern_limit) grow_extern_output(9); + extern_ptr[0] = code; + store64(extern_ptr + 1, val); + extern_ptr += 9; +} +#endif + +/* Marshal the given value in the output buffer */ + +int caml_extern_allow_out_of_heap = 0; + +static void extern_rec(value v) +{ + struct code_fragment * cf; + struct extern_item * sp; + sp = extern_stack; + + while(1) { + if (Is_long(v)) { + intnat n = Long_val(v); + if (n >= 0 && n < 0x40) { + write(PREFIX_SMALL_INT + n); + } else if (n >= -(1 << 7) && n < (1 << 7)) { + writecode8(CODE_INT8, n); + } else if (n >= -(1 << 15) && n < (1 << 15)) { + writecode16(CODE_INT16, n); +#ifdef ARCH_SIXTYFOUR + } else if (n < -((intnat)1 << 30) || n >= ((intnat)1 << 30)) { + if (extern_flags & COMPAT_32) + extern_failwith("output_value: integer cannot be read back on " + "32-bit platform"); + writecode64(CODE_INT64, n); +#endif + } else + writecode32(CODE_INT32, n); + goto next_item; + } + if (Is_in_value_area(v) || caml_extern_allow_out_of_heap) { + header_t hd = Hd_val(v); + tag_t tag = Tag_hd(hd); + mlsize_t sz = Wosize_hd(hd); + + if (tag == Forward_tag) { + value f = Forward_val (v); + if (Is_block (f) + && (!Is_in_value_area(f) || Tag_val (f) == Forward_tag + || Tag_val (f) == Lazy_tag +#ifdef FLAT_FLOAT_ARRAY + || Tag_val (f) == Double_tag +#endif + )){ + /* Do not short-circuit the pointer. */ + }else{ + v = f; + continue; + } + } + /* Atoms are treated specially for two reasons: they are not allocated + in the externed block, and they are automatically shared. */ + if (sz == 0) { + if (tag < 16) { + write(PREFIX_SMALL_BLOCK + tag); + } else { +#ifdef WITH_PROFINFO + writecode32(CODE_BLOCK32, Hd_no_profinfo(hd)); +#else + writecode32(CODE_BLOCK32, hd); +#endif + } + goto next_item; + } + /* Check if already seen */ + if (Color_hd(hd) == Caml_blue) { + uintnat d = obj_counter - (uintnat) Field(v, 0); + if (d < 0x100) { + writecode8(CODE_SHARED8, d); + } else if (d < 0x10000) { + writecode16(CODE_SHARED16, d); +#ifdef ARCH_SIXTYFOUR + } else if (d >= (uintnat)1 << 32) { + writecode64(CODE_SHARED64, d); +#endif + } else { + writecode32(CODE_SHARED32, d); + } + goto next_item; + } + + /* Output the contents of the object */ + switch(tag) { + case String_tag: { + mlsize_t len = caml_string_length(v); + if (len < 0x20) { + write(PREFIX_SMALL_STRING + len); + } else if (len < 0x100) { + writecode8(CODE_STRING8, len); + } else { +#ifdef ARCH_SIXTYFOUR + if (len > 0xFFFFFB && (extern_flags & COMPAT_32)) + extern_failwith("output_value: string cannot be read back on " + "32-bit platform"); + if (len < (uintnat)1 << 32) + writecode32(CODE_STRING32, len); + else + writecode64(CODE_STRING64, len); +#else + writecode32(CODE_STRING32, len); +#endif + } + writeblock(String_val(v), len); + size_32 += 1 + (len + 4) / 4; + size_64 += 1 + (len + 8) / 8; + extern_record_location(v); + break; + } + case Double_tag: { + if (sizeof(double) != 8) + extern_invalid_argument("output_value: non-standard floats"); + write(CODE_DOUBLE_NATIVE); + writeblock_float8((double *) v, 1); + size_32 += 1 + 2; + size_64 += 1 + 1; + extern_record_location(v); + break; + } + case Double_array_tag: { + mlsize_t nfloats; + if (sizeof(double) != 8) + extern_invalid_argument("output_value: non-standard floats"); + nfloats = Wosize_val(v) / Double_wosize; + if (nfloats < 0x100) { + writecode8(CODE_DOUBLE_ARRAY8_NATIVE, nfloats); + } else { +#ifdef ARCH_SIXTYFOUR + if (nfloats > 0x1FFFFF && (extern_flags & COMPAT_32)) + extern_failwith("output_value: float array cannot be read back on " + "32-bit platform"); + if (nfloats < (uintnat) 1 << 32) + writecode32(CODE_DOUBLE_ARRAY32_NATIVE, nfloats); + else + writecode64(CODE_DOUBLE_ARRAY64_NATIVE, nfloats); +#else + writecode32(CODE_DOUBLE_ARRAY32_NATIVE, nfloats); +#endif + } + writeblock_float8((double *) v, nfloats); + size_32 += 1 + nfloats * 2; + size_64 += 1 + nfloats; + extern_record_location(v); + break; + } + case Abstract_tag: + extern_invalid_argument("output_value: abstract value (Abstract)"); + break; + case Infix_tag: + writecode32(CODE_INFIXPOINTER, Infix_offset_hd(hd)); + v = v - Infix_offset_hd(hd); /* PR#5772 */ + continue; + case Custom_tag: { + uintnat sz_32, sz_64; + char * ident = Custom_ops_val(v)->identifier; + void (*serialize)(value v, uintnat * bsize_32, + uintnat * bsize_64) + = Custom_ops_val(v)->serialize; + if (serialize == NULL) + extern_invalid_argument("output_value: abstract value (Custom)"); + write(CODE_CUSTOM); + writeblock(ident, strlen(ident) + 1); + Custom_ops_val(v)->serialize(v, &sz_32, &sz_64); + size_32 += 2 + ((sz_32 + 3) >> 2); /* header + ops + data */ + size_64 += 2 + ((sz_64 + 7) >> 3); + extern_record_location(v); + break; + } + default: { + value field0; + if (tag < 16 && sz < 8) { + write(PREFIX_SMALL_BLOCK + tag + (sz << 4)); + } else { +#ifdef ARCH_SIXTYFOUR +#ifdef WITH_PROFINFO + header_t hd_erased = Hd_no_profinfo(hd); +#else + header_t hd_erased = hd; +#endif + if (sz > 0x3FFFFF && (extern_flags & COMPAT_32)) + extern_failwith("output_value: array cannot be read back on " + "32-bit platform"); + if (hd_erased < (uintnat)1 << 32) + writecode32(CODE_BLOCK32, Whitehd_hd (hd_erased)); + else + writecode64(CODE_BLOCK64, Whitehd_hd (hd_erased)); +#else + writecode32(CODE_BLOCK32, Whitehd_hd (hd)); +#endif + } + size_32 += 1 + sz; + size_64 += 1 + sz; + field0 = Field(v, 0); + extern_record_location(v); + /* Remember that we still have to serialize fields 1 ... sz - 1 */ + if (sz > 1) { + sp++; + if (sp >= extern_stack_limit) sp = extern_resize_stack(sp); + sp->v = &Field(v,1); + sp->count = sz-1; + } + /* Continue serialization with the first field */ + v = field0; + continue; + } + } + } + else if ((cf = caml_extern_find_code((char *) v)) != NULL) { + if ((extern_flags & CLOSURES) == 0) + extern_invalid_argument("output_value: functional value"); + writecode32(CODE_CODEPOINTER, (char *) v - cf->code_start); + writeblock((const char *)cf->digest, 16); + } else { + extern_invalid_argument("output_value: abstract value (outside heap)"); + } + next_item: + /* Pop one more item to marshal, if any */ + if (sp == extern_stack) { + /* We are done. Cleanup the stack and leave the function */ + extern_free_stack(); + return; + } + v = *((sp->v)++); + if (--(sp->count) == 0) sp--; + } + /* Never reached as function leaves with return */ +} + +static int extern_flag_values[] = { NO_SHARING, CLOSURES, COMPAT_32 }; + +static intnat extern_value(value v, value flags, + /*out*/ char header[32], + /*out*/ int * header_len) +{ + intnat res_len; + /* Parse flag list */ + extern_flags = caml_convert_flag_list(flags, extern_flag_values); + /* Initializations */ + init_extern_trail(); + obj_counter = 0; + size_32 = 0; + size_64 = 0; + /* Marshal the object */ + extern_rec(v); + /* Record end of output */ + close_extern_output(); + /* Undo the modifications done on externed blocks */ + extern_replay_trail(); + /* Write the header */ + res_len = extern_output_length(); +#ifdef ARCH_SIXTYFOUR + if (res_len >= ((intnat)1 << 32) || + size_32 >= ((intnat)1 << 32) || size_64 >= ((intnat)1 << 32)) { + /* The object is too big for the small header format. + Fail if we are in compat32 mode, or use big header. */ + if (extern_flags & COMPAT_32) { + free_extern_output(); + caml_failwith("output_value: object too big to be read back on " + "32-bit platform"); + } + store32(header, Intext_magic_number_big); + store32(header + 4, 0); + store64(header + 8, res_len); + store64(header + 16, obj_counter); + store64(header + 24, size_64); + *header_len = 32; + return res_len; + } +#endif + /* Use the small header format */ + store32(header, Intext_magic_number_small); + store32(header + 4, res_len); + store32(header + 8, obj_counter); + store32(header + 12, size_32); + store32(header + 16, size_64); + *header_len = 20; + return res_len; +} + +void caml_output_val(struct channel *chan, value v, value flags) +{ + char header[32]; + int header_len; + struct output_block * blk, * nextblk; + + if (! caml_channel_binary_mode(chan)) + caml_failwith("output_value: not a binary channel"); + init_extern_output(); + extern_value(v, flags, header, &header_len); + /* During [caml_really_putblock], concurrent [caml_output_val] operations + can take place (via signal handlers or context switching in systhreads), + and [extern_output_first] may change. So, save it in a local variable. */ + blk = extern_output_first; + caml_really_putblock(chan, header, header_len); + while (blk != NULL) { + caml_really_putblock(chan, blk->data, blk->end - blk->data); + nextblk = blk->next; + caml_stat_free(blk); + blk = nextblk; + } +} + +CAMLprim value caml_output_value(value vchan, value v, value flags) +{ + CAMLparam3 (vchan, v, flags); + struct channel * channel = Channel(vchan); + + Lock(channel); + caml_output_val(channel, v, flags); + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_output_value_to_bytes(value v, value flags) +{ + char header[32]; + int header_len; + intnat data_len, ofs; + value res; + struct output_block * blk, * nextblk; + + init_extern_output(); + data_len = extern_value(v, flags, header, &header_len); + /* PR#4030: it is prudent to save extern_output_first before allocating + the result, as in caml_output_val */ + blk = extern_output_first; + res = caml_alloc_string(header_len + data_len); + ofs = 0; + memcpy(&Byte(res, ofs), header, header_len); + ofs += header_len; + while (blk != NULL) { + int n = blk->end - blk->data; + memcpy(&Byte(res, ofs), blk->data, n); + ofs += n; + nextblk = blk->next; + caml_stat_free(blk); + blk = nextblk; + } + return res; +} + +CAMLprim value caml_output_value_to_string(value v, value flags) +{ + return caml_output_value_to_bytes(v,flags); +} + +CAMLexport intnat caml_output_value_to_block(value v, value flags, + char * buf, intnat len) +{ + char header[32]; + int header_len; + intnat data_len; + /* At this point we don't know the size of the header. + Guess that it is small, and fix up later if not. */ + extern_userprovided_output = buf + 20; + extern_ptr = extern_userprovided_output; + extern_limit = buf + len; + data_len = extern_value(v, flags, header, &header_len); + if (header_len != 20) { + /* Bad guess! Need to shift the output to make room for big header. + Make sure there is room. */ + if (header_len + data_len > len) + caml_failwith("Marshal.to_buffer: buffer overflow"); + memmove(buf + header_len, buf + 20, data_len); + } + memcpy(buf, header, header_len); + return header_len + data_len; +} + +CAMLprim value caml_output_value_to_buffer(value buf, value ofs, value len, + value v, value flags) +{ + intnat l = + caml_output_value_to_block(v, flags, + &Byte(buf, Long_val(ofs)), Long_val(len)); + return Val_long(l); +} + +CAMLexport void caml_output_value_to_malloc(value v, value flags, + /*out*/ char ** buf, + /*out*/ intnat * len) +{ + char header[32]; + int header_len; + intnat data_len; + char * res; + struct output_block * blk; + + init_extern_output(); + data_len = extern_value(v, flags, header, &header_len); + res = caml_stat_alloc_noexc(header_len + data_len); + if (res == NULL) extern_out_of_memory(); + *buf = res; + *len = header_len + data_len; + memcpy(res, header, header_len); + res += header_len; + for (blk = extern_output_first; blk != NULL; blk = blk->next) { + int n = blk->end - blk->data; + memcpy(res, blk->data, n); + res += n; + } + free_extern_output(); +} + +/* Functions for writing user-defined marshallers */ + +CAMLexport void caml_serialize_int_1(int i) +{ + if (extern_ptr + 1 > extern_limit) grow_extern_output(1); + extern_ptr[0] = i; + extern_ptr += 1; +} + +CAMLexport void caml_serialize_int_2(int i) +{ + if (extern_ptr + 2 > extern_limit) grow_extern_output(2); + store16(extern_ptr, i); + extern_ptr += 2; +} + +CAMLexport void caml_serialize_int_4(int32_t i) +{ + if (extern_ptr + 4 > extern_limit) grow_extern_output(4); + store32(extern_ptr, i); + extern_ptr += 4; +} + +CAMLexport void caml_serialize_int_8(int64_t i) +{ + if (extern_ptr + 8 > extern_limit) grow_extern_output(8); + store64(extern_ptr, i); + extern_ptr += 8; +} + +CAMLexport void caml_serialize_float_4(float f) +{ + caml_serialize_block_4(&f, 1); +} + +CAMLexport void caml_serialize_float_8(double f) +{ + caml_serialize_block_float_8(&f, 1); +} + +CAMLexport void caml_serialize_block_1(void * data, intnat len) +{ + if (extern_ptr + len > extern_limit) grow_extern_output(len); + memcpy(extern_ptr, data, len); + extern_ptr += len; +} + +CAMLexport void caml_serialize_block_2(void * data, intnat len) +{ + if (extern_ptr + 2 * len > extern_limit) grow_extern_output(2 * len); +#ifndef ARCH_BIG_ENDIAN + { + unsigned char * p; + char * q; + for (p = data, q = extern_ptr; len > 0; len--, p += 2, q += 2) + Reverse_16(q, p); + extern_ptr = q; + } +#else + memcpy(extern_ptr, data, len * 2); + extern_ptr += len * 2; +#endif +} + +CAMLexport void caml_serialize_block_4(void * data, intnat len) +{ + if (extern_ptr + 4 * len > extern_limit) grow_extern_output(4 * len); +#ifndef ARCH_BIG_ENDIAN + { + unsigned char * p; + char * q; + for (p = data, q = extern_ptr; len > 0; len--, p += 4, q += 4) + Reverse_32(q, p); + extern_ptr = q; + } +#else + memcpy(extern_ptr, data, len * 4); + extern_ptr += len * 4; +#endif +} + +CAMLexport void caml_serialize_block_8(void * data, intnat len) +{ + if (extern_ptr + 8 * len > extern_limit) grow_extern_output(8 * len); +#ifndef ARCH_BIG_ENDIAN + { + unsigned char * p; + char * q; + for (p = data, q = extern_ptr; len > 0; len--, p += 8, q += 8) + Reverse_64(q, p); + extern_ptr = q; + } +#else + memcpy(extern_ptr, data, len * 8); + extern_ptr += len * 8; +#endif +} + +CAMLexport void caml_serialize_block_float_8(void * data, intnat len) +{ + if (extern_ptr + 8 * len > extern_limit) grow_extern_output(8 * len); +#if ARCH_FLOAT_ENDIANNESS == 0x01234567 + memcpy(extern_ptr, data, len * 8); + extern_ptr += len * 8; +#elif ARCH_FLOAT_ENDIANNESS == 0x76543210 + { + unsigned char * p; + char * q; + for (p = data, q = extern_ptr; len > 0; len--, p += 8, q += 8) + Reverse_64(q, p); + extern_ptr = q; + } +#else + { + unsigned char * p; + char * q; + for (p = data, q = extern_ptr; len > 0; len--, p += 8, q += 8) + Permute_64(q, 0x01234567, p, ARCH_FLOAT_ENDIANNESS); + extern_ptr = q; + } +#endif +} + +/* Find where a code pointer comes from */ + +CAMLexport struct code_fragment * caml_extern_find_code(char *addr) +{ + int i; + for (i = caml_code_fragments_table.size - 1; i >= 0; i--) { + struct code_fragment * cf = caml_code_fragments_table.contents[i]; + if (! cf->digest_computed) { + caml_md5_block(cf->digest, cf->code_start, cf->code_end - cf->code_start); + cf->digest_computed = 1; + } + if (cf->code_start <= addr && addr < cf->code_end) return cf; + } + return NULL; +} diff --git a/test/monniaux/ocaml/byterun/fail.c b/test/monniaux/ocaml/byterun/fail.c new file mode 100644 index 00000000..6b4dee5f --- /dev/null +++ b/test/monniaux/ocaml/byterun/fail.c @@ -0,0 +1,204 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Raising exceptions from C. */ + +#include +#include +#include "caml/alloc.h" +#include "caml/fail.h" +#include "caml/io.h" +#include "caml/gc.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/printexc.h" +#include "caml/signals.h" +#include "caml/stacks.h" + +CAMLexport struct longjmp_buffer * caml_external_raise = NULL; +value caml_exn_bucket; + +CAMLexport void caml_raise(value v) +{ + Unlock_exn(); + caml_exn_bucket = v; + if (caml_external_raise == NULL) caml_fatal_uncaught_exception(v); + siglongjmp(caml_external_raise->buf, 1); +} + +CAMLexport void caml_raise_constant(value tag) +{ + caml_raise(tag); +} + +CAMLexport void caml_raise_with_arg(value tag, value arg) +{ + CAMLparam2 (tag, arg); + CAMLlocal1 (bucket); + + bucket = caml_alloc_small (2, 0); + Field(bucket, 0) = tag; + Field(bucket, 1) = arg; + caml_raise(bucket); + CAMLnoreturn; +} + +CAMLexport void caml_raise_with_args(value tag, int nargs, value args[]) +{ + CAMLparam1 (tag); + CAMLxparamN (args, nargs); + value bucket; + int i; + + CAMLassert(1 + nargs <= Max_young_wosize); + bucket = caml_alloc_small (1 + nargs, 0); + Field(bucket, 0) = tag; + for (i = 0; i < nargs; i++) Field(bucket, 1 + i) = args[i]; + caml_raise(bucket); + CAMLnoreturn; +} + +CAMLexport void caml_raise_with_string(value tag, char const *msg) +{ + CAMLparam1(tag); + value v_msg = caml_copy_string(msg); + caml_raise_with_arg(tag, v_msg); + CAMLnoreturn; +} + +/* PR#5115: Built-in exceptions can be triggered by input_value + while reading the initial value of [caml_global_data]. + + We check against this issue here in byterun/fail.c instead of + byterun/intern.c. Having the check here means that these calls will + be slightly slower for all bytecode programs (not just the calls + coming from intern). Because intern.c is shared between byterun/ + and asmrun/, putting checks there would slow do input_value for + natively-compiled programs that do not need these checks. +*/ +static void check_global_data(char const *exception_name) +{ + if (caml_global_data == 0) { + fprintf(stderr, "Fatal error: exception %s\n", exception_name); + exit(2); + } +} + +static void check_global_data_param(char const *exception_name, char const *msg) +{ + if (caml_global_data == 0) { + fprintf(stderr, "Fatal error: exception %s(\"%s\")\n", exception_name, msg); + exit(2); + } +} + +static inline value caml_get_failwith_tag (char const *msg) +{ + check_global_data_param("Failure", msg); + return Field(caml_global_data, FAILURE_EXN); +} + +CAMLexport void caml_failwith (char const *msg) +{ + caml_raise_with_string(caml_get_failwith_tag(msg), msg); +} + +CAMLexport void caml_failwith_value (value msg) +{ + CAMLparam1(msg); + value tag = caml_get_failwith_tag(String_val(msg)); + caml_raise_with_arg(tag, msg); + CAMLnoreturn; +} + +static inline value caml_get_invalid_argument_tag (char const *msg) +{ + check_global_data_param("Invalid_argument", msg); + return Field(caml_global_data, INVALID_EXN); +} + +CAMLexport void caml_invalid_argument (char const *msg) +{ + caml_raise_with_string(caml_get_invalid_argument_tag(msg), msg); +} + +CAMLexport void caml_invalid_argument_value (value msg) +{ + CAMLparam1(msg); + value tag = caml_get_invalid_argument_tag(String_val(msg)); + caml_raise_with_arg(tag, msg); + CAMLnoreturn; +} + +CAMLexport void caml_array_bound_error(void) +{ + caml_invalid_argument("index out of bounds"); +} + +CAMLexport void caml_raise_out_of_memory(void) +{ + check_global_data("Out_of_memory"); + caml_raise_constant(Field(caml_global_data, OUT_OF_MEMORY_EXN)); +} + +CAMLexport void caml_raise_stack_overflow(void) +{ + check_global_data("Stack_overflow"); + caml_raise_constant(Field(caml_global_data, STACK_OVERFLOW_EXN)); +} + +CAMLexport void caml_raise_sys_error(value msg) +{ + check_global_data_param("Sys_error", String_val(msg)); + caml_raise_with_arg(Field(caml_global_data, SYS_ERROR_EXN), msg); +} + +CAMLexport void caml_raise_end_of_file(void) +{ + check_global_data("End_of_file"); + caml_raise_constant(Field(caml_global_data, END_OF_FILE_EXN)); +} + +CAMLexport void caml_raise_zero_divide(void) +{ + check_global_data("Division_by_zero"); + caml_raise_constant(Field(caml_global_data, ZERO_DIVIDE_EXN)); +} + +CAMLexport void caml_raise_not_found(void) +{ + check_global_data("Not_found"); + caml_raise_constant(Field(caml_global_data, NOT_FOUND_EXN)); +} + +CAMLexport void caml_raise_sys_blocked_io(void) +{ + check_global_data("Sys_blocked_io"); + caml_raise_constant(Field(caml_global_data, SYS_BLOCKED_IO)); +} + +int caml_is_special_exception(value exn) { + /* this function is only used in caml_format_exception to produce + a more readable textual representation of some exceptions. It is + better to fall back to the general, less readable representation + than to abort with a fatal error as above. */ + if (caml_global_data == 0) return 0; + return exn == Field(caml_global_data, MATCH_FAILURE_EXN) + || exn == Field(caml_global_data, ASSERT_FAILURE_EXN) + || exn == Field(caml_global_data, UNDEFINED_RECURSIVE_MODULE_EXN); +} diff --git a/test/monniaux/ocaml/byterun/finalise.c b/test/monniaux/ocaml/byterun/finalise.c new file mode 100644 index 00000000..d34913fb --- /dev/null +++ b/test/monniaux/ocaml/byterun/finalise.c @@ -0,0 +1,445 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Moscova, INRIA Rocquencourt */ +/* */ +/* Copyright 2000 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Handling of finalised values. */ + +#include "caml/callback.h" +#include "caml/compact.h" +#include "caml/fail.h" +#include "caml/finalise.h" +#include "caml/minor_gc.h" +#include "caml/mlvalues.h" +#include "caml/roots.h" +#include "caml/signals.h" +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) +#include "caml/spacetime.h" +#endif + +struct final { + value fun; + value val; + int offset; +}; + +struct finalisable { + struct final *table; + uintnat old; + uintnat young; + uintnat size; +}; +/* [0..old) : finalisable set, the values are in the major heap + [old..young) : recent set, the values could be in the minor heap + [young..size) : free space + + The element of the finalisable set are moved to the finalising set + below when the value are unreachable (for the first or last time). + +*/ + +static struct finalisable finalisable_first = {NULL,0,0,0}; +static struct finalisable finalisable_last = {NULL,0,0,0}; + +struct to_do { + struct to_do *next; + int size; + struct final item[1]; /* variable size */ +}; + +static struct to_do *to_do_hd = NULL; +static struct to_do *to_do_tl = NULL; +/* + to_do_hd: head of the list of finalisation functions that can be run. + to_do_tl: tail of the list of finalisation functions that can be run. + + It is the finalising set. +*/ + + +/* [size] is a number of elements for the [to_do.item] array */ +static void alloc_to_do (int size) +{ + struct to_do *result = caml_stat_alloc_noexc (sizeof (struct to_do) + + size * sizeof (struct final)); + if (result == NULL) caml_fatal_error ("out of memory"); + result->next = NULL; + result->size = size; + if (to_do_tl == NULL){ + to_do_hd = result; + to_do_tl = result; + }else{ + CAMLassert (to_do_tl->next == NULL); + to_do_tl->next = result; + to_do_tl = result; + } +} + +/* Find white finalisable values, move them to the finalising set, and + darken them (if darken_value is true). +*/ +static void generic_final_update (struct finalisable * final, int darken_value) +{ + uintnat i, j, k; + uintnat todo_count = 0; + + CAMLassert (final->old <= final->young); + for (i = 0; i < final->old; i++){ + CAMLassert (Is_block (final->table[i].val)); + CAMLassert (Is_in_heap (final->table[i].val)); + if (Is_white_val (final->table[i].val)){ + ++ todo_count; + } + } + + /** invariant: + - 0 <= j <= i /\ 0 <= k <= i /\ 0 <= k <= todo_count + - i : index in final_table, before i all the values are black + (alive or in the minor heap) or the finalizer have been copied + in to_do_tl. + - j : index in final_table, before j all the values are black + (alive or in the minor heap), next available slot. + - k : index in to_do_tl, next available slot. + */ + if (todo_count > 0){ + alloc_to_do (todo_count); + j = k = 0; + for (i = 0; i < final->old; i++){ + CAMLassert (Is_block (final->table[i].val)); + CAMLassert (Is_in_heap (final->table[i].val)); + CAMLassert (Tag_val (final->table[i].val) != Forward_tag); + if(Is_white_val (final->table[i].val)){ + /** dead */ + to_do_tl->item[k] = final->table[i]; + if(!darken_value){ + /* The value is not darken so the finalisation function + is called with unit not with the value */ + to_do_tl->item[k].val = Val_unit; + to_do_tl->item[k].offset = 0; + }; + k++; + }else{ + /** alive */ + final->table[j++] = final->table[i]; + } + } + CAMLassert (i == final->old); + CAMLassert (k == todo_count); + final->old = j; + for(;i < final->young; i++){ + final->table[j++] = final->table[i]; + } + final->young = j; + to_do_tl->size = k; + if(darken_value){ + for (i = 0; i < k; i++){ + /* Note that item may already be dark due to multiple entries in + the final table. */ + caml_darken (to_do_tl->item[i].val, NULL); + } + } + } +} + +void caml_final_update_mark_phase (){ + generic_final_update(&finalisable_first, /* darken_value */ 1); +} + +void caml_final_update_clean_phase (){ + generic_final_update(&finalisable_last, /* darken_value */ 0); +} + + +static int running_finalisation_function = 0; + +/* Call the finalisation functions for the finalising set. + Note that this function must be reentrant. +*/ +void caml_final_do_calls (void) +{ + struct final f; + value res; +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + void* saved_spacetime_trie_node_ptr; +#endif + + if (running_finalisation_function) return; + if (to_do_hd != NULL){ + if (caml_finalise_begin_hook != NULL) (*caml_finalise_begin_hook) (); + caml_gc_message (0x80, "Calling finalisation functions.\n"); + while (1){ + while (to_do_hd != NULL && to_do_hd->size == 0){ + struct to_do *next_hd = to_do_hd->next; + caml_stat_free (to_do_hd); + to_do_hd = next_hd; + if (to_do_hd == NULL) to_do_tl = NULL; + } + if (to_do_hd == NULL) break; + CAMLassert (to_do_hd->size > 0); + -- to_do_hd->size; + f = to_do_hd->item[to_do_hd->size]; + running_finalisation_function = 1; +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + /* We record the finaliser's execution separately. + (The code of [caml_callback_exn] will do the hard work of finding + the correct place in the trie.) */ + saved_spacetime_trie_node_ptr = caml_spacetime_trie_node_ptr; + caml_spacetime_trie_node_ptr = caml_spacetime_finaliser_trie_root; +#endif + res = caml_callback_exn (f.fun, f.val + f.offset); +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + caml_spacetime_trie_node_ptr = saved_spacetime_trie_node_ptr; +#endif + running_finalisation_function = 0; + if (Is_exception_result (res)) caml_raise (Extract_exception (res)); + } + caml_gc_message (0x80, "Done calling finalisation functions.\n"); + if (caml_finalise_end_hook != NULL) (*caml_finalise_end_hook) (); + } +} + +/* Call a scanning_action [f] on [x]. */ +#define Call_action(f,x) (*(f)) ((x), &(x)) + +/* Call [*f] on the closures of the finalisable set and + the closures and values of the finalising set. + This is called by the major GC [caml_darken_all_roots] + and by the compactor through [caml_do_roots] +*/ +void caml_final_do_roots (scanning_action f) +{ + uintnat i; + struct to_do *todo; + + CAMLassert (finalisable_first.old <= finalisable_first.young); + for (i = 0; i < finalisable_first.young; i++){ + Call_action (f, finalisable_first.table[i].fun); + }; + + CAMLassert (finalisable_last.old <= finalisable_last.young); + for (i = 0; i < finalisable_last.young; i++){ + Call_action (f, finalisable_last.table[i].fun); + }; + + for (todo = to_do_hd; todo != NULL; todo = todo->next){ + for (i = 0; i < todo->size; i++){ + Call_action (f, todo->item[i].fun); + Call_action (f, todo->item[i].val); + } + } +} + +/* Call caml_invert_root on the values of the finalisable set. This is called + directly by the compactor. +*/ +void caml_final_invert_finalisable_values () +{ + uintnat i; + + CAMLassert (finalisable_first.old <= finalisable_first.young); + for (i = 0; i < finalisable_first.young; i++){ + caml_invert_root(finalisable_first.table[i].val, + &finalisable_first.table[i].val); + }; + + CAMLassert (finalisable_last.old <= finalisable_last.young); + for (i = 0; i < finalisable_last.young; i++){ + caml_invert_root(finalisable_last.table[i].val, + &finalisable_last.table[i].val); + }; +} + +/* Call [caml_oldify_one] on the closures and values of the recent set. + This is called by the minor GC through [caml_oldify_local_roots]. +*/ +void caml_final_oldify_young_roots () +{ + uintnat i; + + CAMLassert (finalisable_first.old <= finalisable_first.young); + for (i = finalisable_first.old; i < finalisable_first.young; i++){ + caml_oldify_one(finalisable_first.table[i].fun, + &finalisable_first.table[i].fun); + caml_oldify_one(finalisable_first.table[i].val, + &finalisable_first.table[i].val); + } + + CAMLassert (finalisable_last.old <= finalisable_last.young); + for (i = finalisable_last.old; i < finalisable_last.young; i++){ + caml_oldify_one(finalisable_last.table[i].fun, + &finalisable_last.table[i].fun); + } + +} + +static void generic_final_minor_update (struct finalisable * final) +{ + uintnat i, j, k; + uintnat todo_count = 0; + + CAMLassert (final->old <= final->young); + for (i = final->old; i < final->young; i++){ + CAMLassert (Is_block (final->table[i].val)); + CAMLassert (Is_in_heap_or_young (final->table[i].val)); + if (Is_young(final->table[i].val) && Hd_val(final->table[i].val) != 0){ + ++ todo_count; + } + } + + /** invariant: + - final->old <= j <= i /\ final->old <= k <= i /\ 0 <= k <= todo_count + - i : index in final_table, before i all the values are alive + or the finalizer have been copied in to_do_tl. + - j : index in final_table, before j all the values are alive, + next available slot. + - k : index in to_do_tl, next available slot. + */ + if (todo_count > 0){ + alloc_to_do (todo_count); + k = 0; + j = final->old; + for (i = final->old; i < final->young; i++){ + CAMLassert (Is_block (final->table[i].val)); + CAMLassert (Is_in_heap_or_young (final->table[i].val)); + CAMLassert (Tag_val (final->table[i].val) != Forward_tag); + if(Is_young(final->table[i].val) && Hd_val(final->table[i].val) != 0){ + /** dead */ + to_do_tl->item[k] = final->table[i]; + /* The finalisation function is called with unit not with the value */ + to_do_tl->item[k].val = Val_unit; + to_do_tl->item[k].offset = 0; + k++; + }else{ + /** alive */ + final->table[j++] = final->table[i]; + } + } + CAMLassert (i == final->young); + CAMLassert (k == todo_count); + final->young = j; + to_do_tl->size = todo_count; + } + + /** update the minor value to the copied major value */ + for (i = final->old; i < final->young; i++){ + CAMLassert (Is_block (final->table[i].val)); + CAMLassert (Is_in_heap_or_young (final->table[i].val)); + if (Is_young(final->table[i].val)) { + CAMLassert (Hd_val(final->table[i].val) == 0); + final->table[i].val = Field(final->table[i].val,0); + } + } + + /** check invariant */ + CAMLassert (final->old <= final->young); + for (i = 0; i < final->young; i++){ + CAMLassert( Is_in_heap(final->table[i].val) ); + }; + +} + +/* At the end of minor collection update the finalise_last roots in + minor heap when moved to major heap or moved them to the finalising + set when dead. +*/ +void caml_final_update_minor_roots () +{ + generic_final_minor_update(&finalisable_last); +} + +/* Empty the recent set into the finalisable set. + This is called at the end of each minor collection. + The minor heap must be empty when this is called. +*/ +void caml_final_empty_young (void) +{ + finalisable_first.old = finalisable_first.young; + finalisable_last.old = finalisable_last.young; +} + +/* Put (f,v) in the recent set. */ +static void generic_final_register (struct finalisable *final, value f, value v) +{ + if (!Is_block (v) + || !Is_in_heap_or_young(v) + || Tag_val (v) == Lazy_tag +#ifdef FLAT_FLOAT_ARRAY + || Tag_val (v) == Double_tag +#endif + || Tag_val (v) == Forward_tag) { + caml_invalid_argument ("Gc.finalise"); + } + CAMLassert (final->old <= final->young); + + if (final->young >= final->size){ + if (final->table == NULL){ + uintnat new_size = 30; + final->table = caml_stat_alloc (new_size * sizeof (struct final)); + CAMLassert (final->old == 0); + CAMLassert (final->young == 0); + final->size = new_size; + }else{ + uintnat new_size = final->size * 2; + final->table = caml_stat_resize (final->table, + new_size * sizeof (struct final)); + final->size = new_size; + } + } + CAMLassert (final->young < final->size); + final->table[final->young].fun = f; + if (Tag_val (v) == Infix_tag){ + final->table[final->young].offset = Infix_offset_val (v); + final->table[final->young].val = v - Infix_offset_val (v); + }else{ + final->table[final->young].offset = 0; + final->table[final->young].val = v; + } + ++ final->young; + +} + +CAMLprim value caml_final_register (value f, value v){ + generic_final_register(&finalisable_first, f, v); + return Val_unit; +} + +CAMLprim value caml_final_register_called_without_value (value f, value v){ + generic_final_register(&finalisable_last, f, v); + return Val_unit; +} + + +CAMLprim value caml_final_release (value unit) +{ + running_finalisation_function = 0; + return Val_unit; +} + +static void gen_final_invariant_check(struct finalisable *final){ + uintnat i; + + CAMLassert (final->old <= final->young); + for (i = 0; i < final->old; i++){ + CAMLassert( Is_in_heap(final->table[i].val) ); + }; + for (i = final->old; i < final->young; i++){ + CAMLassert( Is_in_heap_or_young(final->table[i].val) ); + }; +} + +void caml_final_invariant_check(void){ + gen_final_invariant_check(&finalisable_first); + gen_final_invariant_check(&finalisable_last); +} diff --git a/test/monniaux/ocaml/byterun/fix_code.c b/test/monniaux/ocaml/byterun/fix_code.c new file mode 100644 index 00000000..ec2f08cc --- /dev/null +++ b/test/monniaux/ocaml/byterun/fix_code.c @@ -0,0 +1,195 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Handling of blocks of bytecode (endianness switch, threading). */ + +#include "caml/config.h" + +#ifdef HAS_UNISTD +#include +#else +#include +#endif + +#include "caml/debugger.h" +#include "caml/fix_code.h" +#include "caml/instruct.h" +#include "caml/intext.h" +#include "caml/md5.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/reverse.h" + +code_t caml_start_code; +asize_t caml_code_size; +unsigned char * caml_saved_code; +struct ext_table caml_code_fragments_table; + +/* Read the main bytecode block from a file */ + +void caml_init_code_fragments(void) { + struct code_fragment * cf; + /* Register the code in the table of code fragments */ + cf = caml_stat_alloc(sizeof(struct code_fragment)); + cf->code_start = (char *) caml_start_code; + cf->code_end = (char *) caml_start_code + caml_code_size; + caml_md5_block(cf->digest, caml_start_code, caml_code_size); + cf->digest_computed = 1; + caml_ext_table_init(&caml_code_fragments_table, 8); + caml_ext_table_add(&caml_code_fragments_table, cf); +} + +void caml_load_code(int fd, asize_t len) +{ + int i; + + caml_code_size = len; + caml_start_code = (code_t) caml_stat_alloc(caml_code_size); + if (read(fd, (char *) caml_start_code, caml_code_size) != caml_code_size) + caml_fatal_error("Fatal error: truncated bytecode file.\n"); + caml_init_code_fragments(); + /* Prepare the code for execution */ +#ifdef ARCH_BIG_ENDIAN + caml_fixup_endianness(caml_start_code, caml_code_size); +#endif + if (caml_debugger_in_use) { + len /= sizeof(opcode_t); + caml_saved_code = (unsigned char *) caml_stat_alloc(len); + for (i = 0; i < len; i++) caml_saved_code[i] = caml_start_code[i]; + } +#ifdef THREADED_CODE + /* Better to thread now than at the beginning of [caml_interprete], + since the debugger interface needs to perform SET_EVENT requests + on the code. */ + caml_thread_code(caml_start_code, caml_code_size); +#endif +} + +/* This code is needed only if the processor is big endian */ + +#ifdef ARCH_BIG_ENDIAN + +void caml_fixup_endianness(code_t code, asize_t len) +{ + code_t p; + len /= sizeof(opcode_t); + for (p = code; p < code + len; p++) { + Reverse_32(p, p); + } +} + +#endif + +/* This code is needed only if we're using threaded code */ + +#ifdef THREADED_CODE + +char ** caml_instr_table; +char * caml_instr_base; + +static int* opcode_nargs = NULL; +int* caml_init_opcode_nargs(void) +{ + if( opcode_nargs == NULL ){ + int* l = (int*)caml_stat_alloc(sizeof(int) * FIRST_UNIMPLEMENTED_OP); + int i; + + for (i = 0; i < FIRST_UNIMPLEMENTED_OP; i++) { + l [i] = 0; + } + /* Instructions with one operand */ + l[PUSHACC] = l[ACC] = l[POP] = l[ASSIGN] = + l[PUSHENVACC] = l[ENVACC] = l[PUSH_RETADDR] = l[APPLY] = + l[APPTERM1] = l[APPTERM2] = l[APPTERM3] = l[RETURN] = + l[GRAB] = l[PUSHGETGLOBAL] = l[GETGLOBAL] = l[SETGLOBAL] = + l[PUSHATOM] = l[ATOM] = l[MAKEBLOCK1] = l[MAKEBLOCK2] = + l[MAKEBLOCK3] = l[MAKEFLOATBLOCK] = l[GETFIELD] = + l[GETFLOATFIELD] = l[SETFIELD] = l[SETFLOATFIELD] = + l[BRANCH] = l[BRANCHIF] = l[BRANCHIFNOT] = l[PUSHTRAP] = + l[C_CALL1] = l[C_CALL2] = l[C_CALL3] = l[C_CALL4] = l[C_CALL5] = + l[CONSTINT] = l[PUSHCONSTINT] = l[OFFSETINT] = + l[OFFSETREF] = l[OFFSETCLOSURE] = l[PUSHOFFSETCLOSURE] = 1; + + /* Instructions with two operands */ + l[APPTERM] = l[CLOSURE] = l[PUSHGETGLOBALFIELD] = + l[GETGLOBALFIELD] = l[MAKEBLOCK] = l[C_CALLN] = + l[BEQ] = l[BNEQ] = l[BLTINT] = l[BLEINT] = l[BGTINT] = l[BGEINT] = + l[BULTINT] = l[BUGEINT] = l[GETPUBMET] = 2; + + opcode_nargs = l; + } + return opcode_nargs; +} + +void caml_thread_code (code_t code, asize_t len) +{ + code_t p; + int* l = caml_init_opcode_nargs(); + len /= sizeof(opcode_t); + for (p = code; p < code + len; /*nothing*/) { + opcode_t instr = *p; + if (instr < 0 || instr >= FIRST_UNIMPLEMENTED_OP){ + /* FIXME -- should Assert(false) ? + caml_fatal_error_arg ("Fatal error in fix_code: bad opcode (%lx)\n", + (char *)(long)instr); + */ + instr = STOP; + } + *p++ = (opcode_t)(caml_instr_table[instr] - caml_instr_base); + if (instr == SWITCH) { + uint32_t sizes = *p++; + uint32_t const_size = sizes & 0xFFFF; + uint32_t block_size = sizes >> 16; + p += const_size + block_size; + } else if (instr == CLOSUREREC) { + uint32_t nfuncs = *p++; + p++; /* skip nvars */ + p += nfuncs; + } else { + p += l[instr]; + } + } + CAMLassert(p == code + len); +} + +#else + +int* caml_init_opcode_nargs() +{ + return NULL; +} + +#endif /* THREADED_CODE */ + +void caml_set_instruction(code_t pos, opcode_t instr) +{ +#ifdef THREADED_CODE + *pos = (opcode_t)(caml_instr_table[instr] - caml_instr_base); +#else + *pos = instr; +#endif +} + +int caml_is_instruction(opcode_t instr1, opcode_t instr2) +{ +#ifdef THREADED_CODE + return instr1 == (opcode_t)(caml_instr_table[instr2] - caml_instr_base); +#else + return instr1 == instr2; +#endif +} diff --git a/test/monniaux/ocaml/byterun/floats.c b/test/monniaux/ocaml/byterun/floats.c new file mode 100644 index 00000000..4d2494cf --- /dev/null +++ b/test/monniaux/ocaml/byterun/floats.c @@ -0,0 +1,720 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* The interface of this file is in "caml/mlvalues.h" and "caml/alloc.h" */ + +#include +#include +#include +#include +#include +#include + +#include "caml/alloc.h" +#include "caml/fail.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/misc.h" +#include "caml/reverse.h" +#include "caml/stacks.h" + +#ifdef _MSC_VER +#include +#ifndef isnan +#define isnan _isnan +#endif +#ifndef isfinite +#define isfinite _finite +#endif +#endif + +#ifdef ARCH_ALIGN_DOUBLE + +CAMLexport double caml_Double_val(value val) +{ + union { value v[2]; double d; } buffer; + + CAMLassert(sizeof(double) == 2 * sizeof(value)); + buffer.v[0] = Field(val, 0); + buffer.v[1] = Field(val, 1); + return buffer.d; +} + +CAMLexport void caml_Store_double_val(value val, double dbl) +{ + union { value v[2]; double d; } buffer; + + CAMLassert(sizeof(double) == 2 * sizeof(value)); + buffer.d = dbl; + Field(val, 0) = buffer.v[0]; + Field(val, 1) = buffer.v[1]; +} + +#endif + +CAMLexport value caml_copy_double(double d) +{ + value res; + +#define Setup_for_gc +#define Restore_after_gc + Alloc_small(res, Double_wosize, Double_tag); +#undef Setup_for_gc +#undef Restore_after_gc + Store_double_val(res, d); + return res; +} + +#ifndef FLAT_FLOAT_ARRAY +CAMLexport void caml_Store_double_array_field(value val, mlsize_t i, double dbl) +{ + CAMLparam1 (val); + value d = caml_copy_double (dbl); + + CAMLassert (Tag_val (val) != Double_array_tag); + caml_modify (&Field(val, i), d); + CAMLreturn0; +} +#endif /* ! FLAT_FLOAT_ARRAY */ + +CAMLprim value caml_format_float(value fmt, value arg) +{ + value res; + double d = Double_val(arg); + +#ifdef HAS_BROKEN_PRINTF + if (isfinite(d)) { +#endif + res = caml_alloc_sprintf(String_val(fmt), d); +#ifdef HAS_BROKEN_PRINTF + } else { + if (isnan(d)) { + res = caml_copy_string("nan"); + } else { + if (d > 0) + res = caml_copy_string("inf"); + else + res = caml_copy_string("-inf"); + } + } +#endif + return res; +} + +CAMLprim value caml_hexstring_of_float(value arg, value vprec, value vstyle) +{ + union { uint64_t i; double d; } u; + int sign, exp; + uint64_t m; + char buffer[64]; + char * buf, * p; + intnat prec; + int d; + value res; + + /* Allocate output buffer */ + prec = Long_val(vprec); + /* 12 chars for sign, 0x, decimal point, exponent */ + buf = (prec + 12 <= 64 ? buffer : caml_stat_alloc(prec + 12)); + /* Extract sign, mantissa, and exponent */ + u.d = Double_val(arg); + sign = u.i >> 63; + exp = (u.i >> 52) & 0x7FF; + m = u.i & (((uint64_t) 1 << 52) - 1); + /* Put sign */ + p = buf; + if (sign) { + *p++ = '-'; + } else { + switch (Int_val(vstyle)) { + case '+': *p++ = '+'; break; + case ' ': *p++ = ' '; break; + } + } + /* Treat special cases */ + if (exp == 0x7FF) { + char * txt; + if (m == 0) txt = "infinity"; else txt = "nan"; + memcpy(p, txt, strlen(txt)); + p[strlen(txt)] = 0; + res = caml_copy_string(buf); + } else { + /* Output "0x" prefix */ + *p++ = '0'; *p++ = 'x'; + /* Normalize exponent and mantissa */ + if (exp == 0) { + if (m != 0) exp = -1022; /* denormal */ + } else { + exp = exp - 1023; + m = m | ((uint64_t) 1 << 52); + } + /* If a precision is given, and is small, round mantissa accordingly */ + prec = Long_val(vprec); + if (prec >= 0 && prec < 13) { + int i = 52 - prec * 4; + uint64_t unit = (uint64_t) 1 << i; + uint64_t half = unit >> 1; + uint64_t mask = unit - 1; + uint64_t frac = m & mask; + m = m & ~mask; + /* Round to nearest, ties to even */ + if (frac > half || (frac == half && (m & unit) != 0)) { + m += unit; + } + } + /* Leading digit */ + d = m >> 52; + *p++ = (d < 10 ? d + '0' : d - 10 + 'a'); + m = (m << 4) & (((uint64_t) 1 << 56) - 1); + /* Fractional digits. If a precision is given, print that number of + digits. Otherwise, print as many digits as needed to represent + the mantissa exactly. */ + if (prec >= 0 ? prec > 0 : m != 0) { + *p++ = '.'; + while (prec >= 0 ? prec > 0 : m != 0) { + d = m >> 52; + *p++ = (d < 10 ? d + '0' : d - 10 + 'a'); + m = (m << 4) & (((uint64_t) 1 << 56) - 1); + prec--; + } + } + *p = 0; + /* Add exponent */ + res = caml_alloc_sprintf("%sp%+d", buf, exp); + } + if (buf != buffer) caml_stat_free(buf); + return res; +} + +static int caml_float_of_hex(const char * s, double * res) +{ + int64_t m = 0; /* the mantissa - top 60 bits at most */ + int n_bits = 0; /* total number of bits read */ + int m_bits = 0; /* number of bits in mantissa */ + int x_bits = 0; /* number of bits after mantissa */ + int dec_point = -1; /* bit count corresponding to decimal point */ + /* -1 if no decimal point seen */ + int exp = 0; /* exponent */ + char * p; /* for converting the exponent */ + double f; + + while (*s != 0) { + char c = *s++; + switch (c) { + case '_': + break; + case '.': + if (dec_point >= 0) return -1; /* multiple decimal points */ + dec_point = n_bits; + break; + case 'p': case 'P': { + long e; + if (*s == 0) return -1; /* nothing after exponent mark */ + e = strtol(s, &p, 10); + if (*p != 0) return -1; /* ill-formed exponent */ + /* Handle exponents larger than int by returning 0/∞ directly. + Mind that INT_MIN/INT_MAX are included in the test so as to capture + the overflow case of strtol on Win64 — long and int have the same + size there. */ + if (e <= INT_MIN) { + *res = 0.; + return 0; + } + else if (e >= INT_MAX) { + *res = m == 0 ? 0. : HUGE_VAL; + return 0; + } + /* regular exponent value */ + exp = e; + s = p; /* stop at next loop iteration */ + break; + } + default: { /* Nonzero digit */ + int d; + if (c >= '0' && c <= '9') d = c - '0'; + else if (c >= 'A' && c <= 'F') d = c - 'A' + 10; + else if (c >= 'a' && c <= 'f') d = c - 'a' + 10; + else return -1; /* bad digit */ + n_bits += 4; + if (d == 0 && m == 0) break; /* leading zeros are skipped */ + if (m_bits < 60) { + /* There is still room in m. Add this digit to the mantissa. */ + m = (m << 4) + d; + m_bits += 4; + } else { + /* We've already collected 60 significant bits in m. + Now all we care about is whether there is a nonzero bit + after. In this case, round m to odd so that the later + rounding of m to FP produces the correct result. */ + if (d != 0) m |= 1; /* round to odd */ + x_bits += 4; + } + break; + } + } + } + if (n_bits == 0) return -1; + /* Convert mantissa to FP. We use a signed conversion because we can + (m has 60 bits at most) and because it is faster + on several architectures. */ + f = (double) (int64_t) m; + /* Adjust exponent to take decimal point and extra digits into account */ + { + int adj = x_bits; + if (dec_point >= 0) adj = adj + (dec_point - n_bits); + /* saturated addition exp + adj */ + if (adj > 0 && exp > INT_MAX - adj) + exp = INT_MAX; + else if (adj < 0 && exp < INT_MIN - adj) + exp = INT_MIN; + else + exp = exp + adj; + } + /* Apply exponent if needed */ + if (exp != 0) f = ldexp(f, exp); + /* Done! */ + *res = f; + return 0; +} + +CAMLprim value caml_float_of_string(value vs) +{ + char parse_buffer[64]; + char * buf, * dst, * end; + const char *src; + mlsize_t len; + int sign; + double d; + + /* Check for hexadecimal FP constant */ + src = String_val(vs); + sign = 1; + if (*src == '-') { sign = -1; src++; } + else if (*src == '+') { src++; }; + if (src[0] == '0' && (src[1] == 'x' || src[1] == 'X')) { + if (caml_float_of_hex(src + 2, &d) == -1) + caml_failwith("float_of_string"); + return caml_copy_double(sign < 0 ? -d : d); + } + /* Remove '_' characters before calling strtod () */ + len = caml_string_length(vs); + buf = len < sizeof(parse_buffer) ? parse_buffer : caml_stat_alloc(len + 1); + src = String_val(vs); + dst = buf; + while (len--) { + char c = *src++; + if (c != '_') *dst++ = c; + } + *dst = 0; + if (dst == buf) goto error; + /* Convert using strtod */ + d = strtod((const char *) buf, &end); + if (end != dst) goto error; + if (buf != parse_buffer) caml_stat_free(buf); + return caml_copy_double(d); + error: + if (buf != parse_buffer) caml_stat_free(buf); + caml_failwith("float_of_string"); + return Val_unit; /* not reached */ +} + +CAMLprim value caml_int_of_float(value f) +{ + return Val_long((intnat) Double_val(f)); +} + +CAMLprim value caml_float_of_int(value n) +{ + return caml_copy_double((double) Long_val(n)); +} + +CAMLprim value caml_neg_float(value f) +{ + return caml_copy_double(- Double_val(f)); +} + +CAMLprim value caml_abs_float(value f) +{ + return caml_copy_double(fabs(Double_val(f))); +} + +CAMLprim value caml_add_float(value f, value g) +{ + return caml_copy_double(Double_val(f) + Double_val(g)); +} + +CAMLprim value caml_sub_float(value f, value g) +{ + return caml_copy_double(Double_val(f) - Double_val(g)); +} + +CAMLprim value caml_mul_float(value f, value g) +{ + return caml_copy_double(Double_val(f) * Double_val(g)); +} + +CAMLprim value caml_div_float(value f, value g) +{ + return caml_copy_double(Double_val(f) / Double_val(g)); +} + +CAMLprim value caml_exp_float(value f) +{ + return caml_copy_double(exp(Double_val(f))); +} + +CAMLprim value caml_floor_float(value f) +{ + return caml_copy_double(floor(Double_val(f))); +} + +CAMLprim value caml_fmod_float(value f1, value f2) +{ + return caml_copy_double(fmod(Double_val(f1), Double_val(f2))); +} + +CAMLprim value caml_frexp_float(value f) +{ + CAMLparam1 (f); + CAMLlocal2 (res, mantissa); + int exponent; + + mantissa = caml_copy_double(frexp (Double_val(f), &exponent)); + res = caml_alloc_tuple(2); + Field(res, 0) = mantissa; + Field(res, 1) = Val_int(exponent); + CAMLreturn (res); +} + +// Seems dumb but intnat could not correspond to int type. +double caml_ldexp_float_unboxed(double f, intnat i) +{ + return ldexp(f, i); +} + + +CAMLprim value caml_ldexp_float(value f, value i) +{ + return caml_copy_double(ldexp(Double_val(f), Int_val(i))); +} + +CAMLprim value caml_log_float(value f) +{ + return caml_copy_double(log(Double_val(f))); +} + +CAMLprim value caml_log10_float(value f) +{ + return caml_copy_double(log10(Double_val(f))); +} + +CAMLprim value caml_modf_float(value f) +{ + double frem; + + CAMLparam1 (f); + CAMLlocal3 (res, quo, rem); + + quo = caml_copy_double(modf (Double_val(f), &frem)); + rem = caml_copy_double(frem); + res = caml_alloc_tuple(2); + Field(res, 0) = quo; + Field(res, 1) = rem; + CAMLreturn (res); +} + +CAMLprim value caml_sqrt_float(value f) +{ + return caml_copy_double(sqrt(Double_val(f))); +} + +CAMLprim value caml_power_float(value f, value g) +{ + return caml_copy_double(pow(Double_val(f), Double_val(g))); +} + +CAMLprim value caml_sin_float(value f) +{ + return caml_copy_double(sin(Double_val(f))); +} + +CAMLprim value caml_sinh_float(value f) +{ + return caml_copy_double(sinh(Double_val(f))); +} + +CAMLprim value caml_cos_float(value f) +{ + return caml_copy_double(cos(Double_val(f))); +} + +CAMLprim value caml_cosh_float(value f) +{ + return caml_copy_double(cosh(Double_val(f))); +} + +CAMLprim value caml_tan_float(value f) +{ + return caml_copy_double(tan(Double_val(f))); +} + +CAMLprim value caml_tanh_float(value f) +{ + return caml_copy_double(tanh(Double_val(f))); +} + +CAMLprim value caml_asin_float(value f) +{ + return caml_copy_double(asin(Double_val(f))); +} + +CAMLprim value caml_acos_float(value f) +{ + return caml_copy_double(acos(Double_val(f))); +} + +CAMLprim value caml_atan_float(value f) +{ + return caml_copy_double(atan(Double_val(f))); +} + +CAMLprim value caml_atan2_float(value f, value g) +{ + return caml_copy_double(atan2(Double_val(f), Double_val(g))); +} + +CAMLprim value caml_ceil_float(value f) +{ + return caml_copy_double(ceil(Double_val(f))); +} + +CAMLexport double caml_hypot(double x, double y) +{ +#ifdef HAS_C99_FLOAT_OPS + return hypot(x, y); +#else + double tmp, ratio; + x = fabs(x); y = fabs(y); + if (x != x) /* x is NaN */ + return y > DBL_MAX ? y : x; /* PR#6321 */ + if (y != y) /* y is NaN */ + return x > DBL_MAX ? x : y; /* PR#6321 */ + if (x < y) { tmp = x; x = y; y = tmp; } + if (x == 0.0) return 0.0; + ratio = y / x; + return x * sqrt(1.0 + ratio * ratio); +#endif +} + +CAMLprim value caml_hypot_float(value f, value g) +{ + return caml_copy_double(caml_hypot(Double_val(f), Double_val(g))); +} + +/* These emulations of expm1() and log1p() are due to William Kahan. + See http://www.plunk.org/~hatch/rightway.php */ +CAMLexport double caml_expm1(double x) +{ +#ifdef HAS_C99_FLOAT_OPS + return expm1(x); +#else + double u = exp(x); + if (u == 1.) + return x; + if (u - 1. == -1.) + return -1.; + return (u - 1.) * x / log(u); +#endif +} + +CAMLexport double caml_log1p(double x) +{ +#ifdef HAS_C99_FLOAT_OPS + return log1p(x); +#else + double u = 1. + x; + if (u == 1.) + return x; + else + return log(u) * x / (u - 1.); +#endif +} + +CAMLprim value caml_expm1_float(value f) +{ + return caml_copy_double(caml_expm1(Double_val(f))); +} + +CAMLprim value caml_log1p_float(value f) +{ + return caml_copy_double(caml_log1p(Double_val(f))); +} + +union double_as_two_int32 { + double d; +#if defined(ARCH_BIG_ENDIAN) || (defined(__arm__) && !defined(__ARM_EABI__)) + struct { uint32_t h; uint32_t l; } i; +#else + struct { uint32_t l; uint32_t h; } i; +#endif +}; + +CAMLexport double caml_copysign(double x, double y) +{ +#ifdef HAS_C99_FLOAT_OPS + return copysign(x, y); +#else + union double_as_two_int32 ux, uy; + ux.d = x; + uy.d = y; + ux.i.h &= 0x7FFFFFFFU; + ux.i.h |= (uy.i.h & 0x80000000U); + return ux.d; +#endif +} + +CAMLprim value caml_copysign_float(value f, value g) +{ + return caml_copy_double(caml_copysign(Double_val(f), Double_val(g))); +} + +#ifdef LACKS_SANE_NAN + +CAMLprim value caml_neq_float(value vf, value vg) +{ + double f = Double_val(vf); + double g = Double_val(vg); + return Val_bool(isnan(f) || isnan(g) || f != g); +} + +#define DEFINE_NAN_CMP(op) (value vf, value vg) \ +{ \ + double f = Double_val(vf); \ + double g = Double_val(vg); \ + return Val_bool(!isnan(f) && !isnan(g) && f op g); \ +} + +intnat caml_float_compare_unboxed(double f, double g) +{ + /* Insane => nan == everything && nan < everything && nan > everything */ + if (isnan(f) && isnan(g)) return 0; + if (!isnan(g) && f < g) return -1; + if (f != g) return 1; + return 0; +} + +#else + +CAMLprim value caml_neq_float(value f, value g) +{ + return Val_bool(Double_val(f) != Double_val(g)); +} + +#define DEFINE_NAN_CMP(op) (value f, value g) \ +{ \ + return Val_bool(Double_val(f) op Double_val(g)); \ +} + +intnat caml_float_compare_unboxed(double f, double g) +{ + /* If one or both of f and g is NaN, order according to the convention + NaN = NaN and NaN < x for all other floats x. */ + /* This branchless implementation is from GPR#164. + Note that [f == f] if and only if f is not NaN. */ + return (f > g) - (f < g) + (f == f) - (g == g); +} + +#endif + +CAMLprim value caml_eq_float DEFINE_NAN_CMP(==) +CAMLprim value caml_le_float DEFINE_NAN_CMP(<=) +CAMLprim value caml_lt_float DEFINE_NAN_CMP(<) +CAMLprim value caml_ge_float DEFINE_NAN_CMP(>=) +CAMLprim value caml_gt_float DEFINE_NAN_CMP(>) + +CAMLprim value caml_float_compare(value vf, value vg) +{ + return Val_int(caml_float_compare_unboxed(Double_val(vf),Double_val(vg))); +} + +enum { FP_normal, FP_subnormal, FP_zero, FP_infinite, FP_nan }; + +value caml_classify_float_unboxed(double vd) +{ +#ifdef ARCH_SIXTYFOUR + union { double d; uint64_t i; } u; + uint64_t n; + uint32_t e; + + u.d = vd; + n = u.i << 1; /* shift sign bit off */ + if (n == 0) return Val_int(FP_zero); + e = n >> 53; /* extract exponent */ + if (e == 0) return Val_int(FP_subnormal); + if (e == 0x7FF) { + if (n << 11 == 0) /* shift exponent off */ + return Val_int(FP_infinite); + else + return Val_int(FP_nan); + } + return Val_int(FP_normal); +#else + union double_as_two_int32 u; + uint32_t h, l; + + u.d = vd; + h = u.i.h; l = u.i.l; + l = l | (h & 0xFFFFF); + h = h & 0x7FF00000; + if ((h | l) == 0) + return Val_int(FP_zero); + if (h == 0) + return Val_int(FP_subnormal); + if (h == 0x7FF00000) { + if (l == 0) + return Val_int(FP_infinite); + else + return Val_int(FP_nan); + } + return Val_int(FP_normal); +#endif +} + +CAMLprim value caml_classify_float(value vd) +{ + return caml_classify_float_unboxed(Double_val(vd)); +} + +/* The [caml_init_ieee_float] function should initialize floating-point hardware + so that it behaves as much as possible like the IEEE standard. + In particular, return special numbers like Infinity and NaN instead + of signalling exceptions. Currently, everyone is in IEEE mode + at program startup, except FreeBSD prior to 4.0R. */ + +#ifdef __FreeBSD__ +#include +#if (__FreeBSD_version < 400017) +#include +#endif +#endif + +void caml_init_ieee_floats(void) +{ +#if defined(__FreeBSD__) && (__FreeBSD_version < 400017) + fpsetmask(0); +#endif +} diff --git a/test/monniaux/ocaml/byterun/freelist.c b/test/monniaux/ocaml/byterun/freelist.c new file mode 100644 index 00000000..4782800e --- /dev/null +++ b/test/monniaux/ocaml/byterun/freelist.c @@ -0,0 +1,621 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#define FREELIST_DEBUG 0 +#if FREELIST_DEBUG +#include +#endif + +#include + +#include "caml/config.h" +#include "caml/freelist.h" +#include "caml/gc.h" +#include "caml/gc_ctrl.h" +#include "caml/memory.h" +#include "caml/major_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" + +/* The free-list is kept sorted by increasing addresses. + This makes the merging of adjacent free blocks possible. + (See [caml_fl_merge_block].) +*/ + +/* A free list block is a [value] (integer representing a pointer to the + first word after the block's header). The end of the list is NULL. */ +#define Val_NULL ((value) NULL) + +/* The sentinel can be located anywhere in memory, but it must not be + adjacent to any heap object. */ +static struct { + value filler1; /* Make sure the sentinel is never adjacent to any block. */ + header_t h; + value first_field; + value filler2; /* Make sure the sentinel is never adjacent to any block. */ +} sentinel = {0, Make_header (0, 0, Caml_blue), Val_NULL, 0}; + +#define Fl_head (Val_bp (&(sentinel.first_field))) +static value fl_prev = Fl_head; /* Current allocation pointer. */ +static value fl_last = Val_NULL; /* Last block in the list. Only valid + just after [caml_fl_allocate] returns NULL. */ +value caml_fl_merge = Fl_head; /* Current insertion pointer. Managed + jointly with [sweep_slice]. */ +asize_t caml_fl_cur_wsz = 0; /* Number of words in the free list, + including headers but not fragments. */ + +#define FLP_MAX 1000 +static value flp [FLP_MAX]; +static int flp_size = 0; +static value beyond = Val_NULL; + +#define Next(b) (Field (b, 0)) + +#define Policy_next_fit 0 +#define Policy_first_fit 1 +uintnat caml_allocation_policy = Policy_next_fit; +#define policy caml_allocation_policy + +#ifdef DEBUG +static void fl_check (void) +{ + value cur, prev; + int prev_found = 0, flp_found = 0, merge_found = 0; + uintnat size_found = 0; + int sz = 0; + + prev = Fl_head; + cur = Next (prev); + while (cur != Val_NULL){ + size_found += Whsize_bp (cur); + CAMLassert (Is_in_heap (cur)); + if (cur == fl_prev) prev_found = 1; + if (policy == Policy_first_fit && Wosize_bp (cur) > sz){ + sz = Wosize_bp (cur); + if (flp_found < flp_size){ + CAMLassert (Next (flp[flp_found]) == cur); + ++ flp_found; + }else{ + CAMLassert (beyond == Val_NULL || cur >= Next (beyond)); + } + } + if (cur == caml_fl_merge) merge_found = 1; + prev = cur; + cur = Next (prev); + } + if (policy == Policy_next_fit) CAMLassert (prev_found || fl_prev == Fl_head); + if (policy == Policy_first_fit) CAMLassert (flp_found == flp_size); + CAMLassert (merge_found || caml_fl_merge == Fl_head); + CAMLassert (size_found == caml_fl_cur_wsz); +} + +#endif + +/* [allocate_block] is called by [caml_fl_allocate]. Given a suitable free + block and the requested size, it allocates a new block from the free + block. There are three cases: + 0. The free block has the requested size. Detach the block from the + free-list and return it. + 1. The free block is 1 word longer than the requested size. Detach + the block from the free list. The remaining word cannot be linked: + turn it into an empty block (header only), and return the rest. + 2. The free block is large enough. Split it in two and return the right + block. + In all cases, the allocated block is right-justified in the free block: + it is located in the high-address words of the free block, so that + the linking of the free-list does not change in case 2. +*/ +static header_t *allocate_block (mlsize_t wh_sz, int flpi, value prev, + value cur) +{ + header_t h = Hd_bp (cur); + CAMLassert (Whsize_hd (h) >= wh_sz); + if (Wosize_hd (h) < wh_sz + 1){ /* Cases 0 and 1. */ + caml_fl_cur_wsz -= Whsize_hd (h); + Next (prev) = Next (cur); + CAMLassert (Is_in_heap (Next (prev)) || Next (prev) == Val_NULL); + if (caml_fl_merge == cur) caml_fl_merge = prev; +#ifdef DEBUG + fl_last = Val_NULL; +#endif + /* In case 1, the following creates the empty block correctly. + In case 0, it gives an invalid header to the block. The function + calling [caml_fl_allocate] will overwrite it. */ + Hd_op (cur) = Make_header (0, 0, Caml_white); + if (policy == Policy_first_fit){ + if (flpi + 1 < flp_size && flp[flpi + 1] == cur){ + flp[flpi + 1] = prev; + }else if (flpi == flp_size - 1){ + beyond = (prev == Fl_head) ? Val_NULL : prev; + -- flp_size; + } + } + }else{ /* Case 2. */ + caml_fl_cur_wsz -= wh_sz; + Hd_op (cur) = Make_header (Wosize_hd (h) - wh_sz, 0, Caml_blue); + } + if (policy == Policy_next_fit) fl_prev = prev; + return (header_t *) &Field (cur, Wosize_hd (h) - wh_sz); +} + +#ifdef CAML_INSTR +static uintnat instr_size [20] = + {0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0,0}; +static char *instr_name [20] = { + NULL, + "alloc01@", + "alloc02@", + "alloc03@", + "alloc04@", + "alloc05@", + "alloc06@", + "alloc07@", + "alloc08@", + "alloc09@", + "alloc10-19@", + "alloc20-29@", + "alloc30-39@", + "alloc40-49@", + "alloc50-59@", + "alloc60-69@", + "alloc70-79@", + "alloc80-89@", + "alloc90-99@", + "alloc_large@", +}; +uintnat caml_instr_alloc_jump = 0; +/* number of pointers followed to allocate from the free list */ +#endif /*CAML_INSTR*/ + +/* [caml_fl_allocate] does not set the header of the newly allocated block. + The calling function must do it before any GC function gets called. + [caml_fl_allocate] returns a head pointer. +*/ +header_t *caml_fl_allocate (mlsize_t wo_sz) +{ + value cur = Val_NULL, prev; + header_t *result; + int i; + mlsize_t sz, prevsz; + CAMLassert (sizeof (char *) == sizeof (value)); + CAMLassert (wo_sz >= 1); +#ifdef CAML_INSTR + if (wo_sz < 10){ + ++instr_size[wo_sz]; + }else if (wo_sz < 100){ + ++instr_size[wo_sz/10 + 9]; + }else{ + ++instr_size[19]; + } +#endif /* CAML_INSTR */ + + switch (policy){ + case Policy_next_fit: + CAMLassert (fl_prev != Val_NULL); + /* Search from [fl_prev] to the end of the list. */ + prev = fl_prev; + cur = Next (prev); + while (cur != Val_NULL){ + CAMLassert (Is_in_heap (cur)); + if (Wosize_bp (cur) >= wo_sz){ + return allocate_block (Whsize_wosize (wo_sz), 0, prev, cur); + } + prev = cur; + cur = Next (prev); +#ifdef CAML_INSTR + ++ caml_instr_alloc_jump; +#endif + } + fl_last = prev; + /* Search from the start of the list to [fl_prev]. */ + prev = Fl_head; + cur = Next (prev); + while (prev != fl_prev){ + if (Wosize_bp (cur) >= wo_sz){ + return allocate_block (Whsize_wosize (wo_sz), 0, prev, cur); + } + prev = cur; + cur = Next (prev); +#ifdef CAML_INSTR + ++ caml_instr_alloc_jump; +#endif + } + /* No suitable block was found. */ + return NULL; + break; + + case Policy_first_fit: { + /* Search in the flp array. */ + for (i = 0; i < flp_size; i++){ + sz = Wosize_bp (Next (flp[i])); + if (sz >= wo_sz){ +#if FREELIST_DEBUG + if (i > 5) fprintf (stderr, "FLP: found at %d size=%d\n", i, wo_sz); +#endif + result = allocate_block (Whsize_wosize (wo_sz), i, flp[i], + Next (flp[i])); + goto update_flp; + } + } + /* Extend the flp array. */ + if (flp_size == 0){ + prev = Fl_head; + prevsz = 0; + }else{ + prev = Next (flp[flp_size - 1]); + prevsz = Wosize_bp (prev); + if (beyond != Val_NULL) prev = beyond; + } + while (flp_size < FLP_MAX){ + cur = Next (prev); + if (cur == Val_NULL){ + fl_last = prev; + beyond = (prev == Fl_head) ? Val_NULL : prev; + return NULL; + }else{ + sz = Wosize_bp (cur); + if (sz > prevsz){ + flp[flp_size] = prev; + ++ flp_size; + if (sz >= wo_sz){ + beyond = cur; + i = flp_size - 1; +#if FREELIST_DEBUG + if (flp_size > 5){ + fprintf (stderr, "FLP: extended to %d\n", flp_size); + } +#endif + result = allocate_block (Whsize_wosize (wo_sz), flp_size - 1, prev, + cur); + goto update_flp; + } + prevsz = sz; + } + } + prev = cur; + } + beyond = cur; + + /* The flp table is full. Do a slow first-fit search. */ +#if FREELIST_DEBUG + fprintf (stderr, "FLP: table is full -- slow first-fit\n"); +#endif + if (beyond != Val_NULL){ + prev = beyond; + }else{ + prev = flp[flp_size - 1]; + } + prevsz = Wosize_bp (Next (flp[FLP_MAX-1])); + CAMLassert (prevsz < wo_sz); + cur = Next (prev); + while (cur != Val_NULL){ + CAMLassert (Is_in_heap (cur)); + sz = Wosize_bp (cur); + if (sz < prevsz){ + beyond = cur; + }else if (sz >= wo_sz){ + return allocate_block (Whsize_wosize (wo_sz), flp_size, prev, cur); + } + prev = cur; + cur = Next (prev); + } + fl_last = prev; + return NULL; + + update_flp: /* (i, sz) */ + /* The block at [i] was removed or reduced. Update the table. */ + CAMLassert (0 <= i && i < flp_size + 1); + if (i < flp_size){ + if (i > 0){ + prevsz = Wosize_bp (Next (flp[i-1])); + }else{ + prevsz = 0; + } + if (i == flp_size - 1){ + if (Wosize_bp (Next (flp[i])) <= prevsz){ + beyond = Next (flp[i]); + -- flp_size; + }else{ + beyond = Val_NULL; + } + }else{ + value buf [FLP_MAX]; + int j = 0; + mlsize_t oldsz = sz; + + prev = flp[i]; + while (prev != flp[i+1] && j < FLP_MAX - i){ + cur = Next (prev); + sz = Wosize_bp (cur); + if (sz > prevsz){ + buf[j++] = prev; + prevsz = sz; + if (sz >= oldsz){ + CAMLassert (sz == oldsz); + break; + } + } + prev = cur; + } +#if FREELIST_DEBUG + if (j > 2) fprintf (stderr, "FLP: update; buf size = %d\n", j); +#endif + if (FLP_MAX >= flp_size + j - 1){ + if (j != 1){ + memmove (&flp[i+j], &flp[i+1], sizeof (value) * (flp_size-i-1)); + } + if (j > 0) memmove (&flp[i], &buf[0], sizeof (value) * j); + flp_size += j - 1; + }else{ + if (FLP_MAX > i + j){ + if (j != 1){ + memmove (&flp[i+j], &flp[i+1], sizeof (value) * (FLP_MAX-i-j)); + } + if (j > 0) memmove (&flp[i], &buf[0], sizeof (value) * j); + }else{ + if (i != FLP_MAX){ + memmove (&flp[i], &buf[0], sizeof (value) * (FLP_MAX - i)); + } + } + flp_size = FLP_MAX - 1; + beyond = Next (flp[FLP_MAX - 1]); + } + } + } + return result; + } + break; + + default: + CAMLassert (0); /* unknown policy */ + break; + } + return NULL; /* NOT REACHED */ +} + +/* Location of the last fragment seen by the sweeping code. + This is a pointer to the first word after the fragment, which is + the header of the next block. + Note that [last_fragment] doesn't point to the fragment itself, + but to the block after it. +*/ +static header_t *last_fragment; + +void caml_fl_init_merge (void) +{ +#ifdef CAML_INSTR + int i; + for (i = 1; i < 20; i++){ + CAML_INSTR_INT (instr_name[i], instr_size[i]); + instr_size[i] = 0; + } +#endif /* CAML_INSTR */ + last_fragment = NULL; + caml_fl_merge = Fl_head; +#ifdef DEBUG + fl_check (); +#endif +} + +static void truncate_flp (value changed) +{ + if (changed == Fl_head){ + flp_size = 0; + beyond = Val_NULL; + }else{ + while (flp_size > 0 && Next (flp[flp_size - 1]) >= changed) + -- flp_size; + if (beyond >= changed) beyond = Val_NULL; + } +} + +/* This is called by caml_compact_heap. */ +void caml_fl_reset (void) +{ + Next (Fl_head) = Val_NULL; + switch (policy){ + case Policy_next_fit: + fl_prev = Fl_head; + break; + case Policy_first_fit: + truncate_flp (Fl_head); + break; + default: + CAMLassert (0); + break; + } + caml_fl_cur_wsz = 0; + caml_fl_init_merge (); +} + +/* [caml_fl_merge_block] returns the head pointer of the next block after [bp], + because merging blocks may change the size of [bp]. */ +header_t *caml_fl_merge_block (value bp) +{ + value prev, cur; + header_t *adj; + header_t hd = Hd_val (bp); + mlsize_t prev_wosz; + + caml_fl_cur_wsz += Whsize_hd (hd); + +#ifdef DEBUG + caml_set_fields (bp, 0, Debug_free_major); +#endif + prev = caml_fl_merge; + cur = Next (prev); + /* The sweep code makes sure that this is the right place to insert + this block: */ + CAMLassert (prev < bp || prev == Fl_head); + CAMLassert (cur > bp || cur == Val_NULL); + + if (policy == Policy_first_fit) truncate_flp (prev); + + /* If [last_fragment] and [bp] are adjacent, merge them. */ + if (last_fragment == Hp_bp (bp)){ + mlsize_t bp_whsz = Whsize_val (bp); + if (bp_whsz <= Max_wosize){ + hd = Make_header (bp_whsz, 0, Caml_white); + bp = (value) last_fragment; + Hd_val (bp) = hd; + caml_fl_cur_wsz += Whsize_wosize (0); + } + } + + /* If [bp] and [cur] are adjacent, remove [cur] from the free-list + and merge them. */ + adj = (header_t *) &Field (bp, Wosize_hd (hd)); + if (adj == Hp_val (cur)){ + value next_cur = Next (cur); + mlsize_t cur_whsz = Whsize_val (cur); + + if (Wosize_hd (hd) + cur_whsz <= Max_wosize){ + Next (prev) = next_cur; + if (policy == Policy_next_fit && fl_prev == cur) fl_prev = prev; + hd = Make_header (Wosize_hd (hd) + cur_whsz, 0, Caml_blue); + Hd_val (bp) = hd; + adj = (header_t *) &Field (bp, Wosize_hd (hd)); +#ifdef DEBUG + fl_last = Val_NULL; + Next (cur) = (value) Debug_free_major; + Hd_val (cur) = Debug_free_major; +#endif + cur = next_cur; + } + } + /* If [prev] and [bp] are adjacent merge them, else insert [bp] into + the free-list if it is big enough. */ + prev_wosz = Wosize_val (prev); + if ((header_t *) &Field (prev, prev_wosz) == Hp_val (bp) + && prev_wosz + Whsize_hd (hd) < Max_wosize){ + Hd_val (prev) = Make_header (prev_wosz + Whsize_hd (hd), 0,Caml_blue); +#ifdef DEBUG + Hd_val (bp) = Debug_free_major; +#endif + CAMLassert (caml_fl_merge == prev); + }else if (Wosize_hd (hd) != 0){ + Hd_val (bp) = Bluehd_hd (hd); + Next (bp) = cur; + Next (prev) = bp; + caml_fl_merge = bp; + }else{ + /* This is a fragment. Leave it in white but remember it for eventual + merging with the next block. */ + last_fragment = (header_t *) bp; + caml_fl_cur_wsz -= Whsize_wosize (0); + } + return adj; +} + +/* This is a heap extension. We have to insert it in the right place + in the free-list. + [caml_fl_add_blocks] can only be called right after a call to + [caml_fl_allocate] that returned Val_NULL. + Most of the heap extensions are expected to be at the end of the + free list. (This depends on the implementation of [malloc].) + + [bp] must point to a list of blocks chained by their field 0, + terminated by Val_NULL, and field 1 of the first block must point to + the last block. +*/ +void caml_fl_add_blocks (value bp) +{ + value cur = bp; + CAMLassert (fl_last != Val_NULL); + CAMLassert (Next (fl_last) == Val_NULL); + do { + caml_fl_cur_wsz += Whsize_bp (cur); + cur = Field(cur, 0); + } while (cur != Val_NULL); + + if (bp > fl_last){ + Next (fl_last) = bp; + if (fl_last == caml_fl_merge && (char *) bp < caml_gc_sweep_hp){ + caml_fl_merge = Field (bp, 1); + } + if (policy == Policy_first_fit && flp_size < FLP_MAX){ + flp [flp_size++] = fl_last; + } + }else{ + value prev; + + prev = Fl_head; + cur = Next (prev); + while (cur != Val_NULL && cur < bp){ + CAMLassert (prev < bp || prev == Fl_head); + /* XXX TODO: extend flp on the fly */ + prev = cur; + cur = Next (prev); + } + CAMLassert (prev < bp || prev == Fl_head); + CAMLassert (cur > bp || cur == Val_NULL); + Next (Field (bp, 1)) = cur; + Next (prev) = bp; + /* When inserting blocks between [caml_fl_merge] and [caml_gc_sweep_hp], + we must advance [caml_fl_merge] to the new block, so that [caml_fl_merge] + is always the last free-list block before [caml_gc_sweep_hp]. */ + if (prev == caml_fl_merge && (char *) bp < caml_gc_sweep_hp){ + caml_fl_merge = Field (bp, 1); + } + if (policy == Policy_first_fit) truncate_flp (bp); + } +} + +/* Cut a block of memory into Max_wosize pieces, give them headers, + and optionally merge them into the free list. + arguments: + p: pointer to the first word of the block + size: size of the block (in words) + do_merge: 1 -> do merge; 0 -> do not merge + color: which color to give to the pieces; if [do_merge] is 1, this + is overridden by the merge code, but we have historically used + [Caml_white]. +*/ +void caml_make_free_blocks (value *p, mlsize_t size, int do_merge, int color) +{ + mlsize_t sz; + + while (size > 0){ + if (size > Whsize_wosize (Max_wosize)){ + sz = Whsize_wosize (Max_wosize); + }else{ + sz = size; + } + *(header_t *)p = + Make_header (Wosize_whsize (sz), 0, color); + if (do_merge) caml_fl_merge_block (Val_hp (p)); + size -= sz; + p += sz; + } +} + +void caml_set_allocation_policy (uintnat p) +{ + switch (p){ + case Policy_next_fit: + fl_prev = Fl_head; + policy = p; + break; + case Policy_first_fit: + flp_size = 0; + beyond = Val_NULL; + policy = p; + break; + default: + break; + } +} diff --git a/test/monniaux/ocaml/byterun/gc_ctrl.c b/test/monniaux/ocaml/byterun/gc_ctrl.c new file mode 100644 index 00000000..1cec4a25 --- /dev/null +++ b/test/monniaux/ocaml/byterun/gc_ctrl.c @@ -0,0 +1,691 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include "caml/alloc.h" +#include "caml/backtrace.h" +#include "caml/compact.h" +#include "caml/custom.h" +#include "caml/fail.h" +#include "caml/finalise.h" +#include "caml/freelist.h" +#include "caml/gc.h" +#include "caml/gc_ctrl.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/minor_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/signals.h" +#ifdef NATIVE_CODE +#include "caml/stack.h" +#else +#include "caml/stacks.h" +#endif +#include "caml/startup_aux.h" + +#ifndef NATIVE_CODE +extern uintnat caml_max_stack_size; /* defined in stacks.c */ +#endif + +double caml_stat_minor_words = 0.0, + caml_stat_promoted_words = 0.0, + caml_stat_major_words = 0.0; + +intnat caml_stat_minor_collections = 0, + caml_stat_major_collections = 0, + caml_stat_heap_wsz = 0, + caml_stat_top_heap_wsz = 0, + caml_stat_compactions = 0, + caml_stat_heap_chunks = 0; + +extern uintnat caml_major_heap_increment; /* percent or words; see major_gc.c */ +extern uintnat caml_percent_free; /* see major_gc.c */ +extern uintnat caml_percent_max; /* see compact.c */ +extern uintnat caml_allocation_policy; /* see freelist.c */ + +#define Next(hp) ((hp) + Whsize_hp (hp)) + +#ifdef DEBUG + +/* Check that [v]'s header looks good. [v] must be a block in the heap. */ +static void check_head (value v) +{ + CAMLassert (Is_block (v)); + CAMLassert (Is_in_heap (v)); + + CAMLassert (Wosize_val (v) != 0); + CAMLassert (Color_hd (Hd_val (v)) != Caml_blue); + CAMLassert (Is_in_heap (v)); + if (Tag_val (v) == Infix_tag){ + int offset = Wsize_bsize (Infix_offset_val (v)); + value trueval = Val_op (&Field (v, -offset)); + CAMLassert (Tag_val (trueval) == Closure_tag); + CAMLassert (Wosize_val (trueval) > offset); + CAMLassert (Is_in_heap (&Field (trueval, Wosize_val (trueval) - 1))); + }else{ + CAMLassert (Is_in_heap (&Field (v, Wosize_val (v) - 1))); + } + if (Tag_val (v) == Double_tag){ + CAMLassert (Wosize_val (v) == Double_wosize); + }else if (Tag_val (v) == Double_array_tag){ + CAMLassert (Wosize_val (v) % Double_wosize == 0); + } +} + +static void check_block (header_t *hp) +{ + mlsize_t i; + value v = Val_hp (hp); + value f; + + check_head (v); + switch (Tag_hp (hp)){ + case Abstract_tag: break; + case String_tag: + break; + case Double_tag: + CAMLassert (Wosize_val (v) == Double_wosize); + break; + case Double_array_tag: + CAMLassert (Wosize_val (v) % Double_wosize == 0); + break; + case Custom_tag: + CAMLassert (!Is_in_heap (Custom_ops_val (v))); + break; + + case Infix_tag: + CAMLassert (0); + break; + + default: + CAMLassert (Tag_hp (hp) < No_scan_tag); + for (i = 0; i < Wosize_hp (hp); i++){ + f = Field (v, i); + if (Is_block (f) && Is_in_heap (f)){ + check_head (f); + CAMLassert (Color_val (f) != Caml_blue); + } + } + } +} + +#endif /* DEBUG */ + +/* Check the heap structure (if compiled in debug mode) and + gather statistics; return the stats if [returnstats] is true, + otherwise return [Val_unit]. +*/ +static value heap_stats (int returnstats) +{ + CAMLparam0 (); + intnat live_words = 0, live_blocks = 0, + free_words = 0, free_blocks = 0, largest_free = 0, + fragments = 0, heap_chunks = 0; + char *chunk = caml_heap_start, *chunk_end; + header_t *cur_hp; +#ifdef DEBUG + header_t *prev_hp; +#endif + header_t cur_hd; + +#ifdef DEBUG + caml_gc_message (-1, "### OCaml runtime: heap check ###\n"); +#endif + + while (chunk != NULL){ + ++ heap_chunks; + chunk_end = chunk + Chunk_size (chunk); +#ifdef DEBUG + prev_hp = NULL; +#endif + cur_hp = (header_t *) chunk; + while (cur_hp < (header_t *) chunk_end){ + cur_hd = Hd_hp (cur_hp); + CAMLassert (Next (cur_hp) <= (header_t *) chunk_end); + switch (Color_hd (cur_hd)){ + case Caml_white: + if (Wosize_hd (cur_hd) == 0){ + ++ fragments; + CAMLassert (prev_hp == NULL + || Color_hp (prev_hp) != Caml_blue + || cur_hp == (header_t *) caml_gc_sweep_hp); + }else{ + if (caml_gc_phase == Phase_sweep + && cur_hp >= (header_t *) caml_gc_sweep_hp){ + ++ free_blocks; + free_words += Whsize_hd (cur_hd); + if (Whsize_hd (cur_hd) > largest_free){ + largest_free = Whsize_hd (cur_hd); + } + }else{ + ++ live_blocks; + live_words += Whsize_hd (cur_hd); +#ifdef DEBUG + check_block (cur_hp); +#endif + } + } + break; + case Caml_gray: case Caml_black: + CAMLassert (Wosize_hd (cur_hd) > 0); + ++ live_blocks; + live_words += Whsize_hd (cur_hd); +#ifdef DEBUG + check_block (cur_hp); +#endif + break; + case Caml_blue: + CAMLassert (Wosize_hd (cur_hd) > 0); + ++ free_blocks; + free_words += Whsize_hd (cur_hd); + if (Whsize_hd (cur_hd) > largest_free){ + largest_free = Whsize_hd (cur_hd); + } + /* not true any more with big heap chunks + CAMLassert (prev_hp == NULL + || (Color_hp (prev_hp) != Caml_blue + && Wosize_hp (prev_hp) > 0) + || cur_hp == caml_gc_sweep_hp); + CAMLassert (Next (cur_hp) == chunk_end + || (Color_hp (Next (cur_hp)) != Caml_blue + && Wosize_hp (Next (cur_hp)) > 0) + || (Whsize_hd (cur_hd) + Wosize_hp (Next (cur_hp)) + > Max_wosize) + || Next (cur_hp) == caml_gc_sweep_hp); + */ + break; + } +#ifdef DEBUG + prev_hp = cur_hp; +#endif + cur_hp = Next (cur_hp); + } + CAMLassert (cur_hp == (header_t *) chunk_end); + chunk = Chunk_next (chunk); + } + +#ifdef DEBUG + caml_final_invariant_check(); +#endif + + CAMLassert (heap_chunks == caml_stat_heap_chunks); + CAMLassert (live_words + free_words + fragments == caml_stat_heap_wsz); + + if (returnstats){ + CAMLlocal1 (res); + + /* get a copy of these before allocating anything... */ + double minwords = caml_stat_minor_words + + (double) (caml_young_alloc_end - caml_young_ptr); + double prowords = caml_stat_promoted_words; + double majwords = caml_stat_major_words + (double) caml_allocated_words; + intnat mincoll = caml_stat_minor_collections; + intnat majcoll = caml_stat_major_collections; + intnat heap_words = caml_stat_heap_wsz; + intnat cpct = caml_stat_compactions; + intnat top_heap_words = caml_stat_top_heap_wsz; + + res = caml_alloc_tuple (16); + Store_field (res, 0, caml_copy_double (minwords)); + Store_field (res, 1, caml_copy_double (prowords)); + Store_field (res, 2, caml_copy_double (majwords)); + Store_field (res, 3, Val_long (mincoll)); + Store_field (res, 4, Val_long (majcoll)); + Store_field (res, 5, Val_long (heap_words)); + Store_field (res, 6, Val_long (heap_chunks)); + Store_field (res, 7, Val_long (live_words)); + Store_field (res, 8, Val_long (live_blocks)); + Store_field (res, 9, Val_long (free_words)); + Store_field (res, 10, Val_long (free_blocks)); + Store_field (res, 11, Val_long (largest_free)); + Store_field (res, 12, Val_long (fragments)); + Store_field (res, 13, Val_long (cpct)); + Store_field (res, 14, Val_long (top_heap_words)); + Store_field (res, 15, Val_long (caml_stack_usage())); + CAMLreturn (res); + }else{ + CAMLreturn (Val_unit); + } +} + +#ifdef DEBUG +void caml_heap_check (void) +{ + heap_stats (0); +} +#endif + +CAMLprim value caml_gc_stat(value v) +{ + value result; + CAML_INSTR_SETUP (tmr, ""); + CAMLassert (v == Val_unit); + result = heap_stats (1); + CAML_INSTR_TIME (tmr, "explicit/gc_stat"); + return result; +} + +CAMLprim value caml_gc_quick_stat(value v) +{ + CAMLparam0 (); + CAMLlocal1 (res); + + /* get a copy of these before allocating anything... */ + double minwords = caml_stat_minor_words + + (double) (caml_young_alloc_end - caml_young_ptr); + double prowords = caml_stat_promoted_words; + double majwords = caml_stat_major_words + (double) caml_allocated_words; + intnat mincoll = caml_stat_minor_collections; + intnat majcoll = caml_stat_major_collections; + intnat heap_words = caml_stat_heap_wsz; + intnat top_heap_words = caml_stat_top_heap_wsz; + intnat cpct = caml_stat_compactions; + intnat heap_chunks = caml_stat_heap_chunks; + + res = caml_alloc_tuple (16); + Store_field (res, 0, caml_copy_double (minwords)); + Store_field (res, 1, caml_copy_double (prowords)); + Store_field (res, 2, caml_copy_double (majwords)); + Store_field (res, 3, Val_long (mincoll)); + Store_field (res, 4, Val_long (majcoll)); + Store_field (res, 5, Val_long (heap_words)); + Store_field (res, 6, Val_long (heap_chunks)); + Store_field (res, 7, Val_long (0)); + Store_field (res, 8, Val_long (0)); + Store_field (res, 9, Val_long (0)); + Store_field (res, 10, Val_long (0)); + Store_field (res, 11, Val_long (0)); + Store_field (res, 12, Val_long (0)); + Store_field (res, 13, Val_long (cpct)); + Store_field (res, 14, Val_long (top_heap_words)); + Store_field (res, 15, Val_long (caml_stack_usage())); + CAMLreturn (res); +} + +double caml_gc_minor_words_unboxed() +{ + return (caml_stat_minor_words + + (double) (caml_young_alloc_end - caml_young_ptr)); +} + +CAMLprim value caml_gc_minor_words(value v) +{ + CAMLparam0 (); /* v is ignored */ + CAMLreturn(caml_copy_double(caml_gc_minor_words_unboxed())); +} + +CAMLprim value caml_gc_counters(value v) +{ + CAMLparam0 (); /* v is ignored */ + CAMLlocal1 (res); + + /* get a copy of these before allocating anything... */ + double minwords = caml_stat_minor_words + + (double) (caml_young_alloc_end - caml_young_ptr); + double prowords = caml_stat_promoted_words; + double majwords = caml_stat_major_words + (double) caml_allocated_words; + + res = caml_alloc_tuple (3); + Store_field (res, 0, caml_copy_double (minwords)); + Store_field (res, 1, caml_copy_double (prowords)); + Store_field (res, 2, caml_copy_double (majwords)); + CAMLreturn (res); +} + +CAMLprim value caml_gc_huge_fallback_count (value v) +{ + return Val_long (caml_huge_fallback_count); +} + +CAMLprim value caml_gc_get(value v) +{ + CAMLparam0 (); /* v is ignored */ + CAMLlocal1 (res); + + res = caml_alloc_tuple (8); + Store_field (res, 0, Val_long (caml_minor_heap_wsz)); /* s */ + Store_field (res, 1, Val_long (caml_major_heap_increment)); /* i */ + Store_field (res, 2, Val_long (caml_percent_free)); /* o */ + Store_field (res, 3, Val_long (caml_verb_gc)); /* v */ + Store_field (res, 4, Val_long (caml_percent_max)); /* O */ +#ifndef NATIVE_CODE + Store_field (res, 5, Val_long (caml_max_stack_size)); /* l */ +#else + Store_field (res, 5, Val_long (0)); +#endif + Store_field (res, 6, Val_long (caml_allocation_policy)); /* a */ + Store_field (res, 7, Val_long (caml_major_window)); /* w */ + CAMLreturn (res); +} + +#define Max(x,y) ((x) < (y) ? (y) : (x)) + +static uintnat norm_pfree (uintnat p) +{ + return Max (p, 1); +} + +static uintnat norm_pmax (uintnat p) +{ + return p; +} + +static intnat norm_minsize (intnat s) +{ + if (s < Minor_heap_min) s = Minor_heap_min; + if (s > Minor_heap_max) s = Minor_heap_max; + return s; +} + +static uintnat norm_window (intnat w) +{ + if (w < 1) w = 1; + if (w > Max_major_window) w = Max_major_window; + return w; +} + +CAMLprim value caml_gc_set(value v) +{ + uintnat newpf, newpm; + asize_t newheapincr; + asize_t newminwsz; + uintnat oldpolicy; + CAML_INSTR_SETUP (tmr, ""); + + caml_verb_gc = Long_val (Field (v, 3)); + +#ifndef NATIVE_CODE + caml_change_max_stack_size (Long_val (Field (v, 5))); +#endif + + newpf = norm_pfree (Long_val (Field (v, 2))); + if (newpf != caml_percent_free){ + caml_percent_free = newpf; + caml_gc_message (0x20, "New space overhead: %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", caml_percent_free); + } + + newpm = norm_pmax (Long_val (Field (v, 4))); + if (newpm != caml_percent_max){ + caml_percent_max = newpm; + caml_gc_message (0x20, "New max overhead: %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", caml_percent_max); + } + + newheapincr = Long_val (Field (v, 1)); + if (newheapincr != caml_major_heap_increment){ + caml_major_heap_increment = newheapincr; + if (newheapincr > 1000){ + caml_gc_message (0x20, "New heap increment size: %" + ARCH_INTNAT_PRINTF_FORMAT "uk words\n", + caml_major_heap_increment/1024); + }else{ + caml_gc_message (0x20, "New heap increment size: %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", + caml_major_heap_increment); + } + } + oldpolicy = caml_allocation_policy; + caml_set_allocation_policy (Long_val (Field (v, 6))); + if (oldpolicy != caml_allocation_policy){ + caml_gc_message (0x20, "New allocation policy: %" + ARCH_INTNAT_PRINTF_FORMAT "u\n", caml_allocation_policy); + } + + /* This field was added in 4.03.0. */ + if (Wosize_val (v) >= 8){ + int old_window = caml_major_window; + caml_set_major_window (norm_window (Long_val (Field (v, 7)))); + if (old_window != caml_major_window){ + caml_gc_message (0x20, "New smoothing window size: %d\n", + caml_major_window); + } + } + + /* Minor heap size comes last because it will trigger a minor collection + (thus invalidating [v]) and it can raise [Out_of_memory]. */ + newminwsz = norm_minsize (Long_val (Field (v, 0))); + if (newminwsz != caml_minor_heap_wsz){ + caml_gc_message (0x20, "New minor heap size: %" + ARCH_SIZET_PRINTF_FORMAT "uk words\n", newminwsz / 1024); + caml_set_minor_heap_size (Bsize_wsize (newminwsz)); + } + CAML_INSTR_TIME (tmr, "explicit/gc_set"); + return Val_unit; +} + +CAMLprim value caml_gc_minor(value v) +{ + CAML_INSTR_SETUP (tmr, ""); + CAMLassert (v == Val_unit); + caml_request_minor_gc (); + caml_gc_dispatch (); + CAML_INSTR_TIME (tmr, "explicit/gc_minor"); + return Val_unit; +} + +static void test_and_compact (void) +{ + float fp; + + fp = 100.0 * caml_fl_cur_wsz / (caml_stat_heap_wsz - caml_fl_cur_wsz); + if (fp > 999999.0) fp = 999999.0; + caml_gc_message (0x200, "Estimated overhead (lower bound) = %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", + (uintnat) fp); + if (fp >= caml_percent_max){ + caml_gc_message (0x200, "Automatic compaction triggered.\n"); + caml_compact_heap (); + } +} + +CAMLprim value caml_gc_major(value v) +{ + CAML_INSTR_SETUP (tmr, ""); + CAMLassert (v == Val_unit); + caml_gc_message (0x1, "Major GC cycle requested\n"); + caml_empty_minor_heap (); + caml_finish_major_cycle (); + test_and_compact (); + caml_final_do_calls (); + CAML_INSTR_TIME (tmr, "explicit/gc_major"); + return Val_unit; +} + +CAMLprim value caml_gc_full_major(value v) +{ + CAML_INSTR_SETUP (tmr, ""); + CAMLassert (v == Val_unit); + caml_gc_message (0x1, "Full major GC cycle requested\n"); + caml_empty_minor_heap (); + caml_finish_major_cycle (); + caml_final_do_calls (); + caml_empty_minor_heap (); + caml_finish_major_cycle (); + test_and_compact (); + caml_final_do_calls (); + CAML_INSTR_TIME (tmr, "explicit/gc_full_major"); + return Val_unit; +} + +CAMLprim value caml_gc_major_slice (value v) +{ + CAML_INSTR_SETUP (tmr, ""); + CAMLassert (Is_long (v)); + caml_major_collection_slice (Long_val (v)); + CAML_INSTR_TIME (tmr, "explicit/gc_major_slice"); + return Val_long (0); +} + +CAMLprim value caml_gc_compaction(value v) +{ + CAML_INSTR_SETUP (tmr, ""); + CAMLassert (v == Val_unit); + caml_gc_message (0x10, "Heap compaction requested\n"); + caml_empty_minor_heap (); + caml_finish_major_cycle (); + caml_final_do_calls (); + caml_empty_minor_heap (); + caml_finish_major_cycle (); + caml_compact_heap (); + caml_final_do_calls (); + CAML_INSTR_TIME (tmr, "explicit/gc_compact"); + return Val_unit; +} + +CAMLprim value caml_get_minor_free (value v) +{ + return Val_int (caml_young_ptr - caml_young_alloc_start); +} + +CAMLprim value caml_get_major_bucket (value v) +{ + long i = Long_val (v); + if (i < 0) caml_invalid_argument ("Gc.get_bucket"); + if (i < caml_major_window){ + i += caml_major_ring_index; + if (i >= caml_major_window) i -= caml_major_window; + CAMLassert (0 <= i && i < caml_major_window); + return Val_long ((long) (caml_major_ring[i] * 1e6)); + }else{ + return Val_long (0); + } +} + +CAMLprim value caml_get_major_credit (value v) +{ + CAMLassert (v == Val_unit); + return Val_long ((long) (caml_major_work_credit * 1e6)); +} + +uintnat caml_normalize_heap_increment (uintnat i) +{ + if (i < Bsize_wsize (Heap_chunk_min)){ + i = Bsize_wsize (Heap_chunk_min); + } + return ((i + Page_size - 1) >> Page_log) << Page_log; +} + +/* [minor_size] and [major_size] are numbers of words + [major_incr] is either a percentage or a number of words */ +void caml_init_gc (uintnat minor_size, uintnat major_size, + uintnat major_incr, uintnat percent_fr, + uintnat percent_m, uintnat window) +{ + uintnat major_heap_size = + Bsize_wsize (caml_normalize_heap_increment (major_size)); + + CAML_INSTR_INIT (); + if (caml_init_alloc_for_heap () != 0){ + caml_fatal_error ("cannot initialize heap: mmap failed\n"); + } + if (caml_page_table_initialize(Bsize_wsize(minor_size) + major_heap_size)){ + caml_fatal_error ("OCaml runtime error: cannot initialize page table\n"); + } + caml_set_minor_heap_size (Bsize_wsize (norm_minsize (minor_size))); + caml_major_heap_increment = major_incr; + caml_percent_free = norm_pfree (percent_fr); + caml_percent_max = norm_pmax (percent_m); + caml_init_major_heap (major_heap_size); + caml_major_window = norm_window (window); + caml_gc_message (0x20, "Initial minor heap size: %" + ARCH_SIZET_PRINTF_FORMAT "uk words\n", + caml_minor_heap_wsz / 1024); + caml_gc_message (0x20, "Initial major heap size: %" + ARCH_INTNAT_PRINTF_FORMAT "uk bytes\n", + major_heap_size / 1024); + caml_gc_message (0x20, "Initial space overhead: %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", caml_percent_free); + caml_gc_message (0x20, "Initial max overhead: %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", caml_percent_max); + if (caml_major_heap_increment > 1000){ + caml_gc_message (0x20, "Initial heap increment: %" + ARCH_INTNAT_PRINTF_FORMAT "uk words\n", + caml_major_heap_increment / 1024); + }else{ + caml_gc_message (0x20, "Initial heap increment: %" + ARCH_INTNAT_PRINTF_FORMAT "u%%\n", + caml_major_heap_increment); + } + caml_gc_message (0x20, "Initial allocation policy: %" + ARCH_INTNAT_PRINTF_FORMAT "u\n", caml_allocation_policy); + caml_gc_message (0x20, "Initial smoothing window: %d\n", + caml_major_window); +} + + +/* FIXME After the startup_aux.c unification, move these functions there. */ + +CAMLprim value caml_runtime_variant (value unit) +{ + CAMLassert (unit == Val_unit); +#if defined (DEBUG) + return caml_copy_string ("d"); +#elif defined (CAML_INSTR) + return caml_copy_string ("i"); +#else + return caml_copy_string (""); +#endif +} + +extern int caml_parser_trace; + +CAMLprim value caml_runtime_parameters (value unit) +{ +#define F_Z ARCH_INTNAT_PRINTF_FORMAT +#define F_S ARCH_SIZET_PRINTF_FORMAT + + CAMLassert (unit == Val_unit); + return caml_alloc_sprintf + ("a=%d,b=%d,H=%"F_Z"u,i=%"F_Z"u,l=%"F_Z"u,o=%"F_Z"u,O=%"F_Z"u,p=%d,s=%"F_S"u,t=%"F_Z"u,v=%"F_Z"u,w=%d,W=%"F_Z"u", + /* a */ (int) caml_allocation_policy, + /* b */ caml_backtrace_active, + /* h */ /* missing */ /* FIXME add when changed to min_heap_size */ + /* H */ caml_use_huge_pages, + /* i */ caml_major_heap_increment, +#ifdef NATIVE_CODE + /* l */ (uintnat) 0, +#else + /* l */ caml_max_stack_size, +#endif + /* o */ caml_percent_free, + /* O */ caml_percent_max, + /* p */ caml_parser_trace, + /* R */ /* missing */ + /* s */ caml_minor_heap_wsz, + /* t */ caml_trace_level, + /* v */ caml_verb_gc, + /* w */ caml_major_window, + /* W */ caml_runtime_warnings + ); +#undef F_Z +#undef F_S +} + +/* Control runtime warnings */ + +CAMLprim value caml_ml_enable_runtime_warnings(value vbool) +{ + caml_runtime_warnings = Bool_val(vbool); + return Val_unit; +} + +CAMLprim value caml_ml_runtime_warnings_enabled(value unit) +{ + CAMLassert (unit == Val_unit); + return Val_bool(caml_runtime_warnings); +} diff --git a/test/monniaux/ocaml/byterun/globroots.c b/test/monniaux/ocaml/byterun/globroots.c new file mode 100644 index 00000000..f689723c --- /dev/null +++ b/test/monniaux/ocaml/byterun/globroots.c @@ -0,0 +1,291 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2001 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Registration of global memory roots */ + +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/roots.h" +#include "caml/globroots.h" + +/* The sets of global memory roots are represented as skip lists + (see William Pugh, "Skip lists: a probabilistic alternative to + balanced binary trees", Comm. ACM 33(6), 1990). */ + +struct global_root { + value * root; /* the address of the root */ + struct global_root * forward[1]; /* variable-length array */ +}; + +#define NUM_LEVELS 17 + +struct global_root_list { + value * root; /* dummy value for layout compatibility */ + struct global_root * forward[NUM_LEVELS]; /* forward chaining */ + int level; /* max used level */ +}; + +/* Generate a random level for a new node: 0 with probability 3/4, + 1 with probability 3/16, 2 with probability 3/64, etc. + We use a simple linear congruential PRNG (see Knuth vol 2) instead + of random(), because we need exactly 32 bits of pseudo-random data + (i.e. 2 * (NUM_LEVELS - 1)). Moreover, the congruential PRNG + is faster and guaranteed to be deterministic (to reproduce bugs). */ + +static uint32_t random_seed = 0; + +static int random_level(void) +{ + uint32_t r; + int level = 0; + + /* Linear congruence with modulus = 2^32, multiplier = 69069 + (Knuth vol 2 p. 106, line 15 of table 1), additive = 25173. */ + r = random_seed = random_seed * 69069 + 25173; + /* Knuth (vol 2 p. 13) shows that the least significant bits are + "less random" than the most significant bits with a modulus of 2^m, + so consume most significant bits first */ + while ((r & 0xC0000000U) == 0xC0000000U) { level++; r = r << 2; } + CAMLassert(level < NUM_LEVELS); + return level; +} + +/* Insertion in a global root list */ + +static void caml_insert_global_root(struct global_root_list * rootlist, + value * r) +{ + struct global_root * update[NUM_LEVELS]; + struct global_root * e, * f; + int i, new_level; + + CAMLassert(0 <= rootlist->level && rootlist->level < NUM_LEVELS); + + /* Init "cursor" to list head */ + e = (struct global_root *) rootlist; + /* Find place to insert new node */ + for (i = rootlist->level; i >= 0; i--) { + while (1) { + f = e->forward[i]; + if (f == NULL || f->root >= r) break; + e = f; + } + update[i] = e; + } + e = e->forward[0]; + /* If already present, don't do anything */ + if (e != NULL && e->root == r) return; + /* Insert additional element, updating list level if necessary */ + new_level = random_level(); + if (new_level > rootlist->level) { + for (i = rootlist->level + 1; i <= new_level; i++) + update[i] = (struct global_root *) rootlist; + rootlist->level = new_level; + } + e = caml_stat_alloc(sizeof(struct global_root) + + new_level * sizeof(struct global_root *)); + e->root = r; + for (i = 0; i <= new_level; i++) { + e->forward[i] = update[i]->forward[i]; + update[i]->forward[i] = e; + } +} + +/* Deletion in a global root list */ + +static void caml_delete_global_root(struct global_root_list * rootlist, + value * r) +{ + struct global_root * update[NUM_LEVELS]; + struct global_root * e, * f; + int i; + + CAMLassert(0 <= rootlist->level && rootlist->level < NUM_LEVELS); + + /* Init "cursor" to list head */ + e = (struct global_root *) rootlist; + /* Find element in list */ + for (i = rootlist->level; i >= 0; i--) { + while (1) { + f = e->forward[i]; + if (f == NULL || f->root >= r) break; + e = f; + } + update[i] = e; + } + e = e->forward[0]; + /* If not found, nothing to do */ + if (e == NULL || e->root != r) return; + /* Rebuild list without node */ + for (i = 0; i <= rootlist->level; i++) { + if (update[i]->forward[i] == e) + update[i]->forward[i] = e->forward[i]; + } + /* Reclaim list element */ + caml_stat_free(e); + /* Down-correct list level */ + while (rootlist->level > 0 && + rootlist->forward[rootlist->level] == NULL) + rootlist->level--; +} + +/* Iterate over a global root list */ + +static void caml_iterate_global_roots(scanning_action f, + struct global_root_list * rootlist) +{ + struct global_root * gr; + + for (gr = rootlist->forward[0]; gr != NULL; gr = gr->forward[0]) { + f(*(gr->root), gr->root); + } +} + +/* Empty a global root list */ + +static void caml_empty_global_roots(struct global_root_list * rootlist) +{ + struct global_root * gr, * next; + int i; + + CAMLassert(0 <= rootlist->level && rootlist->level < NUM_LEVELS); + + for (gr = rootlist->forward[0]; gr != NULL; /**/) { + next = gr->forward[0]; + caml_stat_free(gr); + gr = next; + } + for (i = 0; i <= rootlist->level; i++) rootlist->forward[i] = NULL; + rootlist->level = 0; +} + +/* The three global root lists */ + +struct global_root_list caml_global_roots = { NULL, { NULL, }, 0 }; + /* mutable roots, don't know whether old or young */ +struct global_root_list caml_global_roots_young = { NULL, { NULL, }, 0 }; + /* generational roots pointing to minor or major heap */ +struct global_root_list caml_global_roots_old = { NULL, { NULL, }, 0 }; + /* generational roots pointing to major heap */ + +/* Register a global C root of the mutable kind */ + +CAMLexport void caml_register_global_root(value *r) +{ + CAMLassert (((intnat) r & 3) == 0); /* compact.c demands this (for now) */ + caml_insert_global_root(&caml_global_roots, r); +} + +/* Un-register a global C root of the mutable kind */ + +CAMLexport void caml_remove_global_root(value *r) +{ + caml_delete_global_root(&caml_global_roots, r); +} + +/* Register a global C root of the generational kind */ + +CAMLexport void caml_register_generational_global_root(value *r) +{ + value v = *r; + CAMLassert (((intnat) r & 3) == 0); /* compact.c demands this (for now) */ + if (Is_block(v)) { + if (Is_young(v)) + caml_insert_global_root(&caml_global_roots_young, r); + else if (Is_in_heap(v)) + caml_insert_global_root(&caml_global_roots_old, r); + } +} + +/* Un-register a global C root of the generational kind */ + +CAMLexport void caml_remove_generational_global_root(value *r) +{ + value v = *r; + if (Is_block(v)) { + if (Is_in_heap_or_young(v)) + caml_delete_global_root(&caml_global_roots_young, r); + if (Is_in_heap(v)) + caml_delete_global_root(&caml_global_roots_old, r); + } +} + +/* Modify the value of a global C root of the generational kind */ + +CAMLexport void caml_modify_generational_global_root(value *r, value newval) +{ + value oldval = *r; + + /* It is OK to have a root in roots_young that suddenly points to + the old generation -- the next minor GC will take care of that. + What needs corrective action is a root in roots_old that suddenly + points to the young generation. */ + if (Is_block(newval) && Is_young(newval) && + Is_block(oldval) && Is_in_heap(oldval)) { + caml_delete_global_root(&caml_global_roots_old, r); + caml_insert_global_root(&caml_global_roots_young, r); + } + /* PR#4704 */ + else if (!Is_block(oldval) && Is_block(newval)) { + /* The previous value in the root was unboxed but now it is boxed. + The root won't appear in any of the root lists thus far (by virtue + of the operation of [caml_register_generational_global_root]), so we + need to make sure it gets in, or else it will never be scanned. */ + if (Is_young(newval)) + caml_insert_global_root(&caml_global_roots_young, r); + else if (Is_in_heap(newval)) + caml_insert_global_root(&caml_global_roots_old, r); + } + else if (Is_block(oldval) && !Is_block(newval)) { + /* The previous value in the root was boxed but now it is unboxed, so + the root should be removed. If [oldval] is young, this will happen + anyway at the next minor collection, but it is safer to delete it + here. */ + if (Is_in_heap_or_young(oldval)) + caml_delete_global_root(&caml_global_roots_young, r); + if (Is_in_heap(oldval)) + caml_delete_global_root(&caml_global_roots_old, r); + } + /* end PR#4704 */ + *r = newval; +} + +/* Scan all global roots */ + +void caml_scan_global_roots(scanning_action f) +{ + caml_iterate_global_roots(f, &caml_global_roots); + caml_iterate_global_roots(f, &caml_global_roots_young); + caml_iterate_global_roots(f, &caml_global_roots_old); +} + +/* Scan global roots for a minor collection */ + +void caml_scan_global_young_roots(scanning_action f) +{ + struct global_root * gr; + + caml_iterate_global_roots(f, &caml_global_roots); + caml_iterate_global_roots(f, &caml_global_roots_young); + /* Move young roots to old roots */ + for (gr = caml_global_roots_young.forward[0]; + gr != NULL; gr = gr->forward[0]) { + caml_insert_global_root(&caml_global_roots_old, gr->root); + } + caml_empty_global_roots(&caml_global_roots_young); +} diff --git a/test/monniaux/ocaml/byterun/hash.c b/test/monniaux/ocaml/byterun/hash.c new file mode 100644 index 00000000..f7d0d222 --- /dev/null +++ b/test/monniaux/ocaml/byterun/hash.c @@ -0,0 +1,419 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* The generic hashing primitive */ + +/* The interface of this file is in "mlvalues.h" (for [caml_hash_variant]) + and in "hash.h" (for the other exported functions). */ + +#include "caml/mlvalues.h" +#include "caml/custom.h" +#include "caml/memory.h" +#include "caml/hash.h" + +/* The new implementation, based on MurmurHash 3, + http://code.google.com/p/smhasher/ */ + +#define ROTL32(x,n) ((x) << n | (x) >> (32-n)) + +#define MIX(h,d) \ + d *= 0xcc9e2d51; \ + d = ROTL32(d, 15); \ + d *= 0x1b873593; \ + h ^= d; \ + h = ROTL32(h, 13); \ + h = h * 5 + 0xe6546b64; + +#define FINAL_MIX(h) \ + h ^= h >> 16; \ + h *= 0x85ebca6b; \ + h ^= h >> 13; \ + h *= 0xc2b2ae35; \ + h ^= h >> 16; + +CAMLexport uint32_t caml_hash_mix_uint32(uint32_t h, uint32_t d) +{ + MIX(h, d); + return h; +} + +/* Mix a platform-native integer. */ + +CAMLexport uint32_t caml_hash_mix_intnat(uint32_t h, intnat d) +{ + uint32_t n; +#ifdef ARCH_SIXTYFOUR + /* Mix the low 32 bits and the high 32 bits, in a way that preserves + 32/64 compatibility: we want n = (uint32_t) d + if d is in the range [-2^31, 2^31-1]. */ + n = (d >> 32) ^ (d >> 63) ^ d; + /* If 0 <= d < 2^31: d >> 32 = 0 d >> 63 = 0 + If -2^31 <= d < 0: d >> 32 = -1 d >> 63 = -1 + In both cases, n = (uint32_t) d. */ +#else + n = d; +#endif + MIX(h, n); + return h; +} + +/* Mix a 64-bit integer. */ + +CAMLexport uint32_t caml_hash_mix_int64(uint32_t h, int64_t d) +{ + uint32_t hi = (uint32_t) (d >> 32), lo = (uint32_t) d; + MIX(h, lo); + MIX(h, hi); + return h; +} + +/* Mix a double-precision float. + Treats +0.0 and -0.0 identically. + Treats all NaNs identically. +*/ + +CAMLexport uint32_t caml_hash_mix_double(uint32_t hash, double d) +{ + union { + double d; +#if defined(ARCH_BIG_ENDIAN) || (defined(__arm__) && !defined(__ARM_EABI__)) + struct { uint32_t h; uint32_t l; } i; +#else + struct { uint32_t l; uint32_t h; } i; +#endif + } u; + uint32_t h, l; + /* Convert to two 32-bit halves */ + u.d = d; + h = u.i.h; l = u.i.l; + /* Normalize NaNs */ + if ((h & 0x7FF00000) == 0x7FF00000 && (l | (h & 0xFFFFF)) != 0) { + h = 0x7FF00000; + l = 0x00000001; + } + /* Normalize -0 into +0 */ + else if (h == 0x80000000 && l == 0) { + h = 0; + } + MIX(hash, l); + MIX(hash, h); + return hash; +} + +/* Mix a single-precision float. + Treats +0.0 and -0.0 identically. + Treats all NaNs identically. +*/ + +CAMLexport uint32_t caml_hash_mix_float(uint32_t hash, float d) +{ + union { + float f; + uint32_t i; + } u; + uint32_t n; + /* Convert to int32_t */ + u.f = d; n = u.i; + /* Normalize NaNs */ + if ((n & 0x7F800000) == 0x7F800000 && (n & 0x007FFFFF) != 0) { + n = 0x7F800001; + } + /* Normalize -0 into +0 */ + else if (n == 0x80000000) { + n = 0; + } + MIX(hash, n); + return hash; +} + +/* Mix an OCaml string */ + +CAMLexport uint32_t caml_hash_mix_string(uint32_t h, value s) +{ + mlsize_t len = caml_string_length(s); + mlsize_t i; + uint32_t w; + + /* Mix by 32-bit blocks (little-endian) */ + for (i = 0; i + 4 <= len; i += 4) { +#ifdef ARCH_BIG_ENDIAN + w = Byte_u(s, i) + | (Byte_u(s, i+1) << 8) + | (Byte_u(s, i+2) << 16) + | (Byte_u(s, i+3) << 24); +#else + w = *((uint32_t *) &Byte_u(s, i)); +#endif + MIX(h, w); + } + /* Finish with up to 3 bytes */ + w = 0; + switch (len & 3) { + case 3: w = Byte_u(s, i+2) << 16; /* fallthrough */ + case 2: w |= Byte_u(s, i+1) << 8; /* fallthrough */ + case 1: w |= Byte_u(s, i); + MIX(h, w); + default: /*skip*/; /* len & 3 == 0, no extra bytes, do nothing */ + } + /* Finally, mix in the length. Ignore the upper 32 bits, generally 0. */ + h ^= (uint32_t) len; + return h; +} + +/* Maximal size of the queue used for breadth-first traversal. */ +#define HASH_QUEUE_SIZE 256 +/* Maximal number of Forward_tag links followed in one step */ +#define MAX_FORWARD_DEREFERENCE 1000 + +/* The generic hash function */ + +CAMLprim value caml_hash(value count, value limit, value seed, value obj) +{ + value queue[HASH_QUEUE_SIZE]; /* Queue of values to examine */ + intnat rd; /* Position of first value in queue */ + intnat wr; /* One past position of last value in queue */ + intnat sz; /* Max number of values to put in queue */ + intnat num; /* Max number of meaningful values to see */ + uint32_t h; /* Rolling hash */ + value v; + mlsize_t i, len; + + sz = Long_val(limit); + if (sz < 0 || sz > HASH_QUEUE_SIZE) sz = HASH_QUEUE_SIZE; + num = Long_val(count); + h = Int_val(seed); + queue[0] = obj; rd = 0; wr = 1; + + while (rd < wr && num > 0) { + v = queue[rd++]; + again: + if (Is_long(v)) { + h = caml_hash_mix_intnat(h, v); + num--; + } + else if (Is_in_value_area(v)) { + switch (Tag_val(v)) { + case String_tag: + h = caml_hash_mix_string(h, v); + num--; + break; + case Double_tag: + h = caml_hash_mix_double(h, Double_val(v)); + num--; + break; + case Double_array_tag: + for (i = 0, len = Wosize_val(v) / Double_wosize; i < len; i++) { + h = caml_hash_mix_double(h, Double_flat_field(v, i)); + num--; + if (num <= 0) break; + } + break; + case Abstract_tag: + /* Block contents unknown. Do nothing. */ + break; + case Infix_tag: + /* Mix in the offset to distinguish different functions from + the same mutually-recursive definition */ + h = caml_hash_mix_uint32(h, Infix_offset_val(v)); + v = v - Infix_offset_val(v); + goto again; + case Forward_tag: + /* PR#6361: we can have a loop here, so limit the number of + Forward_tag links being followed */ + for (i = MAX_FORWARD_DEREFERENCE; i > 0; i--) { + v = Forward_val(v); + if (Is_long(v) || !Is_in_value_area(v) || Tag_val(v) != Forward_tag) + goto again; + } + /* Give up on this object and move to the next */ + break; + case Object_tag: + h = caml_hash_mix_intnat(h, Oid_val(v)); + num--; + break; + case Custom_tag: + /* If no hashing function provided, do nothing. */ + /* Only use low 32 bits of custom hash, for 32/64 compatibility */ + if (Custom_ops_val(v)->hash != NULL) { + uint32_t n = (uint32_t) Custom_ops_val(v)->hash(v); + h = caml_hash_mix_uint32(h, n); + num--; + } + break; + default: + /* Mix in the tag and size, but do not count this towards [num] */ + h = caml_hash_mix_uint32(h, Whitehd_hd(Hd_val(v))); + /* Copy fields into queue, not exceeding the total size [sz] */ + for (i = 0, len = Wosize_val(v); i < len; i++) { + if (wr >= sz) break; + queue[wr++] = Field(v, i); + } + break; + } + } else { + /* v is a pointer outside the heap, probably a code pointer. + Shall we count it? Let's say yes by compatibility with old code. */ + h = caml_hash_mix_intnat(h, v); + num--; + } + } + /* Final mixing of bits */ + FINAL_MIX(h); + /* Fold result to the range [0, 2^30-1] so that it is a nonnegative + OCaml integer both on 32 and 64-bit platforms. */ + return Val_int(h & 0x3FFFFFFFU); +} + +/* The old implementation */ + +struct hash_state { + uintnat accu; + intnat univ_limit, univ_count; +}; + +static void hash_aux(struct hash_state*, value obj); + +CAMLprim value caml_hash_univ_param(value count, value limit, value obj) +{ + struct hash_state h; + h.univ_limit = Long_val(limit); + h.univ_count = Long_val(count); + h.accu = 0; + hash_aux(&h, obj); + return Val_long(h.accu & 0x3FFFFFFF); + /* The & has two purposes: ensure that the return value is positive + and give the same result on 32 bit and 64 bit architectures. */ +} + +#define Alpha 65599 +#define Beta 19 +#define Combine(new) (h->accu = h->accu * Alpha + (new)) +#define Combine_small(new) (h->accu = h->accu * Beta + (new)) + +static void hash_aux(struct hash_state* h, value obj) +{ + unsigned char * p; + mlsize_t i, j; + tag_t tag; + + h->univ_limit--; + if (h->univ_count < 0 || h->univ_limit < 0) return; + + again: + if (Is_long(obj)) { + h->univ_count--; + Combine(Long_val(obj)); + return; + } + + /* Pointers into the heap are well-structured blocks. So are atoms. + We can inspect the block contents. */ + + CAMLassert (Is_block (obj)); + if (Is_in_value_area(obj)) { + tag = Tag_val(obj); + switch (tag) { + case String_tag: + h->univ_count--; + i = caml_string_length(obj); + for (p = &Byte_u(obj, 0); i > 0; i--, p++) + Combine_small(*p); + break; + case Double_tag: + /* For doubles, we inspect their binary representation, LSB first. + The results are consistent among all platforms with IEEE floats. */ + h->univ_count--; +#ifdef ARCH_BIG_ENDIAN + for (p = &Byte_u(obj, sizeof(double) - 1), i = sizeof(double); + i > 0; + p--, i--) +#else + for (p = &Byte_u(obj, 0), i = sizeof(double); + i > 0; + p++, i--) +#endif + Combine_small(*p); + break; + case Double_array_tag: + h->univ_count--; + for (j = 0; j < Bosize_val(obj); j += sizeof(double)) { +#ifdef ARCH_BIG_ENDIAN + for (p = &Byte_u(obj, j + sizeof(double) - 1), i = sizeof(double); + i > 0; + p--, i--) +#else + for (p = &Byte_u(obj, j), i = sizeof(double); + i > 0; + p++, i--) +#endif + Combine_small(*p); + } + break; + case Abstract_tag: + /* We don't know anything about the contents of the block. + Better do nothing. */ + break; + case Infix_tag: + hash_aux(h, obj - Infix_offset_val(obj)); + break; + case Forward_tag: + obj = Forward_val (obj); + goto again; + case Object_tag: + h->univ_count--; + Combine(Oid_val(obj)); + break; + case Custom_tag: + /* If no hashing function provided, do nothing */ + if (Custom_ops_val(obj)->hash != NULL) { + h->univ_count--; + Combine(Custom_ops_val(obj)->hash(obj)); + } + break; + default: + h->univ_count--; + Combine_small(tag); + i = Wosize_val(obj); + while (i != 0) { + i--; + hash_aux(h, Field(obj, i)); + } + break; + } + return; + } + + /* Otherwise, obj is a pointer outside the heap, to an object with + a priori unknown structure. Use its physical address as hash key. */ + Combine((intnat) obj); +} + +/* Hashing variant tags */ + +CAMLexport value caml_hash_variant(char const * tag) +{ + value accu; + /* Same hashing algorithm as in ../typing/btype.ml, function hash_variant */ + for (accu = Val_int(0); *tag != 0; tag++) + accu = Val_int(223 * Int_val(accu) + *((unsigned char *) tag)); +#ifdef ARCH_SIXTYFOUR + accu = accu & Val_long(0x7FFFFFFFL); +#endif + /* Force sign extension of bit 31 for compatibility between 32 and 64-bit + platforms */ + return (int32_t) accu; +} diff --git a/test/monniaux/ocaml/byterun/instrtrace.c b/test/monniaux/ocaml/byterun/instrtrace.c new file mode 100644 index 00000000..824562e1 --- /dev/null +++ b/test/monniaux/ocaml/byterun/instrtrace.c @@ -0,0 +1,269 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Trace the instructions executed */ + +#ifdef DEBUG + +#include +#include +#include + +#include "caml/instrtrace.h" +#include "caml/instruct.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/opnames.h" +#include "caml/prims.h" +#include "caml/stacks.h" +#include "caml/startup_aux.h" + +extern code_t caml_start_code; + +intnat caml_icount = 0; + +void caml_stop_here () {} + +void caml_disasm_instr(pc) + code_t pc; +{ + int instr = *pc; + printf("%6ld %s", (long) (pc - caml_start_code), + instr < 0 || instr > STOP ? "???" : names_of_instructions[instr]); + pc++; + switch(instr) { + /* Instructions with one integer operand */ + case PUSHACC: case ACC: case POP: case ASSIGN: + case PUSHENVACC: case ENVACC: case PUSH_RETADDR: case APPLY: + case APPTERM1: case APPTERM2: case APPTERM3: case RETURN: + case GRAB: case PUSHGETGLOBAL: case GETGLOBAL: case SETGLOBAL: + case PUSHATOM: case ATOM: case MAKEBLOCK1: case MAKEBLOCK2: + case MAKEBLOCK3: case MAKEFLOATBLOCK: + case GETFIELD: case SETFIELD: case GETFLOATFIELD: case SETFLOATFIELD: + case BRANCH: case BRANCHIF: case BRANCHIFNOT: case PUSHTRAP: + case CONSTINT: case PUSHCONSTINT: case OFFSETINT: case OFFSETREF: + case OFFSETCLOSURE: case PUSHOFFSETCLOSURE: + printf(" %d\n", pc[0]); break; + /* Instructions with two operands */ + case APPTERM: case CLOSURE: case CLOSUREREC: case PUSHGETGLOBALFIELD: + case GETGLOBALFIELD: case MAKEBLOCK: + case BEQ: case BNEQ: case BLTINT: case BLEINT: case BGTINT: case BGEINT: + case BULTINT: case BUGEINT: + printf(" %d, %d\n", pc[0], pc[1]); break; + /* Instructions with a C primitive as operand */ + case C_CALLN: + printf(" %d,", pc[0]); pc++; + /* fallthrough */ + case C_CALL1: case C_CALL2: case C_CALL3: case C_CALL4: case C_CALL5: + if (pc[0] < 0 || pc[0] >= caml_prim_name_table.size) + printf(" unknown primitive %d\n", pc[0]); + else + printf(" %s\n", (char *) caml_prim_name_table.contents[pc[0]]); + break; + default: + printf("\n"); + } + fflush (stdout); +} + +char * caml_instr_string (code_t pc) +{ + static char buf[256]; + char nambuf[128]; + int instr = *pc; + char *nam; + + nam = (instr < 0 || instr > STOP) + ? (snprintf (nambuf, sizeof(nambuf), "???%d", instr), nambuf) + : names_of_instructions[instr]; + pc++; + switch (instr) { + /* Instructions with one integer operand */ + case PUSHACC: + case ACC: + case POP: + case ASSIGN: + case PUSHENVACC: + case ENVACC: + case PUSH_RETADDR: + case APPLY: + case APPTERM1: + case APPTERM2: + case APPTERM3: + case RETURN: + case GRAB: + case PUSHGETGLOBAL: + case GETGLOBAL: + case SETGLOBAL: + case PUSHATOM: + case ATOM: + case MAKEBLOCK1: + case MAKEBLOCK2: + case MAKEBLOCK3: + case MAKEFLOATBLOCK: + case GETFIELD: + case SETFIELD: + case GETFLOATFIELD: + case SETFLOATFIELD: + case BRANCH: + case BRANCHIF: + case BRANCHIFNOT: + case PUSHTRAP: + case CONSTINT: + case PUSHCONSTINT: + case OFFSETINT: + case OFFSETREF: + case OFFSETCLOSURE: + case PUSHOFFSETCLOSURE: + snprintf(buf, sizeof(buf), "%s %d", nam, pc[0]); + break; + /* Instructions with two operands */ + case APPTERM: + case CLOSURE: + case CLOSUREREC: + case PUSHGETGLOBALFIELD: + case GETGLOBALFIELD: + case MAKEBLOCK: + case BEQ: + case BNEQ: + case BLTINT: + case BLEINT: + case BGTINT: + case BGEINT: + case BULTINT: + case BUGEINT: + snprintf(buf, sizeof(buf), "%s %d, %d", nam, pc[0], pc[1]); + break; + case SWITCH: + snprintf(buf, sizeof(buf), "SWITCH sz%#lx=%ld::ntag%ld nint%ld", + (long) pc[0], (long) pc[0], (unsigned long) pc[0] >> 16, + (unsigned long) pc[0] & 0xffff); + break; + /* Instructions with a C primitive as operand */ + case C_CALLN: + snprintf(buf, sizeof(buf), "%s %d,", nam, pc[0]); + pc++; + /* fallthrough */ + case C_CALL1: + case C_CALL2: + case C_CALL3: + case C_CALL4: + case C_CALL5: + if (pc[0] < 0 || pc[0] >= caml_prim_name_table.size) + snprintf(buf, sizeof(buf), "%s unknown primitive %d", nam, pc[0]); + else + snprintf(buf, sizeof(buf), "%s %s", + nam, (char *) caml_prim_name_table.contents[pc[0]]); + break; + default: + snprintf(buf, sizeof(buf), "%s", nam); + break; + }; + return buf; +} + + +void +caml_trace_value_file (value v, code_t prog, int proglen, FILE * f) +{ + int i; + fprintf (f, "%#" ARCH_INTNAT_PRINTF_FORMAT "x", v); + if (!v) + return; + if (prog && v % sizeof (int) == 0 + && (code_t) v >= prog + && (code_t) v < (code_t) ((char *) prog + proglen)) + fprintf (f, "=code@%ld", (long) ((code_t) v - prog)); + else if (Is_long (v)) + fprintf (f, "=long%" ARCH_INTNAT_PRINTF_FORMAT "d", Long_val (v)); + else if ((void*)v >= (void*)caml_stack_low + && (void*)v < (void*)caml_stack_high) + fprintf (f, "=stack_%ld", (long) ((intnat*)caml_stack_high - (intnat*)v)); + else if (Is_block (v)) { + int s = Wosize_val (v); + int tg = Tag_val (v); + int l = 0; + switch (tg) { + case Closure_tag: + fprintf (f, "=closure[s%d,cod%ld]", + s, (long) ((code_t) (Code_val (v)) - prog)); + goto displayfields; + case String_tag: + l = caml_string_length (v); + fprintf (f, "=string[s%dL%d]'", s, l); + for (i = 0; i < ((l>0x1f)?0x1f:l) ; i++) { + if (isprint ((int) Byte (v, i))) + putc (Byte (v, i), f); + else + putc ('?', f); + }; + fprintf (f, "'"); + goto displayfields; + case Double_tag: + fprintf (f, "=float[s%d]=%g", s, Double_val (v)); + goto displayfields; + case Double_array_tag: + fprintf (f, "=floatarray[s%d]", s); + for (i = 0; i < ((s>0xf)?0xf:s); i++) + fprintf (f, " %g", Double_flat_field (v, i)); + goto displayfields; + case Abstract_tag: + fprintf (f, "=abstract[s%d]", s); + goto displayfields; + case Custom_tag: + fprintf (f, "=custom[s%d]", s); + goto displayfields; + default: + fprintf (f, "=block", tg, s); + displayfields: + if (s > 0) + fputs ("=(", f); + for (i = 0; i < s; i++) { + if (i > 20) { + fputs ("....", f); + break; + }; + if (i > 0) + putc (' ', f); + fprintf (f, "%#" ARCH_INTNAT_PRINTF_FORMAT "x", Field (v, i)); + }; + if (s > 0) + putc (')', f); + }; + } +} + +void +caml_trace_accu_sp_file (value accu, value * sp, code_t prog, int proglen, + FILE * f) +{ + int i; + value *p; + fprintf (f, "accu="); + caml_trace_value_file (accu, prog, proglen, f); + fprintf (f, "\n sp=%#" ARCH_INTNAT_PRINTF_FORMAT "x @%ld:", + (intnat) sp, (long) (caml_stack_high - sp)); + for (p = sp, i = 0; i < 12 + (1 << caml_trace_level) && p < caml_stack_high; + p++, i++) { + fprintf (f, "\n[%ld] ", (long) (caml_stack_high - p)); + caml_trace_value_file (*p, prog, proglen, f); + }; + putc ('\n', f); + fflush (f); +} + +#endif /* DEBUG */ diff --git a/test/monniaux/ocaml/byterun/intern.c b/test/monniaux/ocaml/byterun/intern.c new file mode 100644 index 00000000..565ed10d --- /dev/null +++ b/test/monniaux/ocaml/byterun/intern.c @@ -0,0 +1,1048 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Structured input, compact format */ + +/* The interface of this file is "caml/intext.h" */ + +#include +#include +#include "caml/alloc.h" +#include "caml/callback.h" +#include "caml/config.h" +#include "caml/custom.h" +#include "caml/fail.h" +#include "caml/gc.h" +#include "caml/intext.h" +#include "caml/io.h" +#include "caml/md5.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/misc.h" +#include "caml/reverse.h" + +static unsigned char * intern_src; +/* Reading pointer in block holding input data. */ + +static unsigned char * intern_input = NULL; +/* Pointer to beginning of block holding input data, + if non-NULL this pointer will be freed by the cleanup function. */ + +static header_t * intern_dest; +/* Writing pointer in destination block */ + +static char * intern_extra_block = NULL; +/* If non-NULL, point to new heap chunk allocated with caml_alloc_for_heap. */ + +static asize_t obj_counter; +/* Count how many objects seen so far */ + +static value * intern_obj_table = NULL; +/* The pointers to objects already seen */ + +static unsigned int intern_color; +/* Color to assign to newly created headers */ + +static header_t intern_header; +/* Original header of the destination block. + Meaningful only if intern_extra_block is NULL. */ + +static value intern_block = 0; +/* Point to the heap block allocated as destination block. + Meaningful only if intern_extra_block is NULL. */ + +static char * intern_resolve_code_pointer(unsigned char digest[16], + asize_t offset); + +CAMLnoreturn_start +static void intern_bad_code_pointer(unsigned char digest[16]) +CAMLnoreturn_end; + +static void intern_free_stack(void); + +static inline unsigned char read8u(void) +{ return *intern_src++; } + +static inline signed char read8s(void) +{ return *intern_src++; } + +static inline uint16_t read16u(void) +{ + uint16_t res = (intern_src[0] << 8) + intern_src[1]; + intern_src += 2; + return res; +} + +static inline int16_t read16s(void) +{ + int16_t res = (intern_src[0] << 8) + intern_src[1]; + intern_src += 2; + return res; +} + +static inline uint32_t read32u(void) +{ + uint32_t res = + ((uint32_t)(intern_src[0]) << 24) + (intern_src[1] << 16) + + (intern_src[2] << 8) + intern_src[3]; + intern_src += 4; + return res; +} + +static inline int32_t read32s(void) +{ + int32_t res = + ((uint32_t)(intern_src[0]) << 24) + (intern_src[1] << 16) + + (intern_src[2] << 8) + intern_src[3]; + intern_src += 4; + return res; +} + +#ifdef ARCH_SIXTYFOUR +static uintnat read64u(void) +{ + uintnat res = + ((uintnat) (intern_src[0]) << 56) + + ((uintnat) (intern_src[1]) << 48) + + ((uintnat) (intern_src[2]) << 40) + + ((uintnat) (intern_src[3]) << 32) + + ((uintnat) (intern_src[4]) << 24) + + ((uintnat) (intern_src[5]) << 16) + + ((uintnat) (intern_src[6]) << 8) + + (uintnat) (intern_src[7]); + intern_src += 8; + return res; +} +#endif + +static inline void readblock(void * dest, intnat len) +{ + memcpy(dest, intern_src, len); + intern_src += len; +} + +static void intern_init(void * src, void * input) +{ + /* This is asserted at the beginning of demarshaling primitives. + If it fails, it probably means that an exception was raised + without calling intern_cleanup() during the previous demarshaling. */ + CAMLassert (intern_input == NULL && intern_obj_table == NULL \ + && intern_extra_block == NULL && intern_block == 0); + intern_src = src; + intern_input = input; +} + +static void intern_cleanup(void) +{ + if (intern_input != NULL) { + caml_stat_free(intern_input); + intern_input = NULL; + } + if (intern_obj_table != NULL) { + caml_stat_free(intern_obj_table); + intern_obj_table = NULL; + } + if (intern_extra_block != NULL) { + /* free newly allocated heap chunk */ + caml_free_for_heap(intern_extra_block); + intern_extra_block = NULL; + } else if (intern_block != 0) { + /* restore original header for heap block, otherwise GC is confused */ + Hd_val(intern_block) = intern_header; + intern_block = 0; + } + /* free the recursion stack */ + intern_free_stack(); +} + +static void readfloat(double * dest, unsigned int code) +{ + if (sizeof(double) != 8) { + intern_cleanup(); + caml_invalid_argument("input_value: non-standard floats"); + } + readblock((char *) dest, 8); + /* Fix up endianness, if needed */ +#if ARCH_FLOAT_ENDIANNESS == 0x76543210 + /* Host is big-endian; fix up if data read is little-endian */ + if (code != CODE_DOUBLE_BIG) Reverse_64(dest, dest); +#elif ARCH_FLOAT_ENDIANNESS == 0x01234567 + /* Host is little-endian; fix up if data read is big-endian */ + if (code != CODE_DOUBLE_LITTLE) Reverse_64(dest, dest); +#else + /* Host is neither big nor little; permute as appropriate */ + if (code == CODE_DOUBLE_LITTLE) + Permute_64(dest, ARCH_FLOAT_ENDIANNESS, dest, 0x01234567) + else + Permute_64(dest, ARCH_FLOAT_ENDIANNESS, dest, 0x76543210); +#endif +} + +/* [len] is a number of floats */ +static void readfloats(double * dest, mlsize_t len, unsigned int code) +{ + mlsize_t i; + if (sizeof(double) != 8) { + intern_cleanup(); + caml_invalid_argument("input_value: non-standard floats"); + } + readblock((char *) dest, len * 8); + /* Fix up endianness, if needed */ +#if ARCH_FLOAT_ENDIANNESS == 0x76543210 + /* Host is big-endian; fix up if data read is little-endian */ + if (code != CODE_DOUBLE_ARRAY8_BIG && + code != CODE_DOUBLE_ARRAY32_BIG) { + for (i = 0; i < len; i++) Reverse_64(dest + i, dest + i); + } +#elif ARCH_FLOAT_ENDIANNESS == 0x01234567 + /* Host is little-endian; fix up if data read is big-endian */ + if (code != CODE_DOUBLE_ARRAY8_LITTLE && + code != CODE_DOUBLE_ARRAY32_LITTLE) { + for (i = 0; i < len; i++) Reverse_64(dest + i, dest + i); + } +#else + /* Host is neither big nor little; permute as appropriate */ + if (code == CODE_DOUBLE_ARRAY8_LITTLE || + code == CODE_DOUBLE_ARRAY32_LITTLE) { + for (i = 0; i < len; i++) + Permute_64(dest + i, ARCH_FLOAT_ENDIANNESS, dest + i, 0x01234567); + } else { + for (i = 0; i < len; i++) + Permute_64(dest + i, ARCH_FLOAT_ENDIANNESS, dest + i, 0x76543210); + } +#endif +} + +/* Item on the stack with defined operation */ +struct intern_item { + value * dest; + intnat arg; + enum { + OReadItems, /* read arg items and store them in dest[0], dest[1], ... */ + OFreshOID, /* generate a fresh OID and store it in *dest */ + OShift /* offset *dest by arg */ + } op; +}; + +/* FIXME: This is duplicated in two other places, with the only difference of + the type of elements stored in the stack. Possible solution in C would + be to instantiate stack these function via. C preprocessor macro. + */ + +#define INTERN_STACK_INIT_SIZE 256 +#define INTERN_STACK_MAX_SIZE (1024*1024*100) + +static struct intern_item intern_stack_init[INTERN_STACK_INIT_SIZE]; + +static struct intern_item * intern_stack = intern_stack_init; +static struct intern_item * intern_stack_limit = intern_stack_init + + INTERN_STACK_INIT_SIZE; + +/* Free the recursion stack if needed */ +static void intern_free_stack(void) +{ + if (intern_stack != intern_stack_init) { + caml_stat_free(intern_stack); + /* Reinitialize the globals for next time around */ + intern_stack = intern_stack_init; + intern_stack_limit = intern_stack + INTERN_STACK_INIT_SIZE; + } +} + +/* Same, then raise Out_of_memory */ +static void intern_stack_overflow(void) +{ + caml_gc_message (0x04, "Stack overflow in un-marshaling value\n"); + intern_free_stack(); + caml_raise_out_of_memory(); +} + +static struct intern_item * intern_resize_stack(struct intern_item * sp) +{ + asize_t newsize = 2 * (intern_stack_limit - intern_stack); + asize_t sp_offset = sp - intern_stack; + struct intern_item * newstack; + + if (newsize >= INTERN_STACK_MAX_SIZE) intern_stack_overflow(); + if (intern_stack == intern_stack_init) { + newstack = caml_stat_alloc_noexc(sizeof(struct intern_item) * newsize); + if (newstack == NULL) intern_stack_overflow(); + memcpy(newstack, intern_stack_init, + sizeof(struct intern_item) * INTERN_STACK_INIT_SIZE); + } else { + newstack = caml_stat_resize_noexc(intern_stack, + sizeof(struct intern_item) * newsize); + if (newstack == NULL) intern_stack_overflow(); + } + intern_stack = newstack; + intern_stack_limit = newstack + newsize; + return newstack + sp_offset; +} + +/* Convenience macros for requesting operation on the stack */ +#define PushItem() \ + do { \ + sp++; \ + if (sp >= intern_stack_limit) sp = intern_resize_stack(sp); \ + } while(0) + +#define ReadItems(_dest,_n) \ + do { \ + if (_n > 0) { \ + PushItem(); \ + sp->op = OReadItems; \ + sp->dest = _dest; \ + sp->arg = _n; \ + } \ + } while(0) + +static void intern_rec(value *dest) +{ + unsigned int code; + tag_t tag; + mlsize_t size, len, ofs_ind; + value v; + asize_t ofs; + header_t header; + unsigned char digest[16]; + struct custom_operations * ops; + char * codeptr; + struct intern_item * sp; + + sp = intern_stack; + + /* Initially let's try to read the first object from the stream */ + ReadItems(dest, 1); + + /* The un-marshaler loop, the recursion is unrolled */ + while(sp != intern_stack) { + + /* Interpret next item on the stack */ + dest = sp->dest; + switch (sp->op) { + case OFreshOID: + /* Refresh the object ID */ + /* but do not do it for predefined exception slots */ + if (Long_val(Field((value)dest, 1)) >= 0) + caml_set_oo_id((value)dest); + /* Pop item and iterate */ + sp--; + break; + case OShift: + /* Shift value by an offset */ + *dest += sp->arg; + /* Pop item and iterate */ + sp--; + break; + case OReadItems: + /* Pop item */ + sp->dest++; + if (--(sp->arg) == 0) sp--; + /* Read a value and set v to this value */ + code = read8u(); + if (code >= PREFIX_SMALL_INT) { + if (code >= PREFIX_SMALL_BLOCK) { + /* Small block */ + tag = code & 0xF; + size = (code >> 4) & 0x7; + read_block: + if (size == 0) { + v = Atom(tag); + } else { + v = Val_hp(intern_dest); + if (intern_obj_table != NULL) intern_obj_table[obj_counter++] = v; + *intern_dest = Make_header_allocated_here(size, tag, intern_color); + intern_dest += 1 + size; + /* For objects, we need to freshen the oid */ + if (tag == Object_tag) { + CAMLassert(size >= 2); + /* Request to read rest of the elements of the block */ + ReadItems(&Field(v, 2), size - 2); + /* Request freshing OID */ + PushItem(); + sp->op = OFreshOID; + sp->dest = (value*) v; + sp->arg = 1; + /* Finally read first two block elements: method table and old OID */ + ReadItems(&Field(v, 0), 2); + } else + /* If it's not an object then read the contents of the block */ + ReadItems(&Field(v, 0), size); + } + } else { + /* Small integer */ + v = Val_int(code & 0x3F); + } + } else { + if (code >= PREFIX_SMALL_STRING) { + /* Small string */ + len = (code & 0x1F); + read_string: + size = (len + sizeof(value)) / sizeof(value); + v = Val_hp(intern_dest); + if (intern_obj_table != NULL) intern_obj_table[obj_counter++] = v; + *intern_dest = Make_header_allocated_here(size, String_tag, intern_color); + intern_dest += 1 + size; + Field(v, size - 1) = 0; + ofs_ind = Bsize_wsize(size) - 1; + Byte(v, ofs_ind) = ofs_ind - len; + readblock((char *)String_val(v), len); + } else { + switch(code) { + case CODE_INT8: + v = Val_long(read8s()); + break; + case CODE_INT16: + v = Val_long(read16s()); + break; + case CODE_INT32: + v = Val_long(read32s()); + break; + case CODE_INT64: +#ifdef ARCH_SIXTYFOUR + v = Val_long((intnat) (read64u())); + break; +#else + intern_cleanup(); + caml_failwith("input_value: integer too large"); + break; +#endif + case CODE_SHARED8: + ofs = read8u(); + read_shared: + CAMLassert (ofs > 0); + CAMLassert (ofs <= obj_counter); + CAMLassert (intern_obj_table != NULL); + v = intern_obj_table[obj_counter - ofs]; + break; + case CODE_SHARED16: + ofs = read16u(); + goto read_shared; + case CODE_SHARED32: + ofs = read32u(); + goto read_shared; +#ifdef ARCH_SIXTYFOUR + case CODE_SHARED64: + ofs = read64u(); + goto read_shared; +#endif + case CODE_BLOCK32: + header = (header_t) read32u(); + tag = Tag_hd(header); + size = Wosize_hd(header); + goto read_block; +#ifdef ARCH_SIXTYFOUR + case CODE_BLOCK64: + header = (header_t) read64u(); + tag = Tag_hd(header); + size = Wosize_hd(header); + goto read_block; +#endif + case CODE_STRING8: + len = read8u(); + goto read_string; + case CODE_STRING32: + len = read32u(); + goto read_string; +#ifdef ARCH_SIXTYFOUR + case CODE_STRING64: + len = read64u(); + goto read_string; +#endif + case CODE_DOUBLE_LITTLE: + case CODE_DOUBLE_BIG: + v = Val_hp(intern_dest); + if (intern_obj_table != NULL) intern_obj_table[obj_counter++] = v; + *intern_dest = Make_header_allocated_here(Double_wosize, Double_tag, + intern_color); + intern_dest += 1 + Double_wosize; + readfloat((double *) v, code); + break; + case CODE_DOUBLE_ARRAY8_LITTLE: + case CODE_DOUBLE_ARRAY8_BIG: + len = read8u(); + read_double_array: + size = len * Double_wosize; + v = Val_hp(intern_dest); + if (intern_obj_table != NULL) intern_obj_table[obj_counter++] = v; + *intern_dest = Make_header_allocated_here(size, Double_array_tag, + intern_color); + intern_dest += 1 + size; + readfloats((double *) v, len, code); + break; + case CODE_DOUBLE_ARRAY32_LITTLE: + case CODE_DOUBLE_ARRAY32_BIG: + len = read32u(); + goto read_double_array; +#ifdef ARCH_SIXTYFOUR + case CODE_DOUBLE_ARRAY64_LITTLE: + case CODE_DOUBLE_ARRAY64_BIG: + len = read64u(); + goto read_double_array; +#endif + case CODE_CODEPOINTER: + ofs = read32u(); + readblock(digest, 16); + codeptr = intern_resolve_code_pointer(digest, ofs); + if (codeptr != NULL) { + v = (value) codeptr; + } else { + value * function_placeholder = + caml_named_value ("Debugger.function_placeholder"); + if (function_placeholder != NULL) { + v = *function_placeholder; + } else { + intern_cleanup(); + intern_bad_code_pointer(digest); + } + } + break; + case CODE_INFIXPOINTER: + ofs = read32u(); + /* Read a value to *dest, then offset *dest by ofs */ + PushItem(); + sp->dest = dest; + sp->op = OShift; + sp->arg = ofs; + ReadItems(dest, 1); + continue; /* with next iteration of main loop, skipping *dest = v */ + case CODE_CUSTOM: + ops = caml_find_custom_operations((char *) intern_src); + if (ops == NULL) { + intern_cleanup(); + caml_failwith("input_value: unknown custom block identifier"); + } + while (*intern_src++ != 0) /*nothing*/; /*skip identifier*/ + size = ops->deserialize((void *) (intern_dest + 2)); + size = 1 + (size + sizeof(value) - 1) / sizeof(value); + v = Val_hp(intern_dest); + if (intern_obj_table != NULL) intern_obj_table[obj_counter++] = v; + *intern_dest = Make_header_allocated_here(size, Custom_tag, + intern_color); + Custom_ops_val(v) = ops; + + if (ops->finalize != NULL && Is_young(v)) { + /* Remember that the block has a finalizer. */ + add_to_custom_table (&caml_custom_table, v, 0, 1); + } + + intern_dest += 1 + size; + break; + default: + intern_cleanup(); + caml_failwith("input_value: ill-formed message"); + } + } + } + /* end of case OReadItems */ + *dest = v; + break; + default: + CAMLassert(0); + } + } + /* We are done. Cleanup the stack and leave the function */ + intern_free_stack(); +} + +static void intern_alloc(mlsize_t whsize, mlsize_t num_objects, + int outside_heap) +{ + mlsize_t wosize; + + if (whsize == 0) { + CAMLassert (intern_extra_block == NULL && intern_block == 0 + && intern_obj_table == NULL); + return; + } + wosize = Wosize_whsize(whsize); + if (outside_heap || wosize > Max_wosize) { + /* Round desired size up to next page */ + asize_t request = + ((Bsize_wsize(whsize) + Page_size - 1) >> Page_log) << Page_log; + intern_extra_block = caml_alloc_for_heap(request); + if (intern_extra_block == NULL) { + intern_cleanup(); + caml_raise_out_of_memory(); + } + intern_color = + outside_heap ? Caml_black : caml_allocation_color(intern_extra_block); + intern_dest = (header_t *) intern_extra_block; + CAMLassert (intern_block == 0); + } else { + /* this is a specialised version of caml_alloc from alloc.c */ + if (wosize <= Max_young_wosize){ + if (wosize == 0){ + intern_block = Atom (String_tag); + } else { + intern_block = caml_alloc_small (wosize, String_tag); + } + }else{ + intern_block = caml_alloc_shr_no_raise (wosize, String_tag); + /* do not do the urgent_gc check here because it might darken + intern_block into gray and break the intern_color assertion below */ + if (intern_block == 0) { + intern_cleanup(); + caml_raise_out_of_memory(); + } + } + intern_header = Hd_val(intern_block); + intern_color = Color_hd(intern_header); + CAMLassert (intern_color == Caml_white || intern_color == Caml_black); + intern_dest = (header_t *) Hp_val(intern_block); + CAMLassert (intern_extra_block == NULL); + } + obj_counter = 0; + if (num_objects > 0) { + intern_obj_table = (value *) caml_stat_alloc_noexc(num_objects * sizeof(value)); + if (intern_obj_table == NULL) { + intern_cleanup(); + caml_raise_out_of_memory(); + } + } else + CAMLassert(intern_obj_table == NULL); +} + +static void intern_add_to_heap(mlsize_t whsize) +{ + /* Add new heap chunk to heap if needed */ + if (intern_extra_block != NULL) { + /* If heap chunk not filled totally, build free block at end */ + asize_t request = Chunk_size (intern_extra_block); + header_t * end_extra_block = + (header_t *) intern_extra_block + Wsize_bsize(request); + CAMLassert(intern_block == 0); + CAMLassert(intern_dest <= end_extra_block); + if (intern_dest < end_extra_block){ + caml_make_free_blocks ((value *) intern_dest, + end_extra_block - intern_dest, 0, Caml_white); + } + caml_allocated_words += + Wsize_bsize ((char *) intern_dest - intern_extra_block); + caml_add_to_heap(intern_extra_block); + intern_extra_block = NULL; // To prevent intern_cleanup freeing it + } else { + intern_block = 0; // To prevent intern_cleanup rewriting its header + } +} + +/* Parsing the header */ + +struct marshal_header { + uint32_t magic; + int header_len; + uintnat data_len; + uintnat num_objects; + uintnat whsize; +}; + +static void caml_parse_header(char * fun_name, + /*out*/ struct marshal_header * h) +{ + char errmsg[100]; + + h->magic = read32u(); + switch(h->magic) { + case Intext_magic_number_small: + h->header_len = 20; + h->data_len = read32u(); + h->num_objects = read32u(); +#ifdef ARCH_SIXTYFOUR + read32u(); + h->whsize = read32u(); +#else + h->whsize = read32u(); + read32u(); +#endif + break; + case Intext_magic_number_big: +#ifdef ARCH_SIXTYFOUR + h->header_len = 32; + read32u(); + h->data_len = read64u(); + h->num_objects = read64u(); + h->whsize = read64u(); +#else + errmsg[sizeof(errmsg) - 1] = 0; + snprintf(errmsg, sizeof(errmsg) - 1, + "%s: object too large to be read back on a 32-bit platform", + fun_name); + caml_failwith(errmsg); +#endif + break; + default: + errmsg[sizeof(errmsg) - 1] = 0; + snprintf(errmsg, sizeof(errmsg) - 1, + "%s: bad object", + fun_name); + caml_failwith(errmsg); + } +} + +/* Reading from a channel */ + +static value caml_input_val_core(struct channel *chan, int outside_heap) +{ + intnat r; + char header[32]; + struct marshal_header h; + char * block; + value res; + + if (! caml_channel_binary_mode(chan)) + caml_failwith("input_value: not a binary channel"); + /* Read and parse the header */ + r = caml_really_getblock(chan, header, 20); + if (r == 0) + caml_raise_end_of_file(); + else if (r < 20) + caml_failwith("input_value: truncated object"); + intern_src = (unsigned char *) header; + if (read32u() == Intext_magic_number_big) { + /* Finish reading the header */ + if (caml_really_getblock(chan, header + 20, 32 - 20) < 32 - 20) + caml_failwith("input_value: truncated object"); + } + intern_src = (unsigned char *) header; + caml_parse_header("input_value", &h); + /* Read block from channel */ + block = caml_stat_alloc(h.data_len); + /* During [caml_really_getblock], concurrent [caml_input_val] operations + can take place (via signal handlers or context switching in systhreads), + and [intern_input] may change. So, wait until [caml_really_getblock] + is over before using [intern_input] and the other global vars. */ + if (caml_really_getblock(chan, block, h.data_len) < h.data_len) { + caml_stat_free(block); + caml_failwith("input_value: truncated object"); + } + /* Initialize global state */ + intern_init(block, block); + intern_alloc(h.whsize, h.num_objects, outside_heap); + /* Fill it in */ + intern_rec(&res); + if (!outside_heap) { + intern_add_to_heap(h.whsize); + } else { + caml_disown_for_heap(intern_extra_block); + intern_extra_block = NULL; + intern_block = 0; + } + /* Free everything */ + intern_cleanup(); + return caml_check_urgent_gc(res); +} + +value caml_input_val(struct channel* chan) +{ + return caml_input_val_core(chan, 0); +} + +CAMLprim value caml_input_value(value vchan) +{ + CAMLparam1 (vchan); + struct channel * chan = Channel(vchan); + CAMLlocal1 (res); + + Lock(chan); + res = caml_input_val(chan); + Unlock(chan); + CAMLreturn (res); +} + +/* Reading from memory-resident blocks */ + +CAMLprim value caml_input_value_to_outside_heap(value vchan) +{ + CAMLparam1 (vchan); + struct channel * chan = Channel(vchan); + CAMLlocal1 (res); + + Lock(chan); + res = caml_input_val_core(chan, 1); + Unlock(chan); + CAMLreturn (res); +} + +CAMLexport value caml_input_val_from_bytes(value str, intnat ofs) +{ + CAMLparam1 (str); + CAMLlocal1 (obj); + struct marshal_header h; + + /* Initialize global state */ + intern_init(&Byte_u(str, ofs), NULL); + caml_parse_header("input_val_from_string", &h); + if (ofs + h.header_len + h.data_len > caml_string_length(str)) + caml_failwith("input_val_from_string: bad length"); + /* Allocate result */ + intern_alloc(h.whsize, h.num_objects, 0); + intern_src = &Byte_u(str, ofs + h.header_len); /* If a GC occurred */ + /* Fill it in */ + intern_rec(&obj); + intern_add_to_heap(h.whsize); + /* Free everything */ + intern_cleanup(); + CAMLreturn (caml_check_urgent_gc(obj)); +} + +CAMLprim value caml_input_value_from_string(value str, value ofs) +{ + return caml_input_val_from_bytes(str, Long_val(ofs)); +} + +CAMLprim value caml_input_value_from_bytes(value str, value ofs) +{ + return caml_input_val_from_bytes(str, Long_val(ofs)); +} + +static value input_val_from_block(struct marshal_header * h) +{ + value obj; + /* Allocate result */ + intern_alloc(h->whsize, h->num_objects, 0); + /* Fill it in */ + intern_rec(&obj); + intern_add_to_heap(h->whsize); + /* Free internal data structures */ + intern_cleanup(); + return caml_check_urgent_gc(obj); +} + +CAMLexport value caml_input_value_from_malloc(char * data, intnat ofs) +{ + struct marshal_header h; + + intern_init(data + ofs, data); + + caml_parse_header("input_value_from_malloc", &h); + + return input_val_from_block(&h); +} + +/* [len] is a number of bytes */ +CAMLexport value caml_input_value_from_block(char * data, intnat len) +{ + struct marshal_header h; + + /* Initialize global state */ + intern_init(data, NULL); + caml_parse_header("input_value_from_block", &h); + if (h.header_len + h.data_len > len) + caml_failwith("input_val_from_block: bad length"); + return input_val_from_block(&h); +} + +/* [ofs] is a [value] that represents a number of bytes + result is a [value] that represents a number of bytes + To handle both the small and the big format, + we assume 20 bytes are available at [buff + ofs], + and we return the data size + the length of the part of the header + that remains to be read. */ + +CAMLprim value caml_marshal_data_size(value buff, value ofs) +{ + uint32_t magic; + int header_len; + uintnat data_len; + + intern_src = &Byte_u(buff, Long_val(ofs)); + magic = read32u(); + switch(magic) { + case Intext_magic_number_small: + header_len = 20; + data_len = read32u(); + break; + case Intext_magic_number_big: +#ifdef ARCH_SIXTYFOUR + header_len = 32; + read32u(); + data_len = read64u(); +#else + caml_failwith("Marshal.data_size: " + "object too large to be read back on a 32-bit platform"); +#endif + break; + default: + caml_failwith("Marshal.data_size: bad object"); + } + return Val_long((header_len - 20) + data_len); +} + +/* Resolution of code pointers */ + +static char * intern_resolve_code_pointer(unsigned char digest[16], + asize_t offset) +{ + int i; + for (i = caml_code_fragments_table.size - 1; i >= 0; i--) { + struct code_fragment * cf = caml_code_fragments_table.contents[i]; + if (! cf->digest_computed) { + caml_md5_block(cf->digest, cf->code_start, cf->code_end - cf->code_start); + cf->digest_computed = 1; + } + if (memcmp(digest, cf->digest, 16) == 0) { + if (cf->code_start + offset < cf->code_end) + return cf->code_start + offset; + else + return NULL; + } + } + return NULL; +} + +static void intern_bad_code_pointer(unsigned char digest[16]) +{ + char msg[256]; + snprintf(msg, sizeof(msg), + "input_value: unknown code module " + "%02X%02X%02X%02X%02X%02X%02X%02X" + "%02X%02X%02X%02X%02X%02X%02X%02X", + digest[0], digest[1], digest[2], digest[3], + digest[4], digest[5], digest[6], digest[7], + digest[8], digest[9], digest[10], digest[11], + digest[12], digest[13], digest[14], digest[15]); + caml_failwith(msg); +} + +/* Functions for writing user-defined marshallers */ + +CAMLexport int caml_deserialize_uint_1(void) +{ + return read8u(); +} + +CAMLexport int caml_deserialize_sint_1(void) +{ + return read8s(); +} + +CAMLexport int caml_deserialize_uint_2(void) +{ + return read16u(); +} + +CAMLexport int caml_deserialize_sint_2(void) +{ + return read16s(); +} + +CAMLexport uint32_t caml_deserialize_uint_4(void) +{ + return read32u(); +} + +CAMLexport int32_t caml_deserialize_sint_4(void) +{ + return read32s(); +} + +CAMLexport uint64_t caml_deserialize_uint_8(void) +{ + uint64_t i; + caml_deserialize_block_8(&i, 1); + return i; +} + +CAMLexport int64_t caml_deserialize_sint_8(void) +{ + int64_t i; + caml_deserialize_block_8(&i, 1); + return i; +} + +CAMLexport float caml_deserialize_float_4(void) +{ + float f; + caml_deserialize_block_4(&f, 1); + return f; +} + +CAMLexport double caml_deserialize_float_8(void) +{ + double f; + caml_deserialize_block_float_8(&f, 1); + return f; +} + +CAMLexport void caml_deserialize_block_1(void * data, intnat len) +{ + memcpy(data, intern_src, len); + intern_src += len; +} + +CAMLexport void caml_deserialize_block_2(void * data, intnat len) +{ +#ifndef ARCH_BIG_ENDIAN + unsigned char * p, * q; + for (p = intern_src, q = data; len > 0; len--, p += 2, q += 2) + Reverse_16(q, p); + intern_src = p; +#else + memcpy(data, intern_src, len * 2); + intern_src += len * 2; +#endif +} + +CAMLexport void caml_deserialize_block_4(void * data, intnat len) +{ +#ifndef ARCH_BIG_ENDIAN + unsigned char * p, * q; + for (p = intern_src, q = data; len > 0; len--, p += 4, q += 4) + Reverse_32(q, p); + intern_src = p; +#else + memcpy(data, intern_src, len * 4); + intern_src += len * 4; +#endif +} + +CAMLexport void caml_deserialize_block_8(void * data, intnat len) +{ +#ifndef ARCH_BIG_ENDIAN + unsigned char * p, * q; + for (p = intern_src, q = data; len > 0; len--, p += 8, q += 8) + Reverse_64(q, p); + intern_src = p; +#else + memcpy(data, intern_src, len * 8); + intern_src += len * 8; +#endif +} + +CAMLexport void caml_deserialize_block_float_8(void * data, intnat len) +{ +#if ARCH_FLOAT_ENDIANNESS == 0x01234567 + memcpy(data, intern_src, len * 8); + intern_src += len * 8; +#elif ARCH_FLOAT_ENDIANNESS == 0x76543210 + unsigned char * p, * q; + for (p = intern_src, q = data; len > 0; len--, p += 8, q += 8) + Reverse_64(q, p); + intern_src = p; +#else + unsigned char * p, * q; + for (p = intern_src, q = data; len > 0; len--, p += 8, q += 8) + Permute_64(q, ARCH_FLOAT_ENDIANNESS, p, 0x01234567); + intern_src = p; +#endif +} + +CAMLexport void caml_deserialize_error(char * msg) +{ + intern_cleanup(); + caml_failwith(msg); +} diff --git a/test/monniaux/ocaml/byterun/interp.c b/test/monniaux/ocaml/byterun/interp.c new file mode 100644 index 00000000..0b74df3d --- /dev/null +++ b/test/monniaux/ocaml/byterun/interp.c @@ -0,0 +1,1172 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* The bytecode interpreter */ +#include +#include "caml/alloc.h" +#include "caml/backtrace.h" +#include "caml/callback.h" +#include "caml/debugger.h" +#include "caml/fail.h" +#include "caml/fix_code.h" +#include "caml/instrtrace.h" +#include "caml/instruct.h" +#include "caml/interp.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/prims.h" +#include "caml/signals.h" +#include "caml/stacks.h" +#include "caml/startup_aux.h" + +/* Registers for the abstract machine: + pc the code pointer + sp the stack pointer (grows downward) + accu the accumulator + env heap-allocated environment + caml_trapsp pointer to the current trap frame + extra_args number of extra arguments provided by the caller + +sp is a local copy of the global variable caml_extern_sp. */ + +/* Instruction decoding */ + +#ifdef THREADED_CODE +# define Instruct(name) lbl_##name +# if defined(ARCH_SIXTYFOUR) && !defined(ARCH_CODE32) +# define Jumptbl_base ((char *) &&lbl_ACC0) +# else +# define Jumptbl_base ((char *) 0) +# define jumptbl_base ((char *) 0) +# endif +# ifdef DEBUG +# define Next goto next_instr +# else +# define Next goto *(void *)(jumptbl_base + *pc++) +# endif +#else +# define Instruct(name) case name +# define Next break +#endif + +/* GC interface */ + +#define Setup_for_gc \ + { sp -= 2; sp[0] = accu; sp[1] = env; caml_extern_sp = sp; } +#define Restore_after_gc \ + { accu = sp[0]; env = sp[1]; sp += 2; } +#define Setup_for_c_call \ + { saved_pc = pc; *--sp = env; caml_extern_sp = sp; } +#define Restore_after_c_call \ + { sp = caml_extern_sp; env = *sp++; saved_pc = NULL; } + +/* An event frame must look like accu + a C_CALL frame + a RETURN 1 frame */ +#define Setup_for_event \ + { sp -= 6; \ + sp[0] = accu; /* accu */ \ + sp[1] = Val_unit; /* C_CALL frame: dummy environment */ \ + sp[2] = Val_unit; /* RETURN frame: dummy local 0 */ \ + sp[3] = (value) pc; /* RETURN frame: saved return address */ \ + sp[4] = env; /* RETURN frame: saved environment */ \ + sp[5] = Val_long(extra_args); /* RETURN frame: saved extra args */ \ + caml_extern_sp = sp; } +#define Restore_after_event \ + { sp = caml_extern_sp; accu = sp[0]; \ + pc = (code_t) sp[3]; env = sp[4]; extra_args = Long_val(sp[5]); \ + sp += 6; } + +/* Debugger interface */ + +#define Setup_for_debugger \ + { sp -= 4; \ + sp[0] = accu; sp[1] = (value)(pc - 1); \ + sp[2] = env; sp[3] = Val_long(extra_args); \ + caml_extern_sp = sp; } +#define Restore_after_debugger { sp += 4; } + +#ifdef THREADED_CODE +#define Restart_curr_instr \ + goto *(jumptable[caml_saved_code[pc - 1 - caml_start_code]]) +#else +#define Restart_curr_instr \ + curr_instr = caml_saved_code[pc - 1 - caml_start_code]; \ + goto dispatch_instr +#endif + +/* Register optimization. + Some compilers underestimate the use of the local variables representing + the abstract machine registers, and don't put them in hardware registers, + which slows down the interpreter considerably. + For GCC, I have hand-assigned hardware registers for several architectures. +*/ + +#if defined(__GNUC__) && !defined(DEBUG) && !defined(__INTEL_COMPILER) \ + && !defined(__llvm__) +#ifdef __mips__ +#define PC_REG asm("$16") +#define SP_REG asm("$17") +#define ACCU_REG asm("$18") +#endif +#ifdef __sparc__ +#define PC_REG asm("%l0") +#define SP_REG asm("%l1") +#define ACCU_REG asm("%l2") +#endif +#ifdef __alpha__ +#ifdef __CRAY__ +#define PC_REG asm("r9") +#define SP_REG asm("r10") +#define ACCU_REG asm("r11") +#define JUMPTBL_BASE_REG asm("r12") +#else +#define PC_REG asm("$9") +#define SP_REG asm("$10") +#define ACCU_REG asm("$11") +#define JUMPTBL_BASE_REG asm("$12") +#endif +#endif +#ifdef __i386__ +#define PC_REG asm("%esi") +#define SP_REG asm("%edi") +#define ACCU_REG +#endif +#if defined(__ppc__) || defined(__ppc64__) +#define PC_REG asm("26") +#define SP_REG asm("27") +#define ACCU_REG asm("28") +#endif +#ifdef __hppa__ +#define PC_REG asm("%r18") +#define SP_REG asm("%r17") +#define ACCU_REG asm("%r16") +#endif +#ifdef __mc68000__ +#define PC_REG asm("a5") +#define SP_REG asm("a4") +#define ACCU_REG asm("d7") +#endif +/* PR#4953: these specific registers not available in Thumb mode */ +#if defined (__arm__) && !defined(__thumb__) +#define PC_REG asm("r6") +#define SP_REG asm("r8") +#define ACCU_REG asm("r7") +#endif +#ifdef __ia64__ +#define PC_REG asm("36") +#define SP_REG asm("37") +#define ACCU_REG asm("38") +#define JUMPTBL_BASE_REG asm("39") +#endif +#ifdef __x86_64__ +#define PC_REG asm("%r15") +#define SP_REG asm("%r14") +#define ACCU_REG asm("%r13") +#endif +#ifdef __aarch64__ +#define PC_REG asm("%x19") +#define SP_REG asm("%x20") +#define ACCU_REG asm("%x21") +#define JUMPTBL_BASE_REG asm("%x22") +#endif +#endif + +#ifdef DEBUG +static intnat caml_bcodcount; +#endif + +/* The interpreter itself */ + +value caml_interprete(code_t prog, asize_t prog_size) +{ +#ifdef PC_REG + register code_t pc PC_REG; + register value * sp SP_REG; + register value accu ACCU_REG; +#else + register code_t pc; + register value * sp; + register value accu; +#endif +#if defined(THREADED_CODE) && defined(ARCH_SIXTYFOUR) && !defined(ARCH_CODE32) +#ifdef JUMPTBL_BASE_REG + register char * jumptbl_base JUMPTBL_BASE_REG; +#else + register char * jumptbl_base; +#endif +#endif + value env; + intnat extra_args; + struct longjmp_buffer * initial_external_raise; + int initial_sp_offset; + /* volatile ensures that initial_local_roots and saved_pc + will keep correct value across longjmp */ + struct caml__roots_block * volatile initial_local_roots; + volatile code_t saved_pc = NULL; + struct longjmp_buffer raise_buf; +#ifndef THREADED_CODE + opcode_t curr_instr; +#endif + +#ifdef THREADED_CODE + static void * jumptable[] = { +# include "caml/jumptbl.h" + }; +#endif + + if (prog == NULL) { /* Interpreter is initializing */ +#ifdef THREADED_CODE + caml_instr_table = (char **) jumptable; + caml_instr_base = Jumptbl_base; +#endif + return Val_unit; + } + +#if defined(THREADED_CODE) && defined(ARCH_SIXTYFOUR) && !defined(ARCH_CODE32) + jumptbl_base = Jumptbl_base; +#endif + initial_local_roots = caml_local_roots; + initial_sp_offset = (char *) caml_stack_high - (char *) caml_extern_sp; + initial_external_raise = caml_external_raise; + caml_callback_depth++; + saved_pc = NULL; + + if (sigsetjmp(raise_buf.buf, 0)) { + caml_local_roots = initial_local_roots; + sp = caml_extern_sp; + accu = caml_exn_bucket; + pc = saved_pc; saved_pc = NULL; + if (pc != NULL) pc += 2; + /* +2 adjustement for the sole purpose of backtraces */ + goto raise_exception; + } + caml_external_raise = &raise_buf; + + sp = caml_extern_sp; + pc = prog; + extra_args = 0; + env = Atom(0); + accu = Val_int(0); + +#ifdef THREADED_CODE +#ifdef DEBUG + next_instr: + if (caml_icount-- == 0) caml_stop_here (); + CAMLassert(sp >= caml_stack_low); + CAMLassert(sp <= caml_stack_high); +#endif + goto *(void *)(jumptbl_base + *pc++); /* Jump to the first instruction */ +#else + while(1) { +#ifdef DEBUG + caml_bcodcount++; + if (caml_icount-- == 0) caml_stop_here (); + if (caml_trace_level>1) printf("\n##%" ARCH_INTNAT_PRINTF_FORMAT "d\n", + caml_bcodcount); + if (caml_trace_level>0) caml_disasm_instr(pc); + if (caml_trace_level>1) { + printf("env="); + caml_trace_value_file(env,prog,prog_size,stdout); + putchar('\n'); + caml_trace_accu_sp_file(accu,sp,prog,prog_size,stdout); + fflush(stdout); + }; + CAMLassert(sp >= caml_stack_low); + CAMLassert(sp <= caml_stack_high); +#endif + curr_instr = *pc++; + + dispatch_instr: + switch(curr_instr) { +#endif + +/* Basic stack operations */ + + Instruct(ACC0): + accu = sp[0]; Next; + Instruct(ACC1): + accu = sp[1]; Next; + Instruct(ACC2): + accu = sp[2]; Next; + Instruct(ACC3): + accu = sp[3]; Next; + Instruct(ACC4): + accu = sp[4]; Next; + Instruct(ACC5): + accu = sp[5]; Next; + Instruct(ACC6): + accu = sp[6]; Next; + Instruct(ACC7): + accu = sp[7]; Next; + + Instruct(PUSH): Instruct(PUSHACC0): + *--sp = accu; Next; + Instruct(PUSHACC1): + *--sp = accu; accu = sp[1]; Next; + Instruct(PUSHACC2): + *--sp = accu; accu = sp[2]; Next; + Instruct(PUSHACC3): + *--sp = accu; accu = sp[3]; Next; + Instruct(PUSHACC4): + *--sp = accu; accu = sp[4]; Next; + Instruct(PUSHACC5): + *--sp = accu; accu = sp[5]; Next; + Instruct(PUSHACC6): + *--sp = accu; accu = sp[6]; Next; + Instruct(PUSHACC7): + *--sp = accu; accu = sp[7]; Next; + + Instruct(PUSHACC): + *--sp = accu; + /* Fallthrough */ + Instruct(ACC): + accu = sp[*pc++]; + Next; + + Instruct(POP): + sp += *pc++; + Next; + Instruct(ASSIGN): + sp[*pc++] = accu; + accu = Val_unit; + Next; + +/* Access in heap-allocated environment */ + + Instruct(ENVACC1): + accu = Field(env, 1); Next; + Instruct(ENVACC2): + accu = Field(env, 2); Next; + Instruct(ENVACC3): + accu = Field(env, 3); Next; + Instruct(ENVACC4): + accu = Field(env, 4); Next; + + Instruct(PUSHENVACC1): + *--sp = accu; accu = Field(env, 1); Next; + Instruct(PUSHENVACC2): + *--sp = accu; accu = Field(env, 2); Next; + Instruct(PUSHENVACC3): + *--sp = accu; accu = Field(env, 3); Next; + Instruct(PUSHENVACC4): + *--sp = accu; accu = Field(env, 4); Next; + + Instruct(PUSHENVACC): + *--sp = accu; + /* Fallthrough */ + Instruct(ENVACC): + accu = Field(env, *pc++); + Next; + +/* Function application */ + + Instruct(PUSH_RETADDR): { + sp -= 3; + sp[0] = (value) (pc + *pc); + sp[1] = env; + sp[2] = Val_long(extra_args); + pc++; + Next; + } + Instruct(APPLY): { + extra_args = *pc - 1; + pc = Code_val(accu); + env = accu; + goto check_stacks; + } + Instruct(APPLY1): { + value arg1 = sp[0]; + sp -= 3; + sp[0] = arg1; + sp[1] = (value)pc; + sp[2] = env; + sp[3] = Val_long(extra_args); + pc = Code_val(accu); + env = accu; + extra_args = 0; + goto check_stacks; + } + Instruct(APPLY2): { + value arg1 = sp[0]; + value arg2 = sp[1]; + sp -= 3; + sp[0] = arg1; + sp[1] = arg2; + sp[2] = (value)pc; + sp[3] = env; + sp[4] = Val_long(extra_args); + pc = Code_val(accu); + env = accu; + extra_args = 1; + goto check_stacks; + } + Instruct(APPLY3): { + value arg1 = sp[0]; + value arg2 = sp[1]; + value arg3 = sp[2]; + sp -= 3; + sp[0] = arg1; + sp[1] = arg2; + sp[2] = arg3; + sp[3] = (value)pc; + sp[4] = env; + sp[5] = Val_long(extra_args); + pc = Code_val(accu); + env = accu; + extra_args = 2; + goto check_stacks; + } + + Instruct(APPTERM): { + int nargs = *pc++; + int slotsize = *pc; + value * newsp; + int i; + /* Slide the nargs bottom words of the current frame to the top + of the frame, and discard the remainder of the frame */ + newsp = sp + slotsize - nargs; + for (i = nargs - 1; i >= 0; i--) newsp[i] = sp[i]; + sp = newsp; + pc = Code_val(accu); + env = accu; + extra_args += nargs - 1; + goto check_stacks; + } + Instruct(APPTERM1): { + value arg1 = sp[0]; + sp = sp + *pc - 1; + sp[0] = arg1; + pc = Code_val(accu); + env = accu; + goto check_stacks; + } + Instruct(APPTERM2): { + value arg1 = sp[0]; + value arg2 = sp[1]; + sp = sp + *pc - 2; + sp[0] = arg1; + sp[1] = arg2; + pc = Code_val(accu); + env = accu; + extra_args += 1; + goto check_stacks; + } + Instruct(APPTERM3): { + value arg1 = sp[0]; + value arg2 = sp[1]; + value arg3 = sp[2]; + sp = sp + *pc - 3; + sp[0] = arg1; + sp[1] = arg2; + sp[2] = arg3; + pc = Code_val(accu); + env = accu; + extra_args += 2; + goto check_stacks; + } + + Instruct(RETURN): { + sp += *pc++; + if (extra_args > 0) { + extra_args--; + pc = Code_val(accu); + env = accu; + } else { + pc = (code_t)(sp[0]); + env = sp[1]; + extra_args = Long_val(sp[2]); + sp += 3; + } + Next; + } + + Instruct(RESTART): { + int num_args = Wosize_val(env) - 2; + int i; + sp -= num_args; + for (i = 0; i < num_args; i++) sp[i] = Field(env, i + 2); + env = Field(env, 1); + extra_args += num_args; + Next; + } + + Instruct(GRAB): { + int required = *pc++; + if (extra_args >= required) { + extra_args -= required; + } else { + mlsize_t num_args, i; + num_args = 1 + extra_args; /* arg1 + extra args */ + Alloc_small(accu, num_args + 2, Closure_tag); + Field(accu, 1) = env; + for (i = 0; i < num_args; i++) Field(accu, i + 2) = sp[i]; + Code_val(accu) = pc - 3; /* Point to the preceding RESTART instr. */ + sp += num_args; + pc = (code_t)(sp[0]); + env = sp[1]; + extra_args = Long_val(sp[2]); + sp += 3; + } + Next; + } + + Instruct(CLOSURE): { + int nvars = *pc++; + int i; + if (nvars > 0) *--sp = accu; + if (nvars < Max_young_wosize) { + /* nvars + 1 <= Max_young_wosize, can allocate in minor heap */ + Alloc_small(accu, 1 + nvars, Closure_tag); + for (i = 0; i < nvars; i++) Field(accu, i + 1) = sp[i]; + } else { + /* PR#6385: must allocate in major heap */ + /* caml_alloc_shr and caml_initialize never trigger a GC, + so no need to Setup_for_gc */ + accu = caml_alloc_shr(1 + nvars, Closure_tag); + for (i = 0; i < nvars; i++) caml_initialize(&Field(accu, i + 1), sp[i]); + } + /* The code pointer is not in the heap, so no need to go through + caml_initialize. */ + Code_val(accu) = pc + *pc; + pc++; + sp += nvars; + Next; + } + + Instruct(CLOSUREREC): { + int nfuncs = *pc++; + int nvars = *pc++; + mlsize_t blksize = nfuncs * 2 - 1 + nvars; + int i; + value * p; + if (nvars > 0) *--sp = accu; + if (blksize <= Max_young_wosize) { + Alloc_small(accu, blksize, Closure_tag); + p = &Field(accu, nfuncs * 2 - 1); + for (i = 0; i < nvars; i++, p++) *p = sp[i]; + } else { + /* PR#6385: must allocate in major heap */ + /* caml_alloc_shr and caml_initialize never trigger a GC, + so no need to Setup_for_gc */ + accu = caml_alloc_shr(blksize, Closure_tag); + p = &Field(accu, nfuncs * 2 - 1); + for (i = 0; i < nvars; i++, p++) caml_initialize(p, sp[i]); + } + sp += nvars; + /* The code pointers and infix headers are not in the heap, + so no need to go through caml_initialize. */ + p = &Field(accu, 0); + *p = (value) (pc + pc[0]); + *--sp = accu; + p++; + for (i = 1; i < nfuncs; i++) { + *p = Make_header(i * 2, Infix_tag, Caml_white); /* color irrelevant. */ + p++; + *p = (value) (pc + pc[i]); + *--sp = (value) p; + p++; + } + pc += nfuncs; + Next; + } + + Instruct(PUSHOFFSETCLOSURE): + *--sp = accu; /* fallthrough */ + Instruct(OFFSETCLOSURE): + accu = env + *pc++ * sizeof(value); Next; + + Instruct(PUSHOFFSETCLOSUREM2): + *--sp = accu; /* fallthrough */ + Instruct(OFFSETCLOSUREM2): + accu = env - 2 * sizeof(value); Next; + Instruct(PUSHOFFSETCLOSURE0): + *--sp = accu; /* fallthrough */ + Instruct(OFFSETCLOSURE0): + accu = env; Next; + Instruct(PUSHOFFSETCLOSURE2): + *--sp = accu; /* fallthrough */ + Instruct(OFFSETCLOSURE2): + accu = env + 2 * sizeof(value); Next; + + +/* Access to global variables */ + + Instruct(PUSHGETGLOBAL): + *--sp = accu; + /* Fallthrough */ + Instruct(GETGLOBAL): + accu = Field(caml_global_data, *pc); + pc++; + Next; + + Instruct(PUSHGETGLOBALFIELD): + *--sp = accu; + /* Fallthrough */ + Instruct(GETGLOBALFIELD): { + accu = Field(caml_global_data, *pc); + pc++; + accu = Field(accu, *pc); + pc++; + Next; + } + + Instruct(SETGLOBAL): + caml_modify(&Field(caml_global_data, *pc), accu); + accu = Val_unit; + pc++; + Next; + +/* Allocation of blocks */ + + Instruct(PUSHATOM0): + *--sp = accu; + /* Fallthrough */ + Instruct(ATOM0): + accu = Atom(0); Next; + + Instruct(PUSHATOM): + *--sp = accu; + /* Fallthrough */ + Instruct(ATOM): + accu = Atom(*pc++); Next; + + Instruct(MAKEBLOCK): { + mlsize_t wosize = *pc++; + tag_t tag = *pc++; + mlsize_t i; + value block; + if (wosize <= Max_young_wosize) { + Alloc_small(block, wosize, tag); + Field(block, 0) = accu; + for (i = 1; i < wosize; i++) Field(block, i) = *sp++; + } else { + block = caml_alloc_shr(wosize, tag); + caml_initialize(&Field(block, 0), accu); + for (i = 1; i < wosize; i++) caml_initialize(&Field(block, i), *sp++); + } + accu = block; + Next; + } + Instruct(MAKEBLOCK1): { + tag_t tag = *pc++; + value block; + Alloc_small(block, 1, tag); + Field(block, 0) = accu; + accu = block; + Next; + } + Instruct(MAKEBLOCK2): { + tag_t tag = *pc++; + value block; + Alloc_small(block, 2, tag); + Field(block, 0) = accu; + Field(block, 1) = sp[0]; + sp += 1; + accu = block; + Next; + } + Instruct(MAKEBLOCK3): { + tag_t tag = *pc++; + value block; + Alloc_small(block, 3, tag); + Field(block, 0) = accu; + Field(block, 1) = sp[0]; + Field(block, 2) = sp[1]; + sp += 2; + accu = block; + Next; + } + Instruct(MAKEFLOATBLOCK): { + mlsize_t size = *pc++; + mlsize_t i; + value block; + if (size <= Max_young_wosize / Double_wosize) { + Alloc_small(block, size * Double_wosize, Double_array_tag); + } else { + block = caml_alloc_shr(size * Double_wosize, Double_array_tag); + } + Store_double_flat_field(block, 0, Double_val(accu)); + for (i = 1; i < size; i++){ + Store_double_flat_field(block, i, Double_val(*sp)); + ++ sp; + } + accu = block; + Next; + } + +/* Access to components of blocks */ + + Instruct(GETFIELD0): + accu = Field(accu, 0); Next; + Instruct(GETFIELD1): + accu = Field(accu, 1); Next; + Instruct(GETFIELD2): + accu = Field(accu, 2); Next; + Instruct(GETFIELD3): + accu = Field(accu, 3); Next; + Instruct(GETFIELD): + accu = Field(accu, *pc); pc++; Next; + Instruct(GETFLOATFIELD): { + double d = Double_flat_field(accu, *pc); + Alloc_small(accu, Double_wosize, Double_tag); + Store_double_val(accu, d); + pc++; + Next; + } + + Instruct(SETFIELD0): + caml_modify(&Field(accu, 0), *sp++); + accu = Val_unit; + Next; + Instruct(SETFIELD1): + caml_modify(&Field(accu, 1), *sp++); + accu = Val_unit; + Next; + Instruct(SETFIELD2): + caml_modify(&Field(accu, 2), *sp++); + accu = Val_unit; + Next; + Instruct(SETFIELD3): + caml_modify(&Field(accu, 3), *sp++); + accu = Val_unit; + Next; + Instruct(SETFIELD): + caml_modify(&Field(accu, *pc), *sp++); + accu = Val_unit; + pc++; + Next; + Instruct(SETFLOATFIELD): + Store_double_flat_field(accu, *pc, Double_val(*sp)); + accu = Val_unit; + sp++; + pc++; + Next; + +/* Array operations */ + + Instruct(VECTLENGTH): { + /* Todo: when FLAT_FLOAT_ARRAY is false, this instruction should + be split into VECTLENGTH and FLOATVECTLENGTH because we know + statically which one it is. */ + mlsize_t size = Wosize_val(accu); + if (Tag_val(accu) == Double_array_tag) size = size / Double_wosize; + accu = Val_long(size); + Next; + } + Instruct(GETVECTITEM): + accu = Field(accu, Long_val(sp[0])); + sp += 1; + Next; + Instruct(SETVECTITEM): + caml_modify(&Field(accu, Long_val(sp[0])), sp[1]); + accu = Val_unit; + sp += 2; + Next; + +/* Bytes/String operations */ + Instruct(GETSTRINGCHAR): + Instruct(GETBYTESCHAR): + accu = Val_int(Byte_u(accu, Long_val(sp[0]))); + sp += 1; + Next; + Instruct(SETBYTESCHAR): + Byte_u(accu, Long_val(sp[0])) = Int_val(sp[1]); + sp += 2; + accu = Val_unit; + Next; + +/* Branches and conditional branches */ + + Instruct(BRANCH): + pc += *pc; + Next; + Instruct(BRANCHIF): + if (accu != Val_false) pc += *pc; else pc++; + Next; + Instruct(BRANCHIFNOT): + if (accu == Val_false) pc += *pc; else pc++; + Next; + Instruct(SWITCH): { + uint32_t sizes = *pc++; + if (Is_block(accu)) { + intnat index = Tag_val(accu); + CAMLassert ((uintnat) index < (sizes >> 16)); + pc += pc[(sizes & 0xFFFF) + index]; + } else { + intnat index = Long_val(accu); + CAMLassert ((uintnat) index < (sizes & 0xFFFF)) ; + pc += pc[index]; + } + Next; + } + Instruct(BOOLNOT): + accu = Val_not(accu); + Next; + +/* Exceptions */ + + Instruct(PUSHTRAP): + sp -= 4; + Trap_pc(sp) = pc + *pc; + Trap_link(sp) = caml_trapsp; + sp[2] = env; + sp[3] = Val_long(extra_args); + caml_trapsp = sp; + pc++; + Next; + + Instruct(POPTRAP): + if (caml_something_to_do) { + /* We must check here so that if a signal is pending and its + handler triggers an exception, the exception is trapped + by the current try...with, not the enclosing one. */ + pc--; /* restart the POPTRAP after processing the signal */ + goto process_signal; + } + caml_trapsp = Trap_link(sp); + sp += 4; + Next; + + Instruct(RAISE_NOTRACE): + if (caml_trapsp >= caml_trap_barrier) caml_debugger(TRAP_BARRIER); + goto raise_notrace; + + Instruct(RERAISE): + if (caml_trapsp >= caml_trap_barrier) caml_debugger(TRAP_BARRIER); + if (caml_backtrace_active) caml_stash_backtrace(accu, pc, sp, 1); + goto raise_notrace; + + Instruct(RAISE): + raise_exception: + if (caml_trapsp >= caml_trap_barrier) caml_debugger(TRAP_BARRIER); + if (caml_backtrace_active) caml_stash_backtrace(accu, pc, sp, 0); + raise_notrace: + if ((char *) caml_trapsp + >= (char *) caml_stack_high - initial_sp_offset) { + caml_external_raise = initial_external_raise; + caml_extern_sp = (value *) ((char *) caml_stack_high + - initial_sp_offset); + caml_callback_depth--; + return Make_exception_result(accu); + } + sp = caml_trapsp; + pc = Trap_pc(sp); + caml_trapsp = Trap_link(sp); + env = sp[2]; + extra_args = Long_val(sp[3]); + sp += 4; + Next; + +/* Stack checks */ + + check_stacks: + if (sp < caml_stack_threshold) { + caml_extern_sp = sp; + caml_realloc_stack(Stack_threshold / sizeof(value)); + sp = caml_extern_sp; + } + /* Fall through CHECK_SIGNALS */ + +/* Signal handling */ + + Instruct(CHECK_SIGNALS): /* accu not preserved */ + if (caml_something_to_do) goto process_signal; + Next; + + process_signal: + caml_something_to_do = 0; + Setup_for_event; + caml_process_event(); + Restore_after_event; + Next; + +/* Calling C functions */ + + Instruct(C_CALL1): + Setup_for_c_call; + accu = Primitive(*pc)(accu); + Restore_after_c_call; + pc++; + Next; + Instruct(C_CALL2): + Setup_for_c_call; + accu = Primitive(*pc)(accu, sp[1]); + Restore_after_c_call; + sp += 1; + pc++; + Next; + Instruct(C_CALL3): + Setup_for_c_call; + accu = Primitive(*pc)(accu, sp[1], sp[2]); + Restore_after_c_call; + sp += 2; + pc++; + Next; + Instruct(C_CALL4): + Setup_for_c_call; + accu = Primitive(*pc)(accu, sp[1], sp[2], sp[3]); + Restore_after_c_call; + sp += 3; + pc++; + Next; + Instruct(C_CALL5): + Setup_for_c_call; + accu = Primitive(*pc)(accu, sp[1], sp[2], sp[3], sp[4]); + Restore_after_c_call; + sp += 4; + pc++; + Next; + Instruct(C_CALLN): { + int nargs = *pc++; + *--sp = accu; + Setup_for_c_call; + accu = Primitive(*pc)(sp + 1, nargs); + Restore_after_c_call; + sp += nargs; + pc++; + Next; + } + +/* Integer constants */ + + Instruct(CONST0): + accu = Val_int(0); Next; + Instruct(CONST1): + accu = Val_int(1); Next; + Instruct(CONST2): + accu = Val_int(2); Next; + Instruct(CONST3): + accu = Val_int(3); Next; + + Instruct(PUSHCONST0): + *--sp = accu; accu = Val_int(0); Next; + Instruct(PUSHCONST1): + *--sp = accu; accu = Val_int(1); Next; + Instruct(PUSHCONST2): + *--sp = accu; accu = Val_int(2); Next; + Instruct(PUSHCONST3): + *--sp = accu; accu = Val_int(3); Next; + + Instruct(PUSHCONSTINT): + *--sp = accu; + /* Fallthrough */ + Instruct(CONSTINT): + accu = Val_int(*pc); + pc++; + Next; + +/* Integer arithmetic */ + + Instruct(NEGINT): + accu = (value)(2 - (intnat)accu); Next; + Instruct(ADDINT): + accu = (value)((intnat) accu + (intnat) *sp++ - 1); Next; + Instruct(SUBINT): + accu = (value)((intnat) accu - (intnat) *sp++ + 1); Next; + Instruct(MULINT): + accu = Val_long(Long_val(accu) * Long_val(*sp++)); Next; + + Instruct(DIVINT): { + intnat divisor = Long_val(*sp++); + if (divisor == 0) { Setup_for_c_call; caml_raise_zero_divide(); } + accu = Val_long(Long_val(accu) / divisor); + Next; + } + Instruct(MODINT): { + intnat divisor = Long_val(*sp++); + if (divisor == 0) { Setup_for_c_call; caml_raise_zero_divide(); } + accu = Val_long(Long_val(accu) % divisor); + Next; + } + Instruct(ANDINT): + accu = (value)((intnat) accu & (intnat) *sp++); Next; + Instruct(ORINT): + accu = (value)((intnat) accu | (intnat) *sp++); Next; + Instruct(XORINT): + accu = (value)(((intnat) accu ^ (intnat) *sp++) | 1); Next; + Instruct(LSLINT): + accu = (value)((((intnat) accu - 1) << Long_val(*sp++)) + 1); Next; + Instruct(LSRINT): + accu = (value)((((uintnat) accu) >> Long_val(*sp++)) | 1); Next; + Instruct(ASRINT): + accu = (value)((((intnat) accu) >> Long_val(*sp++)) | 1); Next; + +#define Integer_comparison(typ,opname,tst) \ + Instruct(opname): \ + accu = Val_int((typ) accu tst (typ) *sp++); Next; + + Integer_comparison(intnat,EQ, ==) + Integer_comparison(intnat,NEQ, !=) + Integer_comparison(intnat,LTINT, <) + Integer_comparison(intnat,LEINT, <=) + Integer_comparison(intnat,GTINT, >) + Integer_comparison(intnat,GEINT, >=) + Integer_comparison(uintnat,ULTINT, <) + Integer_comparison(uintnat,UGEINT, >=) + +#define Integer_branch_comparison(typ,opname,tst,debug) \ + Instruct(opname): \ + if ( *pc++ tst (typ) Long_val(accu)) { \ + pc += *pc ; \ + } else { \ + pc++ ; \ + } ; Next; + + Integer_branch_comparison(intnat,BEQ, ==, "==") + Integer_branch_comparison(intnat,BNEQ, !=, "!=") + Integer_branch_comparison(intnat,BLTINT, <, "<") + Integer_branch_comparison(intnat,BLEINT, <=, "<=") + Integer_branch_comparison(intnat,BGTINT, >, ">") + Integer_branch_comparison(intnat,BGEINT, >=, ">=") + Integer_branch_comparison(uintnat,BULTINT, <, "<") + Integer_branch_comparison(uintnat,BUGEINT, >=, ">=") + + Instruct(OFFSETINT): + accu += *pc << 1; + pc++; + Next; + Instruct(OFFSETREF): + Field(accu, 0) += *pc << 1; + accu = Val_unit; + pc++; + Next; + Instruct(ISINT): + accu = Val_long(accu & 1); + Next; + +/* Object-oriented operations */ + +#define Lookup(obj, lab) Field (Field (obj, 0), Int_val(lab)) + + /* please don't forget to keep below code in sync with the + functions caml_cache_public_method and + caml_cache_public_method2 in obj.c */ + + Instruct(GETMETHOD): + accu = Lookup(sp[0], accu); + Next; + +#define CAML_METHOD_CACHE +#ifdef CAML_METHOD_CACHE + Instruct(GETPUBMET): { + /* accu == object, pc[0] == tag, pc[1] == cache */ + value meths = Field (accu, 0); + value ofs; +#ifdef CAML_TEST_CACHE + static int calls = 0, hits = 0; + if (calls >= 10000000) { + fprintf(stderr, "cache hit = %d%%\n", hits / 100000); + calls = 0; hits = 0; + } + calls++; +#endif + *--sp = accu; + accu = Val_int(*pc++); + ofs = *pc & Field(meths,1); + if (*(value*)(((char*)&Field(meths,3)) + ofs) == accu) { +#ifdef CAML_TEST_CACHE + hits++; +#endif + accu = *(value*)(((char*)&Field(meths,2)) + ofs); + } + else + { + int li = 3, hi = Field(meths,0), mi; + while (li < hi) { + mi = ((li+hi) >> 1) | 1; + if (accu < Field(meths,mi)) hi = mi-2; + else li = mi; + } + *pc = (li-3)*sizeof(value); + accu = Field (meths, li-1); + } + pc++; + Next; + } +#else + Instruct(GETPUBMET): + *--sp = accu; + accu = Val_int(*pc); + pc += 2; + /* Fallthrough */ +#endif + Instruct(GETDYNMET): { + /* accu == tag, sp[0] == object, *pc == cache */ + value meths = Field (sp[0], 0); + int li = 3, hi = Field(meths,0), mi; + while (li < hi) { + mi = ((li+hi) >> 1) | 1; + if (accu < Field(meths,mi)) hi = mi-2; + else li = mi; + } + accu = Field (meths, li-1); + Next; + } + +/* Debugging and machine control */ + + Instruct(STOP): + caml_external_raise = initial_external_raise; + caml_extern_sp = sp; + caml_callback_depth--; + return accu; + + Instruct(EVENT): + if (--caml_event_count == 0) { + Setup_for_debugger; + caml_debugger(EVENT_COUNT); + Restore_after_debugger; + } + Restart_curr_instr; + + Instruct(BREAK): + Setup_for_debugger; + caml_debugger(BREAKPOINT); + Restore_after_debugger; + Restart_curr_instr; + +#ifndef THREADED_CODE + default: +#if _MSC_VER >= 1200 + __assume(0); +#else + caml_fatal_error_arg("Fatal error: bad opcode (%" + ARCH_INTNAT_PRINTF_FORMAT "x)\n", + (char *) (intnat) *(pc-1)); +#endif + } + } +#endif +} + +void caml_prepare_bytecode(code_t prog, asize_t prog_size) { + /* other implementations of the interpreter (such as an hypothetical + JIT translator) might want to do something with a bytecode before + running it */ + CAMLassert(prog); + CAMLassert(prog_size>0); + /* actually, the threading of the bytecode might be done here */ +} + +void caml_release_bytecode(code_t prog, asize_t prog_size) { + /* other implementations of the interpreter (such as an hypothetical + JIT translator) might want to know when a bytecode is removed */ + /* check that we have a program */ + CAMLassert(prog); + CAMLassert(prog_size>0); +} diff --git a/test/monniaux/ocaml/byterun/ints.c b/test/monniaux/ocaml/byterun/ints.c new file mode 100644 index 00000000..76ae11d4 --- /dev/null +++ b/test/monniaux/ocaml/byterun/ints.c @@ -0,0 +1,830 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include +#include +#include "caml/alloc.h" +#include "caml/custom.h" +#include "caml/fail.h" +#include "caml/intext.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" + +static const char * parse_sign_and_base(const char * p, + /*out*/ int * base, + /*out*/ int * signedness, + /*out*/ int * sign) +{ + *sign = 1; + if (*p == '-') { + *sign = -1; + p++; + } else if (*p == '+') + p++; + *base = 10; *signedness = 1; + if (*p == '0') { + switch (p[1]) { + case 'x': case 'X': + *base = 16; *signedness = 0; p += 2; break; + case 'o': case 'O': + *base = 8; *signedness = 0; p += 2; break; + case 'b': case 'B': + *base = 2; *signedness = 0; p += 2; break; + case 'u': case 'U': + *signedness = 0; p += 2; break; + } + } + return p; +} + +static int parse_digit(char c) +{ + if (c >= '0' && c <= '9') + return c - '0'; + else if (c >= 'A' && c <= 'F') + return c - 'A' + 10; + else if (c >= 'a' && c <= 'f') + return c - 'a' + 10; + else + return -1; +} + +#define INT_ERRMSG "int_of_string" +#define INT32_ERRMSG "Int32.of_string" +#define INT64_ERRMSG "Int64.of_string" +#define INTNAT_ERRMSG "Nativeint.of_string" + +static intnat parse_intnat(value s, int nbits, const char *errmsg) +{ + const char * p; + uintnat res, threshold; + int sign, base, signedness, d; + + p = parse_sign_and_base(String_val(s), &base, &signedness, &sign); + threshold = ((uintnat) -1) / base; + d = parse_digit(*p); + if (d < 0 || d >= base) caml_failwith(errmsg); + for (p++, res = d; /*nothing*/; p++) { + char c = *p; + if (c == '_') continue; + d = parse_digit(c); + if (d < 0 || d >= base) break; + /* Detect overflow in multiplication base * res */ + if (res > threshold) caml_failwith(errmsg); + res = base * res + d; + /* Detect overflow in addition (base * res) + d */ + if (res < (uintnat) d) caml_failwith(errmsg); + } + if (p != String_val(s) + caml_string_length(s)){ + caml_failwith(errmsg); + } + if (signedness) { + /* Signed representation expected, allow -2^(nbits-1) to 2^(nbits-1) - 1 */ + if (sign >= 0) { + if (res >= (uintnat)1 << (nbits - 1)) caml_failwith(errmsg); + } else { + if (res > (uintnat)1 << (nbits - 1)) caml_failwith(errmsg); + } + } else { + /* Unsigned representation expected, allow 0 to 2^nbits - 1 + and tolerate -(2^nbits - 1) to 0 */ + if (nbits < sizeof(uintnat) * 8 && res >= (uintnat)1 << nbits) + caml_failwith(errmsg); + } + return sign < 0 ? -((intnat) res) : (intnat) res; +} + +value caml_bswap16_direct(value x) +{ + return ((((x & 0x00FF) << 8) | + ((x & 0xFF00) >> 8))); +} + +CAMLprim value caml_bswap16(value v) +{ + intnat x = Int_val(v); + return (Val_int ((((x & 0x00FF) << 8) | + ((x & 0xFF00) >> 8)))); +} + +/* Tagged integers */ + +CAMLprim value caml_int_compare(value v1, value v2) +{ + int res = (v1 > v2) - (v1 < v2); + return Val_int(res); +} + +CAMLprim value caml_int_of_string(value s) +{ + return Val_long(parse_intnat(s, 8 * sizeof(value) - 1, INT_ERRMSG)); +} + +#define FORMAT_BUFFER_SIZE 32 + +static char parse_format(value fmt, + char * suffix, + char format_string[FORMAT_BUFFER_SIZE]) +{ + char * p; + char lastletter; + mlsize_t len, len_suffix; + + /* Copy OCaml format fmt to format_string, + adding the suffix before the last letter of the format */ + len = caml_string_length(fmt); + len_suffix = strlen(suffix); + if (len + len_suffix + 1 >= FORMAT_BUFFER_SIZE) + caml_invalid_argument("format_int: format too long"); + memmove(format_string, String_val(fmt), len); + p = format_string + len - 1; + lastletter = *p; + /* Compress two-letter formats, ignoring the [lnL] annotation */ + if (p[-1] == 'l' || p[-1] == 'n' || p[-1] == 'L') p--; + memmove(p, suffix, len_suffix); p += len_suffix; + *p++ = lastletter; + *p = 0; + /* Return the conversion type (last letter) */ + return lastletter; +} + +CAMLprim value caml_format_int(value fmt, value arg) +{ + char format_string[FORMAT_BUFFER_SIZE]; + char conv; + value res; + + conv = parse_format(fmt, ARCH_INTNAT_PRINTF_FORMAT, format_string); + switch (conv) { + case 'u': case 'x': case 'X': case 'o': + res = caml_alloc_sprintf(format_string, Unsigned_long_val(arg)); + break; + default: + res = caml_alloc_sprintf(format_string, Long_val(arg)); + break; + } + return res; +} + +/* 32-bit integers */ + +static int int32_cmp(value v1, value v2) +{ + int32_t i1 = Int32_val(v1); + int32_t i2 = Int32_val(v2); + return (i1 > i2) - (i1 < i2); +} + +static intnat int32_hash(value v) +{ + return Int32_val(v); +} + +static void int32_serialize(value v, uintnat * bsize_32, + uintnat * bsize_64) +{ + caml_serialize_int_4(Int32_val(v)); + *bsize_32 = *bsize_64 = 4; +} + +static uintnat int32_deserialize(void * dst) +{ + *((int32_t *) dst) = caml_deserialize_sint_4(); + return 4; +} + +CAMLexport struct custom_operations caml_int32_ops = { + "_i", + custom_finalize_default, + int32_cmp, + int32_hash, + int32_serialize, + int32_deserialize, + custom_compare_ext_default +}; + +CAMLexport value caml_copy_int32(int32_t i) +{ + value res = caml_alloc_custom(&caml_int32_ops, 4, 0, 1); + Int32_val(res) = i; + return res; +} + +CAMLprim value caml_int32_neg(value v) +{ return caml_copy_int32(- Int32_val(v)); } + +CAMLprim value caml_int32_add(value v1, value v2) +{ return caml_copy_int32(Int32_val(v1) + Int32_val(v2)); } + +CAMLprim value caml_int32_sub(value v1, value v2) +{ return caml_copy_int32(Int32_val(v1) - Int32_val(v2)); } + +CAMLprim value caml_int32_mul(value v1, value v2) +{ return caml_copy_int32(Int32_val(v1) * Int32_val(v2)); } + +CAMLprim value caml_int32_div(value v1, value v2) +{ + int32_t dividend = Int32_val(v1); + int32_t divisor = Int32_val(v2); + if (divisor == 0) caml_raise_zero_divide(); + /* PR#4740: on some processors, division crashes on overflow. + Implement the same behavior as for type "int". */ + if (dividend == (1<<31) && divisor == -1) return v1; + return caml_copy_int32(dividend / divisor); +} + +CAMLprim value caml_int32_mod(value v1, value v2) +{ + int32_t dividend = Int32_val(v1); + int32_t divisor = Int32_val(v2); + if (divisor == 0) caml_raise_zero_divide(); + /* PR#4740: on some processors, modulus crashes if division overflows. + Implement the same behavior as for type "int". */ + if (dividend == (1<<31) && divisor == -1) return caml_copy_int32(0); + return caml_copy_int32(dividend % divisor); +} + +CAMLprim value caml_int32_and(value v1, value v2) +{ return caml_copy_int32(Int32_val(v1) & Int32_val(v2)); } + +CAMLprim value caml_int32_or(value v1, value v2) +{ return caml_copy_int32(Int32_val(v1) | Int32_val(v2)); } + +CAMLprim value caml_int32_xor(value v1, value v2) +{ return caml_copy_int32(Int32_val(v1) ^ Int32_val(v2)); } + +CAMLprim value caml_int32_shift_left(value v1, value v2) +{ return caml_copy_int32(Int32_val(v1) << Int_val(v2)); } + +CAMLprim value caml_int32_shift_right(value v1, value v2) +{ return caml_copy_int32(Int32_val(v1) >> Int_val(v2)); } + +CAMLprim value caml_int32_shift_right_unsigned(value v1, value v2) +{ return caml_copy_int32((uint32_t)Int32_val(v1) >> Int_val(v2)); } + +static int32_t caml_swap32(int32_t x) +{ + return (((x & 0x000000FF) << 24) | + ((x & 0x0000FF00) << 8) | + ((x & 0x00FF0000) >> 8) | + ((x & 0xFF000000) >> 24)); +} + +value caml_int32_direct_bswap(value v) +{ return caml_swap32(v); } + +CAMLprim value caml_int32_bswap(value v) +{ return caml_copy_int32(caml_swap32(Int32_val(v))); } + +CAMLprim value caml_int32_of_int(value v) +{ return caml_copy_int32(Long_val(v)); } + +CAMLprim value caml_int32_to_int(value v) +{ return Val_long(Int32_val(v)); } + +int32_t caml_int32_of_float_unboxed(double x) +{ return x; } + +CAMLprim value caml_int32_of_float(value v) +{ return caml_copy_int32((int32_t)(Double_val(v))); } + +double caml_int32_to_float_unboxed(int32_t x) +{ return x; } + +CAMLprim value caml_int32_to_float(value v) +{ return caml_copy_double((double)(Int32_val(v))); } + +intnat caml_int32_compare_unboxed(int32_t i1, int32_t i2) +{ + return (i1 > i2) - (i1 < i2); +} + +CAMLprim value caml_int32_compare(value v1, value v2) +{ + return Val_int(caml_int32_compare_unboxed(Int32_val(v1),Int32_val(v2))); +} + +CAMLprim value caml_int32_format(value fmt, value arg) +{ + char format_string[FORMAT_BUFFER_SIZE]; + + parse_format(fmt, ARCH_INT32_PRINTF_FORMAT, format_string); + return caml_alloc_sprintf(format_string, Int32_val(arg)); +} + +CAMLprim value caml_int32_of_string(value s) +{ + return caml_copy_int32(parse_intnat(s, 32, INT32_ERRMSG)); +} + +int32_t caml_int32_bits_of_float_unboxed(double d) +{ + union { float d; int32_t i; } u; + u.d = d; + return u.i; +} + +double caml_int32_float_of_bits_unboxed(int32_t i) +{ + union { float d; int32_t i; } u; + u.i = i; + return u.d; +} + +CAMLprim value caml_int32_bits_of_float(value vd) +{ + return caml_copy_int32(caml_int32_bits_of_float_unboxed(Double_val(vd))); +} + +CAMLprim value caml_int32_float_of_bits(value vi) +{ + return caml_copy_double(caml_int32_float_of_bits_unboxed(Int32_val(vi))); +} + +/* 64-bit integers */ + +#ifdef ARCH_ALIGN_INT64 + +CAMLexport int64_t caml_Int64_val(value v) +{ + union { int32_t i[2]; int64_t j; } buffer; + buffer.i[0] = ((int32_t *) Data_custom_val(v))[0]; + buffer.i[1] = ((int32_t *) Data_custom_val(v))[1]; + return buffer.j; +} + +#endif + +static int int64_cmp(value v1, value v2) +{ + int64_t i1 = Int64_val(v1); + int64_t i2 = Int64_val(v2); + return (i1 > i2) - (i1 < i2); +} + +static intnat int64_hash(value v) +{ + int64_t x = Int64_val(v); + uint32_t lo = (uint32_t) x, hi = (uint32_t) (x >> 32); + return hi ^ lo; +} + +static void int64_serialize(value v, uintnat * bsize_32, + uintnat * bsize_64) +{ + caml_serialize_int_8(Int64_val(v)); + *bsize_32 = *bsize_64 = 8; +} + +static uintnat int64_deserialize(void * dst) +{ +#ifndef ARCH_ALIGN_INT64 + *((int64_t *) dst) = caml_deserialize_sint_8(); +#else + union { int32_t i[2]; int64_t j; } buffer; + buffer.j = caml_deserialize_sint_8(); + ((int32_t *) dst)[0] = buffer.i[0]; + ((int32_t *) dst)[1] = buffer.i[1]; +#endif + return 8; +} + +CAMLexport struct custom_operations caml_int64_ops = { + "_j", + custom_finalize_default, + int64_cmp, + int64_hash, + int64_serialize, + int64_deserialize, + custom_compare_ext_default +}; + +CAMLexport value caml_copy_int64(int64_t i) +{ + value res = caml_alloc_custom(&caml_int64_ops, 8, 0, 1); +#ifndef ARCH_ALIGN_INT64 + Int64_val(res) = i; +#else + union { int32_t i[2]; int64_t j; } buffer; + buffer.j = i; + ((int32_t *) Data_custom_val(res))[0] = buffer.i[0]; + ((int32_t *) Data_custom_val(res))[1] = buffer.i[1]; +#endif + return res; +} + +CAMLprim value caml_int64_neg(value v) +{ return caml_copy_int64(- Int64_val(v)); } + +CAMLprim value caml_int64_add(value v1, value v2) +{ return caml_copy_int64(Int64_val(v1) + Int64_val(v2)); } + +CAMLprim value caml_int64_sub(value v1, value v2) +{ return caml_copy_int64(Int64_val(v1) - Int64_val(v2)); } + +CAMLprim value caml_int64_mul(value v1, value v2) +{ return caml_copy_int64(Int64_val(v1) * Int64_val(v2)); } + +#define Int64_min_int ((intnat) 1 << (sizeof(intnat) * 8 - 1)) + +CAMLprim value caml_int64_div(value v1, value v2) +{ + int64_t dividend = Int64_val(v1); + int64_t divisor = Int64_val(v2); + if (divisor == 0) caml_raise_zero_divide(); + /* PR#4740: on some processors, division crashes on overflow. + Implement the same behavior as for type "int". */ + if (dividend == ((int64_t)1 << 63) && divisor == -1) return v1; + return caml_copy_int64(Int64_val(v1) / divisor); +} + +CAMLprim value caml_int64_mod(value v1, value v2) +{ + int64_t dividend = Int64_val(v1); + int64_t divisor = Int64_val(v2); + if (divisor == 0) caml_raise_zero_divide(); + /* PR#4740: on some processors, division crashes on overflow. + Implement the same behavior as for type "int". */ + if (dividend == ((int64_t)1 << 63) && divisor == -1){ + return caml_copy_int64(0); + } + return caml_copy_int64(Int64_val(v1) % divisor); +} + +CAMLprim value caml_int64_and(value v1, value v2) +{ return caml_copy_int64(Int64_val(v1) & Int64_val(v2)); } + +CAMLprim value caml_int64_or(value v1, value v2) +{ return caml_copy_int64(Int64_val(v1) | Int64_val(v2)); } + +CAMLprim value caml_int64_xor(value v1, value v2) +{ return caml_copy_int64(Int64_val(v1) ^ Int64_val(v2)); } + +CAMLprim value caml_int64_shift_left(value v1, value v2) +{ return caml_copy_int64(Int64_val(v1) << Int_val(v2)); } + +CAMLprim value caml_int64_shift_right(value v1, value v2) +{ return caml_copy_int64(Int64_val(v1) >> Int_val(v2)); } + +CAMLprim value caml_int64_shift_right_unsigned(value v1, value v2) +{ return caml_copy_int64((uint64_t) (Int64_val(v1)) >> Int_val(v2)); } + +#ifdef ARCH_SIXTYFOUR +static value caml_swap64(value x) +{ + return (((((x) & 0x00000000000000FF) << 56) | + (((x) & 0x000000000000FF00) << 40) | + (((x) & 0x0000000000FF0000) << 24) | + (((x) & 0x00000000FF000000) << 8) | + (((x) & 0x000000FF00000000) >> 8) | + (((x) & 0x0000FF0000000000) >> 24) | + (((x) & 0x00FF000000000000) >> 40) | + (((x) & 0xFF00000000000000) >> 56))); +} + +value caml_int64_direct_bswap(value v) +{ return caml_swap64(v); } +#endif + +CAMLprim value caml_int64_bswap(value v) +{ + int64_t x = Int64_val(v); + return caml_copy_int64 + (((x & INT64_LITERAL(0x00000000000000FFU)) << 56) | + ((x & INT64_LITERAL(0x000000000000FF00U)) << 40) | + ((x & INT64_LITERAL(0x0000000000FF0000U)) << 24) | + ((x & INT64_LITERAL(0x00000000FF000000U)) << 8) | + ((x & INT64_LITERAL(0x000000FF00000000U)) >> 8) | + ((x & INT64_LITERAL(0x0000FF0000000000U)) >> 24) | + ((x & INT64_LITERAL(0x00FF000000000000U)) >> 40) | + ((x & INT64_LITERAL(0xFF00000000000000U)) >> 56)); +} + +CAMLprim value caml_int64_of_int(value v) +{ return caml_copy_int64((int64_t) (Long_val(v))); } + +CAMLprim value caml_int64_to_int(value v) +{ return Val_long((intnat) (Int64_val(v))); } + +int64_t caml_int64_of_float_unboxed(double x) +{ return x; } + +CAMLprim value caml_int64_of_float(value v) +{ return caml_copy_int64((int64_t) (Double_val(v))); } + +double caml_int64_to_float_unboxed(int64_t x) +{ return x; } + +CAMLprim value caml_int64_to_float(value v) +{ return caml_copy_double((double) (Int64_val(v))); } + +CAMLprim value caml_int64_of_int32(value v) +{ return caml_copy_int64((int64_t) (Int32_val(v))); } + +CAMLprim value caml_int64_to_int32(value v) +{ return caml_copy_int32((int32_t) (Int64_val(v))); } + +CAMLprim value caml_int64_of_nativeint(value v) +{ return caml_copy_int64((int64_t) (Nativeint_val(v))); } + +CAMLprim value caml_int64_to_nativeint(value v) +{ return caml_copy_nativeint((intnat) (Int64_val(v))); } + +intnat caml_int64_compare_unboxed(int64_t i1, int64_t i2) +{ + return (i1 > i2) - (i1 < i2); +} + +CAMLprim value caml_int64_compare(value v1, value v2) +{ + return Val_int(caml_int64_compare_unboxed(Int64_val(v1),Int64_val(v2))); +} + +CAMLprim value caml_int64_format(value fmt, value arg) +{ + char format_string[FORMAT_BUFFER_SIZE]; + + parse_format(fmt, ARCH_INT64_PRINTF_FORMAT, format_string); + return caml_alloc_sprintf(format_string, Int64_val(arg)); +} + +CAMLprim value caml_int64_of_string(value s) +{ + const char * p; + uint64_t res, threshold; + int sign, base, signedness, d; + + p = parse_sign_and_base(String_val(s), &base, &signedness, &sign); + threshold = ((uint64_t) -1) / base; + d = parse_digit(*p); + if (d < 0 || d >= base) caml_failwith(INT64_ERRMSG); + res = d; + for (p++; /*nothing*/; p++) { + char c = *p; + if (c == '_') continue; + d = parse_digit(c); + if (d < 0 || d >= base) break; + /* Detect overflow in multiplication base * res */ + if (res > threshold) caml_failwith(INT64_ERRMSG); + res = base * res + d; + /* Detect overflow in addition (base * res) + d */ + if (res < (uint64_t) d) caml_failwith(INT64_ERRMSG); + } + if (p != String_val(s) + caml_string_length(s)){ + caml_failwith(INT64_ERRMSG); + } + if (signedness) { + /* Signed representation expected, allow -2^63 to 2^63 - 1 only */ + if (sign >= 0) { + if (res >= (uint64_t)1 << 63) caml_failwith(INT64_ERRMSG); + } else { + if (res > (uint64_t)1 << 63) caml_failwith(INT64_ERRMSG); + } + } + if (sign < 0) res = - res; + return caml_copy_int64(res); +} + +int64_t caml_int64_bits_of_float_unboxed(double d) +{ + union { double d; int64_t i; int32_t h[2]; } u; + u.d = d; +#if defined(__arm__) && !defined(__ARM_EABI__) + { int32_t t = u.h[0]; u.h[0] = u.h[1]; u.h[1] = t; } +#endif + return u.i; +} + +double caml_int64_float_of_bits_unboxed(int64_t i) +{ + union { double d; int64_t i; int32_t h[2]; } u; + u.i = i; +#if defined(__arm__) && !defined(__ARM_EABI__) + { int32_t t = u.h[0]; u.h[0] = u.h[1]; u.h[1] = t; } +#endif + return u.d; +} + +CAMLprim value caml_int64_bits_of_float(value vd) +{ + return caml_copy_int64(caml_int64_bits_of_float_unboxed(Double_val(vd))); +} + +CAMLprim value caml_int64_float_of_bits(value vi) +{ + return caml_copy_double(caml_int64_float_of_bits_unboxed(Int64_val(vi))); +} + +/* Native integers */ + +static int nativeint_cmp(value v1, value v2) +{ + intnat i1 = Nativeint_val(v1); + intnat i2 = Nativeint_val(v2); + return (i1 > i2) - (i1 < i2); +} + +static intnat nativeint_hash(value v) +{ + intnat n = Nativeint_val(v); +#ifdef ARCH_SIXTYFOUR + /* 32/64 bits compatibility trick. See explanations in file "hash.c", + function caml_hash_mix_intnat. */ + return (n >> 32) ^ (n >> 63) ^ n; +#else + return n; +#endif +} + +static void nativeint_serialize(value v, uintnat * bsize_32, + uintnat * bsize_64) +{ + intnat l = Nativeint_val(v); +#ifdef ARCH_SIXTYFOUR + if (l >= -((intnat)1 << 31) && l < ((intnat)1 << 31)) { + caml_serialize_int_1(1); + caml_serialize_int_4((int32_t) l); + } else { + caml_serialize_int_1(2); + caml_serialize_int_8(l); + } +#else + caml_serialize_int_1(1); + caml_serialize_int_4(l); +#endif + *bsize_32 = 4; + *bsize_64 = 8; +} + +static uintnat nativeint_deserialize(void * dst) +{ + switch (caml_deserialize_uint_1()) { + case 1: + *((intnat *) dst) = caml_deserialize_sint_4(); + break; + case 2: +#ifdef ARCH_SIXTYFOUR + *((intnat *) dst) = caml_deserialize_sint_8(); +#else + caml_deserialize_error("input_value: native integer value too large"); +#endif + break; + default: + caml_deserialize_error("input_value: ill-formed native integer"); + } + return sizeof(intnat); +} + +CAMLexport struct custom_operations caml_nativeint_ops = { + "_n", + custom_finalize_default, + nativeint_cmp, + nativeint_hash, + nativeint_serialize, + nativeint_deserialize, + custom_compare_ext_default +}; + +CAMLexport value caml_copy_nativeint(intnat i) +{ + value res = caml_alloc_custom(&caml_nativeint_ops, sizeof(intnat), 0, 1); + Nativeint_val(res) = i; + return res; +} + +CAMLprim value caml_nativeint_neg(value v) +{ return caml_copy_nativeint(- Nativeint_val(v)); } + +CAMLprim value caml_nativeint_add(value v1, value v2) +{ return caml_copy_nativeint(Nativeint_val(v1) + Nativeint_val(v2)); } + +CAMLprim value caml_nativeint_sub(value v1, value v2) +{ return caml_copy_nativeint(Nativeint_val(v1) - Nativeint_val(v2)); } + +CAMLprim value caml_nativeint_mul(value v1, value v2) +{ return caml_copy_nativeint(Nativeint_val(v1) * Nativeint_val(v2)); } + +#define Nativeint_min_int ((intnat) 1 << (sizeof(intnat) * 8 - 1)) + +CAMLprim value caml_nativeint_div(value v1, value v2) +{ + intnat dividend = Nativeint_val(v1); + intnat divisor = Nativeint_val(v2); + if (divisor == 0) caml_raise_zero_divide(); + /* PR#4740: on some processors, modulus crashes if division overflows. + Implement the same behavior as for type "int". */ + if (dividend == Nativeint_min_int && divisor == -1) return v1; + return caml_copy_nativeint(dividend / divisor); +} + +CAMLprim value caml_nativeint_mod(value v1, value v2) +{ + intnat dividend = Nativeint_val(v1); + intnat divisor = Nativeint_val(v2); + if (divisor == 0) caml_raise_zero_divide(); + /* PR#4740: on some processors, modulus crashes if division overflows. + Implement the same behavior as for type "int". */ + if (dividend == Nativeint_min_int && divisor == -1){ + return caml_copy_nativeint(0); + } + return caml_copy_nativeint(dividend % divisor); +} + +CAMLprim value caml_nativeint_and(value v1, value v2) +{ return caml_copy_nativeint(Nativeint_val(v1) & Nativeint_val(v2)); } + +CAMLprim value caml_nativeint_or(value v1, value v2) +{ return caml_copy_nativeint(Nativeint_val(v1) | Nativeint_val(v2)); } + +CAMLprim value caml_nativeint_xor(value v1, value v2) +{ return caml_copy_nativeint(Nativeint_val(v1) ^ Nativeint_val(v2)); } + +CAMLprim value caml_nativeint_shift_left(value v1, value v2) +{ return caml_copy_nativeint(Nativeint_val(v1) << Int_val(v2)); } + +CAMLprim value caml_nativeint_shift_right(value v1, value v2) +{ return caml_copy_nativeint(Nativeint_val(v1) >> Int_val(v2)); } + +CAMLprim value caml_nativeint_shift_right_unsigned(value v1, value v2) +{ return caml_copy_nativeint((uintnat)Nativeint_val(v1) >> Int_val(v2)); } + +value caml_nativeint_direct_bswap(value v) +{ +#ifdef ARCH_SIXTYFOUR + return caml_swap64(v); +#else + return caml_swap32(v); +#endif +} + +CAMLprim value caml_nativeint_bswap(value v) +{ +#ifdef ARCH_SIXTYFOUR + return caml_copy_nativeint(caml_swap64(Nativeint_val(v))); +#else + return caml_copy_nativeint(caml_swap32(Nativeint_val(v))); +#endif +} + +CAMLprim value caml_nativeint_of_int(value v) +{ return caml_copy_nativeint(Long_val(v)); } + +CAMLprim value caml_nativeint_to_int(value v) +{ return Val_long(Nativeint_val(v)); } + +intnat caml_nativeint_of_float_unboxed(double x) +{ return x; } + +CAMLprim value caml_nativeint_of_float(value v) +{ return caml_copy_nativeint((intnat)(Double_val(v))); } + +double caml_nativeint_to_float_unboxed(intnat x) +{ return x; } + +CAMLprim value caml_nativeint_to_float(value v) +{ return caml_copy_double((double)(Nativeint_val(v))); } + +CAMLprim value caml_nativeint_of_int32(value v) +{ return caml_copy_nativeint(Int32_val(v)); } + +CAMLprim value caml_nativeint_to_int32(value v) +{ return caml_copy_int32(Nativeint_val(v)); } + +intnat caml_nativeint_compare_unboxed(intnat i1, intnat i2) +{ + return (i1 > i2) - (i1 < i2); +} + +CAMLprim value caml_nativeint_compare(value v1, value v2) +{ + return Val_int(caml_nativeint_compare_unboxed(Nativeint_val(v1), + Nativeint_val(v2))); +} + +CAMLprim value caml_nativeint_format(value fmt, value arg) +{ + char format_string[FORMAT_BUFFER_SIZE]; + + parse_format(fmt, ARCH_INTNAT_PRINTF_FORMAT, format_string); + return caml_alloc_sprintf(format_string, Nativeint_val(arg)); +} + +CAMLprim value caml_nativeint_of_string(value s) +{ + return caml_copy_nativeint(parse_intnat(s, 8 * sizeof(value), INTNAT_ERRMSG)); +} diff --git a/test/monniaux/ocaml/byterun/io.c b/test/monniaux/ocaml/byterun/io.c new file mode 100644 index 00000000..d124b56a --- /dev/null +++ b/test/monniaux/ocaml/byterun/io.c @@ -0,0 +1,828 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Buffered input/output. */ + +#include +#include +#include +#include +#include +#include +#include "caml/config.h" +#ifdef HAS_UNISTD +#include +#endif +#ifdef __CYGWIN__ +#include +#endif +#include "caml/alloc.h" +#include "caml/custom.h" +#include "caml/fail.h" +#include "caml/io.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/osdeps.h" +#include "caml/signals.h" +#include "caml/sys.h" + +#ifndef SEEK_SET +#define SEEK_SET 0 +#define SEEK_CUR 1 +#define SEEK_END 2 +#endif + +/* Hooks for locking channels */ + +CAMLexport void (*caml_channel_mutex_free) (struct channel *) = NULL; +CAMLexport void (*caml_channel_mutex_lock) (struct channel *) = NULL; +CAMLexport void (*caml_channel_mutex_unlock) (struct channel *) = NULL; +CAMLexport void (*caml_channel_mutex_unlock_exn) (void) = NULL; + +/* List of opened channels */ +CAMLexport struct channel * caml_all_opened_channels = NULL; + +/* Basic functions over type struct channel *. + These functions can be called directly from C. + No locking is performed. */ + +/* Functions shared between input and output */ + +CAMLexport struct channel * caml_open_descriptor_in(int fd) +{ + struct channel * channel; + + channel = (struct channel *) caml_stat_alloc(sizeof(struct channel)); + channel->fd = fd; + caml_enter_blocking_section(); + channel->offset = lseek(fd, 0, SEEK_CUR); + caml_leave_blocking_section(); + channel->curr = channel->max = channel->buff; + channel->end = channel->buff + IO_BUFFER_SIZE; + channel->mutex = NULL; + channel->revealed = 0; + channel->old_revealed = 0; + channel->refcount = 0; + channel->flags = 0; + channel->next = caml_all_opened_channels; + channel->prev = NULL; + channel->name = NULL; + if (caml_all_opened_channels != NULL) + caml_all_opened_channels->prev = channel; + caml_all_opened_channels = channel; + return channel; +} + +CAMLexport struct channel * caml_open_descriptor_out(int fd) +{ + struct channel * channel; + + channel = caml_open_descriptor_in(fd); + channel->max = NULL; + return channel; +} + +static void unlink_channel(struct channel *channel) +{ + if (channel->prev == NULL) { + CAMLassert (channel == caml_all_opened_channels); + caml_all_opened_channels = caml_all_opened_channels->next; + if (caml_all_opened_channels != NULL) + caml_all_opened_channels->prev = NULL; + } else { + channel->prev->next = channel->next; + if (channel->next != NULL) channel->next->prev = channel->prev; + } +} + +CAMLexport void caml_close_channel(struct channel *channel) +{ + CAML_SYS_CLOSE(channel->fd); + if (channel->refcount > 0) return; + if (caml_channel_mutex_free != NULL) (*caml_channel_mutex_free)(channel); + unlink_channel(channel); + caml_stat_free(channel->name); + caml_stat_free(channel); +} + +CAMLexport file_offset caml_channel_size(struct channel *channel) +{ + file_offset offset; + file_offset end; + int fd; + + /* We extract data from [channel] before dropping the OCaml lock, in case + someone else touches the block. */ + fd = channel->fd; + offset = channel->offset; + caml_enter_blocking_section(); + end = lseek(fd, 0, SEEK_END); + if (end == -1 || lseek(fd, offset, SEEK_SET) != offset) { + caml_leave_blocking_section(); + caml_sys_error(NO_ARG); + } + caml_leave_blocking_section(); + return end; +} + +CAMLexport int caml_channel_binary_mode(struct channel *channel) +{ +#if defined(_WIN32) || defined(__CYGWIN__) + int oldmode = setmode(channel->fd, O_BINARY); + if (oldmode == O_TEXT) setmode(channel->fd, O_TEXT); + return oldmode == O_BINARY; +#else + return 1; +#endif +} + +/* Output */ + +/* Attempt to flush the buffer. This will make room in the buffer for + at least one character. Returns true if the buffer is empty at the + end of the flush, or false if some data remains in the buffer. + */ + +CAMLexport int caml_flush_partial(struct channel *channel) +{ + int towrite, written; + + towrite = channel->curr - channel->buff; + CAMLassert (towrite >= 0); + if (towrite > 0) { + written = caml_write_fd(channel->fd, channel->flags, + channel->buff, towrite); + channel->offset += written; + if (written < towrite) + memmove(channel->buff, channel->buff + written, towrite - written); + channel->curr -= written; + } + return (channel->curr == channel->buff); +} + +/* Flush completely the buffer. */ + +CAMLexport void caml_flush(struct channel *channel) +{ + while (! caml_flush_partial(channel)) /*nothing*/; +} + +/* Output data */ + +CAMLexport void caml_putword(struct channel *channel, uint32_t w) +{ + if (! caml_channel_binary_mode(channel)) + caml_failwith("output_binary_int: not a binary channel"); + caml_putch(channel, w >> 24); + caml_putch(channel, w >> 16); + caml_putch(channel, w >> 8); + caml_putch(channel, w); +} + +CAMLexport int caml_putblock(struct channel *channel, char *p, intnat len) +{ + int n, free, towrite, written; + + n = len >= INT_MAX ? INT_MAX : (int) len; + free = channel->end - channel->curr; + if (n < free) { + /* Write request small enough to fit in buffer: transfer to buffer. */ + memmove(channel->curr, p, n); + channel->curr += n; + return n; + } else { + /* Write request overflows buffer (or just fills it up): transfer whatever + fits to buffer and write the buffer */ + memmove(channel->curr, p, free); + towrite = channel->end - channel->buff; + written = caml_write_fd(channel->fd, channel->flags, + channel->buff, towrite); + if (written < towrite) + memmove(channel->buff, channel->buff + written, towrite - written); + channel->offset += written; + channel->curr = channel->end - written; + return free; + } +} + +CAMLexport void caml_really_putblock(struct channel *channel, + char *p, intnat len) +{ + int written; + while (len > 0) { + written = caml_putblock(channel, p, len); + p += written; + len -= written; + } +} + +CAMLexport void caml_seek_out(struct channel *channel, file_offset dest) +{ + caml_flush(channel); + caml_enter_blocking_section(); + if (lseek(channel->fd, dest, SEEK_SET) != dest) { + caml_leave_blocking_section(); + caml_sys_error(NO_ARG); + } + caml_leave_blocking_section(); + channel->offset = dest; +} + +CAMLexport file_offset caml_pos_out(struct channel *channel) +{ + return channel->offset + (file_offset)(channel->curr - channel->buff); +} + +/* Input */ + +/* caml_do_read is exported for Cash */ +CAMLexport int caml_do_read(int fd, char *p, unsigned int n) +{ + return caml_read_fd(fd, 0, p, n); +} + +CAMLexport unsigned char caml_refill(struct channel *channel) +{ + int n; + + n = caml_read_fd(channel->fd, channel->flags, + channel->buff, channel->end - channel->buff); + if (n == 0) caml_raise_end_of_file(); + channel->offset += n; + channel->max = channel->buff + n; + channel->curr = channel->buff + 1; + return (unsigned char)(channel->buff[0]); +} + +CAMLexport uint32_t caml_getword(struct channel *channel) +{ + int i; + uint32_t res; + + if (! caml_channel_binary_mode(channel)) + caml_failwith("input_binary_int: not a binary channel"); + res = 0; + for(i = 0; i < 4; i++) { + res = (res << 8) + caml_getch(channel); + } + return res; +} + +CAMLexport int caml_getblock(struct channel *channel, char *p, intnat len) +{ + int n, avail, nread; + + n = len >= INT_MAX ? INT_MAX : (int) len; + avail = channel->max - channel->curr; + if (n <= avail) { + memmove(p, channel->curr, n); + channel->curr += n; + return n; + } else if (avail > 0) { + memmove(p, channel->curr, avail); + channel->curr += avail; + return avail; + } else { + nread = caml_read_fd(channel->fd, channel->flags, channel->buff, + channel->end - channel->buff); + channel->offset += nread; + channel->max = channel->buff + nread; + if (n > nread) n = nread; + memmove(p, channel->buff, n); + channel->curr = channel->buff + n; + return n; + } +} + +/* Returns the number of bytes read. */ +CAMLexport intnat caml_really_getblock(struct channel *chan, char *p, intnat n) +{ + intnat k = n; + int r; + while (k > 0) { + r = caml_getblock(chan, p, k); + if (r == 0) break; + p += r; + k -= r; + } + return n - k; +} + +CAMLexport void caml_seek_in(struct channel *channel, file_offset dest) +{ + if (dest >= channel->offset - (channel->max - channel->buff) && + dest <= channel->offset) { + channel->curr = channel->max - (channel->offset - dest); + } else { + caml_enter_blocking_section(); + if (lseek(channel->fd, dest, SEEK_SET) != dest) { + caml_leave_blocking_section(); + caml_sys_error(NO_ARG); + } + caml_leave_blocking_section(); + channel->offset = dest; + channel->curr = channel->max = channel->buff; + } +} + +CAMLexport file_offset caml_pos_in(struct channel *channel) +{ + return channel->offset - (file_offset)(channel->max - channel->curr); +} + +CAMLexport intnat caml_input_scan_line(struct channel *channel) +{ + char * p; + int n; + + p = channel->curr; + do { + if (p >= channel->max) { + /* No more characters available in the buffer */ + if (channel->curr > channel->buff) { + /* Try to make some room in the buffer by shifting the unread + portion at the beginning */ + memmove(channel->buff, channel->curr, channel->max - channel->curr); + n = channel->curr - channel->buff; + channel->curr -= n; + channel->max -= n; + p -= n; + } + if (channel->max >= channel->end) { + /* Buffer is full, no room to read more characters from the input. + Return the number of characters in the buffer, with negative + sign to indicate that no newline was encountered. */ + return -(channel->max - channel->curr); + } + /* Fill the buffer as much as possible */ + n = caml_read_fd(channel->fd, channel->flags, + channel->max, channel->end - channel->max); + if (n == 0) { + /* End-of-file encountered. Return the number of characters in the + buffer, with negative sign since we haven't encountered + a newline. */ + return -(channel->max - channel->curr); + } + channel->offset += n; + channel->max += n; + } + } while (*p++ != '\n'); + /* Found a newline. Return the length of the line, newline included. */ + return (p - channel->curr); +} + +/* OCaml entry points for the I/O functions. Wrap struct channel * + objects into a heap-allocated object. Perform locking + and unlocking around the I/O operations. */ + +/* FIXME CAMLexport, but not in io.h exported for Cash ? */ +CAMLexport void caml_finalize_channel(value vchan) +{ + struct channel * chan = Channel(vchan); + if ((chan->flags & CHANNEL_FLAG_MANAGED_BY_GC) == 0) return; + if (--chan->refcount > 0) return; + if (caml_channel_mutex_free != NULL) (*caml_channel_mutex_free)(chan); + + if (chan->fd != -1 && chan->name && caml_runtime_warnings_active()) + fprintf(stderr, + "[ocaml] channel opened on file '%s' dies without being closed\n", + chan->name + ); + + if (chan->max == NULL && chan->curr != chan->buff){ + /* + This is an unclosed out channel (chan->max == NULL) with a + non-empty buffer: keep it around so the OCaml [at_exit] function + gets a chance to flush it. We would want to simply flush the + channel now, but (i) flushing can raise exceptions, and (ii) it + is potentially a blocking operation. Both are forbidden in a + finalization function. + + Refs: + http://caml.inria.fr/mantis/view.php?id=6902 + https://github.com/ocaml/ocaml/pull/210 + */ + if (chan->name && caml_runtime_warnings_active()) + fprintf(stderr, + "[ocaml] (moreover, it has unflushed data)\n" + ); + } else { + unlink_channel(chan); + caml_stat_free(chan->name); + caml_stat_free(chan); + } +} + +static int compare_channel(value vchan1, value vchan2) +{ + struct channel * chan1 = Channel(vchan1); + struct channel * chan2 = Channel(vchan2); + return (chan1 == chan2) ? 0 : (chan1 < chan2) ? -1 : 1; +} + +static intnat hash_channel(value vchan) +{ + return (intnat) (Channel(vchan)); +} + +static struct custom_operations channel_operations = { + "_chan", + caml_finalize_channel, + compare_channel, + hash_channel, + custom_serialize_default, + custom_deserialize_default, + custom_compare_ext_default +}; + +CAMLexport value caml_alloc_channel(struct channel *chan) +{ + value res; + chan->refcount++; /* prevent finalization during next alloc */ + res = caml_alloc_custom(&channel_operations, sizeof(struct channel *), + 1, 1000); + Channel(res) = chan; + return res; +} + +CAMLprim value caml_ml_open_descriptor_in(value fd) +{ + struct channel * chan = caml_open_descriptor_in(Int_val(fd)); + chan->flags |= CHANNEL_FLAG_MANAGED_BY_GC; + return caml_alloc_channel(chan); +} + +CAMLprim value caml_ml_open_descriptor_out(value fd) +{ + struct channel * chan = caml_open_descriptor_out(Int_val(fd)); + chan->flags |= CHANNEL_FLAG_MANAGED_BY_GC; + return caml_alloc_channel(chan); +} + +CAMLprim value caml_ml_set_channel_name(value vchannel, value vname) +{ + struct channel * channel = Channel(vchannel); + caml_stat_free(channel->name); + if (caml_string_length(vname) > 0) + channel->name = caml_stat_strdup(String_val(vname)); + else + channel->name = NULL; + return Val_unit; +} + +#define Pair_tag 0 + +CAMLprim value caml_ml_out_channels_list (value unit) +{ + CAMLparam0 (); + CAMLlocal3 (res, tail, chan); + struct channel * channel; + + res = Val_emptylist; + for (channel = caml_all_opened_channels; + channel != NULL; + channel = channel->next) + /* Testing channel->fd >= 0 looks unnecessary, as + caml_ml_close_channel changes max when setting fd to -1. */ + if (channel->max == NULL) { + chan = caml_alloc_channel (channel); + tail = res; + res = caml_alloc_small (2, Pair_tag); + Field (res, 0) = chan; + Field (res, 1) = tail; + } + CAMLreturn (res); +} + +CAMLprim value caml_channel_descriptor(value vchannel) +{ + int fd = Channel(vchannel)->fd; + if (fd == -1) { errno = EBADF; caml_sys_error(NO_ARG); } + return Val_int(fd); +} + +CAMLprim value caml_ml_close_channel(value vchannel) +{ + int result; + int do_syscall; + int fd; + + /* For output channels, must have flushed before */ + struct channel * channel = Channel(vchannel); + if (channel->fd != -1){ + fd = channel->fd; + channel->fd = -1; + do_syscall = 1; + }else{ + do_syscall = 0; + result = 0; + } + /* Ensure that every read or write on the channel will cause an + immediate caml_flush_partial or caml_refill, thus raising a Sys_error + exception */ + channel->curr = channel->max = channel->end; + + if (do_syscall) { + caml_enter_blocking_section(); + result = CAML_SYS_CLOSE(fd); + caml_leave_blocking_section(); + } + + if (result == -1) caml_sys_error (NO_ARG); + return Val_unit; +} + +/* EOVERFLOW is the Unix98 error indicating that a file position or file + size is not representable. + ERANGE is the ANSI C error indicating that some argument to some + function is out of range. This is less precise than EOVERFLOW, + but guaranteed to be defined on all ANSI C environments. */ +#ifndef EOVERFLOW +#define EOVERFLOW ERANGE +#endif + +CAMLprim value caml_ml_channel_size(value vchannel) +{ + file_offset size = caml_channel_size(Channel(vchannel)); + if (size > Max_long) { errno = EOVERFLOW; caml_sys_error(NO_ARG); } + return Val_long(size); +} + +CAMLprim value caml_ml_channel_size_64(value vchannel) +{ + return Val_file_offset(caml_channel_size(Channel(vchannel))); +} + +CAMLprim value caml_ml_set_binary_mode(value vchannel, value mode) +{ +#if defined(_WIN32) || defined(__CYGWIN__) + struct channel * channel = Channel(vchannel); +#if defined(_WIN32) + /* The implementation of [caml_read_fd] and [caml_write_fd] in win32.c + doesn't support socket I/O with CRLF conversion. */ + if ((channel->flags & CHANNEL_FLAG_FROM_SOCKET) != 0 + && ! Bool_val(mode)) { + errno = EINVAL; + caml_sys_error(NO_ARG); + } +#endif + if (setmode(channel->fd, Bool_val(mode) ? O_BINARY : O_TEXT) == -1) + caml_sys_error(NO_ARG); +#endif + return Val_unit; +} + +/* + If the channel is closed, DO NOT raise a "bad file descriptor" + exception, but do nothing (the buffer is already empty). + This is because some libraries will flush at exit, even on + file descriptors that may be closed. +*/ + +CAMLprim value caml_ml_flush_partial(value vchannel) +{ + CAMLparam1 (vchannel); + struct channel * channel = Channel(vchannel); + int res; + + if (channel->fd == -1) CAMLreturn(Val_true); + Lock(channel); + res = caml_flush_partial(channel); + Unlock(channel); + CAMLreturn (Val_bool(res)); +} + +CAMLprim value caml_ml_flush(value vchannel) +{ + CAMLparam1 (vchannel); + struct channel * channel = Channel(vchannel); + + if (channel->fd == -1) CAMLreturn(Val_unit); + Lock(channel); + caml_flush(channel); + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_ml_output_char(value vchannel, value ch) +{ + CAMLparam2 (vchannel, ch); + struct channel * channel = Channel(vchannel); + + Lock(channel); + caml_putch(channel, Long_val(ch)); + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_ml_output_int(value vchannel, value w) +{ + CAMLparam2 (vchannel, w); + struct channel * channel = Channel(vchannel); + + Lock(channel); + caml_putword(channel, Long_val(w)); + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_ml_output_partial(value vchannel, value buff, value start, + value length) +{ + CAMLparam4 (vchannel, buff, start, length); + struct channel * channel = Channel(vchannel); + int res; + + Lock(channel); + res = caml_putblock(channel, &Byte(buff, Long_val(start)), Long_val(length)); + Unlock(channel); + CAMLreturn (Val_int(res)); +} + +CAMLprim value caml_ml_output_bytes(value vchannel, value buff, value start, + value length) +{ + CAMLparam4 (vchannel, buff, start, length); + struct channel * channel = Channel(vchannel); + intnat pos = Long_val(start); + intnat len = Long_val(length); + + Lock(channel); + /* We cannot call caml_really_putblock here because buff may move + during caml_write_fd */ + while (len > 0) { + int written = caml_putblock(channel, &Byte(buff, pos), len); + pos += written; + len -= written; + } + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_ml_output(value vchannel, value buff, value start, + value length) +{ + return caml_ml_output_bytes (vchannel, buff, start, length); +} + +CAMLprim value caml_ml_seek_out(value vchannel, value pos) +{ + CAMLparam2 (vchannel, pos); + struct channel * channel = Channel(vchannel); + + Lock(channel); + caml_seek_out(channel, Long_val(pos)); + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_ml_seek_out_64(value vchannel, value pos) +{ + CAMLparam2 (vchannel, pos); + struct channel * channel = Channel(vchannel); + + Lock(channel); + caml_seek_out(channel, File_offset_val(pos)); + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_ml_pos_out(value vchannel) +{ + file_offset pos = caml_pos_out(Channel(vchannel)); + if (pos > Max_long) { errno = EOVERFLOW; caml_sys_error(NO_ARG); } + return Val_long(pos); +} + +CAMLprim value caml_ml_pos_out_64(value vchannel) +{ + return Val_file_offset(caml_pos_out(Channel(vchannel))); +} + +CAMLprim value caml_ml_input_char(value vchannel) +{ + CAMLparam1 (vchannel); + struct channel * channel = Channel(vchannel); + unsigned char c; + + Lock(channel); + c = caml_getch(channel); + Unlock(channel); + CAMLreturn (Val_long(c)); +} + +CAMLprim value caml_ml_input_int(value vchannel) +{ + CAMLparam1 (vchannel); + struct channel * channel = Channel(vchannel); + intnat i; + + Lock(channel); + i = caml_getword(channel); + Unlock(channel); +#ifdef ARCH_SIXTYFOUR + i = (i << 32) >> 32; /* Force sign extension */ +#endif + CAMLreturn (Val_long(i)); +} + +CAMLprim value caml_ml_input(value vchannel, value buff, value vstart, + value vlength) +{ + CAMLparam4 (vchannel, buff, vstart, vlength); + struct channel * channel = Channel(vchannel); + intnat start, len; + int n, avail, nread; + + Lock(channel); + /* We cannot call caml_getblock here because buff may move during + caml_read_fd */ + start = Long_val(vstart); + len = Long_val(vlength); + n = len >= INT_MAX ? INT_MAX : (int) len; + avail = channel->max - channel->curr; + if (n <= avail) { + memmove(&Byte(buff, start), channel->curr, n); + channel->curr += n; + } else if (avail > 0) { + memmove(&Byte(buff, start), channel->curr, avail); + channel->curr += avail; + n = avail; + } else { + nread = caml_read_fd(channel->fd, channel->flags, channel->buff, + channel->end - channel->buff); + channel->offset += nread; + channel->max = channel->buff + nread; + if (n > nread) n = nread; + memmove(&Byte(buff, start), channel->buff, n); + channel->curr = channel->buff + n; + } + Unlock(channel); + CAMLreturn (Val_long(n)); +} + +CAMLprim value caml_ml_seek_in(value vchannel, value pos) +{ + CAMLparam2 (vchannel, pos); + struct channel * channel = Channel(vchannel); + + Lock(channel); + caml_seek_in(channel, Long_val(pos)); + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_ml_seek_in_64(value vchannel, value pos) +{ + CAMLparam2 (vchannel, pos); + struct channel * channel = Channel(vchannel); + + Lock(channel); + caml_seek_in(channel, File_offset_val(pos)); + Unlock(channel); + CAMLreturn (Val_unit); +} + +CAMLprim value caml_ml_pos_in(value vchannel) +{ + file_offset pos = caml_pos_in(Channel(vchannel)); + if (pos > Max_long) { errno = EOVERFLOW; caml_sys_error(NO_ARG); } + return Val_long(pos); +} + +CAMLprim value caml_ml_pos_in_64(value vchannel) +{ + return Val_file_offset(caml_pos_in(Channel(vchannel))); +} + +CAMLprim value caml_ml_input_scan_line(value vchannel) +{ + CAMLparam1 (vchannel); + struct channel * channel = Channel(vchannel); + intnat res; + + Lock(channel); + res = caml_input_scan_line(channel); + Unlock(channel); + CAMLreturn (Val_long(res)); +} + +CAMLprim value caml_terminfo_rows(value vchannel) +{ + return Val_int(caml_num_rows_fd(Channel(vchannel)->fd)); +} diff --git a/test/monniaux/ocaml/byterun/lexing.c b/test/monniaux/ocaml/byterun/lexing.c new file mode 100644 index 00000000..b1049904 --- /dev/null +++ b/test/monniaux/ocaml/byterun/lexing.c @@ -0,0 +1,233 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* The table-driven automaton for lexers generated by camllex. */ + +#include "caml/fail.h" +#include "caml/mlvalues.h" +#include "caml/stacks.h" + +struct lexer_buffer { + value refill_buff; + value lex_buffer; + value lex_buffer_len; + value lex_abs_pos; + value lex_start_pos; + value lex_curr_pos; + value lex_last_pos; + value lex_last_action; + value lex_eof_reached; + value lex_mem; + value lex_start_p; + value lex_curr_p; +}; + +struct lexing_table { + value lex_base; + value lex_backtrk; + value lex_default; + value lex_trans; + value lex_check; + value lex_base_code; + value lex_backtrk_code; + value lex_default_code; + value lex_trans_code; + value lex_check_code; + value lex_code; +}; + +#if defined(ARCH_BIG_ENDIAN) || SIZEOF_SHORT != 2 +#define Short(tbl,n) \ + (*((unsigned char *)((tbl) + (n) * 2)) + \ + (*((signed char *)((tbl) + (n) * 2 + 1)) << 8)) +#else +#define Short(tbl,n) (((short *)(tbl))[(n)]) +#endif + +CAMLprim value caml_lex_engine(struct lexing_table *tbl, value start_state, + struct lexer_buffer *lexbuf) +{ + int state, base, backtrk, c; + + state = Int_val(start_state); + if (state >= 0) { + /* First entry */ + lexbuf->lex_last_pos = lexbuf->lex_start_pos = lexbuf->lex_curr_pos; + lexbuf->lex_last_action = Val_int(-1); + } else { + /* Reentry after refill */ + state = -state - 1; + } + while(1) { + /* Lookup base address or action number for current state */ + base = Short(tbl->lex_base, state); + if (base < 0) return Val_int(-base-1); + /* See if it's a backtrack point */ + backtrk = Short(tbl->lex_backtrk, state); + if (backtrk >= 0) { + lexbuf->lex_last_pos = lexbuf->lex_curr_pos; + lexbuf->lex_last_action = Val_int(backtrk); + } + /* See if we need a refill */ + if (lexbuf->lex_curr_pos >= lexbuf->lex_buffer_len){ + if (lexbuf->lex_eof_reached == Val_bool (0)){ + return Val_int(-state - 1); + }else{ + c = 256; + } + }else{ + /* Read next input char */ + c = Byte_u(lexbuf->lex_buffer, Long_val(lexbuf->lex_curr_pos)); + lexbuf->lex_curr_pos += 2; + } + /* Determine next state */ + if (Short(tbl->lex_check, base + c) == state) + state = Short(tbl->lex_trans, base + c); + else + state = Short(tbl->lex_default, state); + /* If no transition on this char, return to last backtrack point */ + if (state < 0) { + lexbuf->lex_curr_pos = lexbuf->lex_last_pos; + if (lexbuf->lex_last_action == Val_int(-1)) { + caml_failwith("lexing: empty token"); + } else { + return lexbuf->lex_last_action; + } + }else{ + /* Erase the EOF condition only if the EOF pseudo-character was + consumed by the automaton (i.e. there was no backtrack above) + */ + if (c == 256) lexbuf->lex_eof_reached = Val_bool (0); + } + } +} + +/***********************************************/ +/* New lexer engine, with memory of positions */ +/***********************************************/ + +static void run_mem(char *pc, value mem, value curr_pos) { + for (;;) { + unsigned char dst, src ; + + dst = *pc++ ; + if (dst == 0xff) + return ; + src = *pc++ ; + if (src == 0xff) { + /* fprintf(stderr,"[%hhu] <- %d\n",dst,Int_val(curr_pos)) ;*/ + Field(mem,dst) = curr_pos ; + } else { + /* fprintf(stderr,"[%hhu] <- [%hhu]\n",dst,src) ; */ + Field(mem,dst) = Field(mem,src) ; + } + } +} + +static void run_tag(char *pc, value mem) { + for (;;) { + unsigned char dst, src ; + + dst = *pc++ ; + if (dst == 0xff) + return ; + src = *pc++ ; + if (src == 0xff) { + /* fprintf(stderr,"[%hhu] <- -1\n",dst) ; */ + Field(mem,dst) = Val_int(-1) ; + } else { + /* fprintf(stderr,"[%hhu] <- [%hhu]\n",dst,src) ; */ + Field(mem,dst) = Field(mem,src) ; + } + } +} + +CAMLprim value caml_new_lex_engine(struct lexing_table *tbl, value start_state, + struct lexer_buffer *lexbuf) +{ + int state, base, backtrk, c, pstate ; + state = Int_val(start_state); + if (state >= 0) { + /* First entry */ + lexbuf->lex_last_pos = lexbuf->lex_start_pos = lexbuf->lex_curr_pos; + lexbuf->lex_last_action = Val_int(-1); + } else { + /* Reentry after refill */ + state = -state - 1; + } + while(1) { + /* Lookup base address or action number for current state */ + base = Short(tbl->lex_base, state); + if (base < 0) { + int pc_off = Short(tbl->lex_base_code, state) ; + run_tag(Bp_val(tbl->lex_code) + pc_off, lexbuf->lex_mem); + /* fprintf(stderr,"Perform: %d\n",-base-1) ; */ + return Val_int(-base-1); + } + /* See if it's a backtrack point */ + backtrk = Short(tbl->lex_backtrk, state); + if (backtrk >= 0) { + int pc_off = Short(tbl->lex_backtrk_code, state); + run_tag(Bp_val(tbl->lex_code) + pc_off, lexbuf->lex_mem); + lexbuf->lex_last_pos = lexbuf->lex_curr_pos; + lexbuf->lex_last_action = Val_int(backtrk); + + } + /* See if we need a refill */ + if (lexbuf->lex_curr_pos >= lexbuf->lex_buffer_len){ + if (lexbuf->lex_eof_reached == Val_bool (0)){ + return Val_int(-state - 1); + }else{ + c = 256; + } + }else{ + /* Read next input char */ + c = Byte_u(lexbuf->lex_buffer, Long_val(lexbuf->lex_curr_pos)); + lexbuf->lex_curr_pos += 2; + } + /* Determine next state */ + pstate=state ; + if (Short(tbl->lex_check, base + c) == state) + state = Short(tbl->lex_trans, base + c); + else + state = Short(tbl->lex_default, state); + /* If no transition on this char, return to last backtrack point */ + if (state < 0) { + lexbuf->lex_curr_pos = lexbuf->lex_last_pos; + if (lexbuf->lex_last_action == Val_int(-1)) { + caml_failwith("lexing: empty token"); + } else { + return lexbuf->lex_last_action; + } + }else{ + /* If some transition, get and perform memory moves */ + int base_code = Short(tbl->lex_base_code, pstate) ; + int pc_off ; + if (Short(tbl->lex_check_code, base_code + c) == pstate) + pc_off = Short(tbl->lex_trans_code, base_code + c) ; + else + pc_off = Short(tbl->lex_default_code, pstate) ; + if (pc_off > 0) + run_mem(Bp_val(tbl->lex_code) + pc_off, lexbuf->lex_mem, + lexbuf->lex_curr_pos) ; + /* Erase the EOF condition only if the EOF pseudo-character was + consumed by the automaton (i.e. there was no backtrack above) + */ + if (c == 256) lexbuf->lex_eof_reached = Val_bool (0); + } + } +} diff --git a/test/monniaux/ocaml/byterun/main.c b/test/monniaux/ocaml/byterun/main.c new file mode 100644 index 00000000..5e5839ff --- /dev/null +++ b/test/monniaux/ocaml/byterun/main.c @@ -0,0 +1,47 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Main entry point (can be overridden by a user-provided main() + function that calls caml_main() later). */ + +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/sys.h" +#include "caml/osdeps.h" +#ifdef _WIN32 +#include +#endif + +CAMLextern void caml_main (char_os **); + +#ifdef _WIN32 +CAMLextern void caml_expand_command_line (int *, wchar_t ***); + +int wmain(int argc, wchar_t **argv) +#else +int main(int argc, char **argv) +#endif +{ +#ifdef _WIN32 + /* Expand wildcards and diversions in command line */ + caml_expand_command_line(&argc, &argv); +#endif + + caml_main(argv); + caml_sys_exit(Val_int(0)); + return 0; /* not reached */ +} diff --git a/test/monniaux/ocaml/byterun/major_gc.c b/test/monniaux/ocaml/byterun/major_gc.c new file mode 100644 index 00000000..493f2cc6 --- /dev/null +++ b/test/monniaux/ocaml/byterun/major_gc.c @@ -0,0 +1,943 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* DM */ +#include + +#include +#include + +#include "caml/compact.h" +#include "caml/custom.h" +#include "caml/config.h" +#include "caml/fail.h" +#include "caml/finalise.h" +#include "caml/freelist.h" +#include "caml/gc.h" +#include "caml/gc_ctrl.h" +#include "caml/major_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/roots.h" +#include "caml/signals.h" +#include "caml/weak.h" + +#if defined (NATIVE_CODE) && defined (NO_NAKED_POINTERS) +#define NATIVE_CODE_AND_NO_NAKED_POINTERS +#else +#undef NATIVE_CODE_AND_NO_NAKED_POINTERS +#endif + +#ifdef _MSC_VER +static inline double fmin(double a, double b) { + return (a < b) ? a : b; +} +#endif + +uintnat caml_percent_free; +uintnat caml_major_heap_increment; +CAMLexport char *caml_heap_start; +char *caml_gc_sweep_hp; +int caml_gc_phase; /* always Phase_mark, Pase_clean, + Phase_sweep, or Phase_idle */ +static value *gray_vals; +static value *gray_vals_cur, *gray_vals_end; +static asize_t gray_vals_size; +static int heap_is_pure; /* The heap is pure if the only gray objects + below [markhp] are also in [gray_vals]. */ +uintnat caml_allocated_words; +uintnat caml_dependent_size, caml_dependent_allocated; +double caml_extra_heap_resources; +uintnat caml_fl_wsz_at_phase_change = 0; + +extern char *caml_fl_merge; /* Defined in freelist.c. */ + +static char *markhp, *chunk, *limit; + +int caml_gc_subphase; /* Subphase_{mark_roots,mark_main,mark_final} */ + +/** + Ephemerons: + During mark phase the list caml_ephe_list_head of ephemerons + is iterated by different pointers that follow the invariants: + caml_ephe_list_head ->* ephes_checked_if_pure ->* ephes_to_check ->* null + | | | + (1) (2) (3) + + At the start of mark phase, (1) and (2) are empty. + + In mark phase: + - the ephemerons in (1) have a data alive or none + (nb: new ephemerons are added in this part by weak.c) + - the ephemerons in (2) have at least a white key or are white + if ephe_list_pure is true, otherwise they are in an unknown state and + must be checked again. + - the ephemerons in (3) are in an unknown state and must be checked + + At the end of mark phase, (3) is empty and ephe_list_pure is true. + The ephemeron in (1) and (2) will be cleaned (white keys and datas + replaced by none or the ephemeron is removed from the list if it is white) + in clean phase. + + In clean phase: + caml_ephe_list_head ->* ephes_to_check ->* null + | | + (1) (3) + + In clean phase, (2) is not used, ephes_to_check is initialized at + caml_ephe_list_head: + - the ephemerons in (1) are clean. + - the ephemerons in (3) should be cleaned or removed if white. + + */ +static int ephe_list_pure; +/** The ephemerons is pure if since the start of its iteration + no value have been darken. */ +static value *ephes_checked_if_pure; +static value *ephes_to_check; + +int caml_major_window = 1; +double caml_major_ring[Max_major_window] = { 0. }; +int caml_major_ring_index = 0; +double caml_major_work_credit = 0.0; +double caml_gc_clock = 0.0; + +#ifdef DEBUG +static unsigned long major_gc_counter = 0; +#endif + +void (*caml_major_gc_hook)(void) = NULL; + +static void realloc_gray_vals (void) +{ + value *new; + + CAMLassert (gray_vals_cur == gray_vals_end); + if (gray_vals_size < caml_stat_heap_wsz / 32){ + caml_gc_message (0x08, "Growing gray_vals to %" + ARCH_INTNAT_PRINTF_FORMAT "uk bytes\n", + (intnat) gray_vals_size * sizeof (value) / 512); + new = (value *) caml_stat_resize_noexc ((char *) gray_vals, + 2 * gray_vals_size * + sizeof (value)); + if (new == NULL){ + caml_gc_message (0x08, "No room for growing gray_vals\n"); + gray_vals_cur = gray_vals; + heap_is_pure = 0; + }else{ + gray_vals = new; + gray_vals_cur = gray_vals + gray_vals_size; + gray_vals_size *= 2; + gray_vals_end = gray_vals + gray_vals_size; + } + }else{ + gray_vals_cur = gray_vals + gray_vals_size / 2; + heap_is_pure = 0; + } +} + +void caml_darken (value v, value *p /* not used */) +{ +#ifdef NATIVE_CODE_AND_NO_NAKED_POINTERS + if (Is_block (v) && !Is_young (v) && Wosize_val (v) > 0) { +#else + if (Is_block (v) && Is_in_heap (v)) { +#endif + header_t h = Hd_val (v); + tag_t t = Tag_hd (h); + if (t == Infix_tag){ + v -= Infix_offset_val(v); + h = Hd_val (v); + t = Tag_hd (h); + } +#ifdef NATIVE_CODE_AND_NO_NAKED_POINTERS + /* We insist that naked pointers to outside the heap point to things that + look like values with headers coloured black. This isn't always + strictly necessary but is essential in certain cases---in particular + when the value is allocated in a read-only section. (For the values + where it would be safe it is a performance improvement since we avoid + putting them on the grey list.) */ + CAMLassert (Is_in_heap (v) || Is_black_hd (h)); +#endif + CAMLassert (!Is_blue_hd (h)); + if (Is_white_hd (h)){ + ephe_list_pure = 0; + if (t < No_scan_tag){ + Hd_val (v) = Grayhd_hd (h); + *gray_vals_cur++ = v; + if (gray_vals_cur >= gray_vals_end) realloc_gray_vals (); + }else{ + Hd_val (v) = Blackhd_hd (h); + } + } + } +} + +static void start_cycle (void) +{ + CAMLassert (caml_gc_phase == Phase_idle); + CAMLassert (gray_vals_cur == gray_vals); + caml_gc_message (0x01, "Starting new major GC cycle\n"); + caml_darken_all_roots_start (); + caml_gc_phase = Phase_mark; + caml_gc_subphase = Subphase_mark_roots; + markhp = NULL; + ephe_list_pure = 1; + ephes_checked_if_pure = &caml_ephe_list_head; + ephes_to_check = &caml_ephe_list_head; +#ifdef DEBUG + ++ major_gc_counter; + caml_heap_check (); +#endif +} + +/* We may stop the slice inside values, in order to avoid large latencies + on large arrays. In this case, [current_value] is the partially-marked + value and [current_index] is the index of the next field to be marked. +*/ +static value current_value = 0; +static mlsize_t current_index = 0; + +/* For instrumentation */ +#ifdef CAML_INSTR +#define INSTR(x) x +#else +#define INSTR(x) /**/ +#endif + +static void init_sweep_phase(void) +{ + /* Phase_clean is done. */ + /* Initialise the sweep phase. */ + caml_gc_sweep_hp = caml_heap_start; + caml_fl_init_merge (); + caml_gc_phase = Phase_sweep; + chunk = caml_heap_start; + caml_gc_sweep_hp = chunk; + limit = chunk + Chunk_size (chunk); + caml_fl_wsz_at_phase_change = caml_fl_cur_wsz; + if (caml_major_gc_hook) (*caml_major_gc_hook)(); +} + +/* auxillary function of mark_slice */ +static inline value* mark_slice_darken(value *gray_vals_ptr, value v, int i, + int in_ephemeron, int *slice_pointers) +{ + value child; + header_t chd; + + child = Field (v, i); + +#ifdef NATIVE_CODE_AND_NO_NAKED_POINTERS + if (Is_block (child) + && ! Is_young (child) + && Wosize_val (child) > 0 /* Atoms never need to be marked. */ + /* Closure blocks contain code pointers at offsets that cannot + be reliably determined, so we always use the page table when + marking such values. */ + && (!(Tag_val (v) == Closure_tag || Tag_val (v) == Infix_tag) || + Is_in_heap (child))) { +#else + if (Is_block (child) && Is_in_heap (child)) { +#endif + INSTR (++ *slice_pointers;) + chd = Hd_val (child); + if (Tag_hd (chd) == Forward_tag){ + value f = Forward_val (child); + if ((in_ephemeron && Is_long(f)) || + (Is_block (f) + && (!Is_in_value_area(f) || Tag_val (f) == Forward_tag + || Tag_val (f) == Lazy_tag +#ifdef FLAT_FLOAT_ARRAY + || Tag_val (f) == Double_tag +#endif + ))){ + /* Do not short-circuit the pointer. */ + }else{ + /* The variable child is not changed because it must be mark alive */ + Field (v, i) = f; + if (Is_block (f) && Is_young (f) && !Is_young (child)){ + if(in_ephemeron){ + add_to_ephe_ref_table (&caml_ephe_ref_table, v, i); + }else{ + add_to_ref_table (&caml_ref_table, &Field (v, i)); + } + } + } + } + else if (Tag_hd(chd) == Infix_tag) { + child -= Infix_offset_val(child); + chd = Hd_val(child); + } +#ifdef NATIVE_CODE_AND_NO_NAKED_POINTERS + /* See [caml_darken] for a description of this assertion. */ + CAMLassert (Is_in_heap (child) || Is_black_hd (chd)); +#endif + if (Is_white_hd (chd)){ + ephe_list_pure = 0; + Hd_val (child) = Grayhd_hd (chd); + *gray_vals_ptr++ = child; + if (gray_vals_ptr >= gray_vals_end) { + gray_vals_cur = gray_vals_ptr; + realloc_gray_vals (); + gray_vals_ptr = gray_vals_cur; + } + } + } + + return gray_vals_ptr; +} + +static value* mark_ephe_aux (value *gray_vals_ptr, intnat *work, + int *slice_pointers) +{ + value v, data, key; + header_t hd; + mlsize_t size, i; + + v = *ephes_to_check; + hd = Hd_val(v); + CAMLassert(Tag_val (v) == Abstract_tag); + data = Field(v,CAML_EPHE_DATA_OFFSET); + if ( data != caml_ephe_none && + Is_block (data) && Is_in_heap (data) && Is_white_val (data)){ + + int alive_data = 1; + + /* The liveness of the ephemeron is one of the condition */ + if (Is_white_hd (hd)) alive_data = 0; + + /* The liveness of the keys not caml_ephe_none is the other condition */ + size = Wosize_hd (hd); + for (i = CAML_EPHE_FIRST_KEY; alive_data && i < size; i++){ + key = Field (v, i); + ephemeron_again: + if (key != caml_ephe_none && + Is_block (key) && Is_in_heap (key)){ + if (Tag_val (key) == Forward_tag){ + value f = Forward_val (key); + if (Is_long (f) || + (Is_block (f) && + (!Is_in_value_area(f) || Tag_val (f) == Forward_tag + || Tag_val (f) == Lazy_tag +#ifdef FLAT_FLOAT_ARRAY + || Tag_val (f) == Double_tag +#endif + ))){ + /* Do not short-circuit the pointer. */ + }else{ + Field (v, i) = key = f; + goto ephemeron_again; + } + } + if (Is_white_val (key)){ + alive_data = 0; + } + } + } + *work -= Whsize_wosize(i); + + if (alive_data){ + gray_vals_ptr = mark_slice_darken(gray_vals_ptr,v, + CAML_EPHE_DATA_OFFSET, + /*in_ephemeron=*/1, + slice_pointers); + } else { /* not triggered move to the next one */ + ephes_to_check = &Field(v,CAML_EPHE_LINK_OFFSET); + return gray_vals_ptr; + } + } else { /* a simily weak pointer or an already alive data */ + *work -= 1; + } + + /* all keys black or data none or black + move the ephemerons from (3) to the end of (1) */ + if ( ephes_checked_if_pure == ephes_to_check ) { + /* corner case and optim */ + ephes_checked_if_pure = &Field(v,CAML_EPHE_LINK_OFFSET); + ephes_to_check = ephes_checked_if_pure; + } else { + /* - remove v from the list (3) */ + *ephes_to_check = Field(v,CAML_EPHE_LINK_OFFSET); + /* - insert it at the end of (1) */ + Field(v,CAML_EPHE_LINK_OFFSET) = *ephes_checked_if_pure; + *ephes_checked_if_pure = v; + ephes_checked_if_pure = &Field(v,CAML_EPHE_LINK_OFFSET); + } + return gray_vals_ptr; +} + + + +static void mark_slice (intnat work) +{ + value *gray_vals_ptr; /* Local copy of [gray_vals_cur] */ + value v; + header_t hd; + mlsize_t size, i, start, end; /* [start] is a local copy of [current_index] */ +#ifdef CAML_INSTR + int slice_fields = 0; +#endif + int slice_pointers = 0; /** gcc removes it when not in CAML_INSTR */ + + caml_gc_message (0x40, "Marking %"ARCH_INTNAT_PRINTF_FORMAT"d words\n", work); + caml_gc_message (0x40, "Subphase = %d\n", caml_gc_subphase); + gray_vals_ptr = gray_vals_cur; + v = current_value; + start = current_index; + while (work > 0){ + if (v == 0 && gray_vals_ptr > gray_vals){ + CAMLassert (start == 0); + v = *--gray_vals_ptr; + CAMLassert (Is_gray_val (v)); + } + if (v != 0){ + hd = Hd_val(v); + CAMLassert (Is_gray_hd (hd)); + size = Wosize_hd (hd); + end = start + work; + if (Tag_hd (hd) < No_scan_tag){ + start = size < start ? size : start; + end = size < end ? size : end; + CAMLassert (end >= start); + INSTR (slice_fields += end - start;) + INSTR (if (size > end) + CAML_INSTR_INT ("major/mark/slice/remain", size - end);) + for (i = start; i < end; i++){ + gray_vals_ptr = mark_slice_darken(gray_vals_ptr,v,i, + /*in_ephemeron=*/ 0, + &slice_pointers); + } + if (end < size){ + work = 0; + start = end; + /* [v] doesn't change. */ + CAMLassert (Is_gray_val (v)); + }else{ + CAMLassert (end == size); + Hd_val (v) = Blackhd_hd (hd); + work -= Whsize_wosize(end - start); + start = 0; + v = 0; + } + }else{ + /* The block doesn't contain any pointers. */ + CAMLassert (start == 0); + Hd_val (v) = Blackhd_hd (hd); + work -= Whsize_wosize(size); + v = 0; + } + }else if (markhp != NULL){ + if (markhp == limit){ + chunk = Chunk_next (chunk); + if (chunk == NULL){ + markhp = NULL; + }else{ + markhp = chunk; + limit = chunk + Chunk_size (chunk); + } + }else{ + if (Is_gray_val (Val_hp (markhp))){ + CAMLassert (gray_vals_ptr == gray_vals); + CAMLassert (v == 0 && start == 0); + v = Val_hp (markhp); + } + markhp += Bhsize_hp (markhp); + } + }else if (!heap_is_pure){ + heap_is_pure = 1; + chunk = caml_heap_start; + markhp = chunk; + limit = chunk + Chunk_size (chunk); + } else if (caml_gc_subphase == Subphase_mark_roots) { + gray_vals_cur = gray_vals_ptr; + work = caml_darken_all_roots_slice (work); + gray_vals_ptr = gray_vals_cur; + if (work > 0){ + caml_gc_subphase = Subphase_mark_main; + } + } else if (*ephes_to_check != (value) NULL) { + /* Continue to scan the list of ephe */ + gray_vals_ptr = mark_ephe_aux(gray_vals_ptr,&work,&slice_pointers); + } else if (!ephe_list_pure){ + /* We must scan again the list because some value have been darken */ + ephe_list_pure = 1; + ephes_to_check = ephes_checked_if_pure; + }else{ + switch (caml_gc_subphase){ + case Subphase_mark_main: { + /* Subphase_mark_main is done. + Mark finalised values. */ + gray_vals_cur = gray_vals_ptr; + caml_final_update_mark_phase (); + gray_vals_ptr = gray_vals_cur; + if (gray_vals_ptr > gray_vals){ + v = *--gray_vals_ptr; + CAMLassert (start == 0); + } + /* Complete the marking */ + ephes_to_check = ephes_checked_if_pure; + caml_gc_subphase = Subphase_mark_final; + } + break; + case Subphase_mark_final: { + /** The set of unreachable value will not change anymore for + this cycle. Start clean phase. */ + caml_gc_phase = Phase_clean; + caml_final_update_clean_phase (); + if (caml_ephe_list_head != (value) NULL){ + /* Initialise the clean phase. */ + ephes_to_check = &caml_ephe_list_head; + } else { + /* Initialise the sweep phase. */ + init_sweep_phase(); + } + work = 0; + } + break; + default: CAMLassert (0); + } + } + } + gray_vals_cur = gray_vals_ptr; + current_value = v; + current_index = start; + INSTR (CAML_INSTR_INT ("major/mark/slice/fields#", slice_fields);) + INSTR (CAML_INSTR_INT ("major/mark/slice/pointers#", slice_pointers);) +} + +/* Clean ephemerons */ +static void clean_slice (intnat work) +{ + value v; + + caml_gc_message (0x40, "Cleaning %" + ARCH_INTNAT_PRINTF_FORMAT "d words\n", work); + while (work > 0){ + v = *ephes_to_check; + if (v != (value) NULL){ + if (Is_white_val (v)){ + /* The whole array is dead, remove it from the list. */ + *ephes_to_check = Field (v, CAML_EPHE_LINK_OFFSET); + work -= 1; + }else{ + caml_ephe_clean(v); + ephes_to_check = &Field (v, CAML_EPHE_LINK_OFFSET); + work -= Whsize_val (v); + } + }else{ /* End of list reached */ + /* Phase_clean is done. */ + /* Initialise the sweep phase. */ + init_sweep_phase(); + work = 0; + } + } +} + +static void sweep_slice (intnat work) +{ + char *hp; + header_t hd; + + caml_gc_message (0x40, "Sweeping %" + ARCH_INTNAT_PRINTF_FORMAT "d words\n", work); + while (work > 0){ + if (caml_gc_sweep_hp < limit){ + hp = caml_gc_sweep_hp; + hd = Hd_hp (hp); + work -= Whsize_hd (hd); + caml_gc_sweep_hp += Bhsize_hd (hd); + switch (Color_hd (hd)){ + case Caml_white: + if (Tag_hd (hd) == Custom_tag){ + void (*final_fun)(value) = Custom_ops_val(Val_hp(hp))->finalize; + if (final_fun != NULL) final_fun(Val_hp(hp)); + } + caml_gc_sweep_hp = (char *) caml_fl_merge_block (Val_hp (hp)); + break; + case Caml_blue: + /* Only the blocks of the free-list are blue. See [freelist.c]. */ + caml_fl_merge = Bp_hp (hp); + break; + default: /* gray or black */ + CAMLassert (Color_hd (hd) == Caml_black); + Hd_hp (hp) = Whitehd_hd (hd); + break; + } + CAMLassert (caml_gc_sweep_hp <= limit); + }else{ + chunk = Chunk_next (chunk); + if (chunk == NULL){ + /* Sweeping is done. */ + ++ caml_stat_major_collections; + work = 0; + caml_gc_phase = Phase_idle; + caml_request_minor_gc (); + }else{ + caml_gc_sweep_hp = chunk; + limit = chunk + Chunk_size (chunk); + } + } + } +} + +#ifdef CAML_INSTR +static char *mark_slice_name[] = { + /* 0 */ NULL, + /* 1 */ NULL, + /* 2 */ NULL, + /* 3 */ NULL, + /* 4 */ NULL, + /* 5 */ NULL, + /* 6 */ NULL, + /* 7 */ NULL, + /* 8 */ NULL, + /* 9 */ NULL, + /* 10 */ "major/mark_roots", + /* 11 */ "major/mark_main", + /* 12 */ "major/mark_weak1", + /* 13 */ "major/mark_weak2", + /* 14 */ "major/mark_final", +}; +#endif + +/* The main entry point for the major GC. Called about once for each + minor GC. [howmuch] is the amount of work to do: + -1 if the GC is triggered automatically + 0 to let the GC compute the amount of work + [n] to make the GC do enough work to (on average) free [n] words + */ +void caml_major_collection_slice (intnat howmuch) +{ + double p, dp, filt_p, spend; + intnat computed_work; + int i; + /* + Free memory at the start of the GC cycle (garbage + free list) (assumed): + FM = caml_stat_heap_wsz * caml_percent_free + / (100 + caml_percent_free) + + Assuming steady state and enforcing a constant allocation rate, then + FM is divided in 2/3 for garbage and 1/3 for free list. + G = 2 * FM / 3 + G is also the amount of memory that will be used during this cycle + (still assuming steady state). + + Proportion of G consumed since the previous slice: + PH = caml_allocated_words / G + = caml_allocated_words * 3 * (100 + caml_percent_free) + / (2 * caml_stat_heap_wsz * caml_percent_free) + Proportion of extra-heap resources consumed since the previous slice: + PE = caml_extra_heap_resources + Proportion of total work to do in this slice: + P = max (PH, PE) + + Here, we insert a time-based filter on the P variable to avoid large + latency spikes in the GC, so the P below is a smoothed-out version of + the P above. + + Amount of marking work for the GC cycle: + MW = caml_stat_heap_wsz * 100 / (100 + caml_percent_free) + + caml_incremental_roots_count + Amount of sweeping work for the GC cycle: + SW = caml_stat_heap_wsz + + In order to finish marking with a non-empty free list, we will + use 40% of the time for marking, and 60% for sweeping. + + Let MT be the time spent marking, ST the time spent sweeping, and TT + the total time for this cycle. We have: + MT = 40/100 * TT + ST = 60/100 * TT + + Amount of time to spend on this slice: + T = P * TT = P * MT / (40/100) = P * ST / (60/100) + + Since we must do MW work in MT time or SW work in ST time, the amount + of work for this slice is: + MS = P * MW / (40/100) if marking + SS = P * SW / (60/100) if sweeping + + Amount of marking work for a marking slice: + MS = P * MW / (40/100) + MS = P * (caml_stat_heap_wsz * 250 / (100 + caml_percent_free) + + 2.5 * caml_incremental_roots_count) + Amount of sweeping work for a sweeping slice: + SS = P * SW / (60/100) + SS = P * caml_stat_heap_wsz * 5 / 3 + + This slice will either mark MS words or sweep SS words. + */ + + if (caml_major_slice_begin_hook != NULL) (*caml_major_slice_begin_hook) (); + CAML_INSTR_SETUP (tmr, "major"); + + p = (double) caml_allocated_words * 3.0 * (100 + caml_percent_free) + / caml_stat_heap_wsz / caml_percent_free / 2.0; + if (caml_dependent_size > 0){ + dp = (double) caml_dependent_allocated * (100 + caml_percent_free) + / caml_dependent_size / caml_percent_free; + }else{ + dp = 0.0; + } + if (p < dp) p = dp; + if (p < caml_extra_heap_resources) p = caml_extra_heap_resources; + if (p > 0.3) p = 0.3; + CAML_INSTR_INT ("major/work/extra#", + (uintnat) (caml_extra_heap_resources * 1000000)); + + caml_gc_message (0x40, "ordered work = %" + ARCH_INTNAT_PRINTF_FORMAT "d words\n", howmuch); + caml_gc_message (0x40, "allocated_words = %" + ARCH_INTNAT_PRINTF_FORMAT "u\n", + caml_allocated_words); + caml_gc_message (0x40, "extra_heap_resources = %" + ARCH_INTNAT_PRINTF_FORMAT "uu\n", + (uintnat) (caml_extra_heap_resources * 1000000)); + caml_gc_message (0x40, "raw work-to-do = %" + ARCH_INTNAT_PRINTF_FORMAT "du\n", + (intnat) (p * 1000000)); + + for (i = 0; i < caml_major_window; i++){ + caml_major_ring[i] += p / caml_major_window; + } + + if (caml_gc_clock >= 1.0){ + caml_gc_clock -= 1.0; + ++caml_major_ring_index; + if (caml_major_ring_index >= caml_major_window){ + caml_major_ring_index = 0; + } + } + if (howmuch == -1){ + /* auto-triggered GC slice: spend work credit on the current bucket, + then do the remaining work, if any */ + /* Note that the minor GC guarantees that the major slice is called in + automatic mode (with [howmuch] = -1) at least once per clock tick. + This means we never leave a non-empty bucket behind. */ + spend = fmin (caml_major_work_credit, + caml_major_ring[caml_major_ring_index]); + caml_major_work_credit -= spend; + filt_p = caml_major_ring[caml_major_ring_index] - spend; + caml_major_ring[caml_major_ring_index] = 0.0; + }else{ + /* forced GC slice: do work and add it to the credit */ + if (howmuch == 0){ + /* automatic setting: size of next bucket + we do not use the current bucket, as it may be empty */ + int i = caml_major_ring_index + 1; + if (i >= caml_major_window) i = 0; + filt_p = caml_major_ring[i]; + }else{ + /* manual setting */ + filt_p = (double) howmuch * 3.0 * (100 + caml_percent_free) + / caml_stat_heap_wsz / caml_percent_free / 2.0; + } + caml_major_work_credit += filt_p; + } + + p = filt_p; + + caml_gc_message (0x40, "filtered work-to-do = %" + ARCH_INTNAT_PRINTF_FORMAT "du\n", + (intnat) (p * 1000000)); + + if (caml_gc_phase == Phase_idle){ + if (caml_young_ptr == caml_young_alloc_end){ + /* We can only start a major GC cycle if the minor allocation arena + is empty, otherwise we'd have to treat it as a set of roots. */ + start_cycle (); + CAML_INSTR_TIME (tmr, "major/roots"); + } + p = 0; + goto finished; + } + + if (p < 0){ + p = 0; + goto finished; + } + + if (caml_gc_phase == Phase_mark || caml_gc_phase == Phase_clean){ + computed_work = (intnat) (p * ((double) caml_stat_heap_wsz * 250 + / (100 + caml_percent_free) + + caml_incremental_roots_count)); + }else{ + computed_work = (intnat) (p * caml_stat_heap_wsz * 5 / 3); + } + caml_gc_message (0x40, "computed work = %" + ARCH_INTNAT_PRINTF_FORMAT "d words\n", computed_work); + if (caml_gc_phase == Phase_mark){ + CAML_INSTR_INT ("major/work/mark#", computed_work); + mark_slice (computed_work); + CAML_INSTR_TIME (tmr, mark_slice_name[caml_gc_subphase]); + caml_gc_message (0x02, "!"); + }else if (caml_gc_phase == Phase_clean){ + clean_slice (computed_work); + caml_gc_message (0x02, "%%"); + }else{ + CAMLassert (caml_gc_phase == Phase_sweep); + CAML_INSTR_INT ("major/work/sweep#", computed_work); + sweep_slice (computed_work); + CAML_INSTR_TIME (tmr, "major/sweep"); + caml_gc_message (0x02, "$"); + } + + if (caml_gc_phase == Phase_idle){ + caml_compact_heap_maybe (); + CAML_INSTR_TIME (tmr, "major/check_and_compact"); + } + + finished: + caml_gc_message (0x40, "work-done = %" + ARCH_INTNAT_PRINTF_FORMAT "du\n", + (intnat) (p * 1000000)); + + /* if some of the work was not done, take it back from the credit + or spread it over the buckets. */ + p = filt_p - p; + spend = fmin (p, caml_major_work_credit); + caml_major_work_credit -= spend; + if (p > spend){ + p -= spend; + p /= caml_major_window; + for (i = 0; i < caml_major_window; i++) caml_major_ring[i] += p; + } + + caml_stat_major_words += caml_allocated_words; + caml_allocated_words = 0; + caml_dependent_allocated = 0; + caml_extra_heap_resources = 0.0; + if (caml_major_slice_end_hook != NULL) (*caml_major_slice_end_hook) (); +} + +/* This does not call [caml_compact_heap_maybe] because the estimates of + free and live memory are only valid for a cycle done incrementally. + Besides, this function itself is called by [caml_compact_heap_maybe]. +*/ +void caml_finish_major_cycle (void) +{ + if (caml_gc_phase == Phase_idle) start_cycle (); + while (caml_gc_phase == Phase_mark) mark_slice (LONG_MAX); + while (caml_gc_phase == Phase_clean) clean_slice (LONG_MAX); + CAMLassert (caml_gc_phase == Phase_sweep); + while (caml_gc_phase == Phase_sweep) sweep_slice (LONG_MAX); + CAMLassert (caml_gc_phase == Phase_idle); + caml_stat_major_words += caml_allocated_words; + caml_allocated_words = 0; +} + +/* Call this function to make sure [bsz] is greater than or equal + to both [Heap_chunk_min] and the current heap increment. +*/ +asize_t caml_clip_heap_chunk_wsz (asize_t wsz) +{ + asize_t result = wsz; + uintnat incr; + + /* Compute the heap increment as a word size. */ + if (caml_major_heap_increment > 1000){ + incr = caml_major_heap_increment; + }else{ + incr = caml_stat_heap_wsz / 100 * caml_major_heap_increment; + } + + if (result < incr){ + result = incr; + } + if (result < Heap_chunk_min){ + result = Heap_chunk_min; + } + return result; +} + +/* [heap_size] is a number of bytes */ +void caml_init_major_heap (asize_t heap_size) +{ + int i; + + /* DM */ heap_size = 65536; + + caml_stat_heap_wsz = caml_clip_heap_chunk_wsz (Wsize_bsize (heap_size)); + caml_stat_top_heap_wsz = caml_stat_heap_wsz; + CAMLassert (Bsize_wsize (caml_stat_heap_wsz) % Page_size == 0); + fprintf(stderr, "allocating %zd\n", Bsize_wsize (caml_stat_heap_wsz)); + fflush(stderr); + caml_heap_start = + (char *) caml_alloc_for_heap (Bsize_wsize (caml_stat_heap_wsz)); + if (caml_heap_start == NULL) + caml_fatal_error ("Fatal error: cannot allocate initial major heap.\n"); + Chunk_next (caml_heap_start) = NULL; + caml_stat_heap_wsz = Wsize_bsize (Chunk_size (caml_heap_start)); + caml_stat_heap_chunks = 1; + caml_stat_top_heap_wsz = caml_stat_heap_wsz; + + if (caml_page_table_add(In_heap, caml_heap_start, + caml_heap_start + Bsize_wsize (caml_stat_heap_wsz)) + != 0) { + caml_fatal_error ("Fatal error: cannot allocate " + "initial page table.\n"); + } + + caml_fl_init_merge (); + caml_make_free_blocks ((value *) caml_heap_start, + caml_stat_heap_wsz, 1, Caml_white); + caml_gc_phase = Phase_idle; + gray_vals_size = 2048; + gray_vals = (value *) caml_stat_alloc_noexc (gray_vals_size * sizeof (value)); + if (gray_vals == NULL) + caml_fatal_error ("Fatal error: not enough memory for the gray cache.\n"); + gray_vals_cur = gray_vals; + gray_vals_end = gray_vals + gray_vals_size; + heap_is_pure = 1; + caml_allocated_words = 0; + caml_extra_heap_resources = 0.0; + for (i = 0; i < Max_major_window; i++) caml_major_ring[i] = 0.0; +} + +void caml_set_major_window (int w){ + uintnat total = 0; + int i; + if (w == caml_major_window) return; + CAMLassert (w <= Max_major_window); + /* Collect the current work-to-do from the buckets. */ + for (i = 0; i < caml_major_window; i++){ + total += caml_major_ring[i]; + } + /* Redistribute to the new buckets. */ + for (i = 0; i < w; i++){ + caml_major_ring[i] = total / w; + } + caml_major_window = w; +} + +void caml_finalise_heap (void) +{ + /* Finishing major cycle (all values become white) */ + caml_empty_minor_heap (); + caml_finish_major_cycle (); + CAMLassert (caml_gc_phase == Phase_idle); + + /* Finalising all values (by means of forced sweeping) */ + caml_fl_init_merge (); + caml_gc_phase = Phase_sweep; + chunk = caml_heap_start; + caml_gc_sweep_hp = chunk; + limit = chunk + Chunk_size (chunk); + while (caml_gc_phase == Phase_sweep) + sweep_slice (LONG_MAX); +} diff --git a/test/monniaux/ocaml/byterun/make.log b/test/monniaux/ocaml/byterun/make.log new file mode 100644 index 00000000..0aa95a93 --- /dev/null +++ b/test/monniaux/ocaml/byterun/make.log @@ -0,0 +1,17 @@ +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE signals.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE signals_byt.c +rm -f libcamlrun.a && ar rc libcamlrun.a interp.o misc.o stacks.o fix_code.o startup_aux.o startup.o freelist.o major_gc.o minor_gc.o memory.o alloc.o roots.o globroots.o fail.o signals.o signals_byt.o printexc.o backtrace_prim.o backtrace.o compare.o ints.o floats.o str.o array.o io.o extern.o intern.o hash.o sys.o meta.o parsing.o gc_ctrl.o md5.o obj.o lexing.o callback.o debugger.o weak.o compact.o finalise.o custom.o dynlink.o spacetime.o afl.o unix.o bigarray.o main.o && ranlib libcamlrun.a +k1-mbr-gcc -c -O3 -Wall -g -DDEBUG -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE -o signals.d.o signals.c +k1-mbr-gcc -c -O3 -Wall -g -DDEBUG -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE -o signals_byt.d.o signals_byt.c +rm -f libcamlrund.a && ar rc libcamlrund.a interp.d.o misc.d.o stacks.d.o fix_code.d.o startup_aux.d.o startup.d.o freelist.d.o major_gc.d.o minor_gc.d.o memory.d.o alloc.d.o roots.d.o globroots.d.o fail.d.o signals.d.o signals_byt.d.o printexc.d.o backtrace_prim.d.o backtrace.d.o compare.d.o ints.d.o floats.d.o str.d.o array.d.o io.d.o extern.d.o intern.d.o hash.d.o sys.d.o meta.d.o parsing.d.o gc_ctrl.d.o md5.d.o obj.d.o lexing.d.o callback.d.o debugger.d.o weak.d.o compact.d.o finalise.d.o custom.d.o dynlink.d.o spacetime.d.o afl.d.o unix.d.o bigarray.d.o main.d.o instrtrace.d.o && ranlib libcamlrund.a +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE -o signals.i.o signals.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE -o signals_byt.i.o signals_byt.c +rm -f libcamlruni.a && ar rc libcamlruni.a interp.i.o misc.i.o stacks.i.o fix_code.i.o startup_aux.i.o startup.i.o freelist.i.o major_gc.i.o minor_gc.i.o memory.i.o alloc.i.o roots.i.o globroots.i.o fail.i.o signals.i.o signals_byt.i.o printexc.i.o backtrace_prim.i.o backtrace.i.o compare.i.o ints.i.o floats.i.o str.i.o array.i.o io.i.o extern.i.o intern.i.o hash.i.o sys.i.o meta.i.o parsing.i.o gc_ctrl.i.o md5.i.o obj.i.o lexing.i.o callback.i.o debugger.i.o weak.i.o compact.i.o finalise.i.o custom.i.o dynlink.i.o spacetime.i.o afl.i.o unix.i.o bigarray.i.o main.i.o && ranlib libcamlruni.a +k1-mbr-gcc -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE -o ocamlrun prims.o libcamlrun.a -lm +/opt/Kalray/usr/local/k1rdtools/bin/../lib/gcc/k1-mbr/4.9.4/../../../../k1-mbr/lib/libgloss.a(times.o): In function `__gloss_times': +/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/newlib/k1_build_mbr/k1-mbr/libgloss/k1-mbr/../../../../libgloss/k1-mbr/times.c:18: undefined reference to `__bsp_frequency' +/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/newlib/k1_build_mbr/k1-mbr/libgloss/k1-mbr/../../../../libgloss/k1-mbr/times.c:18: undefined reference to `__bsp_frequency' +/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/newlib/k1_build_mbr/k1-mbr/libgloss/k1-mbr/../../../../libgloss/k1-mbr/times.c:18: undefined reference to `__bsp_frequency' +collect2: error: ld returned 1 exit status +Makefile:187: recipe for target 'ocamlrun' failed +make: *** [ocamlrun] Error 1 diff --git a/test/monniaux/ocaml/byterun/make2.log b/test/monniaux/ocaml/byterun/make2.log new file mode 100644 index 00000000..719c4b6d --- /dev/null +++ b/test/monniaux/ocaml/byterun/make2.log @@ -0,0 +1,321 @@ +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE interp.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE misc.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE stacks.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE fix_code.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE startup_aux.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE startup.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE freelist.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE major_gc.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE minor_gc.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE memory.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE alloc.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE roots.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE globroots.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE fail.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE signals.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE signals_byt.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE printexc.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE backtrace_prim.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE backtrace.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE compare.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE ints.c +ints.c: In function ‘caml_int64_bswap’: +ints.c:508:5: warning: implicit declaration of function ‘INT64_LITERAL’ [-Wimplicit-function-declaration] + (((x & INT64_LITERAL(0x00000000000000FFU)) << 56) | + ^ +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE floats.c +floats.c: In function ‘caml_Store_double_val’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c: In function ‘caml_copy_double’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_Store_double_array_field’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_float_of_string’: +floats.c:64:17: warning: ‘buffer.v[1]’ may be used uninitialized in this function [-Wmaybe-uninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c:64:17: warning: ‘buffer.v[1]’ may be used uninitialized in this function [-Wmaybe-uninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_float_of_int’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_neg_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_abs_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_add_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_sub_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_mul_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_div_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_exp_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_floor_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_fmod_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_frexp_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_ldexp_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_log_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_log10_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_modf_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_sqrt_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_power_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_sin_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_sinh_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_cos_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_cosh_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_tan_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_tanh_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_asin_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_acos_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_atan_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_atan2_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_ceil_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_hypot_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_expm1_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_log1p_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +floats.c: In function ‘caml_copysign_float’: +floats.c:64:17: warning: ‘buffer.v[1]’ is used uninitialized in this function [-Wuninitialized] + Field(val, 1) = buffer.v[1]; + ^ +floats.c:59:35: note: ‘buffer.v[1]’ was declared here + union { value v[2]; double d; } buffer; + ^ +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE str.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE array.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE io.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE extern.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE intern.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE hash.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE sys.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE meta.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE parsing.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE gc_ctrl.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE md5.c +md5.c: In function ‘caml_MD5Final’: +md5.c:218:5: warning: dereferencing type-punned pointer will break strict-aliasing rules [-Wstrict-aliasing] + ((uint32_t *) ctx->in)[14] = ctx->bits[0]; + ^ +md5.c:219:5: warning: dereferencing type-punned pointer will break strict-aliasing rules [-Wstrict-aliasing] + ((uint32_t *) ctx->in)[15] = ctx->bits[1]; + ^ +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE obj.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE lexing.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE callback.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE debugger.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE weak.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE compact.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE finalise.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE custom.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE dynlink.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE spacetime.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE afl.c +k1-mbr-gcc -c -O3 -Wall -g -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE unix.c +unix.c:48:21: fatal error: sys/dir.h: No such file or directory + #include + ^ +compilation terminated. +Makefile:211: recipe for target 'unix.o' failed +make: *** [unix.o] Error 1 diff --git a/test/monniaux/ocaml/byterun/md5.c b/test/monniaux/ocaml/byterun/md5.c new file mode 100644 index 00000000..2e128010 --- /dev/null +++ b/test/monniaux/ocaml/byterun/md5.c @@ -0,0 +1,325 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include +#include "caml/alloc.h" +#include "caml/fail.h" +#include "caml/md5.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/io.h" +#include "caml/reverse.h" + +/* MD5 message digest */ + +CAMLprim value caml_md5_string(value str, value ofs, value len) +{ + struct MD5Context ctx; + value res; + caml_MD5Init(&ctx); + caml_MD5Update(&ctx, &Byte_u(str, Long_val(ofs)), Long_val(len)); + res = caml_alloc_string(16); + caml_MD5Final(&Byte_u(res, 0), &ctx); + return res; +} + +CAMLexport value caml_md5_channel(struct channel *chan, intnat toread) +{ + CAMLparam0(); + struct MD5Context ctx; + value res; + intnat read; + char buffer[4096]; + + Lock(chan); + caml_MD5Init(&ctx); + if (toread < 0){ + while (1){ + read = caml_getblock (chan, buffer, sizeof(buffer)); + if (read == 0) break; + caml_MD5Update (&ctx, (unsigned char *) buffer, read); + } + }else{ + while (toread > 0) { + read = caml_getblock(chan, buffer, + toread > sizeof(buffer) ? sizeof(buffer) : toread); + if (read == 0) caml_raise_end_of_file(); + caml_MD5Update(&ctx, (unsigned char *) buffer, read); + toread -= read; + } + } + res = caml_alloc_string(16); + caml_MD5Final(&Byte_u(res, 0), &ctx); + Unlock(chan); + CAMLreturn (res); +} + +CAMLprim value caml_md5_chan(value vchan, value len) +{ + CAMLparam2 (vchan, len); + CAMLreturn (caml_md5_channel(Channel(vchan), Long_val(len))); +} + +CAMLexport void caml_md5_block(unsigned char digest[16], + void * data, uintnat len) +{ + struct MD5Context ctx; + caml_MD5Init(&ctx); + caml_MD5Update(&ctx, data, len); + caml_MD5Final(digest, &ctx); +} + +/* + * This code implements the MD5 message-digest algorithm. + * The algorithm is due to Ron Rivest. This code was + * written by Colin Plumb in 1993, no copyright is claimed. + * This code is in the public domain; do with it what you wish. + * + * Equivalent code is available from RSA Data Security, Inc. + * This code has been tested against that, and is equivalent, + * except that you don't need to include two pages of legalese + * with every copy. + * + * To compute the message digest of a chunk of bytes, declare an + * MD5Context structure, pass it to caml_MD5Init, call caml_MD5Update as + * needed on buffers full of bytes, and then call caml_MD5Final, which + * will fill a supplied 16-byte array with the digest. + */ + +#ifndef ARCH_BIG_ENDIAN +#define byteReverse(buf, len) /* Nothing */ +#else +static void byteReverse(unsigned char * buf, unsigned longs) +{ + uint32_t t; + do { + t = (uint32_t) ((unsigned) buf[3] << 8 | buf[2]) << 16 | + ((unsigned) buf[1] << 8 | buf[0]); + *(uint32_t *) buf = t; + buf += 4; + } while (--longs); +} +#endif + +/* + * Start MD5 accumulation. Set bit count to 0 and buffer to mysterious + * initialization constants. + */ +CAMLexport void caml_MD5Init(struct MD5Context *ctx) +{ + ctx->buf[0] = 0x67452301; + ctx->buf[1] = 0xefcdab89; + ctx->buf[2] = 0x98badcfe; + ctx->buf[3] = 0x10325476; + + ctx->bits[0] = 0; + ctx->bits[1] = 0; +} + +/* + * Update context to reflect the concatenation of another buffer full + * of bytes. + */ +CAMLexport void caml_MD5Update(struct MD5Context *ctx, unsigned char *buf, + uintnat len) +{ + uint32_t t; + + /* Update bitcount */ + + t = ctx->bits[0]; + if ((ctx->bits[0] = t + ((uint32_t) len << 3)) < t) + ctx->bits[1]++; /* Carry from low to high */ + ctx->bits[1] += len >> 29; + + t = (t >> 3) & 0x3f; /* Bytes already in shsInfo->data */ + + /* Handle any leading odd-sized chunks */ + + if (t) { + unsigned char *p = (unsigned char *) ctx->in + t; + + t = 64 - t; + if (len < t) { + memcpy(p, buf, len); + return; + } + memcpy(p, buf, t); + byteReverse(ctx->in, 16); + caml_MD5Transform(ctx->buf, (uint32_t *) ctx->in); + buf += t; + len -= t; + } + /* Process data in 64-byte chunks */ + + while (len >= 64) { + memcpy(ctx->in, buf, 64); + byteReverse(ctx->in, 16); + caml_MD5Transform(ctx->buf, (uint32_t *) ctx->in); + buf += 64; + len -= 64; + } + + /* Handle any remaining bytes of data. */ + + memcpy(ctx->in, buf, len); +} + +/* + * Final wrapup - pad to 64-byte boundary with the bit pattern + * 1 0* (64-bit count of bits processed, MSB-first) + */ +CAMLexport void caml_MD5Final(unsigned char *digest, struct MD5Context *ctx) +{ + unsigned count; + unsigned char *p; + + /* Compute number of bytes mod 64 */ + count = (ctx->bits[0] >> 3) & 0x3F; + + /* Set the first char of padding to 0x80. This is safe since there is + always at least one byte free */ + p = ctx->in + count; + *p++ = 0x80; + + /* Bytes of padding needed to make 64 bytes */ + count = 64 - 1 - count; + + /* Pad out to 56 mod 64 */ + if (count < 8) { + /* Two lots of padding: Pad the first block to 64 bytes */ + memset(p, 0, count); + byteReverse(ctx->in, 16); + caml_MD5Transform(ctx->buf, (uint32_t *) ctx->in); + + /* Now fill the next block with 56 bytes */ + memset(ctx->in, 0, 56); + } else { + /* Pad block to 56 bytes */ + memset(p, 0, count - 8); + } + byteReverse(ctx->in, 14); + + /* Append length in bits and transform */ + ((uint32_t *) ctx->in)[14] = ctx->bits[0]; + ((uint32_t *) ctx->in)[15] = ctx->bits[1]; + + caml_MD5Transform(ctx->buf, (uint32_t *) ctx->in); + byteReverse((unsigned char *) ctx->buf, 4); + memcpy(digest, ctx->buf, 16); + memset(ctx, 0, sizeof(*ctx)); /* In case it's sensitive */ +} + +/* The four core functions - F1 is optimized somewhat */ + +/* #define F1(x, y, z) (x & y | ~x & z) */ +#define F1(x, y, z) (z ^ (x & (y ^ z))) +#define F2(x, y, z) F1(z, x, y) +#define F3(x, y, z) (x ^ y ^ z) +#define F4(x, y, z) (y ^ (x | ~z)) + +/* This is the central step in the MD5 algorithm. */ +#define MD5STEP(f, w, x, y, z, data, s) \ + ( w += f(x, y, z) + data, w = w<>(32-s), w += x ) + +/* + * The core of the MD5 algorithm, this alters an existing MD5 hash to + * reflect the addition of 16 longwords of new data. caml_MD5Update blocks + * the data and converts bytes into longwords for this routine. + */ +CAMLexport void caml_MD5Transform(uint32_t *buf, uint32_t *in) +{ + register uint32_t a, b, c, d; + + a = buf[0]; + b = buf[1]; + c = buf[2]; + d = buf[3]; + + MD5STEP(F1, a, b, c, d, in[0] + 0xd76aa478, 7); + MD5STEP(F1, d, a, b, c, in[1] + 0xe8c7b756, 12); + MD5STEP(F1, c, d, a, b, in[2] + 0x242070db, 17); + MD5STEP(F1, b, c, d, a, in[3] + 0xc1bdceee, 22); + MD5STEP(F1, a, b, c, d, in[4] + 0xf57c0faf, 7); + MD5STEP(F1, d, a, b, c, in[5] + 0x4787c62a, 12); + MD5STEP(F1, c, d, a, b, in[6] + 0xa8304613, 17); + MD5STEP(F1, b, c, d, a, in[7] + 0xfd469501, 22); + MD5STEP(F1, a, b, c, d, in[8] + 0x698098d8, 7); + MD5STEP(F1, d, a, b, c, in[9] + 0x8b44f7af, 12); + MD5STEP(F1, c, d, a, b, in[10] + 0xffff5bb1, 17); + MD5STEP(F1, b, c, d, a, in[11] + 0x895cd7be, 22); + MD5STEP(F1, a, b, c, d, in[12] + 0x6b901122, 7); + MD5STEP(F1, d, a, b, c, in[13] + 0xfd987193, 12); + MD5STEP(F1, c, d, a, b, in[14] + 0xa679438e, 17); + MD5STEP(F1, b, c, d, a, in[15] + 0x49b40821, 22); + + MD5STEP(F2, a, b, c, d, in[1] + 0xf61e2562, 5); + MD5STEP(F2, d, a, b, c, in[6] + 0xc040b340, 9); + MD5STEP(F2, c, d, a, b, in[11] + 0x265e5a51, 14); + MD5STEP(F2, b, c, d, a, in[0] + 0xe9b6c7aa, 20); + MD5STEP(F2, a, b, c, d, in[5] + 0xd62f105d, 5); + MD5STEP(F2, d, a, b, c, in[10] + 0x02441453, 9); + MD5STEP(F2, c, d, a, b, in[15] + 0xd8a1e681, 14); + MD5STEP(F2, b, c, d, a, in[4] + 0xe7d3fbc8, 20); + MD5STEP(F2, a, b, c, d, in[9] + 0x21e1cde6, 5); + MD5STEP(F2, d, a, b, c, in[14] + 0xc33707d6, 9); + MD5STEP(F2, c, d, a, b, in[3] + 0xf4d50d87, 14); + MD5STEP(F2, b, c, d, a, in[8] + 0x455a14ed, 20); + MD5STEP(F2, a, b, c, d, in[13] + 0xa9e3e905, 5); + MD5STEP(F2, d, a, b, c, in[2] + 0xfcefa3f8, 9); + MD5STEP(F2, c, d, a, b, in[7] + 0x676f02d9, 14); + MD5STEP(F2, b, c, d, a, in[12] + 0x8d2a4c8a, 20); + + MD5STEP(F3, a, b, c, d, in[5] + 0xfffa3942, 4); + MD5STEP(F3, d, a, b, c, in[8] + 0x8771f681, 11); + MD5STEP(F3, c, d, a, b, in[11] + 0x6d9d6122, 16); + MD5STEP(F3, b, c, d, a, in[14] + 0xfde5380c, 23); + MD5STEP(F3, a, b, c, d, in[1] + 0xa4beea44, 4); + MD5STEP(F3, d, a, b, c, in[4] + 0x4bdecfa9, 11); + MD5STEP(F3, c, d, a, b, in[7] + 0xf6bb4b60, 16); + MD5STEP(F3, b, c, d, a, in[10] + 0xbebfbc70, 23); + MD5STEP(F3, a, b, c, d, in[13] + 0x289b7ec6, 4); + MD5STEP(F3, d, a, b, c, in[0] + 0xeaa127fa, 11); + MD5STEP(F3, c, d, a, b, in[3] + 0xd4ef3085, 16); + MD5STEP(F3, b, c, d, a, in[6] + 0x04881d05, 23); + MD5STEP(F3, a, b, c, d, in[9] + 0xd9d4d039, 4); + MD5STEP(F3, d, a, b, c, in[12] + 0xe6db99e5, 11); + MD5STEP(F3, c, d, a, b, in[15] + 0x1fa27cf8, 16); + MD5STEP(F3, b, c, d, a, in[2] + 0xc4ac5665, 23); + + MD5STEP(F4, a, b, c, d, in[0] + 0xf4292244, 6); + MD5STEP(F4, d, a, b, c, in[7] + 0x432aff97, 10); + MD5STEP(F4, c, d, a, b, in[14] + 0xab9423a7, 15); + MD5STEP(F4, b, c, d, a, in[5] + 0xfc93a039, 21); + MD5STEP(F4, a, b, c, d, in[12] + 0x655b59c3, 6); + MD5STEP(F4, d, a, b, c, in[3] + 0x8f0ccc92, 10); + MD5STEP(F4, c, d, a, b, in[10] + 0xffeff47d, 15); + MD5STEP(F4, b, c, d, a, in[1] + 0x85845dd1, 21); + MD5STEP(F4, a, b, c, d, in[8] + 0x6fa87e4f, 6); + MD5STEP(F4, d, a, b, c, in[15] + 0xfe2ce6e0, 10); + MD5STEP(F4, c, d, a, b, in[6] + 0xa3014314, 15); + MD5STEP(F4, b, c, d, a, in[13] + 0x4e0811a1, 21); + MD5STEP(F4, a, b, c, d, in[4] + 0xf7537e82, 6); + MD5STEP(F4, d, a, b, c, in[11] + 0xbd3af235, 10); + MD5STEP(F4, c, d, a, b, in[2] + 0x2ad7d2bb, 15); + MD5STEP(F4, b, c, d, a, in[9] + 0xeb86d391, 21); + + buf[0] += a; + buf[1] += b; + buf[2] += c; + buf[3] += d; +} diff --git a/test/monniaux/ocaml/byterun/memory.c b/test/monniaux/ocaml/byterun/memory.c new file mode 100644 index 00000000..f92b23c4 --- /dev/null +++ b/test/monniaux/ocaml/byterun/memory.c @@ -0,0 +1,1016 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include +#include +#include +#include +#include "caml/address_class.h" +#include "caml/config.h" +#include "caml/fail.h" +#include "caml/freelist.h" +#include "caml/gc.h" +#include "caml/gc_ctrl.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/major_gc.h" +#include "caml/minor_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/signals.h" + +int caml_huge_fallback_count = 0; +/* Number of times that mmapping big pages fails and we fell back to small + pages. This counter is available to the program through + [Gc.huge_fallback_count]. +*/ + +uintnat caml_use_huge_pages = 0; +/* True iff the program allocates heap chunks by mmapping huge pages. + This is set when parsing [OCAMLRUNPARAM] and must stay constant + after that. +*/ + +extern uintnat caml_percent_free; /* major_gc.c */ + +/* Page table management */ + +#define Page(p) ((uintnat) (p) >> Page_log) +#define Page_mask ((uintnat) -1 << Page_log) + +#ifdef ARCH_SIXTYFOUR + +/* 64-bit implementation: + The page table is represented sparsely as a hash table + with linear probing */ + +struct page_table { + mlsize_t size; /* size == 1 << (wordsize - shift) */ + int shift; + mlsize_t mask; /* mask == size - 1 */ + mlsize_t occupancy; + uintnat * entries; /* [size] */ +}; + +static struct page_table caml_page_table; + +/* Page table entries are the logical 'or' of + - the key: address of a page (low Page_log bits = 0) + - the data: a 8-bit integer */ + +#define Page_entry_matches(entry,addr) \ + ((((entry) ^ (addr)) & Page_mask) == 0) + +/* Multiplicative Fibonacci hashing + (Knuth, TAOCP vol 3, section 6.4, page 518). + HASH_FACTOR is (sqrt(5) - 1) / 2 * 2^wordsize. */ +#ifdef ARCH_SIXTYFOUR +#define HASH_FACTOR 11400714819323198486UL +#else +#define HASH_FACTOR 2654435769UL +#endif +#define Hash(v) (((v) * HASH_FACTOR) >> caml_page_table.shift) + +int caml_page_table_lookup(void * addr) +{ + uintnat h, e; + + h = Hash(Page(addr)); + /* The first hit is almost always successful, so optimize for this case */ + e = caml_page_table.entries[h]; + if (Page_entry_matches(e, (uintnat)addr)) return e & 0xFF; + while(1) { + if (e == 0) return 0; + h = (h + 1) & caml_page_table.mask; + e = caml_page_table.entries[h]; + if (Page_entry_matches(e, (uintnat)addr)) return e & 0xFF; + } +} + +int caml_page_table_initialize(mlsize_t bytesize) +{ + uintnat pagesize = Page(bytesize); + + caml_page_table.size = 1; + caml_page_table.shift = 8 * sizeof(uintnat); + /* Aim for initial load factor between 1/4 and 1/2 */ + while (caml_page_table.size < 2 * pagesize) { + caml_page_table.size <<= 1; + caml_page_table.shift -= 1; + } + caml_page_table.mask = caml_page_table.size - 1; + caml_page_table.occupancy = 0; + caml_page_table.entries = + caml_stat_calloc_noexc(caml_page_table.size, sizeof(uintnat)); + if (caml_page_table.entries == NULL) + return -1; + else + return 0; +} + +static int caml_page_table_resize(void) +{ + struct page_table old = caml_page_table; + uintnat * new_entries; + uintnat i, h; + + caml_gc_message (0x08, "Growing page table to %" + ARCH_INTNAT_PRINTF_FORMAT "u entries\n", + caml_page_table.size); + + new_entries = caml_stat_calloc_noexc(2 * old.size, sizeof(uintnat)); + if (new_entries == NULL) { + caml_gc_message (0x08, "No room for growing page table\n"); + return -1; + } + + caml_page_table.size = 2 * old.size; + caml_page_table.shift = old.shift - 1; + caml_page_table.mask = caml_page_table.size - 1; + caml_page_table.occupancy = old.occupancy; + caml_page_table.entries = new_entries; + + for (i = 0; i < old.size; i++) { + uintnat e = old.entries[i]; + if (e == 0) continue; + h = Hash(Page(e)); + while (caml_page_table.entries[h] != 0) + h = (h + 1) & caml_page_table.mask; + caml_page_table.entries[h] = e; + } + + caml_stat_free(old.entries); + return 0; +} + +static int caml_page_table_modify(uintnat page, int toclear, int toset) +{ + uintnat h; + + CAMLassert ((page & ~Page_mask) == 0); + + /* Resize to keep load factor below 1/2 */ + if (caml_page_table.occupancy * 2 >= caml_page_table.size) { + if (caml_page_table_resize() != 0) return -1; + } + h = Hash(Page(page)); + while (1) { + if (caml_page_table.entries[h] == 0) { + caml_page_table.entries[h] = page | toset; + caml_page_table.occupancy++; + break; + } + if (Page_entry_matches(caml_page_table.entries[h], page)) { + caml_page_table.entries[h] = + (caml_page_table.entries[h] & ~toclear) | toset; + break; + } + h = (h + 1) & caml_page_table.mask; + } + return 0; +} + +#else + +/* 32-bit implementation: + The page table is represented as a 2-level array of unsigned char */ + +CAMLexport unsigned char * caml_page_table[Pagetable1_size]; +static unsigned char caml_page_table_empty[Pagetable2_size] = { 0, }; + +int caml_page_table_initialize(mlsize_t bytesize) +{ + int i; + for (i = 0; i < Pagetable1_size; i++) + caml_page_table[i] = caml_page_table_empty; + return 0; +} + +static int caml_page_table_modify(uintnat page, int toclear, int toset) +{ + uintnat i = Pagetable_index1(page); + uintnat j = Pagetable_index2(page); + + if (caml_page_table[i] == caml_page_table_empty) { + unsigned char * new_tbl = caml_stat_calloc_noexc(Pagetable2_size, 1); + if (new_tbl == 0) return -1; + caml_page_table[i] = new_tbl; + } + caml_page_table[i][j] = (caml_page_table[i][j] & ~toclear) | toset; + return 0; +} + +#endif + +int caml_page_table_add(int kind, void * start, void * end) +{ + uintnat pstart = (uintnat) start & Page_mask; + uintnat pend = ((uintnat) end - 1) & Page_mask; + uintnat p; + + for (p = pstart; p <= pend; p += Page_size) + if (caml_page_table_modify(p, 0, kind) != 0) return -1; + return 0; +} + +int caml_page_table_remove(int kind, void * start, void * end) +{ + uintnat pstart = (uintnat) start & Page_mask; + uintnat pend = ((uintnat) end - 1) & Page_mask; + uintnat p; + + for (p = pstart; p <= pend; p += Page_size) + if (caml_page_table_modify(p, kind, 0) != 0) return -1; + return 0; +} + + +/* Initialize the [alloc_for_heap] system. + This function must be called exactly once, and it must be called + before the first call to [alloc_for_heap]. + It returns 0 on success and -1 on failure. +*/ +int caml_init_alloc_for_heap (void) +{ + return 0; +} + +/* Allocate a block of the requested size, to be passed to + [caml_add_to_heap] later. + [request] will be rounded up to some implementation-dependent size. + The caller must use [Chunk_size] on the result to recover the actual + size. + Return NULL if the request cannot be satisfied. The returned pointer + is a hp, but the header (and the contents) must be initialized by the + caller. +*/ +char *caml_alloc_for_heap (asize_t request) +{ + if (caml_use_huge_pages){ +#ifdef HAS_HUGE_PAGES + uintnat size = Round_mmap_size (sizeof (heap_chunk_head) + request); + void *block; + char *mem; + block = mmap (NULL, size, PROT_READ | PROT_WRITE, + MAP_PRIVATE | MAP_ANONYMOUS | MAP_HUGETLB, -1, 0); + if (block == MAP_FAILED) return NULL; + mem = (char *) block + sizeof (heap_chunk_head); + Chunk_size (mem) = size - sizeof (heap_chunk_head); + Chunk_block (mem) = block; + return mem; +#else + return NULL; +#endif + }else{ + char *mem; + void *block; + + request = ((request + Page_size - 1) >> Page_log) << Page_log; + mem = caml_stat_alloc_aligned_noexc (request + sizeof (heap_chunk_head), + sizeof (heap_chunk_head), &block); + if (mem == NULL) return NULL; + mem += sizeof (heap_chunk_head); + Chunk_size (mem) = request; + Chunk_block (mem) = block; + return mem; + } +} + +/* Use this function if a block allocated with [caml_alloc_for_heap] is + not actually going to be added to the heap. The caller is responsible + for freeing it. */ +void caml_disown_for_heap (char* mem) +{ + /* Currently a no-op. */ + (void)mem; /* can CAMLunused_{start,end} be used here? */ +} + +/* Use this function to free a block allocated with [caml_alloc_for_heap] + if you don't add it with [caml_add_to_heap]. +*/ +void caml_free_for_heap (char *mem) +{ + if (caml_use_huge_pages){ +#ifdef HAS_HUGE_PAGES + munmap (Chunk_block (mem), Chunk_size (mem) + sizeof (heap_chunk_head)); +#else + CAMLassert (0); +#endif + }else{ + caml_stat_free (Chunk_block (mem)); + } +} + +/* Take a chunk of memory as argument, which must be the result of a + call to [caml_alloc_for_heap], and insert it into the heap chaining. + The contents of the chunk must be a sequence of valid blocks and + fragments: no space between blocks and no trailing garbage. If + some blocks are blue, they must be added to the free list by the + caller. All other blocks must have the color [caml_allocation_color(m)]. + The caller must update [caml_allocated_words] if applicable. + Return value: 0 if no error; -1 in case of error. + + See also: caml_compact_heap, which duplicates most of this function. +*/ +int caml_add_to_heap (char *m) +{ +#ifdef DEBUG + /* Should check the contents of the block. */ +#endif /* DEBUG */ + + caml_gc_message (0x04, "Growing heap to %" + ARCH_INTNAT_PRINTF_FORMAT "uk bytes\n", + (Bsize_wsize (caml_stat_heap_wsz) + Chunk_size (m)) / 1024); + + /* Register block in page table */ + if (caml_page_table_add(In_heap, m, m + Chunk_size(m)) != 0) + return -1; + + /* Chain this heap chunk. */ + { + char **last = &caml_heap_start; + char *cur = *last; + + while (cur != NULL && cur < m){ + last = &(Chunk_next (cur)); + cur = *last; + } + Chunk_next (m) = cur; + *last = m; + + ++ caml_stat_heap_chunks; + } + + caml_stat_heap_wsz += Wsize_bsize (Chunk_size (m)); + if (caml_stat_heap_wsz > caml_stat_top_heap_wsz){ + caml_stat_top_heap_wsz = caml_stat_heap_wsz; + } + return 0; +} + +/* Allocate more memory from malloc for the heap. + Return a blue block of at least the requested size. + The blue block is chained to a sequence of blue blocks (through their + field 0); the last block of the chain is pointed by field 1 of the + first. There may be a fragment after the last block. + The caller must insert the blocks into the free list. + [request] is a number of words and must be less than or equal + to [Max_wosize]. + Return NULL when out of memory. +*/ +static value *expand_heap (mlsize_t request) +{ + /* these point to headers, but we do arithmetic on them, hence [value *]. */ + value *mem, *hp, *prev; + asize_t over_request, malloc_request, remain; + + CAMLassert (request <= Max_wosize); + over_request = request + request / 100 * caml_percent_free; + malloc_request = caml_clip_heap_chunk_wsz (over_request); + mem = (value *) caml_alloc_for_heap (Bsize_wsize (malloc_request)); + if (mem == NULL){ + caml_gc_message (0x04, "No room for growing heap\n"); + return NULL; + } + remain = Wsize_bsize (Chunk_size (mem)); + prev = hp = mem; + /* FIXME find a way to do this with a call to caml_make_free_blocks */ + while (Wosize_whsize (remain) > Max_wosize){ + Hd_hp (hp) = Make_header (Max_wosize, 0, Caml_blue); +#ifdef DEBUG + caml_set_fields (Val_hp (hp), 0, Debug_free_major); +#endif + hp += Whsize_wosize (Max_wosize); + remain -= Whsize_wosize (Max_wosize); + Field (Val_hp (mem), 1) = Field (Val_hp (prev), 0) = Val_hp (hp); + prev = hp; + } + if (remain > 1){ + Hd_hp (hp) = Make_header (Wosize_whsize (remain), 0, Caml_blue); +#ifdef DEBUG + caml_set_fields (Val_hp (hp), 0, Debug_free_major); +#endif + Field (Val_hp (mem), 1) = Field (Val_hp (prev), 0) = Val_hp (hp); + Field (Val_hp (hp), 0) = (value) NULL; + }else{ + Field (Val_hp (prev), 0) = (value) NULL; + if (remain == 1) { + Hd_hp (hp) = Make_header_allocated_here (0, 0, Caml_white); + } + } + CAMLassert (Wosize_hp (mem) >= request); + if (caml_add_to_heap ((char *) mem) != 0){ + caml_free_for_heap ((char *) mem); + return NULL; + } + return Op_hp (mem); +} + +/* Remove the heap chunk [chunk] from the heap and give the memory back + to [free]. +*/ +void caml_shrink_heap (char *chunk) +{ + char **cp; + + /* Never deallocate the first chunk, because caml_heap_start is both the + first block and the base address for page numbers, and we don't + want to shift the page table, it's too messy (see above). + It will never happen anyway, because of the way compaction works. + (see compact.c) + XXX FIXME this has become false with the fix to PR#5389 (see compact.c) + */ + if (chunk == caml_heap_start) return; + + caml_stat_heap_wsz -= Wsize_bsize (Chunk_size (chunk)); + caml_gc_message (0x04, "Shrinking heap to %" + ARCH_INTNAT_PRINTF_FORMAT "uk words\n", + caml_stat_heap_wsz / 1024); + +#ifdef DEBUG + { + mlsize_t i; + for (i = 0; i < Wsize_bsize (Chunk_size (chunk)); i++){ + ((value *) chunk) [i] = Debug_free_shrink; + } + } +#endif + + -- caml_stat_heap_chunks; + + /* Remove [chunk] from the list of chunks. */ + cp = &caml_heap_start; + while (*cp != chunk) cp = &(Chunk_next (*cp)); + *cp = Chunk_next (chunk); + + /* Remove the pages of [chunk] from the page table. */ + caml_page_table_remove(In_heap, chunk, chunk + Chunk_size (chunk)); + + /* Free the [malloc] block that contains [chunk]. */ + caml_free_for_heap (chunk); +} + +color_t caml_allocation_color (void *hp) +{ + if (caml_gc_phase == Phase_mark || caml_gc_phase == Phase_clean + || (caml_gc_phase == Phase_sweep && (addr)hp >= (addr)caml_gc_sweep_hp)){ + return Caml_black; + }else{ + CAMLassert (caml_gc_phase == Phase_idle + || (caml_gc_phase == Phase_sweep + && (addr)hp < (addr)caml_gc_sweep_hp)); + return Caml_white; + } +} + +static inline value caml_alloc_shr_aux (mlsize_t wosize, tag_t tag, + int raise_oom, uintnat profinfo) +{ + header_t *hp; + value *new_block; + + if (wosize > Max_wosize) { + if (raise_oom) + caml_raise_out_of_memory (); + else + return 0; + } + hp = caml_fl_allocate (wosize); + if (hp == NULL){ + new_block = expand_heap (wosize); + if (new_block == NULL) { + if (!raise_oom) + return 0; + else if (caml_in_minor_collection) + caml_fatal_error ("Fatal error: out of memory.\n"); + else + caml_raise_out_of_memory (); + } + caml_fl_add_blocks ((value) new_block); + hp = caml_fl_allocate (wosize); + } + + CAMLassert (Is_in_heap (Val_hp (hp))); + + /* Inline expansion of caml_allocation_color. */ + if (caml_gc_phase == Phase_mark || caml_gc_phase == Phase_clean + || (caml_gc_phase == Phase_sweep && (addr)hp >= (addr)caml_gc_sweep_hp)){ + Hd_hp (hp) = Make_header_with_profinfo (wosize, tag, Caml_black, profinfo); + }else{ + CAMLassert (caml_gc_phase == Phase_idle + || (caml_gc_phase == Phase_sweep + && (addr)hp < (addr)caml_gc_sweep_hp)); + Hd_hp (hp) = Make_header_with_profinfo (wosize, tag, Caml_white, profinfo); + } + CAMLassert (Hd_hp (hp) + == Make_header_with_profinfo (wosize, tag, caml_allocation_color (hp), + profinfo)); + caml_allocated_words += Whsize_wosize (wosize); + if (caml_allocated_words > caml_minor_heap_wsz){ + CAML_INSTR_INT ("request_major/alloc_shr@", 1); + caml_request_major_slice (); + } +#ifdef DEBUG + { + uintnat i; + for (i = 0; i < wosize; i++){ + Field (Val_hp (hp), i) = Debug_uninit_major; + } + } +#endif + return Val_hp (hp); +} + +CAMLexport value caml_alloc_shr_no_raise (mlsize_t wosize, tag_t tag) +{ + return caml_alloc_shr_aux(wosize, tag, 0, 0); +} + +#ifdef WITH_PROFINFO + +/* Use this to debug problems with macros... */ +#define NO_PROFINFO 0xff + +CAMLexport value caml_alloc_shr_with_profinfo (mlsize_t wosize, tag_t tag, + intnat profinfo) +{ + return caml_alloc_shr_aux(wosize, tag, 1, profinfo); +} + +CAMLexport value caml_alloc_shr_preserving_profinfo (mlsize_t wosize, + tag_t tag, header_t old_header) +{ + return caml_alloc_shr_with_profinfo (wosize, tag, Profinfo_hd(old_header)); +} + +#else +#define NO_PROFINFO 0 +#endif /* WITH_PROFINFO */ + +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) +#include "caml/spacetime.h" + +CAMLexport value caml_alloc_shr (mlsize_t wosize, tag_t tag) +{ + return caml_alloc_shr_with_profinfo (wosize, tag, + caml_spacetime_my_profinfo (NULL, wosize)); +} +#else +CAMLexport value caml_alloc_shr (mlsize_t wosize, tag_t tag) +{ + return caml_alloc_shr_aux (wosize, tag, 1, NO_PROFINFO); +} +#endif + +/* Dependent memory is all memory blocks allocated out of the heap + that depend on the GC (and finalizers) for deallocation. + For the GC to take dependent memory into account when computing + its automatic speed setting, + you must call [caml_alloc_dependent_memory] when you alloate some + dependent memory, and [caml_free_dependent_memory] when you + free it. In both cases, you pass as argument the size (in bytes) + of the block being allocated or freed. +*/ +CAMLexport void caml_alloc_dependent_memory (mlsize_t nbytes) +{ + caml_dependent_size += nbytes / sizeof (value); + caml_dependent_allocated += nbytes / sizeof (value); +} + +CAMLexport void caml_free_dependent_memory (mlsize_t nbytes) +{ + if (caml_dependent_size < nbytes / sizeof (value)){ + caml_dependent_size = 0; + }else{ + caml_dependent_size -= nbytes / sizeof (value); + } +} + +/* Use this function to tell the major GC to speed up when you use + finalized blocks to automatically deallocate resources (other + than memory). The GC will do at least one cycle every [max] + allocated resources; [res] is the number of resources allocated + this time. + Note that only [res/max] is relevant. The units (and kind of + resource) can change between calls to [caml_adjust_gc_speed]. +*/ +CAMLexport void caml_adjust_gc_speed (mlsize_t res, mlsize_t max) +{ + if (max == 0) max = 1; + if (res > max) res = max; + caml_extra_heap_resources += (double) res / (double) max; + if (caml_extra_heap_resources > 1.0){ + CAML_INSTR_INT ("request_major/adjust_gc_speed_1@", 1); + caml_extra_heap_resources = 1.0; + caml_request_major_slice (); + } + if (caml_extra_heap_resources + > (double) caml_minor_heap_wsz / 2.0 + / (double) caml_stat_heap_wsz) { + CAML_INSTR_INT ("request_major/adjust_gc_speed_2@", 1); + caml_request_major_slice (); + } +} + +/* You must use [caml_initialize] to store the initial value in a field of + a shared block, unless you are sure the value is not a young block. + A block value [v] is a shared block if and only if [Is_in_heap (v)] + is true. +*/ +/* [caml_initialize] never calls the GC, so you may call it while a block is + unfinished (i.e. just after a call to [caml_alloc_shr].) */ +/* PR#6084 workaround: define it as a weak symbol */ +CAMLexport CAMLweakdef void caml_initialize (value *fp, value val) +{ + CAMLassert(Is_in_heap_or_young(fp)); + *fp = val; + if (!Is_young((value)fp) && Is_block (val) && Is_young (val)) { + add_to_ref_table (&caml_ref_table, fp); + } +} + +/* You must use [caml_modify] to change a field of an existing shared block, + unless you are sure the value being overwritten is not a shared block and + the value being written is not a young block. */ +/* [caml_modify] never calls the GC. */ +/* [caml_modify] can also be used to do assignment on data structures that are + in the minor heap instead of in the major heap. In this case, it + is a bit slower than simple assignment. + In particular, you can use [caml_modify] when you don't know whether the + block being changed is in the minor heap or the major heap. */ +/* PR#6084 workaround: define it as a weak symbol */ + +CAMLexport CAMLweakdef void caml_modify (value *fp, value val) +{ + /* The write barrier implemented by [caml_modify] checks for the + following two conditions and takes appropriate action: + 1- a pointer from the major heap to the minor heap is created + --> add [fp] to the remembered set + 2- a pointer from the major heap to the major heap is overwritten, + while the GC is in the marking phase + --> call [caml_darken] on the overwritten pointer so that the + major GC treats it as an additional root. + */ + value old; + + if (Is_young((value)fp)) { + /* The modified object resides in the minor heap. + Conditions 1 and 2 cannot occur. */ + *fp = val; + } else { + /* The modified object resides in the major heap. */ + CAMLassert(Is_in_heap(fp)); + old = *fp; + *fp = val; + if (Is_block(old)) { + /* If [old] is a pointer within the minor heap, we already + have a major->minor pointer and [fp] is already in the + remembered set. Conditions 1 and 2 cannot occur. */ + if (Is_young(old)) return; + /* Here, [old] can be a pointer within the major heap. + Check for condition 2. */ + if (caml_gc_phase == Phase_mark) caml_darken(old, NULL); + } + /* Check for condition 1. */ + if (Is_block(val) && Is_young(val)) { + add_to_ref_table (&caml_ref_table, fp); + } + } +} + + +/* Global memory pool. + + The pool is structured as a ring of blocks, where each block's header + contains two links: to the previous and to the next block. The data + structure allows for insertions and removals of blocks in constant time, + given that a pointer to the operated block is provided. + + Initially, the pool contains a single block -- a pivot with no data, the + guaranteed existence of which makes for a more concise implementation. + + The API functions that operate on the pool receive not pointers to the + block's header, but rather pointers to the block's "data" field. This + behaviour is required to maintain compatibility with the interfaces of + [malloc], [realloc], and [free] family of functions, as well as to hide + the implementation from the user. +*/ + +/* A type with the most strict alignment requirements */ +union max_align { + char c; + short s; + long l; + int i; + float f; + double d; + void *v; + void (*q)(void); +}; + +struct pool_block { +#ifdef DEBUG + long magic; +#endif + struct pool_block *next; + struct pool_block *prev; + /* Use C99's flexible array types if possible */ +#if (__STDC_VERSION__ >= 199901L) + union max_align data[]; /* not allocated, used for alignment purposes */ +#else + union max_align data[1]; +#endif +}; + +#if (__STDC_VERSION__ >= 199901L) +#define SIZEOF_POOL_BLOCK sizeof(struct pool_block) +#else +#define SIZEOF_POOL_BLOCK offsetof(struct pool_block, data) +#endif + +static struct pool_block *pool = NULL; + + +/* Returns a pointer to the block header, given a pointer to "data" */ +static struct pool_block* get_pool_block(caml_stat_block b) +{ + if (b == NULL) + return NULL; + + else { + struct pool_block *pb = + (struct pool_block*)(((char*)b) - SIZEOF_POOL_BLOCK); +#ifdef DEBUG + CAMLassert(pb->magic == Debug_pool_magic); +#endif + return pb; + } +} + +CAMLexport void caml_stat_create_pool(void) +{ + if (pool == NULL) { + pool = malloc(SIZEOF_POOL_BLOCK); + if (pool == NULL) + caml_fatal_error("Fatal error: out of memory.\n"); +#ifdef DEBUG + pool->magic = Debug_pool_magic; +#endif + pool->next = pool; + pool->prev = pool; + } +} + +CAMLexport void caml_stat_destroy_pool(void) +{ + if (pool != NULL) { + pool->prev->next = NULL; + while (pool != NULL) { + struct pool_block *next = pool->next; + free(pool); + pool = next; + } + pool = NULL; + } +} + +/* [sz] and [modulo] are numbers of bytes */ +CAMLexport void* caml_stat_alloc_aligned_noexc(asize_t sz, int modulo, + caml_stat_block *b) +{ + char *raw_mem; + uintnat aligned_mem; + CAMLassert (modulo < Page_size); + raw_mem = (char *) caml_stat_alloc_noexc(sz + Page_size); + if (raw_mem == NULL) return NULL; + *b = raw_mem; + raw_mem += modulo; /* Address to be aligned */ + aligned_mem = (((uintnat) raw_mem / Page_size + 1) * Page_size); +#ifdef DEBUG + { + uintnat *p; + uintnat *p0 = (void *) *b; + uintnat *p1 = (void *) (aligned_mem - modulo); + uintnat *p2 = (void *) (aligned_mem - modulo + sz); + uintnat *p3 = (void *) ((char *) *b + sz + Page_size); + for (p = p0; p < p1; p++) *p = Debug_filler_align; + for (p = p1; p < p2; p++) *p = Debug_uninit_align; + for (p = p2; p < p3; p++) *p = Debug_filler_align; + } +#endif + return (char *) (aligned_mem - modulo); +} + +/* [sz] and [modulo] are numbers of bytes */ +CAMLexport void* caml_stat_alloc_aligned(asize_t sz, int modulo, + caml_stat_block *b) +{ + void *result = caml_stat_alloc_aligned_noexc(sz, modulo, b); + /* malloc() may return NULL if size is 0 */ + if ((result == NULL) && (sz != 0)) + caml_raise_out_of_memory(); + return result; +} + +/* [sz] is a number of bytes */ +CAMLexport caml_stat_block caml_stat_alloc_noexc(asize_t sz) +{ + /* Backward compatibility mode */ + if (pool == NULL) + return malloc(sz); + else { + struct pool_block *pb = malloc(sz + SIZEOF_POOL_BLOCK); + if (pb == NULL) return NULL; +#ifdef DEBUG + memset(&(pb->data), Debug_uninit_stat, sz); + pb->magic = Debug_pool_magic; +#endif + + /* Linking the block into the ring */ + pb->next = pool->next; + pb->prev = pool; + pool->next->prev = pb; + pool->next = pb; + + return &(pb->data); + } +} + +/* [sz] is a number of bytes */ +CAMLexport caml_stat_block caml_stat_alloc(asize_t sz) +{ + void *result = caml_stat_alloc_noexc(sz); + /* malloc() may return NULL if size is 0 */ + if ((result == NULL) && (sz != 0)) + caml_raise_out_of_memory(); + return result; +} + +CAMLexport void caml_stat_free(caml_stat_block b) +{ + /* Backward compatibility mode */ + if (pool == NULL) + free(b); + else { + struct pool_block *pb = get_pool_block(b); + if (pb == NULL) return; + + /* Unlinking the block from the ring */ + pb->prev->next = pb->next; + pb->next->prev = pb->prev; + + free(pb); + } +} + +/* [sz] is a number of bytes */ +CAMLexport caml_stat_block caml_stat_resize_noexc(caml_stat_block b, asize_t sz) +{ + /* Backward compatibility mode */ + if (pool == NULL) + return realloc(b, sz); + else { + struct pool_block *pb = get_pool_block(b); + struct pool_block *pb_new = realloc(pb, sz + SIZEOF_POOL_BLOCK); + if (pb_new == NULL) return NULL; + + /* Relinking the new block into the ring in place of the old one */ + pb_new->prev->next = pb_new; + pb_new->next->prev = pb_new; + + return &(pb_new->data); + } +} + +/* [sz] is a number of bytes */ +CAMLexport caml_stat_block caml_stat_resize(caml_stat_block b, asize_t sz) +{ + void *result = caml_stat_resize_noexc(b, sz); + if (result == NULL) + caml_raise_out_of_memory(); + return result; +} + +/* [sz] is a number of bytes */ +CAMLexport caml_stat_block caml_stat_calloc_noexc(asize_t num, asize_t sz) +{ + uintnat total; + if (caml_umul_overflow(sz, num, &total)) + return NULL; + else { + caml_stat_block result = caml_stat_alloc_noexc(total); + if (result != NULL) + memset(result, 0, total); + return result; + } +} + +CAMLexport caml_stat_string caml_stat_strdup_noexc(const char *s) +{ + size_t slen = strlen(s); + caml_stat_block result = caml_stat_alloc_noexc(slen + 1); + if (result == NULL) + return NULL; + memcpy(result, s, slen + 1); + return result; +} + +CAMLexport caml_stat_string caml_stat_strdup(const char *s) +{ + caml_stat_string result = caml_stat_strdup_noexc(s); + if (result == NULL) + caml_raise_out_of_memory(); + return result; +} + +#ifdef _WIN32 + +CAMLexport wchar_t * caml_stat_wcsdup(const wchar_t *s) +{ + int slen = wcslen(s); + wchar_t* result = caml_stat_alloc((slen + 1)*sizeof(wchar_t)); + if (result == NULL) + caml_raise_out_of_memory(); + memcpy(result, s, (slen + 1)*sizeof(wchar_t)); + return result; +} + +#endif + +CAMLexport caml_stat_string caml_stat_strconcat(int n, ...) +{ + va_list args; + char *result, *p; + size_t len = 0; + int i; + + va_start(args, n); + for (i = 0; i < n; i++) { + const char *s = va_arg(args, const char*); + len += strlen(s); + } + va_end(args); + + result = caml_stat_alloc(len + 1); + + va_start(args, n); + p = result; + for (i = 0; i < n; i++) { + const char *s = va_arg(args, const char*); + size_t l = strlen(s); + memcpy(p, s, l); + p += l; + } + va_end(args); + + *p = 0; + return result; +} + +#ifdef _WIN32 + +CAMLexport wchar_t* caml_stat_wcsconcat(int n, ...) +{ + va_list args; + wchar_t *result, *p; + size_t len = 0; + int i; + + va_start(args, n); + for (i = 0; i < n; i++) { + const wchar_t *s = va_arg(args, const wchar_t*); + len += wcslen(s); + } + va_end(args); + + result = caml_stat_alloc((len + 1)*sizeof(wchar_t)); + + va_start(args, n); + p = result; + for (i = 0; i < n; i++) { + const wchar_t *s = va_arg(args, const wchar_t*); + size_t l = wcslen(s); + memcpy(p, s, l*sizeof(wchar_t)); + p += l; + } + va_end(args); + + *p = 0; + return result; +} + +#endif diff --git a/test/monniaux/ocaml/byterun/meta.c b/test/monniaux/ocaml/byterun/meta.c new file mode 100644 index 00000000..03e0479d --- /dev/null +++ b/test/monniaux/ocaml/byterun/meta.c @@ -0,0 +1,234 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Primitives for the toplevel */ + +#include +#include "caml/alloc.h" +#include "caml/config.h" +#include "caml/fail.h" +#include "caml/fix_code.h" +#include "caml/interp.h" +#include "caml/intext.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/minor_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/prims.h" +#include "caml/stacks.h" + +#ifndef NATIVE_CODE + +CAMLprim value caml_get_global_data(value unit) +{ + return caml_global_data; +} + +char * caml_section_table = NULL; +asize_t caml_section_table_size; + +CAMLprim value caml_get_section_table(value unit) +{ + if (caml_section_table == NULL) caml_raise_not_found(); + return caml_input_value_from_block(caml_section_table, + caml_section_table_size); +} + +CAMLprim value caml_reify_bytecode(value prog, value len) +{ + struct code_fragment * cf = caml_stat_alloc(sizeof(struct code_fragment)); + value clos; + + cf->code_start = (char *) prog; + cf->code_end = (char *) prog + Long_val(len); + cf->digest_computed = 0; + caml_ext_table_add(&caml_code_fragments_table, cf); + +#ifdef ARCH_BIG_ENDIAN + caml_fixup_endianness((code_t) prog, (asize_t) Long_val(len)); +#endif +#ifdef THREADED_CODE + caml_thread_code((code_t) prog, (asize_t) Long_val(len)); +#endif + caml_prepare_bytecode((code_t) prog, (asize_t) Long_val(len)); + clos = caml_alloc_small (1, Closure_tag); + Code_val(clos) = (code_t) prog; + return clos; +} + +/* signal to the interpreter machinery that a bytecode is no more + needed (before freeing it) - this might be useful for a JIT + implementation */ + +CAMLprim value caml_static_release_bytecode(value prog, value len) +{ + struct code_fragment * cf = NULL, * cfi; + int i; + for (i = 0; i < caml_code_fragments_table.size; i++) { + cfi = (struct code_fragment *) caml_code_fragments_table.contents[i]; + if (cfi->code_start == (char *) prog && + cfi->code_end == (char *) prog + Long_val(len)) { + cf = cfi; + break; + } + } + + if (!cf) { + /* [cf] Not matched with a caml_reify_bytecode call; impossible. */ + CAMLassert (0); + } else { + caml_ext_table_remove(&caml_code_fragments_table, cf); + } + +#ifndef NATIVE_CODE + caml_release_bytecode((code_t) prog, (asize_t) Long_val(len)); +#else + caml_failwith("Meta.static_release_bytecode impossible with native code"); +#endif + return Val_unit; +} + +CAMLprim value caml_register_code_fragment(value prog, value len, value digest) +{ + struct code_fragment * cf = caml_stat_alloc(sizeof(struct code_fragment)); + cf->code_start = (char *) prog; + cf->code_end = (char *) prog + Long_val(len); + memcpy(cf->digest, String_val(digest), 16); + cf->digest_computed = 1; + caml_ext_table_add(&caml_code_fragments_table, cf); + return Val_unit; +} + +CAMLprim value caml_realloc_global(value size) +{ + mlsize_t requested_size, actual_size, i; + value new_global_data; + + requested_size = Long_val(size); + actual_size = Wosize_val(caml_global_data); + if (requested_size >= actual_size) { + requested_size = (requested_size + 0x100) & 0xFFFFFF00; + caml_gc_message (0x08, "Growing global data to %" + ARCH_INTNAT_PRINTF_FORMAT "u entries\n", + requested_size); + new_global_data = caml_alloc_shr(requested_size, 0); + for (i = 0; i < actual_size; i++) + caml_initialize(&Field(new_global_data, i), Field(caml_global_data, i)); + for (i = actual_size; i < requested_size; i++){ + Field (new_global_data, i) = Val_long (0); + } + caml_global_data = new_global_data; + } + return Val_unit; +} + +CAMLprim value caml_get_current_environment(value unit) +{ + return *caml_extern_sp; +} + +CAMLprim value caml_invoke_traced_function(value codeptr, value env, value arg) +{ + /* Stack layout on entry: + return frame into instrument_closure function + arg3 to call_original_code (arg) + arg2 to call_original_code (env) + arg1 to call_original_code (codeptr) + arg3 to call_original_code (arg) + arg2 to call_original_code (env) + saved env */ + + /* Stack layout on exit: + return frame into instrument_closure function + actual arg to code (arg) + pseudo return frame into codeptr: + extra_args = 0 + environment = env + PC = codeptr + arg3 to call_original_code (arg) same 6 bottom words as + arg2 to call_original_code (env) on entrance, but + arg1 to call_original_code (codeptr) shifted down 4 words + arg3 to call_original_code (arg) + arg2 to call_original_code (env) + saved env */ + + value * osp, * nsp; + int i; + + osp = caml_extern_sp; + caml_extern_sp -= 4; + nsp = caml_extern_sp; + for (i = 0; i < 6; i++) nsp[i] = osp[i]; + nsp[6] = codeptr; + nsp[7] = env; + nsp[8] = Val_int(0); + nsp[9] = arg; + return Val_unit; +} + +#else + +/* Dummy definitions to support compilation of ocamlc.opt */ + +value caml_get_global_data(value unit) +{ + caml_invalid_argument("Meta.get_global_data"); + return Val_unit; /* not reached */ +} + +value caml_get_section_table(value unit) +{ + caml_invalid_argument("Meta.get_section_table"); + return Val_unit; /* not reached */ +} + +value caml_realloc_global(value size) +{ + caml_invalid_argument("Meta.realloc_global"); + return Val_unit; /* not reached */ +} + +value caml_invoke_traced_function(value codeptr, value env, value arg) +{ + caml_invalid_argument("Meta.invoke_traced_function"); + return Val_unit; /* not reached */ +} + +value caml_reify_bytecode(value prog, value len) +{ + caml_invalid_argument("Meta.reify_bytecode"); + return Val_unit; /* not reached */ +} + +value caml_static_release_bytecode(value prog, value len) +{ + caml_invalid_argument("Meta.static_release_bytecode"); + return Val_unit; /* not reached */ +} + +value * caml_stack_low; +value * caml_stack_high; +value * caml_stack_threshold; +value * caml_extern_sp; +value * caml_trapsp; +int caml_callback_depth; +int volatile caml_something_to_do; +void (* volatile caml_async_action_hook)(void); +struct longjmp_buffer * caml_external_raise; + +#endif diff --git a/test/monniaux/ocaml/byterun/minor_gc.c b/test/monniaux/ocaml/byterun/minor_gc.c new file mode 100644 index 00000000..6aa5ed72 --- /dev/null +++ b/test/monniaux/ocaml/byterun/minor_gc.c @@ -0,0 +1,558 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include +#include "caml/custom.h" +#include "caml/config.h" +#include "caml/fail.h" +#include "caml/finalise.h" +#include "caml/gc.h" +#include "caml/gc_ctrl.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/minor_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/roots.h" +#include "caml/signals.h" +#include "caml/weak.h" + +/* Pointers into the minor heap. + [caml_young_base] + The [malloc] block that contains the heap. + [caml_young_start] ... [caml_young_end] + The whole range of the minor heap: all young blocks are inside + this interval. + [caml_young_alloc_start]...[caml_young_alloc_end] + The allocation arena: newly-allocated blocks are carved from + this interval, starting at [caml_young_alloc_end]. + [caml_young_alloc_mid] is the mid-point of this interval. + [caml_young_ptr], [caml_young_trigger], [caml_young_limit] + These pointers are all inside the allocation arena. + - [caml_young_ptr] is where the next allocation will take place. + - [caml_young_trigger] is how far we can allocate before triggering + [caml_gc_dispatch]. Currently, it is either [caml_young_alloc_start] + or the mid-point of the allocation arena. + - [caml_young_limit] is the pointer that is compared to + [caml_young_ptr] for allocation. It is either + [caml_young_alloc_end] if a signal is pending and we are in + native code, or [caml_young_trigger]. +*/ + +struct generic_table CAML_TABLE_STRUCT(char); + +asize_t caml_minor_heap_wsz; +static void *caml_young_base = NULL; +CAMLexport value *caml_young_start = NULL, *caml_young_end = NULL; +CAMLexport value *caml_young_alloc_start = NULL, + *caml_young_alloc_mid = NULL, + *caml_young_alloc_end = NULL; +CAMLexport value *caml_young_ptr = NULL, *caml_young_limit = NULL; +CAMLexport value *caml_young_trigger = NULL; + +CAMLexport struct caml_ref_table + caml_ref_table = { NULL, NULL, NULL, NULL, NULL, 0, 0}; + +CAMLexport struct caml_ephe_ref_table + caml_ephe_ref_table = { NULL, NULL, NULL, NULL, NULL, 0, 0}; + +CAMLexport struct caml_custom_table + caml_custom_table = { NULL, NULL, NULL, NULL, NULL, 0, 0}; +/* Table of custom blocks in the minor heap that contain finalizers + or GC speed parameters. */ + +int caml_in_minor_collection = 0; + +double caml_extra_heap_resources_minor = 0; + +/* [sz] and [rsv] are numbers of entries */ +static void alloc_generic_table (struct generic_table *tbl, asize_t sz, + asize_t rsv, asize_t element_size) +{ + void *new_table; + + tbl->size = sz; + tbl->reserve = rsv; + new_table = (void *) caml_stat_alloc_noexc((tbl->size + tbl->reserve) * + element_size); + if (new_table == NULL) caml_fatal_error ("Fatal error: not enough memory\n"); + if (tbl->base != NULL) caml_stat_free (tbl->base); + tbl->base = new_table; + tbl->ptr = tbl->base; + tbl->threshold = tbl->base + tbl->size * element_size; + tbl->limit = tbl->threshold; + tbl->end = tbl->base + (tbl->size + tbl->reserve) * element_size; +} + +void caml_alloc_table (struct caml_ref_table *tbl, asize_t sz, asize_t rsv) +{ + alloc_generic_table ((struct generic_table *) tbl, sz, rsv, sizeof (value *)); +} + +void caml_alloc_ephe_table (struct caml_ephe_ref_table *tbl, asize_t sz, + asize_t rsv) +{ + alloc_generic_table ((struct generic_table *) tbl, sz, rsv, + sizeof (struct caml_ephe_ref_elt)); +} + +void caml_alloc_custom_table (struct caml_custom_table *tbl, asize_t sz, + asize_t rsv) +{ + alloc_generic_table ((struct generic_table *) tbl, sz, rsv, + sizeof (struct caml_custom_elt)); +} + +static void reset_table (struct generic_table *tbl) +{ + tbl->size = 0; + tbl->reserve = 0; + if (tbl->base != NULL) caml_stat_free (tbl->base); + tbl->base = tbl->ptr = tbl->threshold = tbl->limit = tbl->end = NULL; +} + +static void clear_table (struct generic_table *tbl) +{ + tbl->ptr = tbl->base; + tbl->limit = tbl->threshold; +} + +void caml_set_minor_heap_size (asize_t bsz) +{ + char *new_heap; + void *new_heap_base; + + CAMLassert (bsz >= Bsize_wsize(Minor_heap_min)); + CAMLassert (bsz <= Bsize_wsize(Minor_heap_max)); + CAMLassert (bsz % sizeof (value) == 0); + if (caml_young_ptr != caml_young_alloc_end){ + CAML_INSTR_INT ("force_minor/set_minor_heap_size@", 1); + caml_requested_minor_gc = 0; + caml_young_trigger = caml_young_alloc_mid; + caml_young_limit = caml_young_trigger; + caml_empty_minor_heap (); + } + CAMLassert (caml_young_ptr == caml_young_alloc_end); + new_heap = caml_stat_alloc_aligned_noexc(bsz, 0, &new_heap_base); + if (new_heap == NULL) caml_raise_out_of_memory(); + if (caml_page_table_add(In_young, new_heap, new_heap + bsz) != 0) + caml_raise_out_of_memory(); + + if (caml_young_start != NULL){ + caml_page_table_remove(In_young, caml_young_start, caml_young_end); + caml_stat_free (caml_young_base); + } + caml_young_base = new_heap_base; + caml_young_start = (value *) new_heap; + caml_young_end = (value *) (new_heap + bsz); + caml_young_alloc_start = caml_young_start; + caml_young_alloc_mid = caml_young_alloc_start + Wsize_bsize (bsz) / 2; + caml_young_alloc_end = caml_young_end; + caml_young_trigger = caml_young_alloc_start; + caml_young_limit = caml_young_trigger; + caml_young_ptr = caml_young_alloc_end; + caml_minor_heap_wsz = Wsize_bsize (bsz); + + reset_table ((struct generic_table *) &caml_ref_table); + reset_table ((struct generic_table *) &caml_ephe_ref_table); + reset_table ((struct generic_table *) &caml_custom_table); +} + +static value oldify_todo_list = 0; + +/* Note that the tests on the tag depend on the fact that Infix_tag, + Forward_tag, and No_scan_tag are contiguous. */ + +void caml_oldify_one (value v, value *p) +{ + value result; + header_t hd; + mlsize_t sz, i; + tag_t tag; + + tail_call: + if (Is_block (v) && Is_young (v)){ + CAMLassert ((value *) Hp_val (v) >= caml_young_ptr); + hd = Hd_val (v); + if (hd == 0){ /* If already forwarded */ + *p = Field (v, 0); /* then forward pointer is first field. */ + }else{ + tag = Tag_hd (hd); + if (tag < Infix_tag){ + value field0; + + sz = Wosize_hd (hd); + result = caml_alloc_shr_preserving_profinfo (sz, tag, hd); + *p = result; + field0 = Field (v, 0); + Hd_val (v) = 0; /* Set forward flag */ + Field (v, 0) = result; /* and forward pointer. */ + if (sz > 1){ + Field (result, 0) = field0; + Field (result, 1) = oldify_todo_list; /* Add this block */ + oldify_todo_list = v; /* to the "to do" list. */ + }else{ + CAMLassert (sz == 1); + p = &Field (result, 0); + v = field0; + goto tail_call; + } + }else if (tag >= No_scan_tag){ + sz = Wosize_hd (hd); + result = caml_alloc_shr_preserving_profinfo (sz, tag, hd); + for (i = 0; i < sz; i++) Field (result, i) = Field (v, i); + Hd_val (v) = 0; /* Set forward flag */ + Field (v, 0) = result; /* and forward pointer. */ + *p = result; + }else if (tag == Infix_tag){ + mlsize_t offset = Infix_offset_hd (hd); + caml_oldify_one (v - offset, p); /* Cannot recurse deeper than 1. */ + *p += offset; + }else{ + value f = Forward_val (v); + tag_t ft = 0; + int vv = 1; + + CAMLassert (tag == Forward_tag); + if (Is_block (f)){ + if (Is_young (f)){ + vv = 1; + ft = Tag_val (Hd_val (f) == 0 ? Field (f, 0) : f); + }else{ + vv = Is_in_value_area(f); + if (vv){ + ft = Tag_val (f); + } + } + } + if (!vv || ft == Forward_tag || ft == Lazy_tag +#ifdef FLAT_FLOAT_ARRAY + || ft == Double_tag +#endif + ){ + /* Do not short-circuit the pointer. Copy as a normal block. */ + CAMLassert (Wosize_hd (hd) == 1); + result = caml_alloc_shr_preserving_profinfo (1, Forward_tag, hd); + *p = result; + Hd_val (v) = 0; /* Set (GC) forward flag */ + Field (v, 0) = result; /* and forward pointer. */ + p = &Field (result, 0); + v = f; + goto tail_call; + }else{ + v = f; /* Follow the forwarding */ + goto tail_call; /* then oldify. */ + } + } + } + }else{ + *p = v; + } +} + +/* Test if the ephemeron is alive, everything outside minor heap is alive */ +static inline int ephe_check_alive_data(struct caml_ephe_ref_elt *re){ + mlsize_t i; + value child; + for (i = CAML_EPHE_FIRST_KEY; i < Wosize_val(re->ephe); i++){ + child = Field (re->ephe, i); + if(child != caml_ephe_none + && Is_block (child) && Is_young (child) + && Hd_val (child) != 0){ /* Value not copied to major heap */ + return 0; + } + } + return 1; +} + +/* Finish the work that was put off by [caml_oldify_one]. + Note that [caml_oldify_one] itself is called by oldify_mopup, so we + have to be careful to remove the first entry from the list before + oldifying its fields. */ +void caml_oldify_mopup (void) +{ + value v, new_v, f; + mlsize_t i; + struct caml_ephe_ref_elt *re; + int redo = 0; + + while (oldify_todo_list != 0){ + v = oldify_todo_list; /* Get the head. */ + CAMLassert (Hd_val (v) == 0); /* It must be forwarded. */ + new_v = Field (v, 0); /* Follow forward pointer. */ + oldify_todo_list = Field (new_v, 1); /* Remove from list. */ + + f = Field (new_v, 0); + if (Is_block (f) && Is_young (f)){ + caml_oldify_one (f, &Field (new_v, 0)); + } + for (i = 1; i < Wosize_val (new_v); i++){ + f = Field (v, i); + if (Is_block (f) && Is_young (f)){ + caml_oldify_one (f, &Field (new_v, i)); + }else{ + Field (new_v, i) = f; + } + } + } + + /* Oldify the data in the minor heap of alive ephemeron + During minor collection keys outside the minor heap are considered alive */ + for (re = caml_ephe_ref_table.base; + re < caml_ephe_ref_table.ptr; re++){ + /* look only at ephemeron with data in the minor heap */ + if (re->offset == 1){ + value *data = &Field(re->ephe,1); + if (*data != caml_ephe_none && Is_block (*data) && Is_young (*data)){ + if (Hd_val (*data) == 0){ /* Value copied to major heap */ + *data = Field (*data, 0); + } else { + if (ephe_check_alive_data(re)){ + caml_oldify_one(*data,data); + redo = 1; /* oldify_todo_list can still be 0 */ + } + } + } + } + } + + if (redo) caml_oldify_mopup (); +} + +/* Make sure the minor heap is empty by performing a minor collection + if needed. +*/ +void caml_empty_minor_heap (void) +{ + value **r; + struct caml_custom_elt *elt; + uintnat prev_alloc_words; + struct caml_ephe_ref_elt *re; + + if (caml_young_ptr != caml_young_alloc_end){ + if (caml_minor_gc_begin_hook != NULL) (*caml_minor_gc_begin_hook) (); + CAML_INSTR_SETUP (tmr, "minor"); + prev_alloc_words = caml_allocated_words; + caml_in_minor_collection = 1; + caml_gc_message (0x02, "<"); + caml_oldify_local_roots(); + CAML_INSTR_TIME (tmr, "minor/local_roots"); + for (r = caml_ref_table.base; r < caml_ref_table.ptr; r++){ + caml_oldify_one (**r, *r); + } + CAML_INSTR_TIME (tmr, "minor/ref_table"); + caml_oldify_mopup (); + CAML_INSTR_TIME (tmr, "minor/copy"); + /* Update the ephemerons */ + for (re = caml_ephe_ref_table.base; + re < caml_ephe_ref_table.ptr; re++){ + if(re->offset < Wosize_val(re->ephe)){ + /* If it is not the case, the ephemeron has been truncated */ + value *key = &Field(re->ephe,re->offset); + if (*key != caml_ephe_none && Is_block (*key) && Is_young (*key)){ + if (Hd_val (*key) == 0){ /* Value copied to major heap */ + *key = Field (*key, 0); + }else{ /* Value not copied so it's dead */ + CAMLassert(!ephe_check_alive_data(re)); + *key = caml_ephe_none; + Field(re->ephe,1) = caml_ephe_none; + } + } + } + } + /* Update the OCaml finalise_last values */ + caml_final_update_minor_roots(); + /* Run custom block finalisation of dead minor values */ + for (elt = caml_custom_table.base; elt < caml_custom_table.ptr; elt++){ + value v = elt->block; + if (Hd_val (v) == 0){ + /* Block was copied to the major heap: adjust GC speed numbers. */ + caml_adjust_gc_speed(elt->mem, elt->max); + }else{ + /* Block will be freed: call finalization function, if any. */ + void (*final_fun)(value) = Custom_ops_val(v)->finalize; + if (final_fun != NULL) final_fun(v); + } + } + CAML_INSTR_TIME (tmr, "minor/update_weak"); + caml_stat_minor_words += caml_young_alloc_end - caml_young_ptr; + caml_gc_clock += (double) (caml_young_alloc_end - caml_young_ptr) + / caml_minor_heap_wsz; + caml_young_ptr = caml_young_alloc_end; + clear_table ((struct generic_table *) &caml_ref_table); + clear_table ((struct generic_table *) &caml_ephe_ref_table); + clear_table ((struct generic_table *) &caml_custom_table); + caml_extra_heap_resources_minor = 0; + caml_gc_message (0x02, ">"); + caml_in_minor_collection = 0; + caml_final_empty_young (); + CAML_INSTR_TIME (tmr, "minor/finalized"); + caml_stat_promoted_words += caml_allocated_words - prev_alloc_words; + CAML_INSTR_INT ("minor/promoted#", caml_allocated_words - prev_alloc_words); + ++ caml_stat_minor_collections; + if (caml_minor_gc_end_hook != NULL) (*caml_minor_gc_end_hook) (); + }else{ + /* The minor heap is empty nothing to do. */ + caml_final_empty_young (); + } +#ifdef DEBUG + { + value *p; + for (p = caml_young_alloc_start; p < caml_young_alloc_end; ++p){ + *p = Debug_free_minor; + } + } +#endif +} + +#ifdef CAML_INSTR +extern uintnat caml_instr_alloc_jump; +#endif + +/* Do a minor collection or a slice of major collection, call finalisation + functions, etc. + Leave enough room in the minor heap to allocate at least one object. +*/ +CAMLexport void caml_gc_dispatch (void) +{ + value *trigger = caml_young_trigger; /* save old value of trigger */ +#ifdef CAML_INSTR + CAML_INSTR_SETUP(tmr, "dispatch"); + CAML_INSTR_TIME (tmr, "overhead"); + CAML_INSTR_INT ("alloc/jump#", caml_instr_alloc_jump); + caml_instr_alloc_jump = 0; +#endif + + if (trigger == caml_young_alloc_start || caml_requested_minor_gc){ + /* The minor heap is full, we must do a minor collection. */ + /* reset the pointers first because the end hooks might allocate */ + caml_requested_minor_gc = 0; + caml_young_trigger = caml_young_alloc_mid; + caml_young_limit = caml_young_trigger; + caml_empty_minor_heap (); + /* The minor heap is empty, we can start a major collection. */ + if (caml_gc_phase == Phase_idle) caml_major_collection_slice (-1); + CAML_INSTR_TIME (tmr, "dispatch/minor"); + + caml_final_do_calls (); + CAML_INSTR_TIME (tmr, "dispatch/finalizers"); + + while (caml_young_ptr - caml_young_alloc_start < Max_young_whsize){ + /* The finalizers or the hooks have filled up the minor heap, we must + repeat the minor collection. */ + caml_requested_minor_gc = 0; + caml_young_trigger = caml_young_alloc_mid; + caml_young_limit = caml_young_trigger; + caml_empty_minor_heap (); + /* The minor heap is empty, we can start a major collection. */ + if (caml_gc_phase == Phase_idle) caml_major_collection_slice (-1); + CAML_INSTR_TIME (tmr, "dispatch/finalizers_minor"); + } + } + if (trigger != caml_young_alloc_start || caml_requested_major_slice){ + /* The minor heap is half-full, do a major GC slice. */ + caml_requested_major_slice = 0; + caml_young_trigger = caml_young_alloc_start; + caml_young_limit = caml_young_trigger; + caml_major_collection_slice (-1); + CAML_INSTR_TIME (tmr, "dispatch/major"); + } +} + +/* For backward compatibility with Lablgtk: do a minor collection to + ensure that the minor heap is empty. +*/ +CAMLexport void caml_minor_collection (void) +{ + caml_requested_minor_gc = 1; + caml_gc_dispatch (); +} + +CAMLexport value caml_check_urgent_gc (value extra_root) +{ + CAMLparam1 (extra_root); + if (caml_requested_major_slice || caml_requested_minor_gc){ + CAML_INSTR_INT ("force_minor/check_urgent_gc@", 1); + caml_gc_dispatch(); + } + CAMLreturn (extra_root); +} + +static void realloc_generic_table +(struct generic_table *tbl, asize_t element_size, + char * msg_intr_int, char *msg_threshold, char *msg_growing, char *msg_error) +{ + CAMLassert (tbl->ptr == tbl->limit); + CAMLassert (tbl->limit <= tbl->end); + CAMLassert (tbl->limit >= tbl->threshold); + + if (tbl->base == NULL){ + alloc_generic_table (tbl, caml_minor_heap_wsz / 8, 256, + element_size); + }else if (tbl->limit == tbl->threshold){ + CAML_INSTR_INT (msg_intr_int, 1); + caml_gc_message (0x08, msg_threshold, 0); + tbl->limit = tbl->end; + caml_request_minor_gc (); + }else{ + asize_t sz; + asize_t cur_ptr = tbl->ptr - tbl->base; + CAMLassert (caml_requested_minor_gc); + + tbl->size *= 2; + sz = (tbl->size + tbl->reserve) * element_size; + caml_gc_message (0x08, msg_growing, (intnat) sz/1024); + tbl->base = caml_stat_resize_noexc (tbl->base, sz); + if (tbl->base == NULL){ + caml_fatal_error (msg_error); + } + tbl->end = tbl->base + (tbl->size + tbl->reserve) * element_size; + tbl->threshold = tbl->base + tbl->size * element_size; + tbl->ptr = tbl->base + cur_ptr; + tbl->limit = tbl->end; + } +} + +void caml_realloc_ref_table (struct caml_ref_table *tbl) +{ + realloc_generic_table + ((struct generic_table *) tbl, sizeof (value *), + "request_minor/realloc_ref_table@", + "ref_table threshold crossed\n", + "Growing ref_table to %" ARCH_INTNAT_PRINTF_FORMAT "dk bytes\n", + "Fatal error: ref_table overflow\n"); +} + +void caml_realloc_ephe_ref_table (struct caml_ephe_ref_table *tbl) +{ + realloc_generic_table + ((struct generic_table *) tbl, sizeof (struct caml_ephe_ref_elt), + "request_minor/realloc_ephe_ref_table@", + "ephe_ref_table threshold crossed\n", + "Growing ephe_ref_table to %" ARCH_INTNAT_PRINTF_FORMAT "dk bytes\n", + "Fatal error: ephe_ref_table overflow\n"); +} + +void caml_realloc_custom_table (struct caml_custom_table *tbl) +{ + realloc_generic_table + ((struct generic_table *) tbl, sizeof (struct caml_custom_elt), + "request_minor/realloc_custom_table@", + "custom_table threshold crossed\n", + "Growing custom_table to %" ARCH_INTNAT_PRINTF_FORMAT "dk bytes\n", + "Fatal error: custom_table overflow\n"); +} diff --git a/test/monniaux/ocaml/byterun/misc.c b/test/monniaux/ocaml/byterun/misc.c new file mode 100644 index 00000000..46e40992 --- /dev/null +++ b/test/monniaux/ocaml/byterun/misc.c @@ -0,0 +1,276 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +#include +#include +#include +#include "caml/config.h" +#include "caml/misc.h" +#include "caml/memory.h" +#include "caml/osdeps.h" +#include "caml/version.h" + +caml_timing_hook caml_major_slice_begin_hook = NULL; +caml_timing_hook caml_major_slice_end_hook = NULL; +caml_timing_hook caml_minor_gc_begin_hook = NULL; +caml_timing_hook caml_minor_gc_end_hook = NULL; +caml_timing_hook caml_finalise_begin_hook = NULL; +caml_timing_hook caml_finalise_end_hook = NULL; + +#ifdef DEBUG + +int caml_failed_assert (char * expr, char * file, int line) +{ + fprintf (stderr, "file %s; line %d ### Assertion failed: %s\n", + file, line, expr); + fflush (stderr); + abort(); +} + +void caml_set_fields (value v, unsigned long start, unsigned long filler) +{ + mlsize_t i; + for (i = start; i < Wosize_val (v); i++){ + Field (v, i) = (value) filler; + } +} + +#endif /* DEBUG */ + +uintnat caml_verb_gc = 0; + +void caml_gc_message (int level, char *msg, ...) +{ + if ((caml_verb_gc & level) != 0){ + va_list ap; + va_start(ap, msg); + vfprintf (stderr, msg, ap); + va_end(ap); + fflush (stderr); + } +} + +CAMLexport void caml_fatal_error (char *msg) +{ + fprintf (stderr, "%s", msg); + exit(2); +} + +CAMLexport void caml_fatal_error_arg (char *fmt, char *arg) +{ + fprintf (stderr, fmt, arg); + exit(2); +} + +CAMLexport void caml_fatal_error_arg2 (char *fmt1, char *arg1, + char *fmt2, char *arg2) +{ + fprintf (stderr, fmt1, arg1); + fprintf (stderr, fmt2, arg2); + exit(2); +} + +/* If you change the caml_ext_table* functions, also update + asmrun/spacetime.c:find_trie_node_from_libunwind. */ + +void caml_ext_table_init(struct ext_table * tbl, int init_capa) +{ + tbl->size = 0; + tbl->capacity = init_capa; + tbl->contents = caml_stat_alloc(sizeof(void *) * init_capa); +} + +int caml_ext_table_add(struct ext_table * tbl, caml_stat_block data) +{ + int res; + if (tbl->size >= tbl->capacity) { + tbl->capacity *= 2; + tbl->contents = + caml_stat_resize(tbl->contents, sizeof(void *) * tbl->capacity); + } + res = tbl->size; + tbl->contents[res] = data; + tbl->size++; + return res; +} + +void caml_ext_table_remove(struct ext_table * tbl, caml_stat_block data) +{ + int i; + for (i = 0; i < tbl->size; i++) { + if (tbl->contents[i] == data) { + caml_stat_free(tbl->contents[i]); + memmove(&tbl->contents[i], &tbl->contents[i + 1], + (tbl->size - i - 1) * sizeof(void *)); + tbl->size--; + } + } +} + +void caml_ext_table_clear(struct ext_table * tbl, int free_entries) +{ + int i; + if (free_entries) { + for (i = 0; i < tbl->size; i++) caml_stat_free(tbl->contents[i]); + } + tbl->size = 0; +} + +void caml_ext_table_free(struct ext_table * tbl, int free_entries) +{ + caml_ext_table_clear(tbl, free_entries); + caml_stat_free(tbl->contents); +} + +/* Integer arithmetic with overflow detection */ + +#if ! (__GNUC__ >= 5 || Caml_has_builtin(__builtin_mul_overflow)) +CAMLexport int caml_umul_overflow(uintnat a, uintnat b, uintnat * res) +{ +#define HALF_SIZE (sizeof(uintnat) * 4) +#define HALF_MASK (((uintnat)1 << HALF_SIZE) - 1) +#define LOW_HALF(x) ((x) & HALF_MASK) +#define HIGH_HALF(x) ((x) >> HALF_SIZE) + /* Cut in half words */ + uintnat al = LOW_HALF(a); + uintnat ah = HIGH_HALF(a); + uintnat bl = LOW_HALF(b); + uintnat bh = HIGH_HALF(b); + /* Exact product is: + al * bl + + ah * bl << HALF_SIZE + + al * bh << HALF_SIZE + + ah * bh << 2*HALF_SIZE + Overflow occurs if: + ah * bh is not 0, i.e. ah != 0 and bh != 0 + OR ah * bl has high half != 0 + OR al * bh has high half != 0 + OR the sum al * bl + LOW_HALF(ah * bl) << HALF_SIZE + + LOW_HALF(al * bh) << HALF_SIZE overflows. + This sum is equal to p = (a * b) modulo word size. */ + uintnat p = a * b; + uintnat p1 = al * bh; + uintnat p2 = ah * bl; + *res = p; + if (ah == 0 && bh == 0) return 0; + if (ah != 0 && bh != 0) return 1; + if (HIGH_HALF(p1) != 0 || HIGH_HALF(p2) != 0) return 1; + p1 <<= HALF_SIZE; + p2 <<= HALF_SIZE; + p1 += p2; + if (p < p1 || p1 < p2) return 1; /* overflow in sums */ + return 0; +#undef HALF_SIZE +#undef HALF_MASK +#undef LOW_HALF +#undef HIGH_HALF +} +#endif + +/* Runtime warnings */ + +uintnat caml_runtime_warnings = 0; +static int caml_runtime_warnings_first = 1; + +int caml_runtime_warnings_active(void) +{ + if (!caml_runtime_warnings) return 0; + if (caml_runtime_warnings_first) { + fprintf(stderr, "[ocaml] (use Sys.enable_runtime_warnings to control " + "these warnings)\n"); + caml_runtime_warnings_first = 0; + } + return 1; +} + +#ifdef CAML_INSTR +/* Timers for profiling GC and allocation (experimental, Linux-only) */ + +#include +#include +#include + +struct CAML_INSTR_BLOCK *CAML_INSTR_LOG = NULL; +intnat CAML_INSTR_STARTTIME, CAML_INSTR_STOPTIME; + +#define Get_time(p,i) ((p)->ts[(i)].tv_nsec + 1000000000 * (p)->ts[(i)].tv_sec) + +void CAML_INSTR_INIT (void) +{ + char *s; + + CAML_INSTR_STARTTIME = 0; + s = caml_secure_getenv ("OCAML_INSTR_START"); + if (s != NULL) CAML_INSTR_STARTTIME = atol (s); + CAML_INSTR_STOPTIME = LONG_MAX; + s = caml_secure_getenv ("OCAML_INSTR_STOP"); + if (s != NULL) CAML_INSTR_STOPTIME = atol (s); +} + +void CAML_INSTR_ATEXIT (void) +{ + int i; + struct CAML_INSTR_BLOCK *p, *prev, *next; + FILE *f = NULL; + char *fname; + + fname = caml_secure_getenv ("OCAML_INSTR_FILE"); + if (fname != NULL){ + char *mode = "a"; + char buf [1000]; + char *name = fname; + + if (name[0] == '@'){ + snprintf (buf, sizeof(buf), "%s.%d", name + 1, getpid ()); + name = buf; + } + if (name[0] == '+'){ + mode = "a"; + name = name + 1; + }else if (name [0] == '>' || name[0] == '-'){ + mode = "w"; + name = name + 1; + } + f = fopen (name, mode); + } + + if (f != NULL){ + /* reverse the list */ + prev = NULL; + p = CAML_INSTR_LOG; + while (p != NULL){ + next = p->next; + p->next = prev; + prev = p; + p = next; + } + CAML_INSTR_LOG = prev; + fprintf (f, "==== OCAML INSTRUMENTATION DATA %s\n", OCAML_VERSION_STRING); + for (p = CAML_INSTR_LOG; p != NULL; p = p->next){ + for (i = 0; i < p->index; i++){ + fprintf (f, "@@ %19ld %19ld %s\n", + (long) Get_time (p, i), (long) Get_time(p, i+1), p->tag[i+1]); + } + if (p->tag[0][0] != '\000'){ + fprintf (f, "@@ %19ld %19ld %s\n", + (long) Get_time (p, 0), (long) Get_time(p, p->index), p->tag[0]); + } + } + fclose (f); + } +} +#endif /* CAML_INSTR */ diff --git a/test/monniaux/ocaml/byterun/obj.c b/test/monniaux/ocaml/byterun/obj.c new file mode 100644 index 00000000..4567b8ae --- /dev/null +++ b/test/monniaux/ocaml/byterun/obj.c @@ -0,0 +1,390 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Operations on objects */ + +#include +#include "caml/alloc.h" +#include "caml/fail.h" +#include "caml/gc.h" +#include "caml/interp.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/minor_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/prims.h" +#include "caml/spacetime.h" + +/* [size] is a value encoding a number of bytes */ +CAMLprim value caml_static_alloc(value size) +{ + return (value) caml_stat_alloc((asize_t) Long_val(size)); +} + +CAMLprim value caml_static_free(value blk) +{ + caml_stat_free((void *) blk); + return Val_unit; +} + +CAMLprim value caml_static_resize(value blk, value new_size) +{ + return (value) caml_stat_resize((char *) blk, (asize_t) Long_val(new_size)); +} + +/* unused since GPR#427 */ +CAMLprim value caml_obj_is_block(value arg) +{ + return Val_bool(Is_block(arg)); +} + +CAMLprim value caml_obj_tag(value arg) +{ + if (Is_long (arg)){ + return Val_int (1000); /* int_tag */ + }else if ((long) arg & (sizeof (value) - 1)){ + return Val_int (1002); /* unaligned_tag */ + }else if (Is_in_value_area (arg)){ + return Val_int(Tag_val(arg)); + }else{ + return Val_int (1001); /* out_of_heap_tag */ + } +} + +CAMLprim value caml_obj_set_tag (value arg, value new_tag) +{ + Tag_val (arg) = Int_val (new_tag); + return Val_unit; +} + +/* [size] is a value encoding a number of blocks */ +CAMLprim value caml_obj_block(value tag, value size) +{ + value res; + mlsize_t sz, i; + tag_t tg; + + sz = Long_val(size); + tg = Long_val(tag); + if (sz == 0) return Atom(tg); + res = caml_alloc(sz, tg); + for (i = 0; i < sz; i++) + Field(res, i) = Val_long(0); + + return res; +} + +/* Spacetime profiling assumes that this function is only called from OCaml. */ +CAMLprim value caml_obj_dup(value arg) +{ + CAMLparam1 (arg); + CAMLlocal1 (res); + mlsize_t sz, i; + tag_t tg; + + sz = Wosize_val(arg); + if (sz == 0) CAMLreturn (arg); + tg = Tag_val(arg); + if (tg >= No_scan_tag) { + res = caml_alloc(sz, tg); + memcpy(Bp_val(res), Bp_val(arg), sz * sizeof(value)); + } else if (sz <= Max_young_wosize) { + uintnat profinfo; + Get_my_profinfo_with_cached_backtrace(profinfo, sz); + res = caml_alloc_small_with_my_or_given_profinfo(sz, tg, profinfo); + for (i = 0; i < sz; i++) Field(res, i) = Field(arg, i); + } else { + res = caml_alloc_shr(sz, tg); + for (i = 0; i < sz; i++) caml_initialize(&Field(res, i), Field(arg, i)); + } + CAMLreturn (res); +} + +/* Shorten the given block to the given size and return void. + Raise Invalid_argument if the given size is less than or equal + to 0 or greater than the current size. + + algorithm: + Change the length field of the header. Make up a black object + with the leftover part of the object: this is needed in the major + heap and harmless in the minor heap. The object cannot be white + because there may still be references to it in the ref table. By + using a black object we ensure that the ref table will be emptied + before the block is reallocated (since there must be a minor + collection within each major cycle). + + [newsize] is a value encoding a number of fields (words, except + for float arrays on 32-bit architectures). +*/ +CAMLprim value caml_obj_truncate (value v, value newsize) +{ + mlsize_t new_wosize = Long_val (newsize); + header_t hd = Hd_val (v); + tag_t tag = Tag_hd (hd); + color_t color = Color_hd (hd); + mlsize_t wosize = Wosize_hd (hd); + mlsize_t i; + + if (tag == Double_array_tag) new_wosize *= Double_wosize; /* PR#156 */ + + if (new_wosize <= 0 || new_wosize > wosize){ + caml_invalid_argument ("Obj.truncate"); + } + if (new_wosize == wosize) return Val_unit; + /* PR#61: since we're about to lose our references to the elements + beyond new_wosize in v, erase them explicitly so that the GC + can darken them as appropriate. */ + if (tag < No_scan_tag) { + for (i = new_wosize; i < wosize; i++){ + caml_modify(&Field(v, i), Val_unit); +#ifdef DEBUG + Field (v, i) = Debug_free_truncate; +#endif + } + } + /* We must use an odd tag for the header of the leftovers so it does not + look like a pointer because there may be some references to it in + ref_table. */ + Field (v, new_wosize) = + Make_header (Wosize_whsize (wosize-new_wosize), Abstract_tag, Caml_black); + Hd_val (v) = + Make_header_with_profinfo (new_wosize, tag, color, Profinfo_val(v)); + return Val_unit; +} + +CAMLprim value caml_obj_add_offset (value v, value offset) +{ + return v + (unsigned long) Int32_val (offset); +} + +/* The following functions are used in stdlib/lazy.ml. + They are not written in OCaml because they must be atomic with respect + to the GC. + */ + +CAMLprim value caml_lazy_follow_forward (value v) +{ + if (Is_block (v) && Is_in_value_area(v) + && Tag_val (v) == Forward_tag){ + return Forward_val (v); + }else{ + return v; + } +} + +CAMLprim value caml_lazy_make_forward (value v) +{ + CAMLparam1 (v); + CAMLlocal1 (res); + + res = caml_alloc_small (1, Forward_tag); + Field (res, 0) = v; + CAMLreturn (res); +} + +/* For mlvalues.h and camlinternalOO.ml + See also GETPUBMET in interp.c + */ + +CAMLprim value caml_get_public_method (value obj, value tag) +{ + value meths = Field (obj, 0); + int li = 3, hi = Field(meths,0), mi; + while (li < hi) { + mi = ((li+hi) >> 1) | 1; + if (tag < Field(meths,mi)) hi = mi-2; + else li = mi; + } + /* return 0 if tag is not there */ + return (tag == Field(meths,li) ? Field (meths, li-1) : 0); +} + +/* these two functions might be useful to an hypothetical JIT */ + +#ifdef CAML_JIT +#ifdef NATIVE_CODE +#define MARK 1 +#else +#define MARK 0 +#endif +value caml_cache_public_method (value meths, value tag, value *cache) +{ + int li = 3, hi = Field(meths,0), mi; + while (li < hi) { + mi = ((li+hi) >> 1) | 1; + if (tag < Field(meths,mi)) hi = mi-2; + else li = mi; + } + *cache = (li-3)*sizeof(value) + MARK; + return Field (meths, li-1); +} + +value caml_cache_public_method2 (value *meths, value tag, value *cache) +{ + value ofs = *cache & meths[1]; + if (*(value*)(((char*)(meths+3)) + ofs - MARK) == tag) + return *(value*)(((char*)(meths+2)) + ofs - MARK); + { + int li = 3, hi = meths[0], mi; + while (li < hi) { + mi = ((li+hi) >> 1) | 1; + if (tag < meths[mi]) hi = mi-2; + else li = mi; + } + *cache = (li-3)*sizeof(value) + MARK; + return meths[li-1]; + } +} +#endif /*CAML_JIT*/ + +static value oo_last_id = Val_int(0); + +CAMLprim value caml_set_oo_id (value obj) { + Field(obj, 1) = oo_last_id; + oo_last_id += 2; + return obj; +} + +CAMLprim value caml_fresh_oo_id (value v) { + v = oo_last_id; + oo_last_id += 2; + return v; +} + +CAMLprim value caml_int_as_pointer (value n) { + return n - 1; +} + +/* Compute how many words in the heap are occupied by blocks accessible + from a given value */ + +#define ENTRIES_PER_QUEUE_CHUNK 4096 +struct queue_chunk { + struct queue_chunk *next; + value entries[ENTRIES_PER_QUEUE_CHUNK]; +}; + + +CAMLprim value caml_obj_reachable_words(value v) +{ + static struct queue_chunk first_chunk; + struct queue_chunk *read_chunk, *write_chunk; + int write_pos, read_pos, i; + + intnat size = 0; + header_t hd; + mlsize_t sz; + + if (Is_long(v) || !Is_in_heap_or_young(v)) return Val_int(0); + if (Tag_hd(Hd_val(v)) == Infix_tag) v -= Infix_offset_hd(Hd_val(v)); + hd = Hd_val(v); + sz = Wosize_hd(hd); + + read_chunk = write_chunk = &first_chunk; + read_pos = 0; + write_pos = 1; + write_chunk->entries[0] = v | Colornum_hd(hd); + Hd_val(v) = Bluehd_hd(hd); + + /* We maintain a queue of "interesting" blocks that have been seen. + An interesting block is a block in the heap which does not + represent an infix pointer. Infix pointers are normalized to the + beginning of their block. Blocks in the static data area are excluded. + + The function maintains a queue of block pointers. Concretely, + the queue is stored as a linked list of chunks, each chunk + holding a number of pointers to interesting blocks. Initially, + it contains only the "root" value. The first chunk of the queue + is allocated statically. More chunks can be allocated as needed + and released before this function exits. + + When a block is inserted in the queue, it is marked as blue. + This mark is used to avoid a second visit of the same block. + The real color is stored in the last 2 bits of the pointer in the + queue. (Same technique as in extern.c.) + + Note: we make the assumption that there is no pointer + from the static data area to the heap. + */ + + /* First pass: mark accessible blocks and compute their total size */ + while (read_pos != write_pos || read_chunk != write_chunk) { + /* Pop the next element from the queue */ + if (read_pos == ENTRIES_PER_QUEUE_CHUNK) { + read_pos = 0; + read_chunk = read_chunk->next; + } + v = read_chunk->entries[read_pos++] & ~3; + + hd = Hd_val(v); + sz = Wosize_hd(hd); + + size += Whsize_wosize(sz); + + if (Tag_hd(hd) < No_scan_tag) { + /* Push the interesting fields on the queue */ + for (i = 0; i < sz; i++) { + value v2 = Field(v, i); + if (Is_block(v2) && Is_in_heap_or_young(v2)) { + if (Tag_hd(Hd_val(v2)) == Infix_tag){ + v2 -= Infix_offset_hd(Hd_val(v2)); + } + hd = Hd_val(v2); + if (Color_hd(hd) != Caml_blue) { + if (write_pos == ENTRIES_PER_QUEUE_CHUNK) { + struct queue_chunk *new_chunk = + malloc(sizeof(struct queue_chunk)); + if (new_chunk == NULL) { + size = (-1); + goto release; + } + write_chunk->next = new_chunk; + write_pos = 0; + write_chunk = new_chunk; + } + write_chunk->entries[write_pos++] = v2 | Colornum_hd(hd); + Hd_val(v2) = Bluehd_hd(hd); + } + } + } + } + } + + /* Second pass: restore colors and free extra queue chunks */ + release: + read_pos = 0; + read_chunk = &first_chunk; + while (read_pos != write_pos || read_chunk != write_chunk) { + color_t colornum; + if (read_pos == ENTRIES_PER_QUEUE_CHUNK) { + struct queue_chunk *prev = read_chunk; + read_pos = 0; + read_chunk = read_chunk->next; + if (prev != &first_chunk) free(prev); + } + v = read_chunk->entries[read_pos++]; + colornum = v & 3; + v &= ~3; + Hd_val(v) = Coloredhd_hd(Hd_val(v), colornum); + } + if (read_chunk != &first_chunk) free(read_chunk); + + if (size < 0) + caml_raise_out_of_memory(); + return Val_int(size); +} diff --git a/test/monniaux/ocaml/byterun/parsing.c b/test/monniaux/ocaml/byterun/parsing.c new file mode 100644 index 00000000..990eb1f6 --- /dev/null +++ b/test/monniaux/ocaml/byterun/parsing.c @@ -0,0 +1,304 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* The PDA automaton for parsers generated by camlyacc */ + +#include +#include +#include "caml/config.h" +#include "caml/mlvalues.h" +#include "caml/memory.h" +#include "caml/alloc.h" + +#define ERRCODE 256 + +struct parser_tables { /* Mirrors parse_tables in ../stdlib/parsing.mli */ + value actions; + value transl_const; + value transl_block; + char * lhs; + char * len; + char * defred; + char * dgoto; + char * sindex; + char * rindex; + char * gindex; + value tablesize; + char * table; + char * check; + value error_function; + char * names_const; + char * names_block; +}; + +struct parser_env { /* Mirrors parser_env in ../stdlib/parsing.ml */ + value s_stack; + value v_stack; + value symb_start_stack; + value symb_end_stack; + value stacksize; + value stackbase; + value curr_char; + value lval; + value symb_start; + value symb_end; + value asp; + value rule_len; + value rule_number; + value sp; + value state; + value errflag; +}; + +#if defined(ARCH_BIG_ENDIAN) || SIZEOF_SHORT != 2 +#define Short(tbl,n) \ + (*((unsigned char *)((tbl) + (n) * 2)) + \ + (*((signed char *)((tbl) + (n) * 2 + 1)) << 8)) +#else +#define Short(tbl,n) (((short *)(tbl))[n]) +#endif + +int caml_parser_trace = 0; + +/* Input codes */ +/* Mirrors parser_input in ../stdlib/parsing.ml */ +#define START 0 +#define TOKEN_READ 1 +#define STACKS_GROWN_1 2 +#define STACKS_GROWN_2 3 +#define SEMANTIC_ACTION_COMPUTED 4 +#define ERROR_DETECTED 5 + +/* Output codes */ +/* Mirrors parser_output in ../stdlib/parsing.ml */ +#define READ_TOKEN Val_int(0) +#define RAISE_PARSE_ERROR Val_int(1) +#define GROW_STACKS_1 Val_int(2) +#define GROW_STACKS_2 Val_int(3) +#define COMPUTE_SEMANTIC_ACTION Val_int(4) +#define CALL_ERROR_FUNCTION Val_int(5) + +/* To preserve local variables when communicating with the ML code */ + +#define SAVE \ + env->sp = Val_int(sp), \ + env->state = Val_int(state), \ + env->errflag = Val_int(errflag) + +#define RESTORE \ + sp = Int_val(env->sp), \ + state = Int_val(env->state), \ + errflag = Int_val(env->errflag) + +/* Auxiliary for printing token just read */ + +static char * token_name(char * names, int number) +{ + for (/*nothing*/; number > 0; number--) { + if (names[0] == 0) return ""; + names += strlen(names) + 1; + } + return names; +} + +static void print_token(struct parser_tables *tables, int state, value tok) +{ + value v; + + if (Is_long(tok)) { + fprintf(stderr, "State %d: read token %s\n", + state, token_name(tables->names_const, Int_val(tok))); + } else { + fprintf(stderr, "State %d: read token %s(", + state, token_name(tables->names_block, Tag_val(tok))); + v = Field(tok, 0); + if (Is_long(v)) + fprintf(stderr, "%" ARCH_INTNAT_PRINTF_FORMAT "d", Long_val(v)); + else if (Tag_val(v) == String_tag) + fprintf(stderr, "%s", String_val(v)); + else if (Tag_val(v) == Double_tag) + fprintf(stderr, "%g", Double_val(v)); + else + fprintf(stderr, "_"); + fprintf(stderr, ")\n"); + } +} + +/* The pushdown automata */ + +CAMLprim value caml_parse_engine(struct parser_tables *tables, + struct parser_env *env, value cmd, value arg) +{ + int state; + mlsize_t sp, asp; + int errflag; + int n, n1, n2, m, state1; + + switch(Int_val(cmd)) { + + case START: + state = 0; + sp = Int_val(env->sp); + errflag = 0; + + loop: + n = Short(tables->defred, state); + if (n != 0) goto reduce; + if (Int_val(env->curr_char) >= 0) goto testshift; + SAVE; + return READ_TOKEN; + /* The ML code calls the lexer and updates */ + /* symb_start and symb_end */ + case TOKEN_READ: + RESTORE; + if (Is_block(arg)) { + env->curr_char = Field(tables->transl_block, Tag_val(arg)); + caml_modify(&env->lval, Field(arg, 0)); + } else { + env->curr_char = Field(tables->transl_const, Int_val(arg)); + caml_modify(&env->lval, Val_long(0)); + } + if (caml_parser_trace) print_token(tables, state, arg); + + testshift: + n1 = Short(tables->sindex, state); + n2 = n1 + Int_val(env->curr_char); + if (n1 != 0 && n2 >= 0 && n2 <= Int_val(tables->tablesize) && + Short(tables->check, n2) == Int_val(env->curr_char)) goto shift; + n1 = Short(tables->rindex, state); + n2 = n1 + Int_val(env->curr_char); + if (n1 != 0 && n2 >= 0 && n2 <= Int_val(tables->tablesize) && + Short(tables->check, n2) == Int_val(env->curr_char)) { + n = Short(tables->table, n2); + goto reduce; + } + if (errflag > 0) goto recover; + SAVE; + return CALL_ERROR_FUNCTION; + /* The ML code calls the error function */ + case ERROR_DETECTED: + RESTORE; + recover: + if (errflag < 3) { + errflag = 3; + while (1) { + state1 = Int_val(Field(env->s_stack, sp)); + n1 = Short(tables->sindex, state1); + n2 = n1 + ERRCODE; + if (n1 != 0 && n2 >= 0 && n2 <= Int_val(tables->tablesize) && + Short(tables->check, n2) == ERRCODE) { + if (caml_parser_trace) + fprintf(stderr, "Recovering in state %d\n", state1); + goto shift_recover; + } else { + if (caml_parser_trace){ + fprintf(stderr, "Discarding state %d\n", state1); + } + if (sp <= Int_val(env->stackbase)) { + if (caml_parser_trace){ + fprintf(stderr, "No more states to discard\n"); + } + return RAISE_PARSE_ERROR; /* The ML code raises Parse_error */ + } + sp--; + } + } + } else { + if (Int_val(env->curr_char) == 0) + return RAISE_PARSE_ERROR; /* The ML code raises Parse_error */ + if (caml_parser_trace) fprintf(stderr, "Discarding last token read\n"); + env->curr_char = Val_int(-1); + goto loop; + } + + shift: + env->curr_char = Val_int(-1); + if (errflag > 0) errflag--; + shift_recover: + if (caml_parser_trace) + fprintf(stderr, "State %d: shift to state %d\n", + state, Short(tables->table, n2)); + state = Short(tables->table, n2); + sp++; + if (sp < Long_val(env->stacksize)) goto push; + SAVE; + return GROW_STACKS_1; + /* The ML code resizes the stacks */ + case STACKS_GROWN_1: + RESTORE; + push: + Field(env->s_stack, sp) = Val_int(state); + caml_modify(&Field(env->v_stack, sp), env->lval); + Store_field (env->symb_start_stack, sp, env->symb_start); + Store_field (env->symb_end_stack, sp, env->symb_end); + goto loop; + + reduce: + if (caml_parser_trace) + fprintf(stderr, "State %d: reduce by rule %d\n", state, n); + m = Short(tables->len, n); + env->asp = Val_int(sp); + env->rule_number = Val_int(n); + env->rule_len = Val_int(m); + sp = sp - m + 1; + m = Short(tables->lhs, n); + state1 = Int_val(Field(env->s_stack, sp - 1)); + n1 = Short(tables->gindex, m); + n2 = n1 + state1; + if (n1 != 0 && n2 >= 0 && n2 <= Int_val(tables->tablesize) && + Short(tables->check, n2) == state1) { + state = Short(tables->table, n2); + } else { + state = Short(tables->dgoto, m); + } + if (sp < Long_val(env->stacksize)) goto semantic_action; + SAVE; + return GROW_STACKS_2; + /* The ML code resizes the stacks */ + case STACKS_GROWN_2: + RESTORE; + semantic_action: + SAVE; + return COMPUTE_SEMANTIC_ACTION; + /* The ML code calls the semantic action */ + case SEMANTIC_ACTION_COMPUTED: + RESTORE; + Field(env->s_stack, sp) = Val_int(state); + caml_modify(&Field(env->v_stack, sp), arg); + asp = Int_val(env->asp); + Store_field (env->symb_end_stack, sp, Field(env->symb_end_stack, asp)); + if (sp > asp) { + /* This is an epsilon production. Take symb_start equal to symb_end. */ + Store_field (env->symb_start_stack, sp, Field(env->symb_end_stack, asp)); + } + goto loop; + + default: /* Should not happen */ + CAMLassert(0); + return RAISE_PARSE_ERROR; /* Keeps gcc -Wall happy */ + } + +} + +/* Control printing of debugging info */ + +CAMLprim value caml_set_parser_trace(value flag) +{ + value oldflag = Val_bool(caml_parser_trace); + caml_parser_trace = Bool_val(flag); + return oldflag; +} diff --git a/test/monniaux/ocaml/byterun/printexc.c b/test/monniaux/ocaml/byterun/printexc.c new file mode 100644 index 00000000..735c2994 --- /dev/null +++ b/test/monniaux/ocaml/byterun/printexc.c @@ -0,0 +1,155 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Print an uncaught exception and abort */ + +#include +#include +#include +#include "caml/backtrace.h" +#include "caml/callback.h" +#include "caml/debugger.h" +#include "caml/fail.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/printexc.h" +#include "caml/memory.h" + +struct stringbuf { + char * ptr; + char * end; + char data[256]; +}; + +static void add_char(struct stringbuf *buf, char c) +{ + if (buf->ptr < buf->end) *(buf->ptr++) = c; +} + +static void add_string(struct stringbuf *buf, const char *s) +{ + int len = strlen(s); + if (buf->ptr + len > buf->end) len = buf->end - buf->ptr; + if (len > 0) memmove(buf->ptr, s, len); + buf->ptr += len; +} + +CAMLexport char * caml_format_exception(value exn) +{ + mlsize_t start, i; + value bucket, v; + struct stringbuf buf; + char intbuf[64]; + char * res; + + buf.ptr = buf.data; + buf.end = buf.data + sizeof(buf.data) - 1; + if (Tag_val(exn) == 0) { + add_string(&buf, String_val(Field(Field(exn, 0), 0))); + /* Check for exceptions in the style of Match_failure and Assert_failure */ + if (Wosize_val(exn) == 2 && + Is_block(Field(exn, 1)) && + Tag_val(Field(exn, 1)) == 0 && + caml_is_special_exception(Field(exn, 0))) { + bucket = Field(exn, 1); + start = 0; + } else { + bucket = exn; + start = 1; + } + add_char(&buf, '('); + for (i = start; i < Wosize_val(bucket); i++) { + if (i > start) add_string(&buf, ", "); + v = Field(bucket, i); + if (Is_long(v)) { + snprintf(intbuf, sizeof(intbuf), + "%" ARCH_INTNAT_PRINTF_FORMAT "d", Long_val(v)); + add_string(&buf, intbuf); + } else if (Tag_val(v) == String_tag) { + add_char(&buf, '"'); + add_string(&buf, String_val(v)); + add_char(&buf, '"'); + } else { + add_char(&buf, '_'); + } + } + add_char(&buf, ')'); + } else + add_string(&buf, String_val(Field(exn, 0))); + + *buf.ptr = 0; /* Terminate string */ + i = buf.ptr - buf.data + 1; + res = caml_stat_alloc_noexc(i); + if (res == NULL) return NULL; + memmove(res, buf.data, i); + return res; +} + + +#ifdef NATIVE_CODE +# define DEBUGGER_IN_USE 0 +#else +# define DEBUGGER_IN_USE caml_debugger_in_use +#endif + +/* Default C implementation in case the OCaml one is not registered. */ +static void default_fatal_uncaught_exception(value exn) +{ + char * msg; + value * at_exit; + int saved_backtrace_active, saved_backtrace_pos; + + /* Build a string representation of the exception */ + msg = caml_format_exception(exn); + /* Perform "at_exit" processing, ignoring all exceptions that may + be triggered by this */ + saved_backtrace_active = caml_backtrace_active; + saved_backtrace_pos = caml_backtrace_pos; + caml_backtrace_active = 0; + at_exit = caml_named_value("Pervasives.do_at_exit"); + if (at_exit != NULL) caml_callback_exn(*at_exit, Val_unit); + caml_backtrace_active = saved_backtrace_active; + caml_backtrace_pos = saved_backtrace_pos; + /* Display the uncaught exception */ + fprintf(stderr, "Fatal error: exception %s\n", msg); + caml_stat_free(msg); + /* Display the backtrace if available */ + if (caml_backtrace_active && !DEBUGGER_IN_USE) + caml_print_exception_backtrace(); +} + +int caml_abort_on_uncaught_exn = 0; /* see afl.c */ + +void caml_fatal_uncaught_exception(value exn) +{ + value *handle_uncaught_exception; + + handle_uncaught_exception = + caml_named_value("Printexc.handle_uncaught_exception"); + if (handle_uncaught_exception != NULL) + /* [Printexc.handle_uncaught_exception] does not raise exception. */ + caml_callback2(*handle_uncaught_exception, exn, Val_bool(DEBUGGER_IN_USE)); + else + default_fatal_uncaught_exception(exn); + /* Terminate the process */ + if (caml_abort_on_uncaught_exn) { + abort(); + } else { + CAML_SYS_EXIT(2); + exit(2); /* Second exit needed for the Noreturn flag */ + } +} diff --git a/test/monniaux/ocaml/byterun/roots.c b/test/monniaux/ocaml/byterun/roots.c new file mode 100644 index 00000000..1445495a --- /dev/null +++ b/test/monniaux/ocaml/byterun/roots.c @@ -0,0 +1,120 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* To walk the memory roots for garbage collection */ + +#include "caml/finalise.h" +#include "caml/globroots.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/minor_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/roots.h" +#include "caml/stacks.h" + +CAMLexport struct caml__roots_block *caml_local_roots = NULL; + +CAMLexport void (*caml_scan_roots_hook) (scanning_action f) = NULL; + +/* FIXME should rename to [caml_oldify_minor_roots] and synchronise with + asmrun/roots.c */ +/* Call [caml_oldify_one] on (at least) all the roots that point to the minor + heap. */ +void caml_oldify_local_roots (void) +{ + register value * sp; + struct caml__roots_block *lr; + intnat i, j; + + /* The stack */ + for (sp = caml_extern_sp; sp < caml_stack_high; sp++) { + caml_oldify_one (*sp, sp); + } + /* Local C roots */ /* FIXME do the old-frame trick ? */ + for (lr = caml_local_roots; lr != NULL; lr = lr->next) { + for (i = 0; i < lr->ntables; i++){ + for (j = 0; j < lr->nitems; j++){ + sp = &(lr->tables[i][j]); + caml_oldify_one (*sp, sp); + } + } + } + /* Global C roots */ + caml_scan_global_young_roots(&caml_oldify_one); + /* Finalised values */ + caml_final_oldify_young_roots (); + /* Hook */ + if (caml_scan_roots_hook != NULL) (*caml_scan_roots_hook)(&caml_oldify_one); +} + +/* Call [caml_darken] on all roots */ + +void caml_darken_all_roots_start (void) +{ + caml_do_roots (caml_darken, 1); +} + +uintnat caml_incremental_roots_count = 1; + +intnat caml_darken_all_roots_slice (intnat work) +{ + return work; +} + +/* Note, in byte-code there is only one global root, so [do_globals] is + ignored and [caml_darken_all_roots_slice] does nothing. */ +void caml_do_roots (scanning_action f, int do_globals) +{ + CAML_INSTR_SETUP (tmr, "major_roots"); + /* Global variables */ + f(caml_global_data, &caml_global_data); + CAML_INSTR_TIME (tmr, "major_roots/global"); + /* The stack and the local C roots */ + caml_do_local_roots(f, caml_extern_sp, caml_stack_high, caml_local_roots); + CAML_INSTR_TIME (tmr, "major_roots/local"); + /* Global C roots */ + caml_scan_global_roots(f); + CAML_INSTR_TIME (tmr, "major_roots/C"); + /* Finalised values */ + caml_final_do_roots (f); + CAML_INSTR_TIME (tmr, "major_roots/finalised"); + /* Hook */ + if (caml_scan_roots_hook != NULL) (*caml_scan_roots_hook)(f); + CAML_INSTR_TIME (tmr, "major_roots/hook"); +} + +CAMLexport void caml_do_local_roots (scanning_action f, value *stack_low, + value *stack_high, + struct caml__roots_block *local_roots) +{ + register value * sp; + struct caml__roots_block *lr; + int i, j; + + for (sp = stack_low; sp < stack_high; sp++) { + f (*sp, sp); + } + for (lr = local_roots; lr != NULL; lr = lr->next) { + for (i = 0; i < lr->ntables; i++){ + for (j = 0; j < lr->nitems; j++){ + sp = &(lr->tables[i][j]); + f (*sp, sp); + } + } + } +} diff --git a/test/monniaux/ocaml/byterun/signals.c b/test/monniaux/ocaml/byterun/signals.c new file mode 100644 index 00000000..15addf1b --- /dev/null +++ b/test/monniaux/ocaml/byterun/signals.c @@ -0,0 +1,398 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Signal handling, code common to the bytecode and native systems */ + +#include +#include +#include "caml/alloc.h" +#include "caml/callback.h" +#include "caml/config.h" +#include "caml/fail.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/roots.h" +#include "caml/signals.h" +#include "caml/signals_machdep.h" +#include "caml/sys.h" + +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) +#include "caml/spacetime.h" +#endif + +#ifndef NSIG +#define NSIG 64 +#endif + +/* The set of pending signals (received but not yet processed) */ + +CAMLexport intnat volatile caml_signals_are_pending = 0; +CAMLexport intnat volatile caml_pending_signals[NSIG]; + +/* Execute all pending signals */ + +void caml_process_pending_signals(void) +{ + int i; + + if (caml_signals_are_pending) { + caml_signals_are_pending = 0; + for (i = 0; i < NSIG; i++) { + if (caml_pending_signals[i]) { + caml_pending_signals[i] = 0; + caml_execute_signal(i, 0); + } + } + } +} + +/* Record the delivery of a signal, and arrange for it to be processed + as soon as possible: + - in bytecode: via caml_something_to_do, processed in caml_process_event + - in native-code: by playing with the allocation limit, processed + in caml_garbage_collection +*/ + +void caml_record_signal(int signal_number) +{ + caml_pending_signals[signal_number] = 1; + caml_signals_are_pending = 1; +#ifndef NATIVE_CODE + caml_something_to_do = 1; +#else + caml_young_limit = caml_young_alloc_end; +#endif +} + +/* Management of blocking sections. */ + +static intnat volatile caml_async_signal_mode = 0; + +static void caml_enter_blocking_section_default(void) +{ + CAMLassert (caml_async_signal_mode == 0); + caml_async_signal_mode = 1; +} + +static void caml_leave_blocking_section_default(void) +{ + CAMLassert (caml_async_signal_mode == 1); + caml_async_signal_mode = 0; +} + +static int caml_try_leave_blocking_section_default(void) +{ + intnat res; + Read_and_clear(res, caml_async_signal_mode); + return res; +} + +CAMLexport void (*caml_enter_blocking_section_hook)(void) = + caml_enter_blocking_section_default; +CAMLexport void (*caml_leave_blocking_section_hook)(void) = + caml_leave_blocking_section_default; +CAMLexport int (*caml_try_leave_blocking_section_hook)(void) = + caml_try_leave_blocking_section_default; + +CAMLexport void caml_enter_blocking_section(void) +{ + while (1){ + /* Process all pending signals now */ + caml_process_pending_signals(); + caml_enter_blocking_section_hook (); + /* Check again for pending signals. + If none, done; otherwise, try again */ + if (! caml_signals_are_pending) break; + caml_leave_blocking_section_hook (); + } +} + +CAMLexport void caml_leave_blocking_section(void) +{ + int saved_errno; + /* Save the value of errno (PR#5982). */ + saved_errno = errno; + caml_leave_blocking_section_hook (); + caml_process_pending_signals(); + errno = saved_errno; +} + +/* Execute a signal handler immediately */ + +static value caml_signal_handlers = 0; + +void caml_execute_signal(int signal_number, int in_signal_handler) +{ + value res; + value handler; +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + void* saved_spacetime_trie_node_ptr; +#endif +#ifdef POSIX_SIGNALS + sigset_t nsigs, sigs; + /* Block the signal before executing the handler, and record in sigs + the original signal mask */ + sigemptyset(&nsigs); + sigaddset(&nsigs, signal_number); + sigprocmask(SIG_BLOCK, &nsigs, &sigs); +#endif +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + /* We record the signal handler's execution separately, in the same + trie used for finalisers. */ + saved_spacetime_trie_node_ptr + = caml_spacetime_trie_node_ptr; + caml_spacetime_trie_node_ptr + = caml_spacetime_finaliser_trie_root; +#endif +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + /* Handled action may have no associated handler, which we interpret + as meaning the signal should be handled by a call to exit. This is + is used to allow spacetime profiles to be completed on interrupt */ + if (caml_signal_handlers == 0) { + res = caml_sys_exit(Val_int(2)); + } else { + handler = Field(caml_signal_handlers, signal_number); + if (!Is_block(handler)) { + res = caml_sys_exit(Val_int(2)); + } else { +#else + handler = Field(caml_signal_handlers, signal_number); +#endif + res = caml_callback_exn( + handler, + Val_int(caml_rev_convert_signal_number(signal_number))); +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + } + } + caml_spacetime_trie_node_ptr = saved_spacetime_trie_node_ptr; +#endif +#ifdef POSIX_SIGNALS + if (! in_signal_handler) { + /* Restore the original signal mask */ + sigprocmask(SIG_SETMASK, &sigs, NULL); + } else if (Is_exception_result(res)) { + /* Restore the original signal mask and unblock the signal itself */ + sigdelset(&sigs, signal_number); + sigprocmask(SIG_SETMASK, &sigs, NULL); + } +#endif + if (Is_exception_result(res)) caml_raise(Extract_exception(res)); +} + +/* Arrange for a garbage collection to be performed as soon as possible */ + +int volatile caml_requested_major_slice = 0; +int volatile caml_requested_minor_gc = 0; + +void caml_request_major_slice (void) +{ + caml_requested_major_slice = 1; +#ifndef NATIVE_CODE + caml_something_to_do = 1; +#else + caml_young_limit = caml_young_alloc_end; + /* This is only moderately effective on ports that cache [caml_young_limit] + in a register, since [caml_modify] is called directly, not through + [caml_c_call], so it may take a while before the register is reloaded + from [caml_young_limit]. */ +#endif +} + +void caml_request_minor_gc (void) +{ + caml_requested_minor_gc = 1; +#ifndef NATIVE_CODE + caml_something_to_do = 1; +#else + caml_young_limit = caml_young_alloc_end; + /* Same remark as above in [caml_request_major_slice]. */ +#endif +} + +/* OS-independent numbering of signals */ + +#ifndef SIGABRT +#define SIGABRT -1 +#endif +#ifndef SIGALRM +#define SIGALRM -1 +#endif +#ifndef SIGFPE +#define SIGFPE -1 +#endif +#ifndef SIGHUP +#define SIGHUP -1 +#endif +#ifndef SIGILL +#define SIGILL -1 +#endif +#ifndef SIGINT +#define SIGINT -1 +#endif +#ifndef SIGKILL +#define SIGKILL -1 +#endif +#ifndef SIGPIPE +#define SIGPIPE -1 +#endif +#ifndef SIGQUIT +#define SIGQUIT -1 +#endif +#ifndef SIGSEGV +#define SIGSEGV -1 +#endif +#ifndef SIGTERM +#define SIGTERM -1 +#endif +#ifndef SIGUSR1 +#define SIGUSR1 -1 +#endif +#ifndef SIGUSR2 +#define SIGUSR2 -1 +#endif +#ifndef SIGCHLD +#define SIGCHLD -1 +#endif +#ifndef SIGCONT +#define SIGCONT -1 +#endif +#ifndef SIGSTOP +#define SIGSTOP -1 +#endif +#ifndef SIGTSTP +#define SIGTSTP -1 +#endif +#ifndef SIGTTIN +#define SIGTTIN -1 +#endif +#ifndef SIGTTOU +#define SIGTTOU -1 +#endif +#ifndef SIGVTALRM +#define SIGVTALRM -1 +#endif +#ifndef SIGPROF +#define SIGPROF -1 +#endif +#ifndef SIGBUS +#define SIGBUS -1 +#endif +#ifndef SIGPOLL +#define SIGPOLL -1 +#endif +#ifndef SIGSYS +#define SIGSYS -1 +#endif +#ifndef SIGTRAP +#define SIGTRAP -1 +#endif +#ifndef SIGURG +#define SIGURG -1 +#endif +#ifndef SIGXCPU +#define SIGXCPU -1 +#endif +#ifndef SIGXFSZ +#define SIGXFSZ -1 +#endif + +static int posix_signals[] = { + SIGABRT, SIGALRM, SIGFPE, SIGHUP, SIGILL, SIGINT, SIGKILL, SIGPIPE, + SIGQUIT, SIGSEGV, SIGTERM, SIGUSR1, SIGUSR2, SIGCHLD, SIGCONT, + SIGSTOP, SIGTSTP, SIGTTIN, SIGTTOU, SIGVTALRM, SIGPROF, SIGBUS, + SIGPOLL, SIGSYS, SIGTRAP, SIGURG, SIGXCPU, SIGXFSZ +}; + +CAMLexport int caml_convert_signal_number(int signo) +{ + if (signo < 0 && signo >= -(sizeof(posix_signals) / sizeof(int))) + return posix_signals[-signo-1]; + else + return signo; +} + +CAMLexport int caml_rev_convert_signal_number(int signo) +{ + int i; + for (i = 0; i < sizeof(posix_signals) / sizeof(int); i++) + if (signo == posix_signals[i]) return -i - 1; + return signo; +} + +/* Installation of a signal handler (as per [Sys.signal]) */ + +CAMLprim value caml_install_signal_handler(value signal_number, value action) +{ + CAMLparam2 (signal_number, action); + CAMLlocal1 (res); + int sig, act, oldact; + + sig = caml_convert_signal_number(Int_val(signal_number)); + if (sig < 0 || sig >= NSIG) + caml_invalid_argument("Sys.signal: unavailable signal"); + switch(action) { + case Val_int(0): /* Signal_default */ + act = 0; + break; + case Val_int(1): /* Signal_ignore */ + act = 1; + break; + default: /* Signal_handle */ + act = 2; + break; + } + oldact = caml_set_signal_action(sig, act); + switch (oldact) { + case 0: /* was Signal_default */ + res = Val_int(0); + break; + case 1: /* was Signal_ignore */ + res = Val_int(1); + break; + case 2: /* was Signal_handle */ + #if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + /* Handled action may have no associated handler + which we treat as Signal_default */ + if (caml_signal_handlers == 0) { + res = Val_int(0); + } else { + if (!Is_block(Field(caml_signal_handlers, sig))) { + res = Val_int(0); + } else { + res = caml_alloc_small (1, 0); + Field(res, 0) = Field(caml_signal_handlers, sig); + } + } + #else + res = caml_alloc_small (1, 0); + Field(res, 0) = Field(caml_signal_handlers, sig); + #endif + break; + default: /* error in caml_set_signal_action */ + caml_sys_error(NO_ARG); + } + if (Is_block(action)) { + if (caml_signal_handlers == 0) { + caml_signal_handlers = caml_alloc(NSIG, 0); + caml_register_global_root(&caml_signal_handlers); + } + caml_modify(&Field(caml_signal_handlers, sig), Field(action, 0)); + } + caml_process_pending_signals(); + CAMLreturn (res); +} diff --git a/test/monniaux/ocaml/byterun/signals_byt.c b/test/monniaux/ocaml/byterun/signals_byt.c new file mode 100644 index 00000000..bdbcf726 --- /dev/null +++ b/test/monniaux/ocaml/byterun/signals_byt.c @@ -0,0 +1,101 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 2007 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Signal handling, code specific to the bytecode interpreter */ + +#include +#include +#include "caml/config.h" +#include "caml/memory.h" +#include "caml/osdeps.h" +#include "caml/signals.h" +#include "caml/signals_machdep.h" + +#ifndef NSIG +#define NSIG 64 +#endif + +#ifdef _WIN32 +typedef void (*sighandler)(int sig); +extern sighandler caml_win32_signal(int sig, sighandler action); +#define signal(sig,act) caml_win32_signal(sig,act) +#endif + +CAMLexport int volatile caml_something_to_do = 0; +CAMLexport void (* volatile caml_async_action_hook)(void) = NULL; + +void caml_process_event(void) +{ + void (*async_action)(void); + + caml_check_urgent_gc (Val_unit); + caml_process_pending_signals(); + async_action = caml_async_action_hook; + if (async_action != NULL) { + caml_async_action_hook = NULL; + (*async_action)(); + } +} + +static void handle_signal(int signal_number) +{ + int saved_errno; + /* Save the value of errno (PR#5982). */ + saved_errno = errno; +#if !defined(POSIX_SIGNALS) && !defined(BSD_SIGNALS) + signal(signal_number, handle_signal); +#endif + if (signal_number < 0 || signal_number >= NSIG) return; + if (caml_try_leave_blocking_section_hook()) { + caml_execute_signal(signal_number, 1); + caml_enter_blocking_section_hook(); + }else{ + caml_record_signal(signal_number); + } + errno = saved_errno; +} + +int caml_set_signal_action(int signo, int action) +{ + void (*act)(int signo), (*oldact)(int signo); +#ifdef POSIX_SIGNALS + struct sigaction sigact, oldsigact; +#endif + + switch (action) { + case 0: act = SIG_DFL; break; + case 1: act = SIG_IGN; break; + default: act = handle_signal; break; + } + +#ifdef POSIX_SIGNALS + sigact.sa_handler = act; + sigemptyset(&sigact.sa_mask); + sigact.sa_flags = 0; + if (sigaction(signo, &sigact, &oldsigact) == -1) return -1; + oldact = oldsigact.sa_handler; +#else + oldact = signal(signo, act); + if (oldact == SIG_ERR) return -1; +#endif + if (oldact == handle_signal) + return 2; + else if (oldact == SIG_IGN) + return 1; + else + return 0; +} diff --git a/test/monniaux/ocaml/byterun/spacetime.c b/test/monniaux/ocaml/byterun/spacetime.c new file mode 100644 index 00000000..2b0bf1dc --- /dev/null +++ b/test/monniaux/ocaml/byterun/spacetime.c @@ -0,0 +1,38 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Mark Shinwell and Leo White, Jane Street Europe */ +/* */ +/* Copyright 2013--2016, Jane Street Group, LLC */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#include "caml/fail.h" +#include "caml/mlvalues.h" + +int caml_ensure_spacetime_dot_o_is_included = 42; + +CAMLprim value caml_spacetime_only_works_for_native_code(value foo, ...) +{ + caml_failwith("Spacetime profiling only works for native code"); +} + +uintnat caml_spacetime_my_profinfo (void) +{ + return 0; +} + +CAMLprim value caml_spacetime_enabled (value v_unit) +{ + return Val_false; /* running in bytecode */ +} + +CAMLprim value caml_register_channel_for_spacetime (value v_channel) +{ + return Val_unit; +} diff --git a/test/monniaux/ocaml/byterun/stacks.c b/test/monniaux/ocaml/byterun/stacks.c new file mode 100644 index 00000000..d6e7f53c --- /dev/null +++ b/test/monniaux/ocaml/byterun/stacks.c @@ -0,0 +1,119 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* To initialize and resize the stacks */ + +#include +#include "caml/config.h" +#include "caml/fail.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/stacks.h" + +CAMLexport value * caml_stack_low; +CAMLexport value * caml_stack_high; +CAMLexport value * caml_stack_threshold; +CAMLexport value * caml_extern_sp; +CAMLexport value * caml_trapsp; +CAMLexport value * caml_trap_barrier; +value caml_global_data = 0; + +uintnat caml_max_stack_size; /* also used in gc_ctrl.c */ + +void caml_init_stack (uintnat initial_max_size) +{ + caml_stack_low = (value *) caml_stat_alloc(Stack_size); + caml_stack_high = caml_stack_low + Stack_size / sizeof (value); + caml_stack_threshold = caml_stack_low + Stack_threshold / sizeof (value); + caml_extern_sp = caml_stack_high; + caml_trapsp = caml_stack_high; + caml_trap_barrier = caml_stack_high + 1; + caml_max_stack_size = initial_max_size; + caml_gc_message (0x08, "Initial stack limit: %" + ARCH_INTNAT_PRINTF_FORMAT "uk bytes\n", + caml_max_stack_size / 1024 * sizeof (value)); +} + +void caml_realloc_stack(asize_t required_space) +{ + asize_t size; + value * new_low, * new_high, * new_sp; + value * p; + + CAMLassert(caml_extern_sp >= caml_stack_low); + size = caml_stack_high - caml_stack_low; + do { + if (size >= caml_max_stack_size) caml_raise_stack_overflow(); + size *= 2; + } while (size < caml_stack_high - caml_extern_sp + required_space); + caml_gc_message (0x08, "Growing stack to %" + ARCH_INTNAT_PRINTF_FORMAT "uk bytes\n", + (uintnat) size * sizeof(value) / 1024); + new_low = (value *) caml_stat_alloc(size * sizeof(value)); + new_high = new_low + size; + +#define shift(ptr) \ + ((char *) new_high - ((char *) caml_stack_high - (char *) (ptr))) + + new_sp = (value *) shift(caml_extern_sp); + memmove((char *) new_sp, + (char *) caml_extern_sp, + (caml_stack_high - caml_extern_sp) * sizeof(value)); + caml_stat_free(caml_stack_low); + caml_trapsp = (value *) shift(caml_trapsp); + caml_trap_barrier = (value *) shift(caml_trap_barrier); + for (p = caml_trapsp; p < new_high; p = Trap_link(p)) + Trap_link(p) = (value *) shift(Trap_link(p)); + caml_stack_low = new_low; + caml_stack_high = new_high; + caml_stack_threshold = caml_stack_low + Stack_threshold / sizeof (value); + caml_extern_sp = new_sp; + +#undef shift +} + +CAMLprim value caml_ensure_stack_capacity(value required_space) +{ + asize_t req = Long_val(required_space); + if (caml_extern_sp - req < caml_stack_low) caml_realloc_stack(req); + return Val_unit; +} + +void caml_change_max_stack_size (uintnat new_max_size) +{ + asize_t size = caml_stack_high - caml_extern_sp + + Stack_threshold / sizeof (value); + + if (new_max_size < size) new_max_size = size; + if (new_max_size != caml_max_stack_size){ + caml_gc_message (0x08, "Changing stack limit to %" + ARCH_INTNAT_PRINTF_FORMAT "uk bytes\n", + new_max_size * sizeof (value) / 1024); + } + caml_max_stack_size = new_max_size; +} + +CAMLexport uintnat (*caml_stack_usage_hook)(void) = NULL; + +uintnat caml_stack_usage(void) +{ + uintnat sz; + sz = caml_stack_high - caml_extern_sp; + if (caml_stack_usage_hook != NULL) + sz += (*caml_stack_usage_hook)(); + return sz; +} diff --git a/test/monniaux/ocaml/byterun/startup.c b/test/monniaux/ocaml/byterun/startup.c new file mode 100644 index 00000000..037f30d0 --- /dev/null +++ b/test/monniaux/ocaml/byterun/startup.c @@ -0,0 +1,531 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Start-up code */ + +#include +#include +#include +#include +#include "caml/config.h" +#ifdef HAS_UNISTD +#include +#endif +#ifdef _WIN32 +#include +#endif +#include "caml/alloc.h" +#include "caml/backtrace.h" +#include "caml/callback.h" +#include "caml/custom.h" +#include "caml/debugger.h" +#include "caml/dynlink.h" +#include "caml/exec.h" +#include "caml/fail.h" +#include "caml/fix_code.h" +#include "caml/freelist.h" +#include "caml/gc_ctrl.h" +#include "caml/instrtrace.h" +#include "caml/interp.h" +#include "caml/intext.h" +#include "caml/io.h" +#include "caml/memory.h" +#include "caml/minor_gc.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/osdeps.h" +#include "caml/prims.h" +#include "caml/printexc.h" +#include "caml/reverse.h" +#include "caml/signals.h" +#include "caml/stacks.h" +#include "caml/sys.h" +#include "caml/startup.h" +#include "caml/startup_aux.h" +#include "caml/version.h" + +#ifndef O_BINARY +#define O_BINARY 0 +#endif + +#ifndef SEEK_END +#define SEEK_END 2 +#endif + +/* Read the trailer of a bytecode file */ + +static void fixup_endianness_trailer(uint32_t * p) +{ +#ifndef ARCH_BIG_ENDIAN + Reverse_32(p, p); +#endif +} + +static int read_trailer(int fd, struct exec_trailer *trail) +{ + if (lseek(fd, (long) -TRAILER_SIZE, SEEK_END) == -1) + return BAD_BYTECODE; + if (read(fd, (char *) trail, TRAILER_SIZE) < TRAILER_SIZE) + return BAD_BYTECODE; + fixup_endianness_trailer(&trail->num_sections); + if (strncmp(trail->magic, EXEC_MAGIC, 12) == 0) + return 0; + else + return BAD_BYTECODE; +} + +int caml_attempt_open(char_os **name, struct exec_trailer *trail, + int do_open_script) +{ + char_os * truename; + int fd; + int err; + char buf [2], * u8; + + truename = caml_search_exe_in_path(*name); + u8 = caml_stat_strdup_of_os(truename); + caml_gc_message(0x100, "Opening bytecode executable %s\n", u8); + caml_stat_free(u8); + fd = open_os(truename, O_RDONLY | O_BINARY); + if (fd == -1) { + caml_stat_free(truename); + caml_gc_message(0x100, "Cannot open file\n"); + return FILE_NOT_FOUND; + } + if (!do_open_script) { + err = read (fd, buf, 2); + if (err < 2 || (buf [0] == '#' && buf [1] == '!')) { + close(fd); + caml_stat_free(truename); + caml_gc_message(0x100, "Rejected #! script\n"); + return BAD_BYTECODE; + } + } + err = read_trailer(fd, trail); + if (err != 0) { + close(fd); + caml_stat_free(truename); + caml_gc_message(0x100, "Not a bytecode executable\n"); + return err; + } + *name = truename; + return fd; +} + +/* Read the section descriptors */ + +void caml_read_section_descriptors(int fd, struct exec_trailer *trail) +{ + int toc_size, i; + + toc_size = trail->num_sections * 8; + trail->section = caml_stat_alloc(toc_size); + lseek(fd, - (long) (TRAILER_SIZE + toc_size), SEEK_END); + if (read(fd, (char *) trail->section, toc_size) != toc_size) + caml_fatal_error("Fatal error: cannot read section table\n"); + /* Fixup endianness of lengths */ + for (i = 0; i < trail->num_sections; i++) + fixup_endianness_trailer(&(trail->section[i].len)); +} + +/* Position fd at the beginning of the section having the given name. + Return the length of the section data in bytes, or -1 if no section + found with that name. */ + +int32_t caml_seek_optional_section(int fd, struct exec_trailer *trail, + char *name) +{ + long ofs; + int i; + + ofs = TRAILER_SIZE + trail->num_sections * 8; + for (i = trail->num_sections - 1; i >= 0; i--) { + ofs += trail->section[i].len; + if (strncmp(trail->section[i].name, name, 4) == 0) { + lseek(fd, -ofs, SEEK_END); + return trail->section[i].len; + } + } + return -1; +} + +/* Position fd at the beginning of the section having the given name. + Return the length of the section data in bytes. */ + +int32_t caml_seek_section(int fd, struct exec_trailer *trail, char *name) +{ + int32_t len = caml_seek_optional_section(fd, trail, name); + if (len == -1) + caml_fatal_error_arg("Fatal_error: section `%s' is missing\n", name); + return len; +} + +/* Read and return the contents of the section having the given name. + Add a terminating 0. Return NULL if no such section. */ + +static char * read_section(int fd, struct exec_trailer *trail, char *name) +{ + int32_t len; + char * data; + + len = caml_seek_optional_section(fd, trail, name); + if (len == -1) return NULL; + data = caml_stat_alloc(len + 1); + if (read(fd, data, len) != len) + caml_fatal_error_arg("Fatal error: error reading section %s\n", name); + data[len] = 0; + return data; +} + +#ifdef _WIN32 + +static char_os * read_section_to_os(int fd, struct exec_trailer *trail, char *name) +{ + int32_t len, wlen; + char * data; + wchar_t * wdata; + + len = caml_seek_optional_section(fd, trail, name); + if (len == -1) return NULL; + data = caml_stat_alloc(len + 1); + if (read(fd, data, len) != len) + caml_fatal_error_arg("Fatal error: error reading section %s\n", name); + data[len] = 0; + wlen = win_multi_byte_to_wide_char(data, len, NULL, 0); + wdata = caml_stat_alloc((wlen + 1)*sizeof(wchar_t)); + win_multi_byte_to_wide_char(data, len, wdata, wlen); + wdata[wlen] = 0; + caml_stat_free(data); + return wdata; +} + +#else + +#define read_section_to_os read_section + +#endif + +/* Invocation of ocamlrun: 4 cases. + + 1. runtime + bytecode + user types: ocamlrun [options] bytecode args... + arguments: ocamlrun [options] bytecode args... + + 2. bytecode script + user types: bytecode args... + 2a (kernel 1) arguments: ocamlrun ./bytecode args... + 2b (kernel 2) arguments: bytecode bytecode args... + + 3. concatenated runtime and bytecode + user types: composite args... + arguments: composite args... + +Algorithm: + 1- If argument 0 is a valid byte-code file that does not start with #!, + then we are in case 3 and we pass the same command line to the + OCaml program. + 2- In all other cases, we parse the command line as: + (whatever) [options] bytecode args... + and we strip "(whatever) [options]" from the command line. + +*/ + +/* Parse options on the command line */ + +static int parse_command_line(char_os **argv) +{ + int i, j; + + for(i = 1; argv[i] != NULL && argv[i][0] == _T('-'); i++) { + switch(argv[i][1]) { + case _T('t'): + ++ caml_trace_level; /* ignored unless DEBUG mode */ + break; + case _T('v'): + if (!strcmp_os (argv[i], _T("-version"))){ + printf ("The OCaml runtime, version " OCAML_VERSION_STRING "\n"); + exit (0); + }else if (!strcmp_os (argv[i], _T("-vnum"))){ + printf (OCAML_VERSION_STRING "\n"); + exit (0); + }else{ + caml_verb_gc = 0x001+0x004+0x008+0x010+0x020; + } + break; + case _T('p'): + for (j = 0; caml_names_of_builtin_cprim[j] != NULL; j++) + printf("%s\n", caml_names_of_builtin_cprim[j]); + exit(0); + break; + case _T('b'): + caml_record_backtrace(Val_true); + break; + case _T('I'): + if (argv[i + 1] != NULL) { + caml_ext_table_add(&caml_shared_libs_path, argv[i + 1]); + i++; + } + break; + default: + caml_fatal_error_arg("Unknown option %s.\n", caml_stat_strdup_of_os(argv[i])); + } + } + return i; +} + +extern void caml_init_ieee_floats (void); + +#ifdef _WIN32 +extern void caml_signal_thread(void * lpParam); +#endif + +#if defined(_MSC_VER) && __STDC_SECURE_LIB__ >= 200411L + +/* PR 4887: avoid crash box of windows runtime on some system calls */ +extern void caml_install_invalid_parameter_handler(); + +#endif + +extern int caml_ensure_spacetime_dot_o_is_included; + +/* Main entry point when loading code from a file */ + +CAMLexport void caml_main(char_os **argv) +{ + int fd, pos; + struct exec_trailer trail; + struct channel * chan; + value res; + char * req_prims; + char_os * shared_lib_path, * shared_libs; + char_os * exe_name, * proc_self_exe; + + caml_ensure_spacetime_dot_o_is_included++; + + /* Determine options */ +#ifdef DEBUG + caml_verb_gc = 0x3F; +#endif + caml_parse_ocamlrunparam(); +#ifdef DEBUG + caml_gc_message (-1, "### OCaml runtime: debug mode ###\n"); +#endif + if (!caml_startup_aux(/* pooling */ caml_cleanup_on_exit)) + return; + + /* Machine-dependent initialization of the floating-point hardware + so that it behaves as much as possible as specified in IEEE */ + caml_init_ieee_floats(); +#if defined(_MSC_VER) && __STDC_SECURE_LIB__ >= 200411L + caml_install_invalid_parameter_handler(); +#endif + caml_init_custom_operations(); + caml_ext_table_init(&caml_shared_libs_path, 8); + caml_external_raise = NULL; + + /* Determine position of bytecode file */ + pos = 0; + + /* First, try argv[0] (when ocamlrun is called by a bytecode program) */ + exe_name = argv[0]; + fd = caml_attempt_open(&exe_name, &trail, 0); + + /* Little grasshopper wonders why we do that at all, since + "The current executable is ocamlrun itself, it's never a bytecode + program". Little grasshopper "ocamlc -custom" in mind should keep. + With -custom, we have an executable that is ocamlrun itself + concatenated with the bytecode. So, if the attempt with argv[0] + failed, it is worth trying again with executable_name. */ + if (fd < 0 && (proc_self_exe = caml_executable_name()) != NULL) { + exe_name = proc_self_exe; + fd = caml_attempt_open(&exe_name, &trail, 0); + } + + if (fd < 0) { + pos = parse_command_line(argv); + if (argv[pos] == 0) + caml_fatal_error("No bytecode file specified.\n"); + exe_name = argv[pos]; + fd = caml_attempt_open(&exe_name, &trail, 1); + switch(fd) { + case FILE_NOT_FOUND: + caml_fatal_error_arg("Fatal error: cannot find file '%s'\n", caml_stat_strdup_of_os(argv[pos])); + break; + case BAD_BYTECODE: + caml_fatal_error_arg( + "Fatal error: the file '%s' is not a bytecode executable file\n", + caml_stat_strdup_of_os(exe_name)); + break; + } + } + /* Read the table of contents (section descriptors) */ + caml_read_section_descriptors(fd, &trail); + /* Initialize the abstract machine */ + caml_init_gc (caml_init_minor_heap_wsz, caml_init_heap_wsz, + caml_init_heap_chunk_sz, caml_init_percent_free, + caml_init_max_percent_free, caml_init_major_window); + caml_init_stack (caml_init_max_stack_wsz); + caml_init_atom_table(); + caml_init_backtrace(); + /* Initialize the interpreter */ + caml_interprete(NULL, 0); + /* Initialize the debugger, if needed */ + caml_debugger_init(); + /* Load the code */ + caml_code_size = caml_seek_section(fd, &trail, "CODE"); + caml_load_code(fd, caml_code_size); + caml_init_debug_info(); + /* Build the table of primitives */ + shared_lib_path = read_section_to_os(fd, &trail, "DLPT"); + shared_libs = read_section_to_os(fd, &trail, "DLLS"); + req_prims = read_section(fd, &trail, "PRIM"); + if (req_prims == NULL) caml_fatal_error("Fatal error: no PRIM section\n"); + caml_build_primitive_table(shared_lib_path, shared_libs, req_prims); + caml_stat_free(shared_lib_path); + caml_stat_free(shared_libs); + caml_stat_free(req_prims); + /* Load the globals */ + caml_seek_section(fd, &trail, "DATA"); + chan = caml_open_descriptor_in(fd); + caml_global_data = caml_input_val(chan); + caml_close_channel(chan); /* this also closes fd */ + caml_stat_free(trail.section); + /* Ensure that the globals are in the major heap. */ + caml_oldify_one (caml_global_data, &caml_global_data); + caml_oldify_mopup (); + /* Initialize system libraries */ + caml_sys_init(exe_name, argv + pos); +#ifdef _WIN32 + /* Start a thread to handle signals */ + if (caml_secure_getenv(_T("CAMLSIGPIPE"))) + _beginthread(caml_signal_thread, 4096, NULL); +#endif + /* Execute the program */ + caml_debugger(PROGRAM_START); + res = caml_interprete(caml_start_code, caml_code_size); + if (Is_exception_result(res)) { + caml_exn_bucket = Extract_exception(res); + if (caml_debugger_in_use) { + caml_extern_sp = &caml_exn_bucket; /* The debugger needs the + exception value.*/ + caml_debugger(UNCAUGHT_EXC); + } + caml_fatal_uncaught_exception(caml_exn_bucket); + } +} + +/* Main entry point when code is linked in as initialized data */ + +CAMLexport value caml_startup_code_exn( + code_t code, asize_t code_size, + char *data, asize_t data_size, + char *section_table, asize_t section_table_size, + int pooling, + char_os **argv) +{ + char_os * cds_file; + char_os * exe_name; + + /* Determine options */ +#ifdef DEBUG + caml_verb_gc = 0x3F; +#endif + caml_parse_ocamlrunparam(); +#ifdef DEBUG + caml_gc_message (-1, "### OCaml runtime: debug mode ###\n"); +#endif + if (caml_cleanup_on_exit) + pooling = 1; + if (!caml_startup_aux(pooling)) + return Val_unit; + + caml_init_ieee_floats(); +#if defined(_MSC_VER) && __STDC_SECURE_LIB__ >= 200411L + caml_install_invalid_parameter_handler(); +#endif + caml_init_custom_operations(); + cds_file = caml_secure_getenv(_T("CAML_DEBUG_FILE")); + if (cds_file != NULL) { + caml_cds_file = caml_stat_strdup_os(cds_file); + } + exe_name = caml_executable_name(); + if (exe_name == NULL) exe_name = caml_search_exe_in_path(argv[0]); + caml_external_raise = NULL; + /* Initialize the abstract machine */ + caml_init_gc (caml_init_minor_heap_wsz, caml_init_heap_wsz, + caml_init_heap_chunk_sz, caml_init_percent_free, + caml_init_max_percent_free, caml_init_major_window); + caml_init_stack (caml_init_max_stack_wsz); + caml_init_atom_table(); + caml_init_backtrace(); + /* Initialize the interpreter */ + caml_interprete(NULL, 0); + /* Initialize the debugger, if needed */ + caml_debugger_init(); + /* Load the code */ + caml_start_code = code; + caml_code_size = code_size; + caml_init_code_fragments(); + caml_init_debug_info(); + if (caml_debugger_in_use) { + int len, i; + len = code_size / sizeof(opcode_t); + caml_saved_code = (unsigned char *) caml_stat_alloc(len); + for (i = 0; i < len; i++) caml_saved_code[i] = caml_start_code[i]; + } +#ifdef THREADED_CODE + caml_thread_code(caml_start_code, code_size); +#endif + /* Use the builtin table of primitives */ + caml_build_primitive_table_builtin(); + /* Load the globals */ + caml_global_data = caml_input_value_from_block(data, data_size); + /* Ensure that the globals are in the major heap. */ + caml_oldify_one (caml_global_data, &caml_global_data); + caml_oldify_mopup (); + /* Record the sections (for caml_get_section_table in meta.c) */ + caml_section_table = section_table; + caml_section_table_size = section_table_size; + /* Initialize system libraries */ + caml_sys_init(exe_name, argv); + /* Execute the program */ + caml_debugger(PROGRAM_START); + return caml_interprete(caml_start_code, caml_code_size); +} + +CAMLexport void caml_startup_code( + code_t code, asize_t code_size, + char *data, asize_t data_size, + char *section_table, asize_t section_table_size, + int pooling, + char_os **argv) +{ + value res; + + res = caml_startup_code_exn(code, code_size, data, data_size, + section_table, section_table_size, + pooling, argv); + if (Is_exception_result(res)) { + caml_exn_bucket = Extract_exception(res); + if (caml_debugger_in_use) { + caml_extern_sp = &caml_exn_bucket; /* The debugger needs the + exception value.*/ + caml_debugger(UNCAUGHT_EXC); + } + caml_fatal_uncaught_exception(caml_exn_bucket); + } +} diff --git a/test/monniaux/ocaml/byterun/startup_aux.c b/test/monniaux/ocaml/byterun/startup_aux.c new file mode 100644 index 00000000..b913bfae --- /dev/null +++ b/test/monniaux/ocaml/byterun/startup_aux.c @@ -0,0 +1,168 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy and Damien Doligez, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Some runtime initialization functions that are common to bytecode + and native code. */ + +#include +#include "caml/backtrace.h" +#include "caml/memory.h" +#include "caml/callback.h" +#include "caml/major_gc.h" +#ifndef NATIVE_CODE +#include "caml/dynlink.h" +#endif +#include "caml/osdeps.h" +#include "caml/startup_aux.h" + + +/* Initialize the atom table */ + +CAMLexport header_t caml_atom_table[256]; +void caml_init_atom_table(void) +{ + int i; + for(i = 0; i < 256; i++) { +#ifdef NATIVE_CODE + caml_atom_table[i] = Make_header_allocated_here(0, i, Caml_white); +#else + caml_atom_table[i] = Make_header(0, i, Caml_white); +#endif + } + if (caml_page_table_add(In_static_data, + caml_atom_table, caml_atom_table + 256) != 0) { + caml_fatal_error("Fatal error: not enough memory for initial page table"); + } +} + + +/* Parse the OCAMLRUNPARAM environment variable. */ + +uintnat caml_init_percent_free = Percent_free_def; +uintnat caml_init_max_percent_free = Max_percent_free_def; +uintnat caml_init_minor_heap_wsz = Minor_heap_def; +uintnat caml_init_heap_chunk_sz = Heap_chunk_def; +uintnat caml_init_heap_wsz = Init_heap_def; +uintnat caml_init_max_stack_wsz = Max_stack_def; +uintnat caml_init_major_window = Major_window_def; +extern int caml_parser_trace; +uintnat caml_trace_level = 0; +uintnat caml_cleanup_on_exit = 0; + + +static void scanmult (char_os *opt, uintnat *var) +{ + char_os mult = _T(' '); + unsigned int val = 1; + sscanf_os (opt, _T("=%u%c"), &val, &mult); + sscanf_os (opt, _T("=0x%x%c"), &val, &mult); + switch (mult) { + case _T('k'): *var = (uintnat) val * 1024; break; + case _T('M'): *var = (uintnat) val * (1024 * 1024); break; + case _T('G'): *var = (uintnat) val * (1024 * 1024 * 1024); break; + default: *var = (uintnat) val; break; + } +} + +void caml_parse_ocamlrunparam(void) +{ + char_os *opt = caml_secure_getenv (_T("OCAMLRUNPARAM")); + uintnat p; + + if (opt == NULL) opt = caml_secure_getenv (_T("CAMLRUNPARAM")); + + if (opt != NULL){ + while (*opt != _T('\0')){ + switch (*opt++){ + case _T('a'): scanmult (opt, &p); caml_set_allocation_policy (p); break; + case _T('b'): scanmult (opt, &p); caml_record_backtrace(Val_bool (p)); break; + case _T('c'): scanmult (opt, &p); caml_cleanup_on_exit = p; break; + case _T('h'): scanmult (opt, &caml_init_heap_wsz); break; + case _T('H'): scanmult (opt, &caml_use_huge_pages); break; + case _T('i'): scanmult (opt, &caml_init_heap_chunk_sz); break; + case _T('l'): scanmult (opt, &caml_init_max_stack_wsz); break; + case _T('o'): scanmult (opt, &caml_init_percent_free); break; + case _T('O'): scanmult (opt, &caml_init_max_percent_free); break; + case _T('p'): scanmult (opt, &p); caml_parser_trace = p; break; + case _T('R'): break; /* see stdlib/hashtbl.mli */ + case _T('s'): scanmult (opt, &caml_init_minor_heap_wsz); break; + case _T('t'): scanmult (opt, &caml_trace_level); break; + case _T('v'): scanmult (opt, &caml_verb_gc); break; + case _T('w'): scanmult (opt, &caml_init_major_window); break; + case _T('W'): scanmult (opt, &caml_runtime_warnings); break; + } + while (*opt != _T('\0')){ + if (*opt++ == ',') break; + } + } + } +} + + +/* The number of outstanding calls to caml_startup */ +static int startup_count = 0; + +/* Has the runtime been shut down already? */ +static int shutdown_happened = 0; + + +int caml_startup_aux(int pooling) +{ + if (shutdown_happened == 1) + caml_fatal_error("Fatal error: caml_startup was called after the runtime " + "was shut down with caml_shutdown"); + + /* Second and subsequent calls are ignored, + since the runtime has already started */ + startup_count++; + if (startup_count > 1) + return 0; + + if (pooling) + caml_stat_create_pool(); + + return 1; +} + +static void call_registered_value(char* name) +{ + value *f = caml_named_value(name); + if (f != NULL) + caml_callback_exn(*f, Val_unit); +} + +CAMLexport void caml_shutdown(void) +{ + if (startup_count <= 0) + caml_fatal_error("Fatal error: a call to caml_shutdown has no " + "corresponding call to caml_startup"); + + /* Do nothing unless it's the last call remaining */ + startup_count--; + if (startup_count > 0) + return; + + call_registered_value("Pervasives.do_at_exit"); + call_registered_value("Thread.at_shutdown"); + caml_finalise_heap(); +#ifndef NATIVE_CODE + caml_free_shared_libs(); +#endif + caml_stat_destroy_pool(); + + shutdown_happened = 1; +} diff --git a/test/monniaux/ocaml/byterun/str.c b/test/monniaux/ocaml/byterun/str.c new file mode 100644 index 00000000..8e07cb03 --- /dev/null +++ b/test/monniaux/ocaml/byterun/str.c @@ -0,0 +1,474 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Operations on strings */ + +#include +#include +#include +#include +#include "caml/alloc.h" +#include "caml/fail.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/misc.h" + +/* returns a number of bytes (chars) */ +CAMLexport mlsize_t caml_string_length(value s) +{ + mlsize_t temp; + temp = Bosize_val(s) - 1; + CAMLassert (Byte (s, temp - Byte (s, temp)) == 0); + return temp - Byte (s, temp); +} + +/* returns a value that represents a number of bytes (chars) */ +CAMLprim value caml_ml_string_length(value s) +{ + mlsize_t temp; + temp = Bosize_val(s) - 1; + CAMLassert (Byte (s, temp - Byte (s, temp)) == 0); + return Val_long(temp - Byte (s, temp)); +} + +CAMLprim value caml_ml_bytes_length(value s) +{ + return caml_ml_string_length(s); +} + +CAMLexport int caml_string_is_c_safe (value s) +{ + return strlen(String_val(s)) == caml_string_length(s); +} + +/** + * [caml_create_string] is deprecated, + * use [caml_create_bytes] instead + */ +CAMLprim value caml_create_string(value len) +{ + mlsize_t size = Long_val(len); + if (size > Bsize_wsize (Max_wosize) - 1){ + caml_invalid_argument("String.create"); + } + return caml_alloc_string(size); +} + +/* [len] is a value that represents a number of bytes (chars) */ +CAMLprim value caml_create_bytes(value len) +{ + mlsize_t size = Long_val(len); + if (size > Bsize_wsize (Max_wosize) - 1){ + caml_invalid_argument("Bytes.create"); + } + return caml_alloc_string(size); +} + + + +CAMLprim value caml_string_get(value str, value index) +{ + intnat idx = Long_val(index); + if (idx < 0 || idx >= caml_string_length(str)) caml_array_bound_error(); + return Val_int(Byte_u(str, idx)); +} + +CAMLprim value caml_bytes_get(value str, value index) +{ + return caml_string_get(str, index); +} + +CAMLprim value caml_bytes_set(value str, value index, value newval) +{ + intnat idx = Long_val(index); + if (idx < 0 || idx >= caml_string_length(str)) caml_array_bound_error(); + Byte_u(str, idx) = Int_val(newval); + return Val_unit; +} + +/** + * [caml_string_set] is deprecated, + * use [caml_bytes_set] instead + */ +CAMLprim value caml_string_set(value str, value index, value newval) +{ + return caml_bytes_set(str,index,newval); +} + + +CAMLprim value caml_string_get16(value str, value index) +{ + intnat res; + unsigned char b1, b2; + intnat idx = Long_val(index); + if (idx < 0 || idx + 1 >= caml_string_length(str)) caml_array_bound_error(); + b1 = Byte_u(str, idx); + b2 = Byte_u(str, idx + 1); +#ifdef ARCH_BIG_ENDIAN + res = b1 << 8 | b2; +#else + res = b2 << 8 | b1; +#endif + return Val_int(res); +} + +CAMLprim value caml_bytes_get16(value str, value index) +{ + return caml_string_get16(str,index); +} + +CAMLprim value caml_string_get32(value str, value index) +{ + intnat res; + unsigned char b1, b2, b3, b4; + intnat idx = Long_val(index); + if (idx < 0 || idx + 3 >= caml_string_length(str)) caml_array_bound_error(); + b1 = Byte_u(str, idx); + b2 = Byte_u(str, idx + 1); + b3 = Byte_u(str, idx + 2); + b4 = Byte_u(str, idx + 3); +#ifdef ARCH_BIG_ENDIAN + res = b1 << 24 | b2 << 16 | b3 << 8 | b4; +#else + res = b4 << 24 | b3 << 16 | b2 << 8 | b1; +#endif + return caml_copy_int32(res); +} + +CAMLprim value caml_bytes_get32(value str, value index) +{ + return caml_string_get32(str,index); +} + +CAMLprim value caml_string_get64(value str, value index) +{ + uint64_t res; + unsigned char b1, b2, b3, b4, b5, b6, b7, b8; + intnat idx = Long_val(index); + if (idx < 0 || idx + 7 >= caml_string_length(str)) caml_array_bound_error(); + b1 = Byte_u(str, idx); + b2 = Byte_u(str, idx + 1); + b3 = Byte_u(str, idx + 2); + b4 = Byte_u(str, idx + 3); + b5 = Byte_u(str, idx + 4); + b6 = Byte_u(str, idx + 5); + b7 = Byte_u(str, idx + 6); + b8 = Byte_u(str, idx + 7); +#ifdef ARCH_BIG_ENDIAN + res = (uint64_t) b1 << 56 | (uint64_t) b2 << 48 + | (uint64_t) b3 << 40 | (uint64_t) b4 << 32 + | (uint64_t) b5 << 24 | (uint64_t) b6 << 16 + | (uint64_t) b7 << 8 | (uint64_t) b8; +#else + res = (uint64_t) b8 << 56 | (uint64_t) b7 << 48 + | (uint64_t) b6 << 40 | (uint64_t) b5 << 32 + | (uint64_t) b4 << 24 | (uint64_t) b3 << 16 + | (uint64_t) b2 << 8 | (uint64_t) b1; +#endif + return caml_copy_int64(res); +} + +CAMLprim value caml_bytes_get64(value str, value index) +{ + return caml_string_get64(str,index); +} + +CAMLprim value caml_bytes_set16(value str, value index, value newval) +{ + unsigned char b1, b2; + intnat val; + intnat idx = Long_val(index); + if (idx < 0 || idx + 1 >= caml_string_length(str)) caml_array_bound_error(); + val = Long_val(newval); +#ifdef ARCH_BIG_ENDIAN + b1 = 0xFF & val >> 8; + b2 = 0xFF & val; +#else + b2 = 0xFF & val >> 8; + b1 = 0xFF & val; +#endif + Byte_u(str, idx) = b1; + Byte_u(str, idx + 1) = b2; + return Val_unit; +} + +CAMLprim value caml_bytes_set32(value str, value index, value newval) +{ + unsigned char b1, b2, b3, b4; + intnat val; + intnat idx = Long_val(index); + if (idx < 0 || idx + 3 >= caml_string_length(str)) caml_array_bound_error(); + val = Int32_val(newval); +#ifdef ARCH_BIG_ENDIAN + b1 = 0xFF & val >> 24; + b2 = 0xFF & val >> 16; + b3 = 0xFF & val >> 8; + b4 = 0xFF & val; +#else + b4 = 0xFF & val >> 24; + b3 = 0xFF & val >> 16; + b2 = 0xFF & val >> 8; + b1 = 0xFF & val; +#endif + Byte_u(str, idx) = b1; + Byte_u(str, idx + 1) = b2; + Byte_u(str, idx + 2) = b3; + Byte_u(str, idx + 3) = b4; + return Val_unit; +} + +CAMLprim value caml_bytes_set64(value str, value index, value newval) +{ + unsigned char b1, b2, b3, b4, b5, b6, b7, b8; + int64_t val; + intnat idx = Long_val(index); + if (idx < 0 || idx + 7 >= caml_string_length(str)) caml_array_bound_error(); + val = Int64_val(newval); +#ifdef ARCH_BIG_ENDIAN + b1 = 0xFF & val >> 56; + b2 = 0xFF & val >> 48; + b3 = 0xFF & val >> 40; + b4 = 0xFF & val >> 32; + b5 = 0xFF & val >> 24; + b6 = 0xFF & val >> 16; + b7 = 0xFF & val >> 8; + b8 = 0xFF & val; +#else + b8 = 0xFF & val >> 56; + b7 = 0xFF & val >> 48; + b6 = 0xFF & val >> 40; + b5 = 0xFF & val >> 32; + b4 = 0xFF & val >> 24; + b3 = 0xFF & val >> 16; + b2 = 0xFF & val >> 8; + b1 = 0xFF & val; +#endif + Byte_u(str, idx) = b1; + Byte_u(str, idx + 1) = b2; + Byte_u(str, idx + 2) = b3; + Byte_u(str, idx + 3) = b4; + Byte_u(str, idx + 4) = b5; + Byte_u(str, idx + 5) = b6; + Byte_u(str, idx + 6) = b7; + Byte_u(str, idx + 7) = b8; + return Val_unit; +} + +CAMLprim value caml_string_equal(value s1, value s2) +{ + mlsize_t sz1, sz2; + value * p1, * p2; + + if (s1 == s2) return Val_true; + sz1 = Wosize_val(s1); + sz2 = Wosize_val(s2); + if (sz1 != sz2) return Val_false; + for(p1 = Op_val(s1), p2 = Op_val(s2); sz1 > 0; sz1--, p1++, p2++) + if (*p1 != *p2) return Val_false; + return Val_true; +} + +CAMLprim value caml_bytes_equal(value s1, value s2) +{ + return caml_string_equal(s1,s2); +} + +CAMLprim value caml_string_notequal(value s1, value s2) +{ + return Val_not(caml_string_equal(s1, s2)); +} + +CAMLprim value caml_bytes_notequal(value s1, value s2) +{ + return caml_string_notequal(s1,s2); +} + +CAMLprim value caml_string_compare(value s1, value s2) +{ + mlsize_t len1, len2; + int res; + + if (s1 == s2) return Val_int(0); + len1 = caml_string_length(s1); + len2 = caml_string_length(s2); + res = memcmp(String_val(s1), String_val(s2), len1 <= len2 ? len1 : len2); + if (res < 0) return Val_int(-1); + if (res > 0) return Val_int(1); + if (len1 < len2) return Val_int(-1); + if (len1 > len2) return Val_int(1); + return Val_int(0); +} + +CAMLprim value caml_bytes_compare(value s1, value s2) +{ + return caml_string_compare(s1,s2); +} + +CAMLprim value caml_string_lessthan(value s1, value s2) +{ + return caml_string_compare(s1, s2) < Val_int(0) ? Val_true : Val_false; +} + +CAMLprim value caml_bytes_lessthan(value s1, value s2) +{ + return caml_string_lessthan(s1,s2); +} + + +CAMLprim value caml_string_lessequal(value s1, value s2) +{ + return caml_string_compare(s1, s2) <= Val_int(0) ? Val_true : Val_false; +} + +CAMLprim value caml_bytes_lessequal(value s1, value s2) +{ + return caml_string_lessequal(s1,s2); +} + + +CAMLprim value caml_string_greaterthan(value s1, value s2) +{ + return caml_string_compare(s1, s2) > Val_int(0) ? Val_true : Val_false; +} + +CAMLprim value caml_bytes_greaterthan(value s1, value s2) +{ + return caml_string_greaterthan(s1,s2); +} + +CAMLprim value caml_string_greaterequal(value s1, value s2) +{ + return caml_string_compare(s1, s2) >= Val_int(0) ? Val_true : Val_false; +} + +CAMLprim value caml_bytes_greaterequal(value s1, value s2) +{ + return caml_string_greaterequal(s1,s2); +} + +CAMLprim value caml_blit_bytes(value s1, value ofs1, value s2, value ofs2, + value n) +{ + memmove(&Byte(s2, Long_val(ofs2)), &Byte(s1, Long_val(ofs1)), Long_val(n)); + return Val_unit; +} + +CAMLprim value caml_blit_string(value s1, value ofs1, value s2, value ofs2, + value n) +{ + return caml_blit_bytes (s1, ofs1, s2, ofs2, n); +} + +CAMLprim value caml_fill_bytes(value s, value offset, value len, value init) +{ + memset(&Byte(s, Long_val(offset)), Int_val(init), Long_val(len)); + return Val_unit; +} + +/** + * [caml_fill_string] is deprecated, use [caml_fill_bytes] instead + */ +CAMLprim value caml_fill_string(value s, value offset, value len, value init) +{ + return caml_fill_bytes (s, offset, len, init); +} + +CAMLexport value caml_alloc_sprintf(const char * format, ...) +{ + va_list args; + char buf[128]; + int n; + value res; + +#if !defined(_WIN32) || defined(_UCRT) + /* C99-compliant implementation */ + va_start(args, format); + /* "vsnprintf(dest, sz, format, args)" writes at most "sz" characters + into "dest", including the terminating '\0'. + It returns the number of characters of the formatted string, + excluding the terminating '\0'. */ + n = vsnprintf(buf, sizeof(buf), format, args); + va_end(args); + if (n < sizeof(buf)) { + /* All output characters were written to buf, including the + terminating '\0'. Allocate a Caml string with length "n" + as computed by vsnprintf, and copy the output of vsnprintf into it. */ + res = caml_alloc_initialized_string(n, buf); + } else { + /* PR#7568: if the format is in the Caml heap, the following + caml_alloc_string could move or free the format. To prevent + this, take a copy of the format outside the Caml heap. */ + char * saved_format = caml_stat_strdup(format); + /* Allocate a Caml string with length "n" as computed by vsnprintf. */ + res = caml_alloc_string(n); + /* Re-do the formatting, outputting directly in the Caml string. + Note that caml_alloc_string left room for a '\0' at position n, + so the size passed to vsnprintf is n+1. */ + va_start(args, format); + vsnprintf((char *)String_val(res), n + 1, saved_format, args); + va_end(args); + caml_stat_free(saved_format); + } + return res; +#else + /* Implementation specific to the Microsoft CRT library */ + va_start(args, format); + /* "_vsnprintf(dest, sz, format, args)" writes at most "sz" characters + into "dest". Let "len" be the number of characters of the formatted + string. + If "len" < "sz", a null terminator was appended, and "len" is returned. + If "len" == "sz", no null termination, and "len" is returned. + If "len" > "sz", a negative value is returned. */ + n = _vsnprintf(buf, sizeof(buf), format, args); + va_end(args); + if (n >= 0 && n <= sizeof(buf)) { + /* All output characters were written to buf. + "n" is the actual length of the output. + Allocate a Caml string of length "n" and copy the characters into it. */ + res = caml_alloc_string(n); + memcpy(String_val(res), buf, n); + } else { + /* PR#7568: if the format is in the Caml heap, the following + caml_alloc_string could move or free the format. To prevent + this, take a copy of the format outside the Caml heap. */ + char * saved_format = caml_stat_strdup(format); + /* Determine actual length of output, excluding final '\0' */ + va_start(args, format); + n = _vscprintf(format, args); + va_end(args); + res = caml_alloc_string(n); + /* Re-do the formatting, outputting directly in the Caml string. + Note that caml_alloc_string left room for a '\0' at position n, + so the size passed to _vsnprintf is n+1. */ + va_start(args, format); + _vsnprintf(String_val(res), n + 1, saved_format, args); + va_end(args); + caml_stat_free(saved_format); + } + return res; +#endif +} + +CAMLprim value caml_string_of_bytes(value bv) +{ + return bv; +} + +CAMLprim value caml_bytes_of_string(value bv) +{ + return bv; +} diff --git a/test/monniaux/ocaml/byterun/sys.c b/test/monniaux/ocaml/byterun/sys.c new file mode 100644 index 00000000..84c42ddf --- /dev/null +++ b/test/monniaux/ocaml/byterun/sys.c @@ -0,0 +1,723 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* DM HORRIBLE */ +unsigned __bsp_frequency=500000000; + +#define CAML_INTERNALS + +/* Basic system calls */ + +#include +#include +#include +#include +#include +#include +#include +#include +#include +#ifdef _WIN32 +#include /* for _wchdir and _wgetcwd */ +#else +#include +#endif +#include "caml/config.h" +#ifdef HAS_UNISTD +#include +#endif +#ifdef HAS_TIMES +#include +#endif +#ifdef HAS_GETRUSAGE +#include +#include +#endif +#ifdef HAS_GETTIMEOFDAY +#include +#endif +#include "caml/alloc.h" +#include "caml/debugger.h" +#include "caml/fail.h" +#include "caml/gc_ctrl.h" +#include "caml/io.h" +#include "caml/misc.h" +#include "caml/mlvalues.h" +#include "caml/osdeps.h" +#include "caml/signals.h" +#include "caml/stacks.h" +#include "caml/sys.h" +#include "caml/version.h" +#include "caml/callback.h" +#include "caml/startup_aux.h" + +static char * error_message(void) +{ + return strerror(errno); +} + +#ifndef EAGAIN +#define EAGAIN (-1) +#endif +#ifndef EWOULDBLOCK +#define EWOULDBLOCK (-1) +#endif + +CAMLexport void caml_sys_error(value arg) +{ + CAMLparam1 (arg); + char * err; + CAMLlocal1 (str); + + err = error_message(); + if (arg == NO_ARG) { + str = caml_copy_string(err); + } else { + int err_len = strlen(err); + int arg_len = caml_string_length(arg); + str = caml_alloc_string(arg_len + 2 + err_len); + memmove(&Byte(str, 0), String_val(arg), arg_len); + memmove(&Byte(str, arg_len), ": ", 2); + memmove(&Byte(str, arg_len + 2), err, err_len); + } + caml_raise_sys_error(str); + CAMLnoreturn; +} + +CAMLexport void caml_sys_io_error(value arg) +{ + if (errno == EAGAIN || errno == EWOULDBLOCK) { + caml_raise_sys_blocked_io(); + } else { + caml_sys_error(arg); + } +} + +/* Check that [name] can safely be used as a file path */ + +static void caml_sys_check_path(value name) +{ + if (! caml_string_is_c_safe(name)) { + errno = ENOENT; + caml_sys_error(name); + } +} + +CAMLprim value caml_sys_exit(value retcode_v) +{ + int retcode = Int_val(retcode_v); + + if ((caml_verb_gc & 0x400) != 0) { + /* cf caml_gc_counters */ + double minwords = caml_stat_minor_words + + (double) (caml_young_end - caml_young_ptr); + double prowords = caml_stat_promoted_words; + double majwords = caml_stat_major_words + (double) caml_allocated_words; + double allocated_words = minwords + majwords - prowords; + intnat mincoll = caml_stat_minor_collections; + intnat majcoll = caml_stat_major_collections; + intnat heap_words = caml_stat_heap_wsz; + intnat heap_chunks = caml_stat_heap_chunks; + intnat top_heap_words = caml_stat_top_heap_wsz; + intnat cpct = caml_stat_compactions; + caml_gc_message(0x400, "allocated_words: %.0f\n", allocated_words); + caml_gc_message(0x400, "minor_words: %.0f\n", minwords); + caml_gc_message(0x400, "promoted_words: %.0f\n", prowords); + caml_gc_message(0x400, "major_words: %.0f\n", majwords); + caml_gc_message(0x400, "minor_collections: %"ARCH_INTNAT_PRINTF_FORMAT"d\n", + mincoll); + caml_gc_message(0x400, "major_collections: %"ARCH_INTNAT_PRINTF_FORMAT"d\n", + majcoll); + caml_gc_message(0x400, "heap_words: %"ARCH_INTNAT_PRINTF_FORMAT"d\n", + heap_words); + caml_gc_message(0x400, "heap_chunks: %"ARCH_INTNAT_PRINTF_FORMAT"d\n", + heap_chunks); + caml_gc_message(0x400, "top_heap_words: %"ARCH_INTNAT_PRINTF_FORMAT"d\n", + top_heap_words); + caml_gc_message(0x400, "compactions: %"ARCH_INTNAT_PRINTF_FORMAT"d\n", + cpct); + } + +#ifndef NATIVE_CODE + caml_debugger(PROGRAM_EXIT); +#endif + CAML_INSTR_ATEXIT (); + if (caml_cleanup_on_exit) + caml_shutdown(); +#ifdef _WIN32 + caml_restore_win32_terminal(); +#endif + CAML_SYS_EXIT(retcode); + return Val_unit; +} + +#ifndef O_BINARY +#define O_BINARY 0 +#endif +#ifndef O_TEXT +#define O_TEXT 0 +#endif +#ifndef O_NONBLOCK +#ifdef O_NDELAY +#define O_NONBLOCK O_NDELAY +#else +#define O_NONBLOCK 0 +#endif +#endif + +static int sys_open_flags[] = { + O_RDONLY, O_WRONLY, O_APPEND | O_WRONLY, O_CREAT, O_TRUNC, O_EXCL, + O_BINARY, O_TEXT, O_NONBLOCK +}; + +CAMLprim value caml_sys_open(value path, value vflags, value vperm) +{ + CAMLparam3(path, vflags, vperm); + int fd, flags, perm; + char_os * p; + +#if defined(O_CLOEXEC) + flags = O_CLOEXEC; +#elif defined(_WIN32) + flags = _O_NOINHERIT; +#else + flags = 0; +#endif + + caml_sys_check_path(path); + p = caml_stat_strdup_to_os(String_val(path)); + flags |= caml_convert_flag_list(vflags, sys_open_flags); + perm = Int_val(vperm); + /* open on a named FIFO can block (PR#1533) */ + caml_enter_blocking_section(); + fd = CAML_SYS_OPEN(p, flags, perm); + /* fcntl on a fd can block (PR#5069)*/ +#if defined(F_SETFD) && defined(FD_CLOEXEC) && !defined(_WIN32) \ + && !defined(O_CLOEXEC) + if (fd != -1) + fcntl(fd, F_SETFD, FD_CLOEXEC); +#endif + caml_leave_blocking_section(); + caml_stat_free(p); + if (fd == -1) caml_sys_error(path); + CAMLreturn(Val_long(fd)); +} + +CAMLprim value caml_sys_close(value fd_v) +{ + int fd = Int_val(fd_v); + caml_enter_blocking_section(); + CAML_SYS_CLOSE(fd); + caml_leave_blocking_section(); + return Val_unit; +} + +CAMLprim value caml_sys_file_exists(value name) +{ +#ifdef _WIN32 + struct _stati64 st; +#else + struct stat st; +#endif + char_os * p; + int ret; + + if (! caml_string_is_c_safe(name)) return Val_false; + p = caml_stat_strdup_to_os(String_val(name)); + caml_enter_blocking_section(); + ret = CAML_SYS_STAT(p, &st); + caml_leave_blocking_section(); + caml_stat_free(p); + + return Val_bool(ret == 0); +} + +CAMLprim value caml_sys_is_directory(value name) +{ + CAMLparam1(name); +#ifdef _WIN32 + struct _stati64 st; +#else + struct stat st; +#endif + char_os * p; + int ret; + + caml_sys_check_path(name); + p = caml_stat_strdup_to_os(String_val(name)); + caml_enter_blocking_section(); + ret = CAML_SYS_STAT(p, &st); + caml_leave_blocking_section(); + caml_stat_free(p); + + if (ret == -1) caml_sys_error(name); +#ifdef S_ISDIR + CAMLreturn(Val_bool(S_ISDIR(st.st_mode))); +#else + CAMLreturn(Val_bool(st.st_mode & S_IFDIR)); +#endif +} + +CAMLprim value caml_sys_remove(value name) +{ + CAMLparam1(name); + char_os * p; + int ret; + caml_sys_check_path(name); + p = caml_stat_strdup_to_os(String_val(name)); + caml_enter_blocking_section(); + ret = CAML_SYS_UNLINK(p); + caml_leave_blocking_section(); + caml_stat_free(p); + if (ret != 0) caml_sys_error(name); + CAMLreturn(Val_unit); +} + +CAMLprim value caml_sys_rename(value oldname, value newname) +{ + char_os * p_old; + char_os * p_new; + int ret; + caml_sys_check_path(oldname); + caml_sys_check_path(newname); + p_old = caml_stat_strdup_to_os(String_val(oldname)); + p_new = caml_stat_strdup_to_os(String_val(newname)); + caml_enter_blocking_section(); + ret = CAML_SYS_RENAME(p_old, p_new); + caml_leave_blocking_section(); + caml_stat_free(p_new); + caml_stat_free(p_old); + if (ret != 0) + caml_sys_error(NO_ARG); + return Val_unit; +} + +CAMLprim value caml_sys_chdir(value dirname) +{ + CAMLparam1(dirname); + char_os * p; + int ret; + caml_sys_check_path(dirname); + p = caml_stat_strdup_to_os(String_val(dirname)); + caml_enter_blocking_section(); + ret = CAML_SYS_CHDIR(p); + caml_leave_blocking_section(); + caml_stat_free(p); + if (ret != 0) caml_sys_error(dirname); + CAMLreturn(Val_unit); +} + +CAMLprim value caml_sys_getcwd(value unit) +{ + char_os buff[4096]; + char_os * ret; +#ifdef HAS_GETCWD + ret = getcwd_os(buff, sizeof(buff)/sizeof(*buff)); +#else + caml_invalid_argument("Sys.getcwd not implemented"); +#endif /* HAS_GETCWD */ + if (ret == 0) caml_sys_error(NO_ARG); + return caml_copy_string_of_os(buff); +} + +CAMLprim value caml_sys_unsafe_getenv(value var) +{ + char_os * res, * p; + value val; + + if (! caml_string_is_c_safe(var)) caml_raise_not_found(); + p = caml_stat_strdup_to_os(String_val(var)); +#ifdef _WIN32 + res = caml_win32_getenv(p); +#else + res = CAML_SYS_GETENV(p); +#endif + caml_stat_free(p); + if (res == 0) caml_raise_not_found(); + val = caml_copy_string_of_os(res); +#ifdef _WIN32 + caml_stat_free(res); +#endif + return val; +} + +CAMLprim value caml_sys_getenv(value var) +{ + char_os * res, * p; + value val; + + if (! caml_string_is_c_safe(var)) caml_raise_not_found(); + p = caml_stat_strdup_to_os(String_val(var)); +#ifdef _WIN32 + res = caml_win32_getenv(p); +#else + res = caml_secure_getenv(p); +#endif + caml_stat_free(p); + if (res == 0) caml_raise_not_found(); + val = caml_copy_string_of_os(res); +#ifdef _WIN32 + caml_stat_free(res); +#endif + return val; +} + +char_os * caml_exe_name; +char_os ** caml_main_argv; + +CAMLprim value caml_sys_get_argv(value unit) +{ + CAMLparam0 (); /* unit is unused */ + CAMLlocal3 (exe_name, argv, res); + exe_name = caml_copy_string_of_os(caml_exe_name); + argv = caml_alloc_array((void *)caml_copy_string_of_os, (char const **) caml_main_argv); + res = caml_alloc_small(2, 0); + Field(res, 0) = exe_name; + Field(res, 1) = argv; + CAMLreturn(res); +} + +void caml_sys_init(char_os * exe_name, char_os **argv) +{ +#ifdef _WIN32 + /* Initialises the caml_win32_* globals on Windows with the version of + Windows which is running */ + caml_probe_win32_version(); +#if WINDOWS_UNICODE + caml_setup_win32_terminal(); +#endif +#endif +#ifdef CAML_WITH_CPLUGINS + caml_cplugins_init(exe_name, argv); +#endif + caml_exe_name = exe_name; + caml_main_argv = argv; +} + +#ifdef _WIN32 +#define WIFEXITED(status) 1 +#define WEXITSTATUS(status) (status) +#else +#if !(defined(WIFEXITED) && defined(WEXITSTATUS)) +/* Assume old-style V7 status word */ +#define WIFEXITED(status) (((status) & 0xFF) == 0) +#define WEXITSTATUS(status) (((status) >> 8) & 0xFF) +#endif +#endif + +CAMLprim value caml_sys_system_command(value command) +{ + CAMLparam1 (command); + int status, retcode; + char_os *buf; + + if (! caml_string_is_c_safe (command)) { + errno = EINVAL; + caml_sys_error(command); + } + buf = caml_stat_strdup_to_os(String_val(command)); + caml_enter_blocking_section (); + status = CAML_SYS_SYSTEM(buf); + caml_leave_blocking_section (); + caml_stat_free(buf); + if (status == -1) caml_sys_error(command); + if (WIFEXITED(status)) + retcode = WEXITSTATUS(status); + else + retcode = 255; + CAMLreturn (Val_int(retcode)); +} + +double caml_sys_time_include_children_unboxed(value include_children) +{ +#ifdef HAS_GETRUSAGE + struct rusage ru; + double acc = 0.; + + getrusage (RUSAGE_SELF, &ru); + acc += ru.ru_utime.tv_sec + ru.ru_utime.tv_usec / 1e6 + + ru.ru_stime.tv_sec + ru.ru_stime.tv_usec / 1e6; + + if (Bool_val(include_children)) { + getrusage (RUSAGE_CHILDREN, &ru); + acc += ru.ru_utime.tv_sec + ru.ru_utime.tv_usec / 1e6 + + ru.ru_stime.tv_sec + ru.ru_stime.tv_usec / 1e6; + } + + return acc; +#else + #ifdef HAS_TIMES + #ifndef CLK_TCK + #ifdef HZ + #define CLK_TCK HZ + #else + #define CLK_TCK 60 + #endif + #endif + struct tms t; + clock_t acc = 0; + times(&t); + acc += t.tms_utime + t.tms_stime; + if (Bool_val(include_children)) { + acc += t.tms_cutime + t.tms_cstime; + } + return (double)acc / CLK_TCK; + #else + /* clock() is standard ANSI C. We have no way of getting + subprocess times in this branch. */ + return (double)clock() / CLOCKS_PER_SEC; + #endif +#endif +} + +CAMLprim value caml_sys_time_include_children(value include_children) +{ + return caml_copy_double(caml_sys_time_include_children_unboxed(include_children)); +} + +double caml_sys_time_unboxed(value unit) { + return caml_sys_time_include_children_unboxed(Val_false); +} + +CAMLprim value caml_sys_time(value unit) +{ + return caml_copy_double(caml_sys_time_unboxed(unit)); +} + +#ifdef _WIN32 +extern int caml_win32_random_seed (intnat data[16]); +#endif + +CAMLprim value caml_sys_random_seed (value unit) +{ + intnat data[16]; + int n, i; + value res; +#ifdef _WIN32 + n = caml_win32_random_seed(data); +#else + int fd; + n = 0; + /* Try /dev/urandom first */ + fd = open("/dev/urandom", O_RDONLY, 0); + if (fd != -1) { + unsigned char buffer[12]; + int nread = read(fd, buffer, 12); + close(fd); + while (nread > 0) data[n++] = buffer[--nread]; + } + /* If the read from /dev/urandom fully succeeded, we now have 96 bits + of good random data and can stop here. Otherwise, complement + whatever we got (probably nothing) with some not-very-random data. */ + if (n < 12) { +#ifdef HAS_GETTIMEOFDAY + struct timeval tv; + gettimeofday(&tv, NULL); + data[n++] = tv.tv_usec; + data[n++] = tv.tv_sec; +#else + data[n++] = time(NULL); +#endif +#ifdef HAS_UNISTD + data[n++] = getpid(); +#if HAS_PPID + data[n++] = getppid(); +#endif +#endif + } +#endif + /* Convert to an OCaml array of ints */ + res = caml_alloc_small(n, 0); + for (i = 0; i < n; i++) Field(res, i) = Val_long(data[i]); + return res; +} + +CAMLprim value caml_sys_const_big_endian(value unit) +{ +#ifdef ARCH_BIG_ENDIAN + return Val_true; +#else + return Val_false; +#endif +} + +/* returns a value that represents a number of bits */ +CAMLprim value caml_sys_const_word_size(value unit) +{ + return Val_long(8 * sizeof(value)); +} + +/* returns a value that represents a number of bits */ +CAMLprim value caml_sys_const_int_size(value unit) +{ + return Val_long(8 * sizeof(value) - 1) ; +} + +/* returns a value that represents a number of words */ +CAMLprim value caml_sys_const_max_wosize(value unit) +{ + return Val_long(Max_wosize) ; +} + +CAMLprim value caml_sys_const_ostype_unix(value unit) +{ + return Val_bool(0 == strcmp(OCAML_OS_TYPE,"Unix")); +} + +CAMLprim value caml_sys_const_ostype_win32(value unit) +{ + return Val_bool(0 == strcmp(OCAML_OS_TYPE,"Win32")); +} + +CAMLprim value caml_sys_const_ostype_cygwin(value unit) +{ + return Val_bool(0 == strcmp(OCAML_OS_TYPE,"Cygwin")); +} + +CAMLprim value caml_sys_const_backend_type(value unit) +{ + return Val_int(1); /* Bytecode backed */ +} +CAMLprim value caml_sys_get_config(value unit) +{ + CAMLparam0 (); /* unit is unused */ + CAMLlocal2 (result, ostype); + + ostype = caml_copy_string(OCAML_OS_TYPE); + result = caml_alloc_small (3, 0); + Field(result, 0) = ostype; + Field(result, 1) = Val_long (8 * sizeof(value)); +#ifdef ARCH_BIG_ENDIAN + Field(result, 2) = Val_true; +#else + Field(result, 2) = Val_false; +#endif + CAMLreturn (result); +} + +CAMLprim value caml_sys_read_directory(value path) +{ + CAMLparam1(path); + CAMLlocal1(result); + struct ext_table tbl; + char_os * p; + int ret; + + caml_sys_check_path(path); + caml_ext_table_init(&tbl, 50); + p = caml_stat_strdup_to_os(String_val(path)); + caml_enter_blocking_section(); + ret = CAML_SYS_READ_DIRECTORY(p, &tbl); + caml_leave_blocking_section(); + caml_stat_free(p); + if (ret == -1){ + caml_ext_table_free(&tbl, 1); + caml_sys_error(path); + } + caml_ext_table_add(&tbl, NULL); + result = caml_copy_string_array((char const **) tbl.contents); + caml_ext_table_free(&tbl, 1); + CAMLreturn(result); +} + +/* Return true if the value is a filedescriptor (int) that is + * (presumably) open on an interactive terminal */ +CAMLprim value caml_sys_isatty(value chan) +{ + int fd; + value ret; + + fd = (Channel(chan))->fd; +#ifdef _WIN32 + ret = Val_bool(caml_win32_isatty(fd)); +#else + ret = Val_bool(isatty(fd)); +#endif + + return ret; +} + +/* Load dynamic plugins indicated in the CAML_CPLUGINS environment + variable. These plugins can be used to set currently existing + hooks, such as GC hooks and system calls tracing (see misc.h). + */ + +#ifdef CAML_WITH_CPLUGINS + +value (*caml_cplugins_prim)(int,value,value,value) = NULL; + +#define DLL_EXECUTABLE 1 +#define DLL_NOT_GLOBAL 0 + +static struct cplugin_context cplugin_context; + +void caml_load_plugin(char_os *plugin) +{ + void* dll_handle = NULL; + char* u8; + + dll_handle = caml_dlopen(plugin, DLL_EXECUTABLE, DLL_NOT_GLOBAL); + if( dll_handle != NULL ){ + void (* dll_init)(struct cplugin_context*) = + caml_dlsym(dll_handle, "caml_cplugin_init"); + if( dll_init != NULL ){ + cplugin_context.plugin=plugin; + dll_init(&cplugin_context); + } else { + caml_dlclose(dll_handle); + } + } else { + u8 = caml_stat_strdup_of_os(plugin); + fprintf(stderr, "Cannot load C plugin %s\nReason: %s\n", + u8, caml_dlerror()); + caml_stat_free(u8); + } +} + +void caml_cplugins_load(char_os *env_variable) +{ + char_os *plugins = caml_secure_getenv(env_variable); + if(plugins != NULL){ + char_os* curs = plugins; + while(*curs != 0){ + if(*curs == _T(',')){ + if(curs > plugins){ + *curs = 0; + caml_load_plugin(plugins); + } + plugins = curs+1; + } + curs++; + } + if(curs > plugins) caml_load_plugin(plugins); + } +} + +void caml_cplugins_init(char_os * exe_name, char_os **argv) +{ + cplugin_context.api_version = CAML_CPLUGIN_CONTEXT_API; + cplugin_context.prims_bitmap = CAML_CPLUGINS_PRIMS_BITMAP; + cplugin_context.exe_name = exe_name; + cplugin_context.argv = argv; + cplugin_context.ocaml_version = OCAML_VERSION_STRING; + caml_cplugins_load(_T("CAML_CPLUGINS")); +#ifdef NATIVE_CODE + caml_cplugins_load(_T("CAML_NATIVE_CPLUGINS")); +#else + caml_cplugins_load(_T("CAML_BYTE_CPLUGINS")); +#endif +} + +#endif /* CAML_WITH_CPLUGINS */ diff --git a/test/monniaux/ocaml/byterun/toto b/test/monniaux/ocaml/byterun/toto new file mode 100644 index 00000000..ac54a2b6 --- /dev/null +++ b/test/monniaux/ocaml/byterun/toto @@ -0,0 +1,13996 @@ +12429 execve("/opt/Kalray/usr/local/k1rdtools/bin/k1-cluster", ["k1-cluster", "--syscall=libstd_scalls.so", "--", "./ocamlrun", "/home/monniaux/work/caml/quickso"...], 0x7ffec7d0ab80 /* 71 vars */) = 0 +12429 brk(NULL) = 0x1ef3000 +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 4096) = 46 +12429 mmap(NULL, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c78a90000 +12429 access("/etc/ld.so.preload", R_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/tls/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/tls/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/tls/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/tls/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/tls/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/tls/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/tls/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/tls", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao", {st_mode=S_IFDIR|0755, st_size=4096, ...}) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/tls/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/tls/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/tls/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/tls/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/tls/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/tls/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/tls/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/tls", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../lib64", {st_mode=S_IFDIR|0755, st_size=4096, ...}) = 0 +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/tls/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/tls/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/tls/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/tls/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/tls/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/tls/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/tls/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/tls", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./tls/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./tls/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./tls/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./tls/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./tls/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./tls/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./tls/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./tls", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/.", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/tls/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/tls/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/tls/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/tls/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/tls/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/tls/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/tls/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/tls", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/local/k1rdtools/bin/../../../devimage/toolchain_default/toolroot/usr/local/k1rdtools/lib64/./lao", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/tls/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/lib64/tls/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/tls/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/lib64/tls/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/tls/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/lib64/tls/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/tls/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/lib64/tls", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/haswell/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/lib64/haswell/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/haswell/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/lib64/haswell", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/x86_64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/lib64/x86_64", 0x7ffeb1c260e0) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/librt.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/opt/Kalray/usr/lib64", {st_mode=S_IFDIR|0755, st_size=4096, ...}) = 0 +12429 openat(AT_FDCWD, "/etc/ld.so.cache", O_RDONLY|O_CLOEXEC) = 3 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=248338, ...}) = 0 +12429 mmap(NULL, 248338, PROT_READ, MAP_PRIVATE, 3, 0) = 0x7f4c78a53000 +12429 close(3) = 0 +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/lib/x86_64-linux-gnu/librt.so.1", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\0\"\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=31680, ...}) = 0 +12429 mmap(NULL, 2128864, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c78663000 +12429 mprotect(0x7f4c7866a000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c78869000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x6000) = 0x7f4c78869000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libdl.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libdl.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libdl.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/lib/x86_64-linux-gnu/libdl.so.2", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0P\16\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=14560, ...}) = 0 +12429 mmap(NULL, 2109712, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c7845f000 +12429 mprotect(0x7f4c78462000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c78661000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x2000) = 0x7f4c78661000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libmppa_components.so", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_components.so", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0 \212\f\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=11681992, ...}) = 0 +12429 mmap(NULL, 13816744, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c77731000 +12429 mprotect(0x7f4c78238000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c78437000, 122880, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0xb06000) = 0x7f4c78437000 +12429 mmap(0x7f4c78455000, 37800, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c78455000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libxml2.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libxml2.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libxml2.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/usr/lib/x86_64-linux-gnu/libxml2.so.2", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0p\331\2\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=1834232, ...}) = 0 +12429 mmap(NULL, 3934648, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c77370000 +12429 mprotect(0x7f4c77526000, 2097152, PROT_NONE) = 0 +12429 mmap(0x7f4c77726000, 40960, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x1b6000) = 0x7f4c77726000 +12429 mmap(0x7f4c77730000, 2488, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c77730000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libz.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libz.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libz.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/lib/x86_64-linux-gnu/libz.so.1", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\220\37\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=116960, ...}) = 0 +12429 mmap(NULL, 2212016, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c77153000 +12429 mprotect(0x7f4c7716f000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c7736e000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x1b000) = 0x7f4c7736e000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libstdc++.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstdc++.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libstdc++.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/usr/lib/x86_64-linux-gnu/libstdc++.so.6", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\3\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\360\303\10\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=1594832, ...}) = 0 +12429 mmap(NULL, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c78a51000 +12429 mmap(NULL, 3702816, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c76dca000 +12429 mprotect(0x7f4c76f43000, 2097152, PROT_NONE) = 0 +12429 mmap(0x7f4c77143000, 49152, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x179000) = 0x7f4c77143000 +12429 mmap(0x7f4c7714f000, 12320, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c7714f000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libm.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libm.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libm.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/lib/x86_64-linux-gnu/libm.so.6", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\3\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\200\272\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=1700792, ...}) = 0 +12429 mmap(NULL, 3789144, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c76a2c000 +12429 mprotect(0x7f4c76bc9000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c76dc8000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x19c000) = 0x7f4c76dc8000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libgcc_s.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libgcc_s.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libgcc_s.so.1", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/lib/x86_64-linux-gnu/libgcc_s.so.1", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\300*\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=96616, ...}) = 0 +12429 mmap(NULL, 2192432, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c76814000 +12429 mprotect(0x7f4c7682b000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c76a2a000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x16000) = 0x7f4c76a2a000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libpthread.so.0", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libpthread.so.0", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libpthread.so.0", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/lib/x86_64-linux-gnu/libpthread.so.0", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0000b\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=144976, ...}) = 0 +12429 mmap(NULL, 2221184, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c765f5000 +12429 mprotect(0x7f4c7660f000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c7680e000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x19000) = 0x7f4c7680e000 +12429 mmap(0x7f4c76810000, 13440, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c76810000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libc.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libc.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libc.so.6", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/lib/x86_64-linux-gnu/libc.so.6", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\3\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\260\34\2\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=2030544, ...}) = 0 +12429 mmap(NULL, 4131552, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c76204000 +12429 mprotect(0x7f4c763eb000, 2097152, PROT_NONE) = 0 +12429 mmap(0x7f4c765eb000, 24576, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x1e7000) = 0x7f4c765eb000 +12429 mmap(0x7f4c765f1000, 15072, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c765f1000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libicuuc.so.60", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libicuuc.so.60", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libicuuc.so.60", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/usr/lib/x86_64-linux-gnu/libicuuc.so.60", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\0\v\6\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=1792008, ...}) = 0 +12429 mmap(NULL, 3894144, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c75e4d000 +12429 mprotect(0x7f4c75ff0000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c761ef000, 81920, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x1a2000) = 0x7f4c761ef000 +12429 mmap(0x7f4c76203000, 2944, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c76203000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/liblzma.so.5", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/liblzma.so.5", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/liblzma.so.5", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/lib/x86_64-linux-gnu/liblzma.so.5", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\340(\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=153984, ...}) = 0 +12429 mmap(NULL, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c78a4f000 +12429 mmap(NULL, 2248968, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c75c27000 +12429 mprotect(0x7f4c75c4b000, 2097152, PROT_NONE) = 0 +12429 mmap(0x7f4c75e4b000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x24000) = 0x7f4c75e4b000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libicudata.so.60", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libicudata.so.60", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libicudata.so.60", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 access("/etc/ld.so.nohwcap", F_OK) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/usr/lib/x86_64-linux-gnu/libicudata.so.60", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\320\4\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0644, st_size=26904264, ...}) = 0 +12429 mmap(NULL, 28999696, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c7407e000 +12429 mprotect(0x7f4c75a26000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c75c25000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x19a7000) = 0x7f4c75c25000 +12429 close(3) = 0 +12429 mmap(NULL, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c78a4d000 +12429 mmap(NULL, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c78a4b000 +12429 arch_prctl(ARCH_SET_FS, 0x7f4c78a4e6c0) = 0 +12429 mprotect(0x7f4c765eb000, 16384, PROT_READ) = 0 +12429 mprotect(0x7f4c75c25000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c78661000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c7680e000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c75e4b000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c76dc8000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c76a2a000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c77143000, 40960, PROT_READ) = 0 +12429 mprotect(0x7f4c761ef000, 77824, PROT_READ) = 0 +12429 mprotect(0x7f4c7736e000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c77726000, 32768, PROT_READ) = 0 +12429 mprotect(0x7f4c78869000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c78437000, 94208, PROT_READ) = 0 +12429 mprotect(0x602000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c78a92000, 4096, PROT_READ) = 0 +12429 munmap(0x7f4c78a53000, 248338) = 0 +12429 set_tid_address(0x7f4c78a4e990) = 12429 +12429 set_robust_list(0x7f4c78a4e9a0, 24) = 0 +12429 rt_sigaction(SIGRTMIN, {sa_handler=0x7f4c765facb0, sa_mask=[], sa_flags=SA_RESTORER|SA_SIGINFO, sa_restorer=0x7f4c76607890}, NULL, 8) = 0 +12429 rt_sigaction(SIGRT_1, {sa_handler=0x7f4c765fad50, sa_mask=[], sa_flags=SA_RESTORER|SA_RESTART|SA_SIGINFO, sa_restorer=0x7f4c76607890}, NULL, 8) = 0 +12429 rt_sigprocmask(SIG_UNBLOCK, [RTMIN RT_1], NULL, 8) = 0 +12429 prlimit64(0, RLIMIT_STACK, NULL, {rlim_cur=8192*1024, rlim_max=RLIM64_INFINITY}) = 0 +12429 brk(NULL) = 0x1ef3000 +12429 brk(0x1f14000) = 0x1f14000 +12429 futex(0x7f4c7715007c, FUTEX_WAKE_PRIVATE, 2147483647) = 0 +12429 futex(0x7f4c77150088, FUTEX_WAKE_PRIVATE, 2147483647) = 0 +12429 prctl(PR_SET_NAME, "main") = 0 +12429 rt_sigprocmask(SIG_BLOCK, ~[TERM TSTP RTMIN RT_1], NULL, 8) = 0 +12429 mmap(NULL, 8392704, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_STACK, -1, 0) = 0x7f4c7387d000 +12429 mprotect(0x7f4c7387e000, 8388608, PROT_READ|PROT_WRITE) = 0 +12429 clone(child_stack=0x7f4c7407cfb0, flags=CLONE_VM|CLONE_FS|CLONE_FILES|CLONE_SIGHAND|CLONE_THREAD|CLONE_SYSVSEM|CLONE_SETTLS|CLONE_PARENT_SETTID|CLONE_CHILD_CLEARTID, parent_tidptr=0x7f4c7407d9d0, tls=0x7f4c7407d700, child_tidptr=0x7f4c7407d9d0) = 12430 +12430 set_robust_list(0x7f4c7407d9e0, 24) = 0 +12429 readlink("/proc/self/exe", +12430 prctl(PR_SET_NAME, "sigwait" +12429 <... readlink resumed> "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12430 <... prctl resumed> ) = 0 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_common.so", F_OK +12430 mmap(NULL, 134217728, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0 +12429 <... access resumed> ) = -1 ENOENT (No such file or directory) +12430 <... mmap resumed> ) = 0x7f4c6b87d000 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_common.so", F_OK +12430 munmap(0x7f4c6b87d000, 7876608 +12429 <... access resumed> ) = -1 ENOENT (No such file or directory) +12430 <... munmap resumed> ) = 0 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_common.so", F_OK +12430 munmap(0x7f4c70000000, 59232256 +12429 <... access resumed> ) = -1 ENOENT (No such file or directory) +12430 <... munmap resumed> ) = 0 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_common.so", F_OK +12430 mprotect(0x7f4c6c000000, 135168, PROT_READ|PROT_WRITE +12429 <... access resumed> ) = 0 +12430 <... mprotect resumed> ) = 0 +12429 futex(0x7f4c786620c8, FUTEX_WAKE_PRIVATE, 2147483647 +12430 rt_sigtimedwait([INT USR1 ALRM], +12429 <... futex resumed> ) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_common.so", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\20p\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=185512, ...}) = 0 +12429 mmap(NULL, 2318056, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c73647000 +12429 mprotect(0x7f4c73672000, 2097152, PROT_NONE) = 0 +12429 mmap(0x7f4c73872000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x2b000) = 0x7f4c73872000 +12429 mmap(0x7f4c73874000, 36584, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c73874000 +12429 close(3) = 0 +12429 mprotect(0x7f4c73872000, 4096, PROT_READ) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0 g\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=276856, ...}) = 0 +12429 mmap(NULL, 2293640, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c73417000 +12429 mprotect(0x7f4c73444000, 2097152, PROT_NONE) = 0 +12429 mmap(0x7f4c73644000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0x2d000) = 0x7f4c73644000 +12429 mmap(0x7f4c73646000, 3976, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c73646000 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "BEFORE/tls/haswell/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "BEFORE/tls/haswell/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "BEFORE/tls/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "BEFORE/tls/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "BEFORE/haswell/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "BEFORE/haswell/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "BEFORE/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "BEFORE/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/tls/haswell/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/tls/haswell/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/tls/haswell/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/tls/haswell", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/tls/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/tls/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/tls/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/tls", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/haswell/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/haswell/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/haswell/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/haswell", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/multiloader/k1_build/k1_Release_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/tls/haswell/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/tls/haswell/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/tls/haswell/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/tls/haswell", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/tls/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/tls/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/tls/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/tls", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/haswell/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/haswell/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/haswell/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/haswell", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/./work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/kalray_internal/lib64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/tls/haswell/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/tls/haswell/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/tls/haswell/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/tls/haswell", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/tls/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/tls/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/tls/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/tls", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/haswell/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/haswell/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/haswell/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/haswell", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/x86_64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/x86_64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 stat("/work1/hudson/workspace/_RDTools_integration_2/label/centos7-64/rdtools/devimage/toolchain_default/k1-multiloader_host/devimage/lib64", 0x7ffeb1c24410) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libzip.so.2", O_RDONLY|O_CLOEXEC) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/lib64/libzip.so.2", O_RDONLY|O_CLOEXEC) = 3 +12429 read(3, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\220(\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=70136, ...}) = 0 +12429 mmap(NULL, 2163424, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 3, 0) = 0x7f4c73206000 +12429 mprotect(0x7f4c73216000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c73415000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 3, 0xf000) = 0x7f4c73415000 +12429 close(3) = 0 +12429 mprotect(0x7f4c73415000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c73644000, 4096, PROT_READ) = 0 +12429 openat(AT_FDCWD, "/etc/kalray", O_RDONLY|O_NONBLOCK|O_CLOEXEC|O_DIRECTORY) = 3 +12429 fstat(3, {st_mode=S_IFDIR|0755, st_size=4096, ...}) = 0 +12429 getdents(3, /* 3 entries */, 32768) = 80 +12429 brk(0x1f35000) = 0x1f35000 +12429 openat(AT_FDCWD, "/etc/kalray/license.xml", O_RDONLY) = 4 +12429 futex(0x7f4c777306a8, FUTEX_WAKE_PRIVATE, 2147483647) = 0 +12429 read(4, "\0\1\0\0\0P#\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(4, {st_mode=S_IFREG|0644, st_size=47568, ...}) = 0 +12429 mmap(NULL, 2168632, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 4, 0) = 0x7f4c72ff4000 +12429 mprotect(0x7f4c72fff000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c731fe000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 4, 0xa000) = 0x7f4c731fe000 +12429 mmap(0x7f4c73200000, 22328, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c73200000 +12429 close(4) = 0 +12429 mprotect(0x7f4c731fe000, 4096, PROT_READ) = 0 +12429 munmap(0x7f4c78a53000, 248338) = 0 +12429 openat(AT_FDCWD, "/etc/hosts", O_RDONLY|O_CLOEXEC) = 4 +12429 fstat(4, {st_mode=S_IFREG|0644, st_size=369, ...}) = 0 +12429 read(4, "127.0.0.1\tlocalhost\n127.0.1.1\tva"..., 4096) = 369 +12429 read(4, "", 4096) = 0 +12429 close(4) = 0 +12429 getdents(3, /* 0 entries */, 32768) = 0 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "/home/monniaux/.kalray", O_RDONLY|O_NONBLOCK|O_CLOEXEC|O_DIRECTORY) = -1 ENOENT (No such file or directory) +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/mppa_config.yaml", O_RDONLY) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib/mppa_config.yaml", O_RDONLY) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib32/mppa_config.yaml", O_RDONLY) = -1 ENOENT (No such file or directory) +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/mppa_config.yaml", O_RDONLY) = 3 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=616729, ...}) = 0 +12429 read(3, "# {{{ MPPA:\n- MPPA:\n\n # MPPA "..., 16384) = 16384 +12429 brk(0x1f56000) = 0x1f56000 +12429 read(3, "0\n\n # PRO"..., 16384) = 16384 +12429 mmap(NULL, 200704, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c78a5f000 +12429 read(3, " Latency_ECC: 4\n"..., 16384) = 16384 +12429 read(3, " # PROC7 sub-configurati"..., 16384) = 16384 +12429 mremap(0x7f4c78a5f000, 200704, 397312, MREMAP_MAYMOVE) = 0x7f4c789ea000 +12429 read(3, "atency_ECC: 4\n "..., 16384) = 16384 +12429 read(3, " # PROC12 sub-conf"..., 16384) = 16384 +12429 brk(0x1f77000) = 0x1f77000 +12429 read(3, " Late"..., 16384) = 16384 +12429 read(3, " # BANK par"..., 16384) = 16384 +12429 mremap(0x7f4c789ea000, 397312, 790528, MREMAP_MAYMOVE) = 0x7f4c78929000 +12429 read(3, " TYPE: generic_cache\n "..., 16384) = 16384 +12429 brk(0x1f98000) = 0x1f98000 +12429 read(3, " { Debug: Disable, "..., 16384) = 16384 +12429 read(3, "eneric_cache\n "..., 16384) = 16384 +12429 read(3, "bug: Disable, Warning: Enable, L"..., 16384) = 16384 +12429 brk(0x1fb9000) = 0x1fb9000 +12429 read(3, "he\n N"..., 16384) = 16384 +12429 read(3, " { Debug: Disable,"..., 16384) = 16384 +12429 read(3, " TYPE: generic_cac"..., 16384) = 16384 +12429 mremap(0x7f4c78929000, 790528, 1576960, MREMAP_MAYMOVE) = 0x7f4c72e73000 +12429 read(3, " { NA"..., 16384) = 16384 +12429 brk(0x1fda000) = 0x1fda000 +12429 read(3, " DEBUG:\n "..., 16384) = 16384 +12429 read(3, " Latency_MultiBank: 1\n "..., 16384) = 16384 +12429 read(3, " # DEBUG param"..., 16384) = 16384 +12429 brk(0x1ffb000) = 0x1ffb000 +12429 read(3, "Bank: 1\n "..., 16384) = 16384 +12429 read(3, " # DEBU"..., 16384) = 16384 +12429 read(3, " Latency_MultiBan"..., 16384) = 16384 +12429 brk(0x201c000) = 0x201c000 +12429 read(3, "ing_Scheme: Sequential\n "..., 16384) = 16384 +12429 read(3, " CacheSize:"..., 16384) = 16384 +12429 read(3, "}}} DEBUG:\n "..., 16384) = 16384 +12429 read(3, " CacheSize: 16384\n "..., 16384) = 16384 +12429 brk(0x203d000) = 0x203d000 +12429 read(3, " # {{{ C"..., 16384) = 16384 +12429 read(3, "Size: 16384\n "..., 16384) = 16384 +12429 read(3, " }}} DEBUG:\n "..., 16384) = 16384 +12429 brk(0x205e000) = 0x205e000 +12429 read(3, " Cache"..., 16384) = 16384 +12429 mremap(0x7f4c72e73000, 1576960, 3149824, MREMAP_MAYMOVE) = 0x7f4c72b72000 +12429 read(3, " # {{{ DEBUG:\n "..., 16384) = 16384 +12429 read(3, "nable, Level: 0, IOStream: stder"..., 16384) = 16384 +12429 brk(0x207f000) = 0x207f000 +12429 read(3, " DEBUG:\n "..., 16384) = 16384 +12429 read(3, "OStream: stderr }\n "..., 16384) = 16384 +12429 read(3, " DEBUG:\n "..., 16384) = 16384 +12429 brk(0x20a0000) = 0x20a0000 +12429 read(3, "l: 0, IOStream: stderr }\n "..., 16384) = 16384 +12429 read(3, "EBUG:\n "..., 16384) = 16384 +12429 read(3, "ng_Parameter: none\n "..., 16384) = 10521 +12429 read(3, "", 4096) = 0 +12429 read(3, "", 16384) = 0 +12429 brk(0x20c1000) = 0x20c1000 +12429 brk(0x20e2000) = 0x20e2000 +12429 brk(0x2103000) = 0x2103000 +12429 brk(0x2124000) = 0x2124000 +12429 brk(0x2145000) = 0x2145000 +12429 brk(0x2166000) = 0x2166000 +12429 brk(0x2187000) = 0x2187000 +12429 brk(0x21a8000) = 0x21a8000 +12429 brk(0x21c9000) = 0x21c9000 +12429 brk(0x21ea000) = 0x21ea000 +12429 brk(0x220b000) = 0x220b000 +12429 brk(0x222c000) = 0x222c000 +12429 brk(0x224d000) = 0x224d000 +12429 brk(0x226e000) = 0x226e000 +12429 brk(0x228f000) = 0x228f000 +12429 brk(0x22b0000) = 0x22b0000 +12429 brk(0x22d1000) = 0x22d1000 +12429 brk(0x22f2000) = 0x22f2000 +12429 brk(0x2313000) = 0x2313000 +12429 brk(0x2334000) = 0x2334000 +12429 munmap(0x7f4c72b72000, 3149824) = 0 +12429 close(3) = 0 +12429 ioctl(1, TCGETS, {B38400 opost isig icanon echo ...}) = 0 +12429 access("./ocamlrun", F_OK) = 0 +12429 stat("./ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 openat(AT_FDCWD, "./ocamlrun", O_RDONLY) = 3 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 lseek(3, 2289664, SEEK_SET) = 2289664 +12429 read(3, "\3\0\0\0\0\0\0\0\0p\10\0\0\0\0\0\350@\10\0\0\0\0\0\0\200\2\0\0\0\0\0"..., 2104) = 2104 +12429 fstat(3, {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 lseek(3, 2224128, SEEK_SET) = 2224128 +12429 read(3, "\360\235\6\0\0\0\0\0`\0\0\0\0\0\0\0\24\17\0\0\4\0\361\377\0\0\0\0\0\0\0\0"..., 4096) = 4096 +12429 read(3, "F(\0\0\22\0\6\0\350Q\5\0\0\0\0\0D\1\0\0\0\0\0\0\263\31\0\0\22\0\6\0"..., 61440) = 61440 +12429 read(3, "\3\0\0\0\0\0\0\0\0p\10\0\0\0\0\0\350@\10\0\0\0\0\0\0\200\2\0\0\0\0\0"..., 4096) = 2104 +12429 close(3) = 0 +12429 openat(AT_FDCWD, "./ocamlrun", O_RDONLY) = 3 +12429 lseek(3, 0, SEEK_END) = 2291768 +12429 mmap(NULL, 2291768, PROT_READ|PROT_WRITE, MAP_PRIVATE, 3, 0) = 0x7f4c72dc4000 +12429 munmap(0x7f4c72dc4000, 2291768) = 0 +12429 close(3) = 0 +12429 getpid() = 12429 +12429 getpid() = 12429 +12429 unlink("/tmp/mppa-tlm-input-12429") = -1 ENOENT (No such file or directory) +12429 unlink("/tmp/mppa-tlm-output-12429") = -1 ENOENT (No such file or directory) +12429 mknod("/tmp/mppa-tlm-input-12429", S_IFIFO|0666) = 0 +12429 mknod("/tmp/mppa-tlm-output-12429", S_IFIFO|0666) = 0 +12429 mmap(NULL, 8392704, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_STACK, -1, 0) = 0x7f4c727f3000 +12429 mprotect(0x7f4c727f4000, 8388608, PROT_READ|PROT_WRITE) = 0 +12429 clone(child_stack=0x7f4c72ff2fb0, flags=CLONE_VM|CLONE_FS|CLONE_FILES|CLONE_SIGHAND|CLONE_THREAD|CLONE_SYSVSEM|CLONE_SETTLS|CLONE_PARENT_SETTID|CLONE_CHILD_CLEARTID, parent_tidptr=0x7f4c72ff39d0, tls=0x7f4c72ff3700, child_tidptr=0x7f4c72ff39d0) = 12431 +12431 set_robust_list(0x7f4c72ff39e0, 24) = 0 +12431 prctl(PR_SET_NAME, "COM") = 0 +12431 mmap(NULL, 134217728, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0) = 0x7f4c64000000 +12431 munmap(0x7f4c68000000, 67108864) = 0 +12431 mprotect(0x7f4c64000000, 135168, PROT_READ|PROT_WRITE) = 0 +12431 openat(AT_FDCWD, "/tmp/mppa-tlm-input-12429", O_RDONLY +12429 brk(0x2355000) = 0x2355000 +12429 brk(0x2376000) = 0x2376000 +12429 brk(0x2397000) = 0x2397000 +12429 brk(0x23b8000) = 0x23b8000 +12429 brk(0x23d9000) = 0x23d9000 +12429 brk(0x23fa000) = 0x23fa000 +12429 brk(0x241b000) = 0x241b000 +12429 brk(0x243c000) = 0x243c000 +12429 brk(0x245d000) = 0x245d000 +12429 brk(0x247e000) = 0x247e000 +12429 brk(0x249f000) = 0x249f000 +12429 brk(0x24c0000) = 0x24c0000 +12429 brk(0x24e1000) = 0x24e1000 +12429 brk(0x2502000) = 0x2502000 +12429 brk(0x2523000) = 0x2523000 +12429 brk(0x2544000) = 0x2544000 +12429 brk(0x2565000) = 0x2565000 +12429 brk(0x2586000) = 0x2586000 +12429 brk(0x25a7000) = 0x25a7000 +12429 brk(0x25c8000) = 0x25c8000 +12429 brk(0x25e9000) = 0x25e9000 +12429 brk(0x260a000) = 0x260a000 +12429 brk(0x262b000) = 0x262b000 +12429 brk(0x264f000) = 0x264f000 +12429 brk(0x2670000) = 0x2670000 +12429 brk(0x2691000) = 0x2691000 +12429 brk(0x26b2000) = 0x26b2000 +12429 brk(0x26d3000) = 0x26d3000 +12429 brk(0x26f4000) = 0x26f4000 +12429 brk(0x2715000) = 0x2715000 +12429 brk(0x2736000) = 0x2736000 +12429 brk(0x2757000) = 0x2757000 +12429 brk(0x2778000) = 0x2778000 +12429 brk(0x2799000) = 0x2799000 +12429 brk(0x27ba000) = 0x27ba000 +12429 brk(0x27db000) = 0x27db000 +12429 brk(0x27fc000) = 0x27fc000 +12429 brk(0x281d000) = 0x281d000 +12429 brk(0x283e000) = 0x283e000 +12429 brk(0x285f000) = 0x285f000 +12429 brk(0x2880000) = 0x2880000 +12429 brk(0x28a1000) = 0x28a1000 +12429 brk(0x28c2000) = 0x28c2000 +12429 brk(0x28e3000) = 0x28e3000 +12429 brk(0x2904000) = 0x2904000 +12429 brk(0x2925000) = 0x2925000 +12429 brk(0x2946000) = 0x2946000 +12429 brk(0x2967000) = 0x2967000 +12429 brk(0x2988000) = 0x2988000 +12429 brk(0x29a9000) = 0x29a9000 +12429 brk(0x29ca000) = 0x29ca000 +12429 brk(0x29eb000) = 0x29eb000 +12429 brk(0x2a0c000) = 0x2a0c000 +12429 brk(0x2a2d000) = 0x2a2d000 +12429 brk(0x2a4e000) = 0x2a4e000 +12429 brk(0x2a6f000) = 0x2a6f000 +12429 brk(0x2a90000) = 0x2a90000 +12429 brk(0x2ab1000) = 0x2ab1000 +12429 brk(0x2ad2000) = 0x2ad2000 +12429 brk(0x2af3000) = 0x2af3000 +12429 brk(0x2b14000) = 0x2b14000 +12429 brk(0x2b35000) = 0x2b35000 +12429 brk(0x2b56000) = 0x2b56000 +12429 brk(0x2b77000) = 0x2b77000 +12429 brk(0x2b98000) = 0x2b98000 +12429 brk(0x2bb9000) = 0x2bb9000 +12429 brk(0x2bda000) = 0x2bda000 +12429 brk(0x2bfb000) = 0x2bfb000 +12429 brk(0x2c1c000) = 0x2c1c000 +12429 brk(0x2c3d000) = 0x2c3d000 +12429 brk(0x2c5e000) = 0x2c5e000 +12429 brk(0x2c7f000) = 0x2c7f000 +12429 brk(0x2ca0000) = 0x2ca0000 +12429 brk(0x2cc1000) = 0x2cc1000 +12429 brk(0x2ce2000) = 0x2ce2000 +12429 brk(0x2d03000) = 0x2d03000 +12429 brk(0x2d24000) = 0x2d24000 +12429 brk(0x2d45000) = 0x2d45000 +12429 brk(0x2d66000) = 0x2d66000 +12429 brk(0x2d87000) = 0x2d87000 +12429 brk(0x2da8000) = 0x2da8000 +12429 brk(0x2dce000) = 0x2dce000 +12429 brk(0x2def000) = 0x2def000 +12429 brk(0x2e10000) = 0x2e10000 +12429 brk(0x2e31000) = 0x2e31000 +12429 brk(0x2e52000) = 0x2e52000 +12429 brk(0x2e73000) = 0x2e73000 +12429 brk(0x2e94000) = 0x2e94000 +12429 brk(0x2eb5000) = 0x2eb5000 +12429 brk(0x2ed6000) = 0x2ed6000 +12429 brk(0x2ef7000) = 0x2ef7000 +12429 brk(0x2f18000) = 0x2f18000 +12429 brk(0x2f39000) = 0x2f39000 +12429 brk(0x2f5a000) = 0x2f5a000 +12429 brk(0x2f7b000) = 0x2f7b000 +12429 brk(0x2f9c000) = 0x2f9c000 +12429 brk(0x2fbd000) = 0x2fbd000 +12429 brk(0x2ffe000) = 0x2ffe000 +12429 brk(0x301f000) = 0x301f000 +12429 brk(0x3040000) = 0x3040000 +12429 brk(0x3061000) = 0x3061000 +12429 brk(0x3082000) = 0x3082000 +12429 brk(0x30a3000) = 0x30a3000 +12429 brk(0x30c4000) = 0x30c4000 +12429 brk(0x30e5000) = 0x30e5000 +12429 brk(0x3106000) = 0x3106000 +12429 brk(0x3127000) = 0x3127000 +12429 brk(0x3148000) = 0x3148000 +12429 brk(0x3169000) = 0x3169000 +12429 brk(0x318a000) = 0x318a000 +12429 brk(0x31ab000) = 0x31ab000 +12429 brk(0x31cc000) = 0x31cc000 +12429 brk(0x31ed000) = 0x31ed000 +12429 brk(0x320e000) = 0x320e000 +12429 brk(0x322f000) = 0x322f000 +12429 brk(0x3250000) = 0x3250000 +12429 brk(0x3271000) = 0x3271000 +12429 brk(0x3292000) = 0x3292000 +12429 brk(0x32b3000) = 0x32b3000 +12429 brk(0x32d4000) = 0x32d4000 +12429 brk(0x32f5000) = 0x32f5000 +12429 brk(0x3316000) = 0x3316000 +12429 brk(0x3337000) = 0x3337000 +12429 brk(0x3358000) = 0x3358000 +12429 brk(0x3379000) = 0x3379000 +12429 brk(0x339a000) = 0x339a000 +12429 brk(0x33bb000) = 0x33bb000 +12429 brk(0x33dc000) = 0x33dc000 +12429 brk(0x33fd000) = 0x33fd000 +12429 brk(0x341e000) = 0x341e000 +12429 brk(0x343f000) = 0x343f000 +12429 brk(0x3460000) = 0x3460000 +12429 brk(0x3481000) = 0x3481000 +12429 brk(0x34a2000) = 0x34a2000 +12429 brk(0x34c3000) = 0x34c3000 +12429 brk(0x34e4000) = 0x34e4000 +12429 brk(0x3505000) = 0x3505000 +12429 brk(0x3526000) = 0x3526000 +12429 brk(0x3547000) = 0x3547000 +12429 brk(0x3568000) = 0x3568000 +12429 brk(0x3589000) = 0x3589000 +12429 brk(0x35aa000) = 0x35aa000 +12429 brk(0x35cb000) = 0x35cb000 +12429 brk(0x35ec000) = 0x35ec000 +12429 brk(0x360d000) = 0x360d000 +12429 brk(0x362e000) = 0x362e000 +12429 mmap(NULL, 34607104, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c706f2000 +12429 mmap(NULL, 34607104, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c69eff000 +12429 mmap(NULL, 34607104, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c61eff000 +12429 mmap(NULL, 4294971392, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4b61efe000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", O_RDONLY|O_CLOEXEC) = 4 +12429 read(4, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\260\214\2\0\0\0\0\0"..., 832) = 832 +12429 fstat(4, {st_mode=S_IFREG|0755, st_size=1461464, ...}) = 0 +12429 mmap(NULL, 3556520, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 4, 0) = 0x7f4c7038d000 +12429 mprotect(0x7f4c704e1000, 2097152, PROT_NONE) = 0 +12429 mmap(0x7f4c706e1000, 69632, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 4, 0x154000) = 0x7f4c706e1000 +12429 close(4) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libAIR.so", O_RDONLY|O_CLOEXEC) = 4 +12429 read(4, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\0\234\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(4, {st_mode=S_IFREG|0755, st_size=135456, ...}) = 0 +12429 mmap(NULL, 2202280, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 4, 0) = 0x7f4c70173000 +12429 mprotect(0x7f4c7018c000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c7038b000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 4, 0x18000) = 0x7f4c7038b000 +12429 close(4) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libMDT.so", O_RDONLY|O_CLOEXEC) = 4 +12429 read(4, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\320\353\r\0\0\0\0\0"..., 832) = 832 +12429 fstat(4, {st_mode=S_IFREG|0755, st_size=2836832, ...}) = 0 +12429 mmap(NULL, 4885520, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 4, 0) = 0x7f4c69a56000 +12429 mprotect(0x7f4c69c87000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c69e86000, 475136, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 4, 0x230000) = 0x7f4c69e86000 +12429 mmap(0x7f4c69efa000, 19472, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_ANONYMOUS, -1, 0) = 0x7f4c69efa000 +12429 close(4) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libCCL.so", O_RDONLY|O_CLOEXEC) = 4 +12429 read(4, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0`4\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(4, {st_mode=S_IFREG|0755, st_size=65408, ...}) = 0 +12429 mmap(NULL, 2151080, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 4, 0) = 0x7f4c69848000 +12429 mprotect(0x7f4c69854000, 2097152, PROT_NONE) = 0 +12429 mmap(0x7f4c69a54000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 4, 0xc000) = 0x7f4c69a54000 +12429 close(4) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/lao/libBSL.so", O_RDONLY|O_CLOEXEC) = 4 +12429 read(4, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\360x\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(4, {st_mode=S_IFREG|0755, st_size=178120, ...}) = 0 +12429 mmap(NULL, 2259944, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 4, 0) = 0x7f4c69620000 +12429 mprotect(0x7f4c69646000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c69845000, 12288, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 4, 0x25000) = 0x7f4c69845000 +12429 close(4) = 0 +12429 mprotect(0x7f4c69845000, 8192, PROT_READ) = 0 +12429 mprotect(0x7f4c69a54000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c69e86000, 208896, PROT_READ) = 0 +12429 mprotect(0x7f4c7038b000, 4096, PROT_READ) = 0 +12429 mprotect(0x7f4c706e1000, 65536, PROT_READ) = 0 +12429 brk(0x364f000) = 0x364f000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 mmap(NULL, 6819840, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c68f9f000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 openat(AT_FDCWD, "/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", O_RDONLY|O_CLOEXEC) = 4 +12429 read(4, "\177ELF\2\1\1\0\0\0\0\0\0\0\0\0\3\0>\0\1\0\0\0\300\23\0\0\0\0\0\0"..., 832) = 832 +12429 fstat(4, {st_mode=S_IFREG|0755, st_size=19216, ...}) = 0 +12429 mmap(NULL, 2114360, PROT_READ|PROT_EXEC, MAP_PRIVATE|MAP_DENYWRITE, 4, 0) = 0x7f4c68d9a000 +12429 mprotect(0x7f4c68d9e000, 2093056, PROT_NONE) = 0 +12429 mmap(0x7f4c68f9d000, 8192, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_FIXED|MAP_DENYWRITE, 4, 0x3000) = 0x7f4c68f9d000 +12429 close(4) = 0 +12429 mprotect(0x7f4c68f9d000, 4096, PROT_READ) = 0 +12429 brk(0x3670000) = 0x3670000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 mmap(NULL, 6819840, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c68719000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 brk(0x3694000) = 0x3694000 +12429 munmap(0x7f4c68719000, 6819840) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x36b8000) = 0x36b8000 +12429 brk(0x3d39000) = 0x3d39000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x43cb000) = 0x43cb000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 mmap(NULL, 67112960, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4b5defd000 +12429 brk(0x444b000) = 0x444b000 +12429 brk(0x446c000) = 0x446c000 +12429 brk(0x448d000) = 0x448d000 +12429 brk(0x44ae000) = 0x44ae000 +12429 brk(0x44cf000) = 0x44cf000 +12429 brk(0x44f0000) = 0x44f0000 +12429 brk(0x4511000) = 0x4511000 +12429 brk(0x4532000) = 0x4532000 +12429 brk(0x4553000) = 0x4553000 +12429 brk(0x4574000) = 0x4574000 +12429 brk(0x4595000) = 0x4595000 +12429 brk(0x45b6000) = 0x45b6000 +12429 brk(0x45d7000) = 0x45d7000 +12429 brk(0x45f8000) = 0x45f8000 +12429 brk(0x4619000) = 0x4619000 +12429 brk(0x463a000) = 0x463a000 +12429 brk(0x465b000) = 0x465b000 +12429 brk(0x467c000) = 0x467c000 +12429 brk(0x469d000) = 0x469d000 +12429 brk(0x46be000) = 0x46be000 +12429 brk(0x46df000) = 0x46df000 +12429 brk(0x4700000) = 0x4700000 +12429 brk(0x4721000) = 0x4721000 +12429 brk(0x4742000) = 0x4742000 +12429 brk(0x4763000) = 0x4763000 +12429 brk(0x4784000) = 0x4784000 +12429 brk(0x47a5000) = 0x47a5000 +12429 brk(0x47c6000) = 0x47c6000 +12429 brk(0x47e7000) = 0x47e7000 +12429 brk(0x4808000) = 0x4808000 +12429 brk(0x4829000) = 0x4829000 +12429 brk(0x484a000) = 0x484a000 +12429 brk(0x486b000) = 0x486b000 +12429 brk(0x488c000) = 0x488c000 +12429 brk(0x48ad000) = 0x48ad000 +12429 brk(0x48ce000) = 0x48ce000 +12429 brk(0x48ef000) = 0x48ef000 +12429 brk(0x4910000) = 0x4910000 +12429 brk(0x4931000) = 0x4931000 +12429 brk(0x4952000) = 0x4952000 +12429 brk(0x4973000) = 0x4973000 +12429 brk(0x4994000) = 0x4994000 +12429 brk(0x49b5000) = 0x49b5000 +12429 brk(0x49d6000) = 0x49d6000 +12429 brk(0x49f7000) = 0x49f7000 +12429 brk(0x4a18000) = 0x4a18000 +12429 brk(0x4a39000) = 0x4a39000 +12429 brk(0x4a5a000) = 0x4a5a000 +12429 brk(0x4a7b000) = 0x4a7b000 +12429 brk(0x4a9c000) = 0x4a9c000 +12429 brk(0x4abd000) = 0x4abd000 +12429 brk(0x4ade000) = 0x4ade000 +12429 brk(0x4b0c000) = 0x4b0c000 +12429 brk(0x4b30000) = 0x4b30000 +12429 brk(0x4b51000) = 0x4b51000 +12429 brk(0x4b72000) = 0x4b72000 +12429 brk(0x4b93000) = 0x4b93000 +12429 brk(0x4bb4000) = 0x4bb4000 +12429 brk(0x4bd5000) = 0x4bd5000 +12429 brk(0x4bf6000) = 0x4bf6000 +12429 brk(0x4c17000) = 0x4c17000 +12429 brk(0x4c38000) = 0x4c38000 +12429 brk(0x4c59000) = 0x4c59000 +12429 brk(0x4c7a000) = 0x4c7a000 +12429 brk(0x4c9b000) = 0x4c9b000 +12429 brk(0x4cbc000) = 0x4cbc000 +12429 brk(0x4cdd000) = 0x4cdd000 +12429 brk(0x4cfe000) = 0x4cfe000 +12429 brk(0x4d1f000) = 0x4d1f000 +12429 brk(0x4d40000) = 0x4d40000 +12429 brk(0x4d61000) = 0x4d61000 +12429 brk(0x4d82000) = 0x4d82000 +12429 brk(0x4da3000) = 0x4da3000 +12429 brk(0x4dc4000) = 0x4dc4000 +12429 brk(0x4de5000) = 0x4de5000 +12429 brk(0x4e06000) = 0x4e06000 +12429 brk(0x4e27000) = 0x4e27000 +12429 brk(0x4e48000) = 0x4e48000 +12429 brk(0x4e69000) = 0x4e69000 +12429 brk(0x4e8a000) = 0x4e8a000 +12429 brk(0x4eab000) = 0x4eab000 +12429 brk(0x4ecc000) = 0x4ecc000 +12429 brk(0x4eed000) = 0x4eed000 +12429 brk(0x4f0e000) = 0x4f0e000 +12429 brk(0x4f2f000) = 0x4f2f000 +12429 brk(0x4f50000) = 0x4f50000 +12429 brk(0x4f71000) = 0x4f71000 +12429 brk(0x4f92000) = 0x4f92000 +12429 brk(0x4fb3000) = 0x4fb3000 +12429 brk(0x4fd4000) = 0x4fd4000 +12429 brk(0x4ff5000) = 0x4ff5000 +12429 brk(0x5016000) = 0x5016000 +12429 brk(0x5037000) = 0x5037000 +12429 brk(0x5058000) = 0x5058000 +12429 brk(0x5079000) = 0x5079000 +12429 brk(0x509a000) = 0x509a000 +12429 brk(0x50bb000) = 0x50bb000 +12429 brk(0x50dc000) = 0x50dc000 +12429 brk(0x50fd000) = 0x50fd000 +12429 brk(0x511e000) = 0x511e000 +12429 brk(0x513f000) = 0x513f000 +12429 brk(0x5160000) = 0x5160000 +12429 brk(0x5181000) = 0x5181000 +12429 brk(0x51a2000) = 0x51a2000 +12429 brk(0x51c3000) = 0x51c3000 +12429 brk(0x51e4000) = 0x51e4000 +12429 brk(0x5205000) = 0x5205000 +12429 brk(0x5226000) = 0x5226000 +12429 brk(0x5247000) = 0x5247000 +12429 brk(0x5268000) = 0x5268000 +12429 brk(0x5289000) = 0x5289000 +12429 brk(0x52aa000) = 0x52aa000 +12429 brk(0x52cb000) = 0x52cb000 +12429 brk(0x52ec000) = 0x52ec000 +12429 brk(0x530d000) = 0x530d000 +12429 brk(0x532e000) = 0x532e000 +12429 brk(0x534f000) = 0x534f000 +12429 brk(0x5370000) = 0x5370000 +12429 brk(0x5391000) = 0x5391000 +12429 brk(0x53b2000) = 0x53b2000 +12429 brk(0x53d3000) = 0x53d3000 +12429 brk(0x53f4000) = 0x53f4000 +12429 brk(0x5415000) = 0x5415000 +12429 brk(0x5436000) = 0x5436000 +12429 brk(0x5457000) = 0x5457000 +12429 brk(0x5478000) = 0x5478000 +12429 brk(0x5499000) = 0x5499000 +12429 brk(0x54ba000) = 0x54ba000 +12429 brk(0x54db000) = 0x54db000 +12429 brk(0x54fc000) = 0x54fc000 +12429 brk(0x551d000) = 0x551d000 +12429 brk(0x553e000) = 0x553e000 +12429 brk(0x555f000) = 0x555f000 +12429 brk(0x5580000) = 0x5580000 +12429 brk(0x55a1000) = 0x55a1000 +12429 brk(0x55c2000) = 0x55c2000 +12429 brk(0x55e3000) = 0x55e3000 +12429 brk(0x5604000) = 0x5604000 +12429 brk(0x5625000) = 0x5625000 +12429 brk(0x5646000) = 0x5646000 +12429 brk(0x5667000) = 0x5667000 +12429 brk(0x5688000) = 0x5688000 +12429 brk(0x56a9000) = 0x56a9000 +12429 brk(0x56ca000) = 0x56ca000 +12429 brk(0x56eb000) = 0x56eb000 +12429 brk(0x570c000) = 0x570c000 +12429 brk(0x572d000) = 0x572d000 +12429 brk(0x574e000) = 0x574e000 +12429 brk(0x576f000) = 0x576f000 +12429 brk(0x5790000) = 0x5790000 +12429 brk(0x57b1000) = 0x57b1000 +12429 brk(0x57d2000) = 0x57d2000 +12429 brk(0x57f3000) = 0x57f3000 +12429 brk(0x5814000) = 0x5814000 +12429 brk(0x5835000) = 0x5835000 +12429 brk(0x5856000) = 0x5856000 +12429 brk(0x5877000) = 0x5877000 +12429 brk(0x5898000) = 0x5898000 +12429 brk(0x58b9000) = 0x58b9000 +12429 brk(0x58da000) = 0x58da000 +12429 brk(0x58fb000) = 0x58fb000 +12429 brk(0x591c000) = 0x591c000 +12429 brk(0x593d000) = 0x593d000 +12429 brk(0x595e000) = 0x595e000 +12429 brk(0x597f000) = 0x597f000 +12429 brk(0x59a0000) = 0x59a0000 +12429 brk(0x59c1000) = 0x59c1000 +12429 brk(0x59e2000) = 0x59e2000 +12429 brk(0x5a03000) = 0x5a03000 +12429 brk(0x5a24000) = 0x5a24000 +12429 brk(0x5a45000) = 0x5a45000 +12429 brk(0x5a66000) = 0x5a66000 +12429 brk(0x5a87000) = 0x5a87000 +12429 brk(0x5aa8000) = 0x5aa8000 +12429 brk(0x5ac9000) = 0x5ac9000 +12429 brk(0x5aea000) = 0x5aea000 +12429 brk(0x5b0b000) = 0x5b0b000 +12429 brk(0x5b2c000) = 0x5b2c000 +12429 brk(0x5b4d000) = 0x5b4d000 +12429 brk(0x5b6e000) = 0x5b6e000 +12429 brk(0x5b8f000) = 0x5b8f000 +12429 brk(0x5bb0000) = 0x5bb0000 +12429 brk(0x5bd1000) = 0x5bd1000 +12429 brk(0x5bf2000) = 0x5bf2000 +12429 brk(0x5c13000) = 0x5c13000 +12429 brk(0x5c34000) = 0x5c34000 +12429 brk(0x5c55000) = 0x5c55000 +12429 brk(0x5c76000) = 0x5c76000 +12429 brk(0x5c97000) = 0x5c97000 +12429 brk(0x5cb8000) = 0x5cb8000 +12429 brk(0x5cd9000) = 0x5cd9000 +12429 brk(0x5cfa000) = 0x5cfa000 +12429 brk(0x5d1b000) = 0x5d1b000 +12429 brk(0x5d3c000) = 0x5d3c000 +12429 brk(0x5d60000) = 0x5d60000 +12429 brk(0x5d81000) = 0x5d81000 +12429 brk(0x5da2000) = 0x5da2000 +12429 brk(0x5dc3000) = 0x5dc3000 +12429 brk(0x5de4000) = 0x5de4000 +12429 brk(0x5e05000) = 0x5e05000 +12429 brk(0x5e26000) = 0x5e26000 +12429 brk(0x5e47000) = 0x5e47000 +12429 brk(0x5e68000) = 0x5e68000 +12429 brk(0x5e89000) = 0x5e89000 +12429 brk(0x5eaa000) = 0x5eaa000 +12429 brk(0x5ed0000) = 0x5ed0000 +12429 brk(0x5ef1000) = 0x5ef1000 +12429 brk(0x5f12000) = 0x5f12000 +12429 brk(0x5f33000) = 0x5f33000 +12429 brk(0x5f54000) = 0x5f54000 +12429 brk(0x5f75000) = 0x5f75000 +12429 brk(0x5f96000) = 0x5f96000 +12429 brk(0x5fb7000) = 0x5fb7000 +12429 brk(0x5fd8000) = 0x5fd8000 +12429 brk(0x5ff9000) = 0x5ff9000 +12429 brk(0x601a000) = 0x601a000 +12429 brk(0x603b000) = 0x603b000 +12429 brk(0x605c000) = 0x605c000 +12429 brk(0x607d000) = 0x607d000 +12429 brk(0x609e000) = 0x609e000 +12429 brk(0x60bf000) = 0x60bf000 +12429 brk(0x60e0000) = 0x60e0000 +12429 brk(0x6101000) = 0x6101000 +12429 brk(0x6122000) = 0x6122000 +12429 brk(0x6143000) = 0x6143000 +12429 brk(0x6164000) = 0x6164000 +12429 brk(0x6185000) = 0x6185000 +12429 brk(0x61b0000) = 0x61b0000 +12429 brk(0x61d1000) = 0x61d1000 +12429 brk(0x61f2000) = 0x61f2000 +12429 brk(0x6213000) = 0x6213000 +12429 brk(0x6234000) = 0x6234000 +12429 brk(0x6255000) = 0x6255000 +12429 brk(0x6276000) = 0x6276000 +12429 brk(0x6297000) = 0x6297000 +12429 brk(0x62b8000) = 0x62b8000 +12429 brk(0x62d9000) = 0x62d9000 +12429 brk(0x62fa000) = 0x62fa000 +12429 brk(0x631b000) = 0x631b000 +12429 brk(0x633c000) = 0x633c000 +12429 brk(0x635d000) = 0x635d000 +12429 brk(0x637e000) = 0x637e000 +12429 brk(0x639f000) = 0x639f000 +12429 brk(0x63c0000) = 0x63c0000 +12429 brk(0x63e1000) = 0x63e1000 +12429 brk(0x6402000) = 0x6402000 +12429 brk(0x6423000) = 0x6423000 +12429 brk(0x6444000) = 0x6444000 +12429 brk(0x6465000) = 0x6465000 +12429 brk(0x6486000) = 0x6486000 +12429 brk(0x64a7000) = 0x64a7000 +12429 brk(0x64c8000) = 0x64c8000 +12429 brk(0x64e9000) = 0x64e9000 +12429 brk(0x650a000) = 0x650a000 +12429 brk(0x652b000) = 0x652b000 +12429 brk(0x654c000) = 0x654c000 +12429 brk(0x656d000) = 0x656d000 +12429 brk(0x658e000) = 0x658e000 +12429 brk(0x65af000) = 0x65af000 +12429 brk(0x65d0000) = 0x65d0000 +12429 brk(0x65f1000) = 0x65f1000 +12429 brk(0x6612000) = 0x6612000 +12429 brk(0x6633000) = 0x6633000 +12429 brk(0x6654000) = 0x6654000 +12429 brk(0x6675000) = 0x6675000 +12429 brk(0x6696000) = 0x6696000 +12429 brk(0x66b7000) = 0x66b7000 +12429 brk(0x66d8000) = 0x66d8000 +12429 brk(0x66f9000) = 0x66f9000 +12429 brk(0x671a000) = 0x671a000 +12429 brk(0x673b000) = 0x673b000 +12429 brk(0x6770000) = 0x6770000 +12429 brk(0x6791000) = 0x6791000 +12429 brk(0x67b2000) = 0x67b2000 +12429 brk(0x67d3000) = 0x67d3000 +12429 brk(0x67f4000) = 0x67f4000 +12429 brk(0x6815000) = 0x6815000 +12429 brk(0x6836000) = 0x6836000 +12429 brk(0x6857000) = 0x6857000 +12429 brk(0x6878000) = 0x6878000 +12429 brk(0x6899000) = 0x6899000 +12429 brk(0x68ba000) = 0x68ba000 +12429 brk(0x68db000) = 0x68db000 +12429 brk(0x68fc000) = 0x68fc000 +12429 brk(0x691d000) = 0x691d000 +12429 brk(0x693e000) = 0x693e000 +12429 brk(0x695f000) = 0x695f000 +12429 brk(0x6980000) = 0x6980000 +12429 brk(0x69a1000) = 0x69a1000 +12429 brk(0x69c2000) = 0x69c2000 +12429 brk(0x69e3000) = 0x69e3000 +12429 brk(0x6a04000) = 0x6a04000 +12429 brk(0x6a25000) = 0x6a25000 +12429 brk(0x6a46000) = 0x6a46000 +12429 brk(0x6a67000) = 0x6a67000 +12429 brk(0x6a88000) = 0x6a88000 +12429 brk(0x6aa9000) = 0x6aa9000 +12429 brk(0x6aca000) = 0x6aca000 +12429 brk(0x6aeb000) = 0x6aeb000 +12429 brk(0x6b0c000) = 0x6b0c000 +12429 brk(0x6b2d000) = 0x6b2d000 +12429 brk(0x6b4e000) = 0x6b4e000 +12429 brk(0x6b6f000) = 0x6b6f000 +12429 brk(0x6b90000) = 0x6b90000 +12429 brk(0x6bb1000) = 0x6bb1000 +12429 brk(0x6bd2000) = 0x6bd2000 +12429 brk(0x6bf3000) = 0x6bf3000 +12429 brk(0x6c14000) = 0x6c14000 +12429 brk(0x6c35000) = 0x6c35000 +12429 brk(0x6c56000) = 0x6c56000 +12429 brk(0x6c77000) = 0x6c77000 +12429 brk(0x6c98000) = 0x6c98000 +12429 brk(0x6cb9000) = 0x6cb9000 +12429 brk(0x6cda000) = 0x6cda000 +12429 brk(0x6cfb000) = 0x6cfb000 +12429 brk(0x6d1c000) = 0x6d1c000 +12429 brk(0x6d3d000) = 0x6d3d000 +12429 brk(0x6d5e000) = 0x6d5e000 +12429 brk(0x6d7f000) = 0x6d7f000 +12429 brk(0x6da0000) = 0x6da0000 +12429 brk(0x6dc1000) = 0x6dc1000 +12429 brk(0x6de2000) = 0x6de2000 +12429 brk(0x6e03000) = 0x6e03000 +12429 brk(0x6e24000) = 0x6e24000 +12429 brk(0x6e45000) = 0x6e45000 +12429 brk(0x6e66000) = 0x6e66000 +12429 brk(0x6e87000) = 0x6e87000 +12429 brk(0x6ea8000) = 0x6ea8000 +12429 brk(0x6ec9000) = 0x6ec9000 +12429 brk(0x6eea000) = 0x6eea000 +12429 brk(0x6f0b000) = 0x6f0b000 +12429 brk(0x6f2c000) = 0x6f2c000 +12429 brk(0x6f4d000) = 0x6f4d000 +12429 brk(0x6f6e000) = 0x6f6e000 +12429 brk(0x6f8f000) = 0x6f8f000 +12429 brk(0x6fb0000) = 0x6fb0000 +12429 brk(0x6fd1000) = 0x6fd1000 +12429 brk(0x6ff2000) = 0x6ff2000 +12429 brk(0x7013000) = 0x7013000 +12429 brk(0x7034000) = 0x7034000 +12429 brk(0x7055000) = 0x7055000 +12429 brk(0x7076000) = 0x7076000 +12429 brk(0x7097000) = 0x7097000 +12429 brk(0x70b8000) = 0x70b8000 +12429 brk(0x70d9000) = 0x70d9000 +12429 brk(0x70fa000) = 0x70fa000 +12429 brk(0x711b000) = 0x711b000 +12429 brk(0x713c000) = 0x713c000 +12429 brk(0x715d000) = 0x715d000 +12429 brk(0x717e000) = 0x717e000 +12429 brk(0x719f000) = 0x719f000 +12429 brk(0x71c0000) = 0x71c0000 +12429 brk(0x71e1000) = 0x71e1000 +12429 brk(0x7202000) = 0x7202000 +12429 brk(0x7223000) = 0x7223000 +12429 brk(0x7244000) = 0x7244000 +12429 brk(0x7265000) = 0x7265000 +12429 brk(0x7286000) = 0x7286000 +12429 brk(0x72a7000) = 0x72a7000 +12429 brk(0x74b8000) = 0x74b8000 +12429 brk(0x76b8000) = 0x76b8000 +12429 brk(0x78b8000) = 0x78b8000 +12429 brk(0x78d9000) = 0x78d9000 +12429 brk(0x78fa000) = 0x78fa000 +12429 brk(0x7921000) = 0x7921000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x7fa1000) = 0x7fa1000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 brk(0x7fc2000) = 0x7fc2000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x8650000) = 0x8650000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x8ce3000) = 0x8ce3000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 mmap(NULL, 67112960, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4b59efc000 +12429 brk(0x8d04000) = 0x8d04000 +12429 brk(0x8d25000) = 0x8d25000 +12429 brk(0x8d46000) = 0x8d46000 +12429 brk(0x8d67000) = 0x8d67000 +12429 brk(0x8d88000) = 0x8d88000 +12429 brk(0x8da9000) = 0x8da9000 +12429 brk(0x8dca000) = 0x8dca000 +12429 brk(0x8deb000) = 0x8deb000 +12429 brk(0x8e0c000) = 0x8e0c000 +12429 brk(0x8e2d000) = 0x8e2d000 +12429 brk(0x8e4e000) = 0x8e4e000 +12429 brk(0x8e6f000) = 0x8e6f000 +12429 brk(0x8e90000) = 0x8e90000 +12429 brk(0x8eb1000) = 0x8eb1000 +12429 brk(0x8ed2000) = 0x8ed2000 +12429 brk(0x8ef3000) = 0x8ef3000 +12429 brk(0x8f14000) = 0x8f14000 +12429 brk(0x8f35000) = 0x8f35000 +12429 brk(0x8f56000) = 0x8f56000 +12429 brk(0x8f77000) = 0x8f77000 +12429 brk(0x8f98000) = 0x8f98000 +12429 brk(0x8fb9000) = 0x8fb9000 +12429 brk(0x8fda000) = 0x8fda000 +12429 brk(0x8ffb000) = 0x8ffb000 +12429 brk(0x901c000) = 0x901c000 +12429 brk(0x903d000) = 0x903d000 +12429 brk(0x905e000) = 0x905e000 +12429 brk(0x907f000) = 0x907f000 +12429 brk(0x90a0000) = 0x90a0000 +12429 brk(0x90c1000) = 0x90c1000 +12429 brk(0x90e2000) = 0x90e2000 +12429 brk(0x9103000) = 0x9103000 +12429 brk(0x9124000) = 0x9124000 +12429 brk(0x9145000) = 0x9145000 +12429 brk(0x9166000) = 0x9166000 +12429 brk(0x9187000) = 0x9187000 +12429 brk(0x91a8000) = 0x91a8000 +12429 brk(0x91c9000) = 0x91c9000 +12429 brk(0x91ea000) = 0x91ea000 +12429 brk(0x920b000) = 0x920b000 +12429 brk(0x922c000) = 0x922c000 +12429 brk(0x924d000) = 0x924d000 +12429 brk(0x926e000) = 0x926e000 +12429 brk(0x928f000) = 0x928f000 +12429 brk(0x92b0000) = 0x92b0000 +12429 brk(0x92d1000) = 0x92d1000 +12429 brk(0x92f2000) = 0x92f2000 +12429 brk(0x9324000) = 0x9324000 +12429 brk(0x9345000) = 0x9345000 +12429 brk(0x9366000) = 0x9366000 +12429 brk(0x9387000) = 0x9387000 +12429 brk(0x93a8000) = 0x93a8000 +12429 brk(0x93c9000) = 0x93c9000 +12429 brk(0x93ea000) = 0x93ea000 +12429 brk(0x940b000) = 0x940b000 +12429 brk(0x942d000) = 0x942d000 +12429 brk(0x944e000) = 0x944e000 +12429 brk(0x946f000) = 0x946f000 +12429 brk(0x9490000) = 0x9490000 +12429 brk(0x94b1000) = 0x94b1000 +12429 brk(0x94d2000) = 0x94d2000 +12429 brk(0x94f3000) = 0x94f3000 +12429 brk(0x9514000) = 0x9514000 +12429 brk(0x9535000) = 0x9535000 +12429 brk(0x9556000) = 0x9556000 +12429 brk(0x9577000) = 0x9577000 +12429 brk(0x9598000) = 0x9598000 +12429 brk(0x95b9000) = 0x95b9000 +12429 brk(0x95da000) = 0x95da000 +12429 brk(0x95fb000) = 0x95fb000 +12429 brk(0x961c000) = 0x961c000 +12429 brk(0x963d000) = 0x963d000 +12429 brk(0x965e000) = 0x965e000 +12429 brk(0x967f000) = 0x967f000 +12429 brk(0x96a0000) = 0x96a0000 +12429 brk(0x96c1000) = 0x96c1000 +12429 brk(0x96e2000) = 0x96e2000 +12429 brk(0x9703000) = 0x9703000 +12429 brk(0x9724000) = 0x9724000 +12429 brk(0x9745000) = 0x9745000 +12429 brk(0x976b000) = 0x976b000 +12429 brk(0x978c000) = 0x978c000 +12429 brk(0x97ad000) = 0x97ad000 +12429 brk(0x97ce000) = 0x97ce000 +12429 brk(0x97ef000) = 0x97ef000 +12429 brk(0x9810000) = 0x9810000 +12429 brk(0x9831000) = 0x9831000 +12429 brk(0x9852000) = 0x9852000 +12429 brk(0x9873000) = 0x9873000 +12429 brk(0x9894000) = 0x9894000 +12429 brk(0x98b5000) = 0x98b5000 +12429 brk(0x98d6000) = 0x98d6000 +12429 brk(0x98f7000) = 0x98f7000 +12429 brk(0x9918000) = 0x9918000 +12429 brk(0x9939000) = 0x9939000 +12429 brk(0x995a000) = 0x995a000 +12429 brk(0x997b000) = 0x997b000 +12429 brk(0x999c000) = 0x999c000 +12429 brk(0x99bd000) = 0x99bd000 +12429 brk(0x99de000) = 0x99de000 +12429 brk(0x99ff000) = 0x99ff000 +12429 brk(0x9a20000) = 0x9a20000 +12429 brk(0x9a41000) = 0x9a41000 +12429 brk(0x9a62000) = 0x9a62000 +12429 brk(0x9a83000) = 0x9a83000 +12429 brk(0x9aa4000) = 0x9aa4000 +12429 brk(0x9ac5000) = 0x9ac5000 +12429 brk(0x9ae6000) = 0x9ae6000 +12429 brk(0x9b07000) = 0x9b07000 +12429 brk(0x9b28000) = 0x9b28000 +12429 brk(0x9b49000) = 0x9b49000 +12429 brk(0x9b6a000) = 0x9b6a000 +12429 brk(0x9b8b000) = 0x9b8b000 +12429 brk(0x9bac000) = 0x9bac000 +12429 brk(0x9bcd000) = 0x9bcd000 +12429 brk(0x9bee000) = 0x9bee000 +12429 brk(0x9c0f000) = 0x9c0f000 +12429 brk(0x9c30000) = 0x9c30000 +12429 brk(0x9c51000) = 0x9c51000 +12429 brk(0x9c72000) = 0x9c72000 +12429 brk(0x9c93000) = 0x9c93000 +12429 brk(0x9cb4000) = 0x9cb4000 +12429 brk(0x9cd5000) = 0x9cd5000 +12429 brk(0x9cf6000) = 0x9cf6000 +12429 brk(0x9d17000) = 0x9d17000 +12429 brk(0x9d38000) = 0x9d38000 +12429 brk(0x9d59000) = 0x9d59000 +12429 brk(0x9d7a000) = 0x9d7a000 +12429 brk(0x9d9b000) = 0x9d9b000 +12429 brk(0x9dbc000) = 0x9dbc000 +12429 brk(0x9ddd000) = 0x9ddd000 +12429 brk(0x9dfe000) = 0x9dfe000 +12429 brk(0x9e1f000) = 0x9e1f000 +12429 brk(0x9e40000) = 0x9e40000 +12429 brk(0x9e61000) = 0x9e61000 +12429 brk(0x9e82000) = 0x9e82000 +12429 brk(0x9ea3000) = 0x9ea3000 +12429 brk(0x9ec4000) = 0x9ec4000 +12429 brk(0x9ee5000) = 0x9ee5000 +12429 brk(0x9f06000) = 0x9f06000 +12429 brk(0x9f27000) = 0x9f27000 +12429 brk(0x9f48000) = 0x9f48000 +12429 brk(0x9f69000) = 0x9f69000 +12429 brk(0x9f8a000) = 0x9f8a000 +12429 brk(0x9fab000) = 0x9fab000 +12429 brk(0x9fcc000) = 0x9fcc000 +12429 brk(0x9fed000) = 0x9fed000 +12429 brk(0xa00e000) = 0xa00e000 +12429 brk(0xa02f000) = 0xa02f000 +12429 brk(0xa050000) = 0xa050000 +12429 brk(0xa071000) = 0xa071000 +12429 brk(0xa092000) = 0xa092000 +12429 brk(0xa0b3000) = 0xa0b3000 +12429 brk(0xa0d4000) = 0xa0d4000 +12429 brk(0xa0f5000) = 0xa0f5000 +12429 brk(0xa116000) = 0xa116000 +12429 brk(0xa137000) = 0xa137000 +12429 brk(0xa158000) = 0xa158000 +12429 brk(0xa179000) = 0xa179000 +12429 brk(0xa19a000) = 0xa19a000 +12429 brk(0xa1bb000) = 0xa1bb000 +12429 brk(0xa1dc000) = 0xa1dc000 +12429 brk(0xa1fd000) = 0xa1fd000 +12429 brk(0xa21e000) = 0xa21e000 +12429 brk(0xa242000) = 0xa242000 +12429 brk(0xa263000) = 0xa263000 +12429 brk(0xa284000) = 0xa284000 +12429 brk(0xa2a5000) = 0xa2a5000 +12429 brk(0xa2c6000) = 0xa2c6000 +12429 brk(0xa2e7000) = 0xa2e7000 +12429 brk(0xa308000) = 0xa308000 +12429 brk(0xa329000) = 0xa329000 +12429 brk(0xa34a000) = 0xa34a000 +12429 brk(0xa36b000) = 0xa36b000 +12429 brk(0xa38c000) = 0xa38c000 +12429 brk(0xa3b2000) = 0xa3b2000 +12429 brk(0xa3d3000) = 0xa3d3000 +12429 brk(0xa3f4000) = 0xa3f4000 +12429 brk(0xa415000) = 0xa415000 +12429 brk(0xa436000) = 0xa436000 +12429 brk(0xa457000) = 0xa457000 +12429 brk(0xa478000) = 0xa478000 +12429 brk(0xa499000) = 0xa499000 +12429 brk(0xa4ba000) = 0xa4ba000 +12429 brk(0xa4db000) = 0xa4db000 +12429 brk(0xa4fc000) = 0xa4fc000 +12429 brk(0xa51d000) = 0xa51d000 +12429 brk(0xa53e000) = 0xa53e000 +12429 brk(0xa55f000) = 0xa55f000 +12429 brk(0xa580000) = 0xa580000 +12429 brk(0xa5a1000) = 0xa5a1000 +12429 brk(0xa5c2000) = 0xa5c2000 +12429 brk(0xa5e3000) = 0xa5e3000 +12429 brk(0xa604000) = 0xa604000 +12429 brk(0xa625000) = 0xa625000 +12429 brk(0xa646000) = 0xa646000 +12429 brk(0xa667000) = 0xa667000 +12429 brk(0xa692000) = 0xa692000 +12429 brk(0xa6b3000) = 0xa6b3000 +12429 brk(0xa6d4000) = 0xa6d4000 +12429 brk(0xa6f5000) = 0xa6f5000 +12429 brk(0xa716000) = 0xa716000 +12429 brk(0xa737000) = 0xa737000 +12429 brk(0xa758000) = 0xa758000 +12429 brk(0xa779000) = 0xa779000 +12429 brk(0xa79a000) = 0xa79a000 +12429 brk(0xa7bb000) = 0xa7bb000 +12429 brk(0xa7dc000) = 0xa7dc000 +12429 brk(0xa7fd000) = 0xa7fd000 +12429 brk(0xa81e000) = 0xa81e000 +12429 brk(0xa83f000) = 0xa83f000 +12429 brk(0xa860000) = 0xa860000 +12429 brk(0xa881000) = 0xa881000 +12429 brk(0xa8a2000) = 0xa8a2000 +12429 brk(0xa8c3000) = 0xa8c3000 +12429 brk(0xa8e4000) = 0xa8e4000 +12429 brk(0xa905000) = 0xa905000 +12429 brk(0xa926000) = 0xa926000 +12429 brk(0xa947000) = 0xa947000 +12429 brk(0xa968000) = 0xa968000 +12429 brk(0xa989000) = 0xa989000 +12429 brk(0xa9aa000) = 0xa9aa000 +12429 brk(0xa9cb000) = 0xa9cb000 +12429 brk(0xa9ec000) = 0xa9ec000 +12429 brk(0xaa0d000) = 0xaa0d000 +12429 brk(0xaa2e000) = 0xaa2e000 +12429 brk(0xaa4f000) = 0xaa4f000 +12429 brk(0xaa70000) = 0xaa70000 +12429 brk(0xaa91000) = 0xaa91000 +12429 brk(0xaab2000) = 0xaab2000 +12429 brk(0xaad3000) = 0xaad3000 +12429 brk(0xaaf4000) = 0xaaf4000 +12429 brk(0xab15000) = 0xab15000 +12429 brk(0xab36000) = 0xab36000 +12429 brk(0xab57000) = 0xab57000 +12429 brk(0xab78000) = 0xab78000 +12429 brk(0xab99000) = 0xab99000 +12429 brk(0xabba000) = 0xabba000 +12429 brk(0xabdb000) = 0xabdb000 +12429 brk(0xabfc000) = 0xabfc000 +12429 brk(0xac1d000) = 0xac1d000 +12429 brk(0xac52000) = 0xac52000 +12429 brk(0xac73000) = 0xac73000 +12429 brk(0xac94000) = 0xac94000 +12429 brk(0xacb5000) = 0xacb5000 +12429 brk(0xacd6000) = 0xacd6000 +12429 brk(0xacf7000) = 0xacf7000 +12429 brk(0xad18000) = 0xad18000 +12429 brk(0xad39000) = 0xad39000 +12429 brk(0xad5a000) = 0xad5a000 +12429 brk(0xad7b000) = 0xad7b000 +12429 brk(0xad9c000) = 0xad9c000 +12429 brk(0xadbd000) = 0xadbd000 +12429 brk(0xadde000) = 0xadde000 +12429 brk(0xadff000) = 0xadff000 +12429 brk(0xae20000) = 0xae20000 +12429 brk(0xae41000) = 0xae41000 +12429 brk(0xae62000) = 0xae62000 +12429 brk(0xae83000) = 0xae83000 +12429 brk(0xaea4000) = 0xaea4000 +12429 brk(0xaec5000) = 0xaec5000 +12429 brk(0xaee6000) = 0xaee6000 +12429 brk(0xaf07000) = 0xaf07000 +12429 brk(0xaf28000) = 0xaf28000 +12429 brk(0xaf49000) = 0xaf49000 +12429 brk(0xaf6a000) = 0xaf6a000 +12429 brk(0xaf8b000) = 0xaf8b000 +12429 brk(0xafac000) = 0xafac000 +12429 brk(0xafcd000) = 0xafcd000 +12429 brk(0xafee000) = 0xafee000 +12429 brk(0xb00f000) = 0xb00f000 +12429 brk(0xb030000) = 0xb030000 +12429 brk(0xb051000) = 0xb051000 +12429 brk(0xb072000) = 0xb072000 +12429 brk(0xb093000) = 0xb093000 +12429 brk(0xb0b4000) = 0xb0b4000 +12429 brk(0xb0d5000) = 0xb0d5000 +12429 brk(0xb0f6000) = 0xb0f6000 +12429 brk(0xb117000) = 0xb117000 +12429 brk(0xb138000) = 0xb138000 +12429 brk(0xb159000) = 0xb159000 +12429 brk(0xb17a000) = 0xb17a000 +12429 brk(0xb19b000) = 0xb19b000 +12429 brk(0xb1bc000) = 0xb1bc000 +12429 brk(0xb1dd000) = 0xb1dd000 +12429 brk(0xb1fe000) = 0xb1fe000 +12429 brk(0xb21f000) = 0xb21f000 +12429 brk(0xb240000) = 0xb240000 +12429 brk(0xb261000) = 0xb261000 +12429 brk(0xb282000) = 0xb282000 +12429 brk(0xb2a3000) = 0xb2a3000 +12429 brk(0xb2c4000) = 0xb2c4000 +12429 brk(0xb2e5000) = 0xb2e5000 +12429 brk(0xb306000) = 0xb306000 +12429 brk(0xb327000) = 0xb327000 +12429 brk(0xb348000) = 0xb348000 +12429 brk(0xb369000) = 0xb369000 +12429 brk(0xb38a000) = 0xb38a000 +12429 brk(0xb3ab000) = 0xb3ab000 +12429 brk(0xb3cc000) = 0xb3cc000 +12429 brk(0xb3ed000) = 0xb3ed000 +12429 brk(0xb40e000) = 0xb40e000 +12429 brk(0xb42f000) = 0xb42f000 +12429 brk(0xb450000) = 0xb450000 +12429 brk(0xb471000) = 0xb471000 +12429 brk(0xb492000) = 0xb492000 +12429 brk(0xb4b3000) = 0xb4b3000 +12429 brk(0xb4d4000) = 0xb4d4000 +12429 brk(0xb4f5000) = 0xb4f5000 +12429 brk(0xb516000) = 0xb516000 +12429 brk(0xb537000) = 0xb537000 +12429 brk(0xb558000) = 0xb558000 +12429 brk(0xb579000) = 0xb579000 +12429 brk(0xb59a000) = 0xb59a000 +12429 brk(0xb5bb000) = 0xb5bb000 +12429 brk(0xb5dc000) = 0xb5dc000 +12429 brk(0xb5fd000) = 0xb5fd000 +12429 brk(0xb61e000) = 0xb61e000 +12429 brk(0xb63f000) = 0xb63f000 +12429 brk(0xb660000) = 0xb660000 +12429 brk(0xb681000) = 0xb681000 +12429 brk(0xb6a2000) = 0xb6a2000 +12429 brk(0xb6c3000) = 0xb6c3000 +12429 brk(0xb6e4000) = 0xb6e4000 +12429 brk(0xb705000) = 0xb705000 +12429 brk(0xb726000) = 0xb726000 +12429 brk(0xb747000) = 0xb747000 +12429 brk(0xb768000) = 0xb768000 +12429 brk(0xb789000) = 0xb789000 +12429 brk(0xb99a000) = 0xb99a000 +12429 brk(0xbb9b000) = 0xbb9b000 +12429 brk(0xbd9b000) = 0xbd9b000 +12429 brk(0xbdbc000) = 0xbdbc000 +12429 brk(0xbddd000) = 0xbddd000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0xc469000) = 0xc469000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 brk(0xc4c7000) = 0xc4c7000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0xcb47000) = 0xcb47000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0xd1cd000) = 0xd1cd000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 mmap(NULL, 67112960, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4b55efb000 +12429 brk(0xd1ee000) = 0xd1ee000 +12429 brk(0xd20f000) = 0xd20f000 +12429 brk(0xd230000) = 0xd230000 +12429 brk(0xd251000) = 0xd251000 +12429 brk(0xd272000) = 0xd272000 +12429 brk(0xd293000) = 0xd293000 +12429 brk(0xd2b4000) = 0xd2b4000 +12429 brk(0xd2d5000) = 0xd2d5000 +12429 brk(0xd2f6000) = 0xd2f6000 +12429 brk(0xd317000) = 0xd317000 +12429 brk(0xd338000) = 0xd338000 +12429 brk(0xd359000) = 0xd359000 +12429 brk(0xd37a000) = 0xd37a000 +12429 brk(0xd39b000) = 0xd39b000 +12429 brk(0xd3bc000) = 0xd3bc000 +12429 brk(0xd3dd000) = 0xd3dd000 +12429 brk(0xd3fe000) = 0xd3fe000 +12429 brk(0xd41f000) = 0xd41f000 +12429 brk(0xd440000) = 0xd440000 +12429 brk(0xd461000) = 0xd461000 +12429 brk(0xd482000) = 0xd482000 +12429 brk(0xd4a3000) = 0xd4a3000 +12429 brk(0xd4c4000) = 0xd4c4000 +12429 brk(0xd4e5000) = 0xd4e5000 +12429 brk(0xd506000) = 0xd506000 +12429 brk(0xd527000) = 0xd527000 +12429 brk(0xd548000) = 0xd548000 +12429 brk(0xd569000) = 0xd569000 +12429 brk(0xd58a000) = 0xd58a000 +12429 brk(0xd5ab000) = 0xd5ab000 +12429 brk(0xd5cc000) = 0xd5cc000 +12429 brk(0xd5ed000) = 0xd5ed000 +12429 brk(0xd60e000) = 0xd60e000 +12429 brk(0xd62f000) = 0xd62f000 +12429 brk(0xd650000) = 0xd650000 +12429 brk(0xd671000) = 0xd671000 +12429 brk(0xd692000) = 0xd692000 +12429 brk(0xd6b3000) = 0xd6b3000 +12429 brk(0xd6d4000) = 0xd6d4000 +12429 brk(0xd6f5000) = 0xd6f5000 +12429 brk(0xd716000) = 0xd716000 +12429 brk(0xd737000) = 0xd737000 +12429 brk(0xd758000) = 0xd758000 +12429 brk(0xd779000) = 0xd779000 +12429 brk(0xd79a000) = 0xd79a000 +12429 brk(0xd7bb000) = 0xd7bb000 +12429 brk(0xd7dc000) = 0xd7dc000 +12429 brk(0xd80d000) = 0xd80d000 +12429 brk(0xd82e000) = 0xd82e000 +12429 brk(0xd84f000) = 0xd84f000 +12429 brk(0xd870000) = 0xd870000 +12429 brk(0xd891000) = 0xd891000 +12429 brk(0xd8b2000) = 0xd8b2000 +12429 brk(0xd8d3000) = 0xd8d3000 +12429 brk(0xd8f4000) = 0xd8f4000 +12429 brk(0xd916000) = 0xd916000 +12429 brk(0xd937000) = 0xd937000 +12429 brk(0xd958000) = 0xd958000 +12429 brk(0xd979000) = 0xd979000 +12429 brk(0xd99a000) = 0xd99a000 +12429 brk(0xd9bb000) = 0xd9bb000 +12429 brk(0xd9dc000) = 0xd9dc000 +12429 brk(0xd9fd000) = 0xd9fd000 +12429 brk(0xda1e000) = 0xda1e000 +12429 brk(0xda3f000) = 0xda3f000 +12429 brk(0xda60000) = 0xda60000 +12429 brk(0xda81000) = 0xda81000 +12429 brk(0xdaa2000) = 0xdaa2000 +12429 brk(0xdac3000) = 0xdac3000 +12429 brk(0xdae4000) = 0xdae4000 +12429 brk(0xdb05000) = 0xdb05000 +12429 brk(0xdb26000) = 0xdb26000 +12429 brk(0xdb47000) = 0xdb47000 +12429 brk(0xdb68000) = 0xdb68000 +12429 brk(0xdb89000) = 0xdb89000 +12429 brk(0xdbaa000) = 0xdbaa000 +12429 brk(0xdbcb000) = 0xdbcb000 +12429 brk(0xdbec000) = 0xdbec000 +12429 brk(0xdc0d000) = 0xdc0d000 +12429 brk(0xdc2e000) = 0xdc2e000 +12429 brk(0xdc54000) = 0xdc54000 +12429 brk(0xdc75000) = 0xdc75000 +12429 brk(0xdc96000) = 0xdc96000 +12429 brk(0xdcb7000) = 0xdcb7000 +12429 brk(0xdcd8000) = 0xdcd8000 +12429 brk(0xdcf9000) = 0xdcf9000 +12429 brk(0xdd1a000) = 0xdd1a000 +12429 brk(0xdd3b000) = 0xdd3b000 +12429 brk(0xdd5c000) = 0xdd5c000 +12429 brk(0xdd7d000) = 0xdd7d000 +12429 brk(0xdd9e000) = 0xdd9e000 +12429 brk(0xddbf000) = 0xddbf000 +12429 brk(0xdde0000) = 0xdde0000 +12429 brk(0xde01000) = 0xde01000 +12429 brk(0xde22000) = 0xde22000 +12429 brk(0xde43000) = 0xde43000 +12429 brk(0xde64000) = 0xde64000 +12429 brk(0xde85000) = 0xde85000 +12429 brk(0xdea6000) = 0xdea6000 +12429 brk(0xdec7000) = 0xdec7000 +12429 brk(0xdee8000) = 0xdee8000 +12429 brk(0xdf09000) = 0xdf09000 +12429 brk(0xdf2a000) = 0xdf2a000 +12429 brk(0xdf4b000) = 0xdf4b000 +12429 brk(0xdf6c000) = 0xdf6c000 +12429 brk(0xdf8d000) = 0xdf8d000 +12429 brk(0xdfae000) = 0xdfae000 +12429 brk(0xdfcf000) = 0xdfcf000 +12429 brk(0xdff0000) = 0xdff0000 +12429 brk(0xe011000) = 0xe011000 +12429 brk(0xe032000) = 0xe032000 +12429 brk(0xe053000) = 0xe053000 +12429 brk(0xe074000) = 0xe074000 +12429 brk(0xe095000) = 0xe095000 +12429 brk(0xe0b6000) = 0xe0b6000 +12429 brk(0xe0d7000) = 0xe0d7000 +12429 brk(0xe0f8000) = 0xe0f8000 +12429 brk(0xe119000) = 0xe119000 +12429 brk(0xe13a000) = 0xe13a000 +12429 brk(0xe15b000) = 0xe15b000 +12429 brk(0xe17c000) = 0xe17c000 +12429 brk(0xe19d000) = 0xe19d000 +12429 brk(0xe1be000) = 0xe1be000 +12429 brk(0xe1df000) = 0xe1df000 +12429 brk(0xe200000) = 0xe200000 +12429 brk(0xe221000) = 0xe221000 +12429 brk(0xe242000) = 0xe242000 +12429 brk(0xe263000) = 0xe263000 +12429 brk(0xe284000) = 0xe284000 +12429 brk(0xe2a5000) = 0xe2a5000 +12429 brk(0xe2c6000) = 0xe2c6000 +12429 brk(0xe2e7000) = 0xe2e7000 +12429 brk(0xe308000) = 0xe308000 +12429 brk(0xe329000) = 0xe329000 +12429 brk(0xe34a000) = 0xe34a000 +12429 brk(0xe36b000) = 0xe36b000 +12429 brk(0xe38c000) = 0xe38c000 +12429 brk(0xe3ad000) = 0xe3ad000 +12429 brk(0xe3ce000) = 0xe3ce000 +12429 brk(0xe3ef000) = 0xe3ef000 +12429 brk(0xe410000) = 0xe410000 +12429 brk(0xe431000) = 0xe431000 +12429 brk(0xe452000) = 0xe452000 +12429 brk(0xe473000) = 0xe473000 +12429 brk(0xe494000) = 0xe494000 +12429 brk(0xe4b5000) = 0xe4b5000 +12429 brk(0xe4d6000) = 0xe4d6000 +12429 brk(0xe4f7000) = 0xe4f7000 +12429 brk(0xe518000) = 0xe518000 +12429 brk(0xe539000) = 0xe539000 +12429 brk(0xe55a000) = 0xe55a000 +12429 brk(0xe57b000) = 0xe57b000 +12429 brk(0xe59c000) = 0xe59c000 +12429 brk(0xe5bd000) = 0xe5bd000 +12429 brk(0xe5de000) = 0xe5de000 +12429 brk(0xe5ff000) = 0xe5ff000 +12429 brk(0xe620000) = 0xe620000 +12429 brk(0xe641000) = 0xe641000 +12429 brk(0xe662000) = 0xe662000 +12429 brk(0xe683000) = 0xe683000 +12429 brk(0xe6a4000) = 0xe6a4000 +12429 brk(0xe6c5000) = 0xe6c5000 +12429 brk(0xe6e6000) = 0xe6e6000 +12429 brk(0xe707000) = 0xe707000 +12429 brk(0xe72b000) = 0xe72b000 +12429 brk(0xe74c000) = 0xe74c000 +12429 brk(0xe76d000) = 0xe76d000 +12429 brk(0xe78e000) = 0xe78e000 +12429 brk(0xe7af000) = 0xe7af000 +12429 brk(0xe7d0000) = 0xe7d0000 +12429 brk(0xe7f1000) = 0xe7f1000 +12429 brk(0xe812000) = 0xe812000 +12429 brk(0xe833000) = 0xe833000 +12429 brk(0xe854000) = 0xe854000 +12429 brk(0xe875000) = 0xe875000 +12429 brk(0xe89b000) = 0xe89b000 +12429 brk(0xe8bc000) = 0xe8bc000 +12429 brk(0xe8dd000) = 0xe8dd000 +12429 brk(0xe8fe000) = 0xe8fe000 +12429 brk(0xe91f000) = 0xe91f000 +12429 brk(0xe940000) = 0xe940000 +12429 brk(0xe961000) = 0xe961000 +12429 brk(0xe982000) = 0xe982000 +12429 brk(0xe9a3000) = 0xe9a3000 +12429 brk(0xe9c4000) = 0xe9c4000 +12429 brk(0xe9e5000) = 0xe9e5000 +12429 brk(0xea06000) = 0xea06000 +12429 brk(0xea27000) = 0xea27000 +12429 brk(0xea48000) = 0xea48000 +12429 brk(0xea69000) = 0xea69000 +12429 brk(0xea8a000) = 0xea8a000 +12429 brk(0xeaab000) = 0xeaab000 +12429 brk(0xeacc000) = 0xeacc000 +12429 brk(0xeaed000) = 0xeaed000 +12429 brk(0xeb0e000) = 0xeb0e000 +12429 brk(0xeb2f000) = 0xeb2f000 +12429 brk(0xeb50000) = 0xeb50000 +12429 brk(0xeb7b000) = 0xeb7b000 +12429 brk(0xeb9c000) = 0xeb9c000 +12429 brk(0xebbd000) = 0xebbd000 +12429 brk(0xebde000) = 0xebde000 +12429 brk(0xebff000) = 0xebff000 +12429 brk(0xec20000) = 0xec20000 +12429 brk(0xec41000) = 0xec41000 +12429 brk(0xec62000) = 0xec62000 +12429 brk(0xec83000) = 0xec83000 +12429 brk(0xeca4000) = 0xeca4000 +12429 brk(0xecc5000) = 0xecc5000 +12429 brk(0xece6000) = 0xece6000 +12429 brk(0xed07000) = 0xed07000 +12429 brk(0xed28000) = 0xed28000 +12429 brk(0xed49000) = 0xed49000 +12429 brk(0xed6a000) = 0xed6a000 +12429 brk(0xed8b000) = 0xed8b000 +12429 brk(0xedac000) = 0xedac000 +12429 brk(0xedcd000) = 0xedcd000 +12429 brk(0xedee000) = 0xedee000 +12429 brk(0xee0f000) = 0xee0f000 +12429 brk(0xee30000) = 0xee30000 +12429 brk(0xee51000) = 0xee51000 +12429 brk(0xee72000) = 0xee72000 +12429 brk(0xee93000) = 0xee93000 +12429 brk(0xeeb4000) = 0xeeb4000 +12429 brk(0xeed5000) = 0xeed5000 +12429 brk(0xeef6000) = 0xeef6000 +12429 brk(0xef17000) = 0xef17000 +12429 brk(0xef38000) = 0xef38000 +12429 brk(0xef59000) = 0xef59000 +12429 brk(0xef7a000) = 0xef7a000 +12429 brk(0xef9b000) = 0xef9b000 +12429 brk(0xefbc000) = 0xefbc000 +12429 brk(0xefdd000) = 0xefdd000 +12429 brk(0xeffe000) = 0xeffe000 +12429 brk(0xf01f000) = 0xf01f000 +12429 brk(0xf040000) = 0xf040000 +12429 brk(0xf061000) = 0xf061000 +12429 brk(0xf082000) = 0xf082000 +12429 brk(0xf0a3000) = 0xf0a3000 +12429 brk(0xf0c4000) = 0xf0c4000 +12429 brk(0xf0e5000) = 0xf0e5000 +12429 brk(0xf106000) = 0xf106000 +12429 brk(0xf13b000) = 0xf13b000 +12429 brk(0xf15c000) = 0xf15c000 +12429 brk(0xf17d000) = 0xf17d000 +12429 brk(0xf19e000) = 0xf19e000 +12429 brk(0xf1bf000) = 0xf1bf000 +12429 brk(0xf1e0000) = 0xf1e0000 +12429 brk(0xf201000) = 0xf201000 +12429 brk(0xf222000) = 0xf222000 +12429 brk(0xf243000) = 0xf243000 +12429 brk(0xf264000) = 0xf264000 +12429 brk(0xf285000) = 0xf285000 +12429 brk(0xf2a6000) = 0xf2a6000 +12429 brk(0xf2c7000) = 0xf2c7000 +12429 brk(0xf2e8000) = 0xf2e8000 +12429 brk(0xf309000) = 0xf309000 +12429 brk(0xf32a000) = 0xf32a000 +12429 brk(0xf34b000) = 0xf34b000 +12429 brk(0xf36c000) = 0xf36c000 +12429 brk(0xf38d000) = 0xf38d000 +12429 brk(0xf3ae000) = 0xf3ae000 +12429 brk(0xf3cf000) = 0xf3cf000 +12429 brk(0xf3f0000) = 0xf3f0000 +12429 brk(0xf411000) = 0xf411000 +12429 brk(0xf432000) = 0xf432000 +12429 brk(0xf453000) = 0xf453000 +12429 brk(0xf474000) = 0xf474000 +12429 brk(0xf495000) = 0xf495000 +12429 brk(0xf4b6000) = 0xf4b6000 +12429 brk(0xf4d7000) = 0xf4d7000 +12429 brk(0xf4f8000) = 0xf4f8000 +12429 brk(0xf519000) = 0xf519000 +12429 brk(0xf53a000) = 0xf53a000 +12429 brk(0xf55b000) = 0xf55b000 +12429 brk(0xf57c000) = 0xf57c000 +12429 brk(0xf59d000) = 0xf59d000 +12429 brk(0xf5be000) = 0xf5be000 +12429 brk(0xf5df000) = 0xf5df000 +12429 brk(0xf600000) = 0xf600000 +12429 brk(0xf621000) = 0xf621000 +12429 brk(0xf642000) = 0xf642000 +12429 brk(0xf663000) = 0xf663000 +12429 brk(0xf684000) = 0xf684000 +12429 brk(0xf6a5000) = 0xf6a5000 +12429 brk(0xf6c6000) = 0xf6c6000 +12429 brk(0xf6e7000) = 0xf6e7000 +12429 brk(0xf708000) = 0xf708000 +12429 brk(0xf729000) = 0xf729000 +12429 brk(0xf74a000) = 0xf74a000 +12429 brk(0xf76b000) = 0xf76b000 +12429 brk(0xf78c000) = 0xf78c000 +12429 brk(0xf7ad000) = 0xf7ad000 +12429 brk(0xf7ce000) = 0xf7ce000 +12429 brk(0xf7ef000) = 0xf7ef000 +12429 brk(0xf810000) = 0xf810000 +12429 brk(0xf831000) = 0xf831000 +12429 brk(0xf852000) = 0xf852000 +12429 brk(0xf873000) = 0xf873000 +12429 brk(0xf894000) = 0xf894000 +12429 brk(0xf8b5000) = 0xf8b5000 +12429 brk(0xf8d6000) = 0xf8d6000 +12429 brk(0xf8f7000) = 0xf8f7000 +12429 brk(0xf918000) = 0xf918000 +12429 brk(0xf939000) = 0xf939000 +12429 brk(0xf95a000) = 0xf95a000 +12429 brk(0xf97b000) = 0xf97b000 +12429 brk(0xf99c000) = 0xf99c000 +12429 brk(0xf9bd000) = 0xf9bd000 +12429 brk(0xf9de000) = 0xf9de000 +12429 brk(0xf9ff000) = 0xf9ff000 +12429 brk(0xfa20000) = 0xfa20000 +12429 brk(0xfa41000) = 0xfa41000 +12429 brk(0xfa62000) = 0xfa62000 +12429 brk(0xfa83000) = 0xfa83000 +12429 brk(0xfaa4000) = 0xfaa4000 +12429 brk(0xfac5000) = 0xfac5000 +12429 brk(0xfae6000) = 0xfae6000 +12429 brk(0xfb07000) = 0xfb07000 +12429 brk(0xfb28000) = 0xfb28000 +12429 brk(0xfb49000) = 0xfb49000 +12429 brk(0xfb6a000) = 0xfb6a000 +12429 brk(0xfb8b000) = 0xfb8b000 +12429 brk(0xfbac000) = 0xfbac000 +12429 brk(0xfbcd000) = 0xfbcd000 +12429 brk(0xfbee000) = 0xfbee000 +12429 brk(0xfc0f000) = 0xfc0f000 +12429 brk(0xfc30000) = 0xfc30000 +12429 brk(0xfc51000) = 0xfc51000 +12429 brk(0xfc72000) = 0xfc72000 +12429 brk(0xfe84000) = 0xfe84000 +12429 brk(0x10084000) = 0x10084000 +12429 brk(0x10284000) = 0x10284000 +12429 brk(0x102a5000) = 0x102a5000 +12429 brk(0x102c6000) = 0x102c6000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x10952000) = 0x10952000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 brk(0x10973000) = 0x10973000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x11003000) = 0x11003000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x11696000) = 0x11696000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 mmap(NULL, 67112960, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4b51efa000 +12429 brk(0x116b7000) = 0x116b7000 +12429 brk(0x116d8000) = 0x116d8000 +12429 brk(0x116f9000) = 0x116f9000 +12429 brk(0x1171a000) = 0x1171a000 +12429 brk(0x1173b000) = 0x1173b000 +12429 brk(0x1175c000) = 0x1175c000 +12429 brk(0x1177d000) = 0x1177d000 +12429 brk(0x1179e000) = 0x1179e000 +12429 brk(0x117bf000) = 0x117bf000 +12429 brk(0x117e0000) = 0x117e0000 +12429 brk(0x11801000) = 0x11801000 +12429 brk(0x11822000) = 0x11822000 +12429 brk(0x11843000) = 0x11843000 +12429 brk(0x11864000) = 0x11864000 +12429 brk(0x11885000) = 0x11885000 +12429 brk(0x118a6000) = 0x118a6000 +12429 brk(0x118c7000) = 0x118c7000 +12429 brk(0x118e8000) = 0x118e8000 +12429 brk(0x11909000) = 0x11909000 +12429 brk(0x1192a000) = 0x1192a000 +12429 brk(0x1194b000) = 0x1194b000 +12429 brk(0x1196c000) = 0x1196c000 +12429 brk(0x1198d000) = 0x1198d000 +12429 brk(0x119ae000) = 0x119ae000 +12429 brk(0x119cf000) = 0x119cf000 +12429 brk(0x119f0000) = 0x119f0000 +12429 brk(0x11a11000) = 0x11a11000 +12429 brk(0x11a32000) = 0x11a32000 +12429 brk(0x11a53000) = 0x11a53000 +12429 brk(0x11a74000) = 0x11a74000 +12429 brk(0x11a95000) = 0x11a95000 +12429 brk(0x11ab6000) = 0x11ab6000 +12429 brk(0x11ad7000) = 0x11ad7000 +12429 brk(0x11af8000) = 0x11af8000 +12429 brk(0x11b19000) = 0x11b19000 +12429 brk(0x11b3a000) = 0x11b3a000 +12429 brk(0x11b5b000) = 0x11b5b000 +12429 brk(0x11b7c000) = 0x11b7c000 +12429 brk(0x11b9d000) = 0x11b9d000 +12429 brk(0x11bbe000) = 0x11bbe000 +12429 brk(0x11bdf000) = 0x11bdf000 +12429 brk(0x11c00000) = 0x11c00000 +12429 brk(0x11c21000) = 0x11c21000 +12429 brk(0x11c42000) = 0x11c42000 +12429 brk(0x11c63000) = 0x11c63000 +12429 brk(0x11c84000) = 0x11c84000 +12429 brk(0x11ca5000) = 0x11ca5000 +12429 brk(0x11cd6000) = 0x11cd6000 +12429 brk(0x11cf7000) = 0x11cf7000 +12429 brk(0x11d18000) = 0x11d18000 +12429 brk(0x11d39000) = 0x11d39000 +12429 brk(0x11d5a000) = 0x11d5a000 +12429 brk(0x11d7b000) = 0x11d7b000 +12429 brk(0x11d9c000) = 0x11d9c000 +12429 brk(0x11dbd000) = 0x11dbd000 +12429 brk(0x11ddf000) = 0x11ddf000 +12429 brk(0x11e00000) = 0x11e00000 +12429 brk(0x11e21000) = 0x11e21000 +12429 brk(0x11e42000) = 0x11e42000 +12429 brk(0x11e63000) = 0x11e63000 +12429 brk(0x11e84000) = 0x11e84000 +12429 brk(0x11ea5000) = 0x11ea5000 +12429 brk(0x11ec6000) = 0x11ec6000 +12429 brk(0x11ee7000) = 0x11ee7000 +12429 brk(0x11f08000) = 0x11f08000 +12429 brk(0x11f29000) = 0x11f29000 +12429 brk(0x11f4a000) = 0x11f4a000 +12429 brk(0x11f6b000) = 0x11f6b000 +12429 brk(0x11f8c000) = 0x11f8c000 +12429 brk(0x11fad000) = 0x11fad000 +12429 brk(0x11fce000) = 0x11fce000 +12429 brk(0x11fef000) = 0x11fef000 +12429 brk(0x12010000) = 0x12010000 +12429 brk(0x12031000) = 0x12031000 +12429 brk(0x12052000) = 0x12052000 +12429 brk(0x12073000) = 0x12073000 +12429 brk(0x12094000) = 0x12094000 +12429 brk(0x120b5000) = 0x120b5000 +12429 brk(0x120d6000) = 0x120d6000 +12429 brk(0x120f7000) = 0x120f7000 +12429 brk(0x1211e000) = 0x1211e000 +12429 brk(0x1213f000) = 0x1213f000 +12429 brk(0x12160000) = 0x12160000 +12429 brk(0x12181000) = 0x12181000 +12429 brk(0x121a2000) = 0x121a2000 +12429 brk(0x121c3000) = 0x121c3000 +12429 brk(0x121e4000) = 0x121e4000 +12429 brk(0x12205000) = 0x12205000 +12429 brk(0x12226000) = 0x12226000 +12429 brk(0x12247000) = 0x12247000 +12429 brk(0x12268000) = 0x12268000 +12429 brk(0x12289000) = 0x12289000 +12429 brk(0x122aa000) = 0x122aa000 +12429 brk(0x122cb000) = 0x122cb000 +12429 brk(0x122ec000) = 0x122ec000 +12429 brk(0x1230d000) = 0x1230d000 +12429 brk(0x1232e000) = 0x1232e000 +12429 brk(0x1234f000) = 0x1234f000 +12429 brk(0x12370000) = 0x12370000 +12429 brk(0x12391000) = 0x12391000 +12429 brk(0x123b2000) = 0x123b2000 +12429 brk(0x123d3000) = 0x123d3000 +12429 brk(0x123f4000) = 0x123f4000 +12429 brk(0x12415000) = 0x12415000 +12429 brk(0x12436000) = 0x12436000 +12429 brk(0x12457000) = 0x12457000 +12429 brk(0x12478000) = 0x12478000 +12429 brk(0x12499000) = 0x12499000 +12429 brk(0x124ba000) = 0x124ba000 +12429 brk(0x124db000) = 0x124db000 +12429 brk(0x124fc000) = 0x124fc000 +12429 brk(0x1251d000) = 0x1251d000 +12429 brk(0x1253e000) = 0x1253e000 +12429 brk(0x1255f000) = 0x1255f000 +12429 brk(0x12580000) = 0x12580000 +12429 brk(0x125a1000) = 0x125a1000 +12429 brk(0x125c2000) = 0x125c2000 +12429 brk(0x125e3000) = 0x125e3000 +12429 brk(0x12604000) = 0x12604000 +12429 brk(0x12625000) = 0x12625000 +12429 brk(0x12646000) = 0x12646000 +12429 brk(0x12667000) = 0x12667000 +12429 brk(0x12688000) = 0x12688000 +12429 brk(0x126a9000) = 0x126a9000 +12429 brk(0x126ca000) = 0x126ca000 +12429 brk(0x126eb000) = 0x126eb000 +12429 brk(0x1270c000) = 0x1270c000 +12429 brk(0x1272d000) = 0x1272d000 +12429 brk(0x1274e000) = 0x1274e000 +12429 brk(0x1276f000) = 0x1276f000 +12429 brk(0x12790000) = 0x12790000 +12429 brk(0x127b1000) = 0x127b1000 +12429 brk(0x127d2000) = 0x127d2000 +12429 brk(0x127f3000) = 0x127f3000 +12429 brk(0x12814000) = 0x12814000 +12429 brk(0x12835000) = 0x12835000 +12429 brk(0x12856000) = 0x12856000 +12429 brk(0x12877000) = 0x12877000 +12429 brk(0x12898000) = 0x12898000 +12429 brk(0x128b9000) = 0x128b9000 +12429 brk(0x128da000) = 0x128da000 +12429 brk(0x128fb000) = 0x128fb000 +12429 brk(0x1291c000) = 0x1291c000 +12429 brk(0x1293d000) = 0x1293d000 +12429 brk(0x1295e000) = 0x1295e000 +12429 brk(0x1297f000) = 0x1297f000 +12429 brk(0x129a0000) = 0x129a0000 +12429 brk(0x129c1000) = 0x129c1000 +12429 brk(0x129e2000) = 0x129e2000 +12429 brk(0x12a03000) = 0x12a03000 +12429 brk(0x12a24000) = 0x12a24000 +12429 brk(0x12a45000) = 0x12a45000 +12429 brk(0x12a66000) = 0x12a66000 +12429 brk(0x12a87000) = 0x12a87000 +12429 brk(0x12aa8000) = 0x12aa8000 +12429 brk(0x12ac9000) = 0x12ac9000 +12429 brk(0x12aea000) = 0x12aea000 +12429 brk(0x12b0b000) = 0x12b0b000 +12429 brk(0x12b2c000) = 0x12b2c000 +12429 brk(0x12b4d000) = 0x12b4d000 +12429 brk(0x12b6e000) = 0x12b6e000 +12429 brk(0x12b8f000) = 0x12b8f000 +12429 brk(0x12bb0000) = 0x12bb0000 +12429 brk(0x12bd1000) = 0x12bd1000 +12429 brk(0x12bf5000) = 0x12bf5000 +12429 brk(0x12c16000) = 0x12c16000 +12429 brk(0x12c37000) = 0x12c37000 +12429 brk(0x12c58000) = 0x12c58000 +12429 brk(0x12c79000) = 0x12c79000 +12429 brk(0x12c9a000) = 0x12c9a000 +12429 brk(0x12cbb000) = 0x12cbb000 +12429 brk(0x12cdc000) = 0x12cdc000 +12429 brk(0x12cfd000) = 0x12cfd000 +12429 brk(0x12d1e000) = 0x12d1e000 +12429 brk(0x12d3f000) = 0x12d3f000 +12429 brk(0x12d65000) = 0x12d65000 +12429 brk(0x12d86000) = 0x12d86000 +12429 brk(0x12da7000) = 0x12da7000 +12429 brk(0x12dc8000) = 0x12dc8000 +12429 brk(0x12de9000) = 0x12de9000 +12429 brk(0x12e0a000) = 0x12e0a000 +12429 brk(0x12e2b000) = 0x12e2b000 +12429 brk(0x12e4c000) = 0x12e4c000 +12429 brk(0x12e6d000) = 0x12e6d000 +12429 brk(0x12e8e000) = 0x12e8e000 +12429 brk(0x12eaf000) = 0x12eaf000 +12429 brk(0x12ed0000) = 0x12ed0000 +12429 brk(0x12ef1000) = 0x12ef1000 +12429 brk(0x12f12000) = 0x12f12000 +12429 brk(0x12f33000) = 0x12f33000 +12429 brk(0x12f54000) = 0x12f54000 +12429 brk(0x12f75000) = 0x12f75000 +12429 brk(0x12f96000) = 0x12f96000 +12429 brk(0x12fb7000) = 0x12fb7000 +12429 brk(0x12fd8000) = 0x12fd8000 +12429 brk(0x12ff9000) = 0x12ff9000 +12429 brk(0x1301a000) = 0x1301a000 +12429 brk(0x13045000) = 0x13045000 +12429 brk(0x13066000) = 0x13066000 +12429 brk(0x13087000) = 0x13087000 +12429 brk(0x130a8000) = 0x130a8000 +12429 brk(0x130c9000) = 0x130c9000 +12429 brk(0x130ea000) = 0x130ea000 +12429 brk(0x1310b000) = 0x1310b000 +12429 brk(0x1312c000) = 0x1312c000 +12429 brk(0x1314d000) = 0x1314d000 +12429 brk(0x1316e000) = 0x1316e000 +12429 brk(0x1318f000) = 0x1318f000 +12429 brk(0x131b0000) = 0x131b0000 +12429 brk(0x131d1000) = 0x131d1000 +12429 brk(0x131f2000) = 0x131f2000 +12429 brk(0x13213000) = 0x13213000 +12429 brk(0x13234000) = 0x13234000 +12429 brk(0x13255000) = 0x13255000 +12429 brk(0x13276000) = 0x13276000 +12429 brk(0x13297000) = 0x13297000 +12429 brk(0x132b8000) = 0x132b8000 +12429 brk(0x132d9000) = 0x132d9000 +12429 brk(0x132fa000) = 0x132fa000 +12429 brk(0x1331b000) = 0x1331b000 +12429 brk(0x1333c000) = 0x1333c000 +12429 brk(0x1335d000) = 0x1335d000 +12429 brk(0x1337e000) = 0x1337e000 +12429 brk(0x1339f000) = 0x1339f000 +12429 brk(0x133c0000) = 0x133c0000 +12429 brk(0x133e1000) = 0x133e1000 +12429 brk(0x13402000) = 0x13402000 +12429 brk(0x13423000) = 0x13423000 +12429 brk(0x13444000) = 0x13444000 +12429 brk(0x13465000) = 0x13465000 +12429 brk(0x13486000) = 0x13486000 +12429 brk(0x134a7000) = 0x134a7000 +12429 brk(0x134c8000) = 0x134c8000 +12429 brk(0x134e9000) = 0x134e9000 +12429 brk(0x1350a000) = 0x1350a000 +12429 brk(0x1352b000) = 0x1352b000 +12429 brk(0x1354c000) = 0x1354c000 +12429 brk(0x1356d000) = 0x1356d000 +12429 brk(0x1358e000) = 0x1358e000 +12429 brk(0x135af000) = 0x135af000 +12429 brk(0x135d0000) = 0x135d0000 +12429 brk(0x13605000) = 0x13605000 +12429 brk(0x13626000) = 0x13626000 +12429 brk(0x13647000) = 0x13647000 +12429 brk(0x13668000) = 0x13668000 +12429 brk(0x13689000) = 0x13689000 +12429 brk(0x136aa000) = 0x136aa000 +12429 brk(0x136cb000) = 0x136cb000 +12429 brk(0x136ec000) = 0x136ec000 +12429 brk(0x1370d000) = 0x1370d000 +12429 brk(0x1372e000) = 0x1372e000 +12429 brk(0x1374f000) = 0x1374f000 +12429 brk(0x13770000) = 0x13770000 +12429 brk(0x13791000) = 0x13791000 +12429 brk(0x137b2000) = 0x137b2000 +12429 brk(0x137d3000) = 0x137d3000 +12429 brk(0x137f4000) = 0x137f4000 +12429 brk(0x13815000) = 0x13815000 +12429 brk(0x13836000) = 0x13836000 +12429 brk(0x13857000) = 0x13857000 +12429 brk(0x13878000) = 0x13878000 +12429 brk(0x13899000) = 0x13899000 +12429 brk(0x138ba000) = 0x138ba000 +12429 brk(0x138db000) = 0x138db000 +12429 brk(0x138fc000) = 0x138fc000 +12429 brk(0x1391d000) = 0x1391d000 +12429 brk(0x1393e000) = 0x1393e000 +12429 brk(0x1395f000) = 0x1395f000 +12429 brk(0x13980000) = 0x13980000 +12429 brk(0x139a1000) = 0x139a1000 +12429 brk(0x139c2000) = 0x139c2000 +12429 brk(0x139e3000) = 0x139e3000 +12429 brk(0x13a04000) = 0x13a04000 +12429 brk(0x13a25000) = 0x13a25000 +12429 brk(0x13a46000) = 0x13a46000 +12429 brk(0x13a67000) = 0x13a67000 +12429 brk(0x13a88000) = 0x13a88000 +12429 brk(0x13aa9000) = 0x13aa9000 +12429 brk(0x13aca000) = 0x13aca000 +12429 brk(0x13aeb000) = 0x13aeb000 +12429 brk(0x13b0c000) = 0x13b0c000 +12429 brk(0x13b2d000) = 0x13b2d000 +12429 brk(0x13b4e000) = 0x13b4e000 +12429 brk(0x13b6f000) = 0x13b6f000 +12429 brk(0x13b90000) = 0x13b90000 +12429 brk(0x13bb1000) = 0x13bb1000 +12429 brk(0x13bd2000) = 0x13bd2000 +12429 brk(0x13bf3000) = 0x13bf3000 +12429 brk(0x13c14000) = 0x13c14000 +12429 brk(0x13c35000) = 0x13c35000 +12429 brk(0x13c56000) = 0x13c56000 +12429 brk(0x13c77000) = 0x13c77000 +12429 brk(0x13c98000) = 0x13c98000 +12429 brk(0x13cb9000) = 0x13cb9000 +12429 brk(0x13cda000) = 0x13cda000 +12429 brk(0x13cfb000) = 0x13cfb000 +12429 brk(0x13d1c000) = 0x13d1c000 +12429 brk(0x13d3d000) = 0x13d3d000 +12429 brk(0x13d5e000) = 0x13d5e000 +12429 brk(0x13d7f000) = 0x13d7f000 +12429 brk(0x13da0000) = 0x13da0000 +12429 brk(0x13dc1000) = 0x13dc1000 +12429 brk(0x13de2000) = 0x13de2000 +12429 brk(0x13e03000) = 0x13e03000 +12429 brk(0x13e24000) = 0x13e24000 +12429 brk(0x13e45000) = 0x13e45000 +12429 brk(0x13e66000) = 0x13e66000 +12429 brk(0x13e87000) = 0x13e87000 +12429 brk(0x13ea8000) = 0x13ea8000 +12429 brk(0x13ec9000) = 0x13ec9000 +12429 brk(0x13eea000) = 0x13eea000 +12429 brk(0x13f0b000) = 0x13f0b000 +12429 brk(0x13f2c000) = 0x13f2c000 +12429 brk(0x13f4d000) = 0x13f4d000 +12429 brk(0x13f6e000) = 0x13f6e000 +12429 brk(0x13f8f000) = 0x13f8f000 +12429 brk(0x13fb0000) = 0x13fb0000 +12429 brk(0x13fd1000) = 0x13fd1000 +12429 brk(0x13ff2000) = 0x13ff2000 +12429 brk(0x14013000) = 0x14013000 +12429 brk(0x14034000) = 0x14034000 +12429 brk(0x14055000) = 0x14055000 +12429 brk(0x14076000) = 0x14076000 +12429 brk(0x14097000) = 0x14097000 +12429 brk(0x140b8000) = 0x140b8000 +12429 brk(0x140d9000) = 0x140d9000 +12429 brk(0x140fa000) = 0x140fa000 +12429 brk(0x1411b000) = 0x1411b000 +12429 brk(0x1413c000) = 0x1413c000 +12429 brk(0x1434d000) = 0x1434d000 +12429 brk(0x1454d000) = 0x1454d000 +12429 brk(0x1474d000) = 0x1474d000 +12429 brk(0x1476e000) = 0x1476e000 +12429 brk(0x1478f000) = 0x1478f000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x14e1c000) = 0x14e1c000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 brk(0x14e3d000) = 0x14e3d000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x154cc000) = 0x154cc000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 brk(0x15b5f000) = 0x15b5f000 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libiss_k1c.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libiss_k1c.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libmppa_multiloader.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libmppa_multiloader.so", F_OK) = 0 +12429 readlink("/proc/self/exe", "/opt/Kalray/usr/local/k1rdtools/"..., 1024) = 46 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 1024) = 54 +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib32/libstd_scalls.so", F_OK) = -1 ENOENT (No such file or directory) +12429 access("/opt/Kalray/usr/local/k1rdtools/bin/../lib64/libstd_scalls.so", F_OK) = 0 +12429 mmap(NULL, 67112960, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4b4def9000 +12429 brk(0x15bdf000) = 0x15bdf000 +12429 brk(0x15c00000) = 0x15c00000 +12429 brk(0x15c21000) = 0x15c21000 +12429 brk(0x15c42000) = 0x15c42000 +12429 brk(0x15c63000) = 0x15c63000 +12429 brk(0x15c84000) = 0x15c84000 +12429 brk(0x15ca5000) = 0x15ca5000 +12429 brk(0x15cc6000) = 0x15cc6000 +12429 brk(0x15ce7000) = 0x15ce7000 +12429 brk(0x15d08000) = 0x15d08000 +12429 brk(0x15d29000) = 0x15d29000 +12429 brk(0x15d4a000) = 0x15d4a000 +12429 brk(0x15d6b000) = 0x15d6b000 +12429 brk(0x15d8c000) = 0x15d8c000 +12429 brk(0x15dad000) = 0x15dad000 +12429 brk(0x15dce000) = 0x15dce000 +12429 brk(0x15def000) = 0x15def000 +12429 brk(0x15e10000) = 0x15e10000 +12429 brk(0x15e31000) = 0x15e31000 +12429 brk(0x15e52000) = 0x15e52000 +12429 brk(0x15e73000) = 0x15e73000 +12429 brk(0x15e94000) = 0x15e94000 +12429 brk(0x15eb5000) = 0x15eb5000 +12429 brk(0x15ed6000) = 0x15ed6000 +12429 brk(0x15ef7000) = 0x15ef7000 +12429 brk(0x15f18000) = 0x15f18000 +12429 brk(0x15f39000) = 0x15f39000 +12429 brk(0x15f5a000) = 0x15f5a000 +12429 brk(0x15f7b000) = 0x15f7b000 +12429 brk(0x15f9c000) = 0x15f9c000 +12429 brk(0x15fbd000) = 0x15fbd000 +12429 brk(0x15fde000) = 0x15fde000 +12429 brk(0x15fff000) = 0x15fff000 +12429 brk(0x16020000) = 0x16020000 +12429 brk(0x16041000) = 0x16041000 +12429 brk(0x16062000) = 0x16062000 +12429 brk(0x16083000) = 0x16083000 +12429 brk(0x160a4000) = 0x160a4000 +12429 brk(0x160c5000) = 0x160c5000 +12429 brk(0x160e6000) = 0x160e6000 +12429 brk(0x16107000) = 0x16107000 +12429 brk(0x16128000) = 0x16128000 +12429 brk(0x16149000) = 0x16149000 +12429 brk(0x1616a000) = 0x1616a000 +12429 brk(0x1618b000) = 0x1618b000 +12429 brk(0x161ac000) = 0x161ac000 +12429 brk(0x161cd000) = 0x161cd000 +12429 brk(0x161f1000) = 0x161f1000 +12429 brk(0x16212000) = 0x16212000 +12429 brk(0x16233000) = 0x16233000 +12429 brk(0x16254000) = 0x16254000 +12429 brk(0x16275000) = 0x16275000 +12429 brk(0x16296000) = 0x16296000 +12429 brk(0x162b7000) = 0x162b7000 +12429 brk(0x162d8000) = 0x162d8000 +12429 brk(0x162f9000) = 0x162f9000 +12429 brk(0x1631a000) = 0x1631a000 +12429 brk(0x1633b000) = 0x1633b000 +12429 brk(0x1635c000) = 0x1635c000 +12429 brk(0x1637d000) = 0x1637d000 +12429 brk(0x1639e000) = 0x1639e000 +12429 brk(0x163bf000) = 0x163bf000 +12429 brk(0x163e0000) = 0x163e0000 +12429 brk(0x16401000) = 0x16401000 +12429 brk(0x16422000) = 0x16422000 +12429 brk(0x16443000) = 0x16443000 +12429 brk(0x16464000) = 0x16464000 +12429 brk(0x1648b000) = 0x1648b000 +12429 brk(0x164ac000) = 0x164ac000 +12429 brk(0x164cd000) = 0x164cd000 +12429 brk(0x164ee000) = 0x164ee000 +12429 brk(0x1650f000) = 0x1650f000 +12429 brk(0x16530000) = 0x16530000 +12429 brk(0x16551000) = 0x16551000 +12429 brk(0x16572000) = 0x16572000 +12429 brk(0x16593000) = 0x16593000 +12429 brk(0x165b4000) = 0x165b4000 +12429 brk(0x165d5000) = 0x165d5000 +12429 brk(0x165f6000) = 0x165f6000 +12429 brk(0x16617000) = 0x16617000 +12429 brk(0x16638000) = 0x16638000 +12429 brk(0x16659000) = 0x16659000 +12429 brk(0x1667a000) = 0x1667a000 +12429 brk(0x1669b000) = 0x1669b000 +12429 brk(0x166bc000) = 0x166bc000 +12429 brk(0x166dd000) = 0x166dd000 +12429 brk(0x166fe000) = 0x166fe000 +12429 brk(0x1671f000) = 0x1671f000 +12429 brk(0x16740000) = 0x16740000 +12429 brk(0x16761000) = 0x16761000 +12429 brk(0x16782000) = 0x16782000 +12429 brk(0x167a3000) = 0x167a3000 +12429 brk(0x167c4000) = 0x167c4000 +12429 brk(0x167e5000) = 0x167e5000 +12429 brk(0x16806000) = 0x16806000 +12429 brk(0x16827000) = 0x16827000 +12429 brk(0x16848000) = 0x16848000 +12429 brk(0x16869000) = 0x16869000 +12429 brk(0x1688a000) = 0x1688a000 +12429 brk(0x168ab000) = 0x168ab000 +12429 brk(0x168cc000) = 0x168cc000 +12429 brk(0x168ed000) = 0x168ed000 +12429 brk(0x1690e000) = 0x1690e000 +12429 brk(0x1692f000) = 0x1692f000 +12429 brk(0x1695f000) = 0x1695f000 +12429 brk(0x16980000) = 0x16980000 +12429 brk(0x169a1000) = 0x169a1000 +12429 brk(0x169c2000) = 0x169c2000 +12429 brk(0x169e3000) = 0x169e3000 +12429 brk(0x16a04000) = 0x16a04000 +12429 brk(0x16a25000) = 0x16a25000 +12429 brk(0x16a46000) = 0x16a46000 +12429 brk(0x16a67000) = 0x16a67000 +12429 brk(0x16a88000) = 0x16a88000 +12429 brk(0x16aa9000) = 0x16aa9000 +12429 brk(0x16aca000) = 0x16aca000 +12429 brk(0x16aeb000) = 0x16aeb000 +12429 brk(0x16b0c000) = 0x16b0c000 +12429 brk(0x16b2d000) = 0x16b2d000 +12429 brk(0x16b4e000) = 0x16b4e000 +12429 brk(0x16b6f000) = 0x16b6f000 +12429 brk(0x16b90000) = 0x16b90000 +12429 brk(0x16bb1000) = 0x16bb1000 +12429 brk(0x16bd2000) = 0x16bd2000 +12429 brk(0x16bf3000) = 0x16bf3000 +12429 brk(0x16c14000) = 0x16c14000 +12429 brk(0x16c35000) = 0x16c35000 +12429 brk(0x16c56000) = 0x16c56000 +12429 brk(0x16c77000) = 0x16c77000 +12429 brk(0x16c98000) = 0x16c98000 +12429 brk(0x16cb9000) = 0x16cb9000 +12429 brk(0x16cda000) = 0x16cda000 +12429 brk(0x16cfb000) = 0x16cfb000 +12429 brk(0x16d1c000) = 0x16d1c000 +12429 brk(0x16d3d000) = 0x16d3d000 +12429 brk(0x16d5e000) = 0x16d5e000 +12429 brk(0x16d7f000) = 0x16d7f000 +12429 brk(0x16da0000) = 0x16da0000 +12429 brk(0x16dc1000) = 0x16dc1000 +12429 brk(0x16de2000) = 0x16de2000 +12429 brk(0x16e03000) = 0x16e03000 +12429 brk(0x16e24000) = 0x16e24000 +12429 brk(0x16e45000) = 0x16e45000 +12429 brk(0x16e66000) = 0x16e66000 +12429 brk(0x16e87000) = 0x16e87000 +12429 brk(0x16ea8000) = 0x16ea8000 +12429 brk(0x16ec9000) = 0x16ec9000 +12429 brk(0x16eea000) = 0x16eea000 +12429 brk(0x16f0b000) = 0x16f0b000 +12429 brk(0x16f2c000) = 0x16f2c000 +12429 brk(0x16f4d000) = 0x16f4d000 +12429 brk(0x16f6e000) = 0x16f6e000 +12429 brk(0x16f8f000) = 0x16f8f000 +12429 brk(0x16fb0000) = 0x16fb0000 +12429 brk(0x16fd1000) = 0x16fd1000 +12429 brk(0x16ff2000) = 0x16ff2000 +12429 brk(0x17013000) = 0x17013000 +12429 brk(0x17034000) = 0x17034000 +12429 brk(0x17055000) = 0x17055000 +12429 brk(0x17076000) = 0x17076000 +12429 brk(0x17097000) = 0x17097000 +12429 brk(0x170b8000) = 0x170b8000 +12429 brk(0x170d9000) = 0x170d9000 +12429 brk(0x170fa000) = 0x170fa000 +12429 brk(0x1711b000) = 0x1711b000 +12429 brk(0x1713c000) = 0x1713c000 +12429 brk(0x1715d000) = 0x1715d000 +12429 brk(0x1717e000) = 0x1717e000 +12429 brk(0x1719f000) = 0x1719f000 +12429 brk(0x171c0000) = 0x171c0000 +12429 brk(0x171e1000) = 0x171e1000 +12429 brk(0x17202000) = 0x17202000 +12429 brk(0x17223000) = 0x17223000 +12429 brk(0x17244000) = 0x17244000 +12429 brk(0x17265000) = 0x17265000 +12429 brk(0x17286000) = 0x17286000 +12429 brk(0x172a7000) = 0x172a7000 +12429 brk(0x172c8000) = 0x172c8000 +12429 brk(0x172e9000) = 0x172e9000 +12429 brk(0x1730a000) = 0x1730a000 +12429 brk(0x1732b000) = 0x1732b000 +12429 brk(0x1734c000) = 0x1734c000 +12429 brk(0x1736d000) = 0x1736d000 +12429 brk(0x1738e000) = 0x1738e000 +12429 brk(0x173af000) = 0x173af000 +12429 brk(0x173d0000) = 0x173d0000 +12429 brk(0x173f1000) = 0x173f1000 +12429 brk(0x17412000) = 0x17412000 +12429 brk(0x17433000) = 0x17433000 +12429 brk(0x17454000) = 0x17454000 +12429 brk(0x17475000) = 0x17475000 +12429 brk(0x17496000) = 0x17496000 +12429 brk(0x174b7000) = 0x174b7000 +12429 brk(0x174d8000) = 0x174d8000 +12429 brk(0x174f9000) = 0x174f9000 +12429 brk(0x1751a000) = 0x1751a000 +12429 brk(0x1753b000) = 0x1753b000 +12429 brk(0x1755c000) = 0x1755c000 +12429 brk(0x1757d000) = 0x1757d000 +12429 brk(0x1759e000) = 0x1759e000 +12429 brk(0x175bf000) = 0x175bf000 +12429 brk(0x175e0000) = 0x175e0000 +12429 brk(0x17601000) = 0x17601000 +12429 brk(0x17622000) = 0x17622000 +12429 brk(0x17643000) = 0x17643000 +12429 brk(0x17664000) = 0x17664000 +12429 brk(0x17685000) = 0x17685000 +12429 brk(0x176a6000) = 0x176a6000 +12429 brk(0x176c7000) = 0x176c7000 +12429 brk(0x176e8000) = 0x176e8000 +12429 brk(0x17709000) = 0x17709000 +12429 brk(0x1772a000) = 0x1772a000 +12429 brk(0x1774b000) = 0x1774b000 +12429 brk(0x1776c000) = 0x1776c000 +12429 brk(0x1778d000) = 0x1778d000 +12429 brk(0x177ae000) = 0x177ae000 +12429 brk(0x177cf000) = 0x177cf000 +12429 brk(0x177f0000) = 0x177f0000 +12429 brk(0x17811000) = 0x17811000 +12429 brk(0x17832000) = 0x17832000 +12429 brk(0x17853000) = 0x17853000 +12429 brk(0x17874000) = 0x17874000 +12429 brk(0x17895000) = 0x17895000 +12429 brk(0x178b6000) = 0x178b6000 +12429 brk(0x178d7000) = 0x178d7000 +12429 brk(0x178f8000) = 0x178f8000 +12429 brk(0x17919000) = 0x17919000 +12429 brk(0x1793a000) = 0x1793a000 +12429 brk(0x1795b000) = 0x1795b000 +12429 brk(0x1797c000) = 0x1797c000 +12429 brk(0x1799d000) = 0x1799d000 +12429 brk(0x179be000) = 0x179be000 +12429 brk(0x179df000) = 0x179df000 +12429 brk(0x17a00000) = 0x17a00000 +12429 brk(0x17a21000) = 0x17a21000 +12429 brk(0x17a42000) = 0x17a42000 +12429 brk(0x17a63000) = 0x17a63000 +12429 brk(0x17a84000) = 0x17a84000 +12429 brk(0x17aa5000) = 0x17aa5000 +12429 brk(0x17ac6000) = 0x17ac6000 +12429 brk(0x17ae7000) = 0x17ae7000 +12429 brk(0x17b0e000) = 0x17b0e000 +12429 brk(0x17b2f000) = 0x17b2f000 +12429 brk(0x17b50000) = 0x17b50000 +12429 brk(0x17b71000) = 0x17b71000 +12429 brk(0x17b92000) = 0x17b92000 +12429 brk(0x17bb3000) = 0x17bb3000 +12429 brk(0x17bd4000) = 0x17bd4000 +12429 brk(0x17bf5000) = 0x17bf5000 +12429 brk(0x17c16000) = 0x17c16000 +12429 brk(0x17c37000) = 0x17c37000 +12429 brk(0x17c58000) = 0x17c58000 +12429 brk(0x17c79000) = 0x17c79000 +12429 brk(0x17c9a000) = 0x17c9a000 +12429 brk(0x17cbb000) = 0x17cbb000 +12429 brk(0x17cdc000) = 0x17cdc000 +12429 brk(0x17cfd000) = 0x17cfd000 +12429 brk(0x17d1e000) = 0x17d1e000 +12429 brk(0x17d3f000) = 0x17d3f000 +12429 brk(0x17d60000) = 0x17d60000 +12429 brk(0x17d81000) = 0x17d81000 +12429 brk(0x17da2000) = 0x17da2000 +12429 brk(0x17dc3000) = 0x17dc3000 +12429 brk(0x17de4000) = 0x17de4000 +12429 brk(0x17e05000) = 0x17e05000 +12429 brk(0x17e26000) = 0x17e26000 +12429 brk(0x17e47000) = 0x17e47000 +12429 brk(0x17e68000) = 0x17e68000 +12429 brk(0x17e89000) = 0x17e89000 +12429 brk(0x17eaa000) = 0x17eaa000 +12429 brk(0x17ecb000) = 0x17ecb000 +12429 brk(0x17eec000) = 0x17eec000 +12429 brk(0x17f0d000) = 0x17f0d000 +12429 brk(0x17f2e000) = 0x17f2e000 +12429 brk(0x17f4f000) = 0x17f4f000 +12429 brk(0x17f70000) = 0x17f70000 +12429 brk(0x17f91000) = 0x17f91000 +12429 brk(0x17fb2000) = 0x17fb2000 +12429 brk(0x17fd3000) = 0x17fd3000 +12429 brk(0x17ff4000) = 0x17ff4000 +12429 brk(0x18015000) = 0x18015000 +12429 brk(0x18036000) = 0x18036000 +12429 brk(0x18057000) = 0x18057000 +12429 brk(0x18078000) = 0x18078000 +12429 brk(0x18099000) = 0x18099000 +12429 brk(0x180ba000) = 0x180ba000 +12429 brk(0x180db000) = 0x180db000 +12429 brk(0x180fc000) = 0x180fc000 +12429 brk(0x1811d000) = 0x1811d000 +12429 brk(0x1813e000) = 0x1813e000 +12429 brk(0x1815f000) = 0x1815f000 +12429 brk(0x18180000) = 0x18180000 +12429 brk(0x181a1000) = 0x181a1000 +12429 brk(0x181c2000) = 0x181c2000 +12429 brk(0x181e3000) = 0x181e3000 +12429 brk(0x18204000) = 0x18204000 +12429 brk(0x18225000) = 0x18225000 +12429 brk(0x18246000) = 0x18246000 +12429 brk(0x18267000) = 0x18267000 +12429 brk(0x18288000) = 0x18288000 +12429 brk(0x182a9000) = 0x182a9000 +12429 brk(0x182ca000) = 0x182ca000 +12429 brk(0x182eb000) = 0x182eb000 +12429 brk(0x1830c000) = 0x1830c000 +12429 brk(0x1832d000) = 0x1832d000 +12429 brk(0x1834e000) = 0x1834e000 +12429 brk(0x1836f000) = 0x1836f000 +12429 brk(0x18390000) = 0x18390000 +12429 brk(0x183b1000) = 0x183b1000 +12429 brk(0x183d2000) = 0x183d2000 +12429 brk(0x183f3000) = 0x183f3000 +12429 brk(0x18414000) = 0x18414000 +12429 brk(0x18435000) = 0x18435000 +12429 brk(0x18456000) = 0x18456000 +12429 brk(0x18477000) = 0x18477000 +12429 brk(0x18498000) = 0x18498000 +12429 brk(0x184b9000) = 0x184b9000 +12429 brk(0x184da000) = 0x184da000 +12429 brk(0x184fb000) = 0x184fb000 +12429 brk(0x1851c000) = 0x1851c000 +12429 brk(0x1853d000) = 0x1853d000 +12429 brk(0x1855e000) = 0x1855e000 +12429 brk(0x1857f000) = 0x1857f000 +12429 brk(0x185a0000) = 0x185a0000 +12429 brk(0x185c1000) = 0x185c1000 +12429 brk(0x185e2000) = 0x185e2000 +12429 brk(0x18603000) = 0x18603000 +12429 brk(0x18624000) = 0x18624000 +12429 brk(0x18645000) = 0x18645000 +12429 brk(0x18856000) = 0x18856000 +12429 brk(0x18a56000) = 0x18a56000 +12429 brk(0x18c56000) = 0x18c56000 +12429 brk(0x18c77000) = 0x18c77000 +12429 brk(0x18ca5000) = 0x18ca5000 +12429 brk(0x18d05000) = 0x18d05000 +12429 brk(0x18d85000) = 0x18d85000 +12429 brk(0x18e85000) = 0x18e85000 +12429 brk(0x19085000) = 0x19085000 +12429 brk(0x19485000) = 0x19485000 +12429 mmap(NULL, 8392704, PROT_READ|PROT_WRITE, MAP_PRIVATE|MAP_ANONYMOUS, -1, 0) = 0x7f4c68599000 +12429 munmap(0x7f4c68599000, 8392704) = 0 +12429 brk(0x1955a000) = 0x1955a000 +12429 brk(0x1975a000) = 0x1975a000 +12429 brk(0x19b5a000) = 0x19b5a000 +12429 brk(0x1a35a000) = 0x1a35a000 +12429 brk(0x1935a000) = 0x1935a000 +12429 brk(0x1937b000) = 0x1937b000 +12429 brk(0x1939c000) = 0x1939c000 +12429 brk(0x193bd000) = 0x193bd000 +12429 brk(0x193de000) = 0x193de000 +12429 mmap(NULL, 8392704, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_STACK, -1, 0) = 0x7f4c68599000 +12429 mprotect(0x7f4c6859a000, 8388608, PROT_READ|PROT_WRITE) = 0 +12429 clone(child_stack=0x7f4c68d98fb0, flags=CLONE_VM|CLONE_FS|CLONE_FILES|CLONE_SIGHAND|CLONE_THREAD|CLONE_SYSVSEM|CLONE_SETTLS|CLONE_PARENT_SETTID|CLONE_CHILD_CLEARTID, parent_tidptr=0x7f4c68d999d0, tls=0x7f4c68d99700, child_tidptr=0x7f4c68d999d0) = 12432 +12429 mmap(NULL, 8392704, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_STACK, -1, 0) = 0x7f4b4d6f8000 +12432 set_robust_list(0x7f4c68d999e0, 24 +12429 mprotect(0x7f4b4d6f9000, 8388608, PROT_READ|PROT_WRITE +12432 <... set_robust_list resumed> ) = 0 +12429 <... mprotect resumed> ) = 0 +12432 prctl(PR_SET_NAME, "node0_SIM" +12429 clone( +12432 <... prctl resumed> ) = 0 +12429 <... clone resumed> child_stack=0x7f4b4def7fb0, flags=CLONE_VM|CLONE_FS|CLONE_FILES|CLONE_SIGHAND|CLONE_THREAD|CLONE_SYSVSEM|CLONE_SETTLS|CLONE_PARENT_SETTID|CLONE_CHILD_CLEARTID, parent_tidptr=0x7f4b4def89d0, tls=0x7f4b4def8700, child_tidptr=0x7f4b4def89d0) = 12433 +12433 set_robust_list(0x7f4b4def89e0, 24 +12429 mmap(NULL, 8392704, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_STACK, -1, 0 +12433 <... set_robust_list resumed> ) = 0 +12429 <... mmap resumed> ) = 0x7f4b4cef7000 +12433 prctl(PR_SET_NAME, "node1_SIM" +12429 mprotect(0x7f4b4cef8000, 8388608, PROT_READ|PROT_WRITE +12433 <... prctl resumed> ) = 0 +12429 <... mprotect resumed> ) = 0 +12433 mmap(0x7f4c68000000, 67108864, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0 +12429 clone( +12432 mmap(0x7f4c68000000, 67108864, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0 +12433 <... mmap resumed> ) = 0x7f4b48ef7000 +12432 <... mmap resumed> ) = 0x7f4b44ef7000 +12429 <... clone resumed> child_stack=0x7f4b4d6f6fb0, flags=CLONE_VM|CLONE_FS|CLONE_FILES|CLONE_SIGHAND|CLONE_THREAD|CLONE_SYSVSEM|CLONE_SETTLS|CLONE_PARENT_SETTID|CLONE_CHILD_CLEARTID, parent_tidptr=0x7f4b4d6f79d0, tls=0x7f4b4d6f7700, child_tidptr=0x7f4b4d6f79d0) = 12434 +12433 munmap(0x7f4b48ef7000, 67108864 +12429 mmap(NULL, 8392704, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_STACK, -1, 0 +12434 set_robust_list(0x7f4b4d6f79e0, 24 +12429 <... mmap resumed> ) = 0x7f4b4c6f6000 +12434 <... set_robust_list resumed> ) = 0 +12429 mprotect(0x7f4b4c6f7000, 8388608, PROT_READ|PROT_WRITE +12434 prctl(PR_SET_NAME, "node2_SIM" +12429 <... mprotect resumed> ) = 0 +12434 <... prctl resumed> ) = 0 +12433 <... munmap resumed> ) = 0 +12429 clone( +12434 mmap(NULL, 134217728, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0 +12433 mmap(NULL, 134217728, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0 +12434 <... mmap resumed> ) = 0x7f4b3cef7000 +12435 set_robust_list(0x7f4b4cef69e0, 24 +12434 munmap(0x7f4b3cef7000, 51417088 +12435 <... set_robust_list resumed> ) = 0 +12433 <... mmap resumed> ) = 0x7f4b34ef7000 +12435 prctl(PR_SET_NAME, "node3_SIM" +12429 <... clone resumed> child_stack=0x7f4b4cef5fb0, flags=CLONE_VM|CLONE_FS|CLONE_FILES|CLONE_SIGHAND|CLONE_THREAD|CLONE_SYSVSEM|CLONE_SETTLS|CLONE_PARENT_SETTID|CLONE_CHILD_CLEARTID, parent_tidptr=0x7f4b4cef69d0, tls=0x7f4b4cef6700, child_tidptr=0x7f4b4cef69d0) = 12435 +12435 <... prctl resumed> ) = 0 +12429 mmap(NULL, 8392704, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_STACK, -1, 0 +12435 mmap(NULL, 134217728, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0 +12429 <... mmap resumed> ) = 0x7f4b4bef5000 +12435 <... mmap resumed> ) = 0x7f4b2cef7000 +12429 mprotect(0x7f4b4bef6000, 8388608, PROT_READ|PROT_WRITE +12435 munmap(0x7f4b2cef7000, 51417088 +12434 <... munmap resumed> ) = 0 +12435 <... munmap resumed> ) = 0 +12434 munmap(0x7f4b44000000, 15691776 +12429 <... mprotect resumed> ) = 0 +12435 munmap(0x7f4b34000000, 15691776 +12434 <... munmap resumed> ) = 0 +12435 <... munmap resumed> ) = 0 +12429 clone( +12435 mprotect(0x7f4b30000000, 135168, PROT_READ|PROT_WRITE +12434 mprotect(0x7f4b40000000, 135168, PROT_READ|PROT_WRITE +12435 <... mprotect resumed> ) = 0 +12429 <... clone resumed> child_stack=0x7f4b4c6f4fb0, flags=CLONE_VM|CLONE_FS|CLONE_FILES|CLONE_SIGHAND|CLONE_THREAD|CLONE_SYSVSEM|CLONE_SETTLS|CLONE_PARENT_SETTID|CLONE_CHILD_CLEARTID, parent_tidptr=0x7f4b4c6f59d0, tls=0x7f4b4c6f5700, child_tidptr=0x7f4b4c6f59d0) = 12436 +12434 <... mprotect resumed> ) = 0 +12429 futex(0x208778c, FUTEX_WAIT_PRIVATE, 0, NULL +12435 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, +12433 munmap(0x7f4b34ef7000, 51417088 +12435 <... clock_gettime resumed> {tv_sec=2, tv_nsec=548489603}) = 0 +12435 futex(0x208778c, FUTEX_WAIT_PRIVATE, 0, NULL +12436 set_robust_list(0x7f4b4c6f59e0, 24 +12434 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, +12436 <... set_robust_list resumed> ) = 0 +12434 <... clock_gettime resumed> {tv_sec=2, tv_nsec=548554155}) = 0 +12436 prctl(PR_SET_NAME, "node4_SIM" +12434 futex(0x208778c, FUTEX_WAIT_PRIVATE, 0, NULL +12436 <... prctl resumed> ) = 0 +12433 <... munmap resumed> ) = 0 +12436 mmap(NULL, 134217728, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0 +12433 munmap(0x7f4b3c000000, 15691776 +12436 <... mmap resumed> ) = 0x7f4b28000000 +12433 <... munmap resumed> ) = 0 +12436 munmap(0x7f4b2c000000, 67108864 +12433 mprotect(0x7f4b38000000, 135168, PROT_READ|PROT_WRITE +12436 <... munmap resumed> ) = 0 +12433 <... mprotect resumed> ) = 0 +12436 mprotect(0x7f4b28000000, 135168, PROT_READ|PROT_WRITE +12432 munmap(0x7f4b44ef7000, 67108864 +12436 <... mprotect resumed> ) = 0 +12433 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, +12432 <... munmap resumed> ) = 0 +12436 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, +12433 <... clock_gettime resumed> {tv_sec=2, tv_nsec=548717349}) = 0 +12436 <... clock_gettime resumed> {tv_sec=2, tv_nsec=548765067}) = 0 +12433 futex(0x208778c, FUTEX_WAIT_PRIVATE, 0, NULL +12436 futex(0x208778c, FUTEX_WAIT_PRIVATE, 0, NULL +12432 mmap(NULL, 134217728, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0) = 0x7f4b20000000 +12432 munmap(0x7f4b24000000, 67108864) = 0 +12432 mprotect(0x7f4b20000000, 135168, PROT_READ|PROT_WRITE) = 0 +12432 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, {tv_sec=2, tv_nsec=548872449}) = 0 +12432 futex(0x208778c, FUTEX_WAKE_PRIVATE, 2147483647 +12429 <... futex resumed> ) = 0 +12436 <... futex resumed> ) = 0 +12429 futex(0x208778c, FUTEX_WAIT_PRIVATE, 6, NULL +12436 futex(0x208778c, FUTEX_WAIT_PRIVATE, 6, NULL +12435 <... futex resumed> ) = 0 +12434 <... futex resumed> ) = 0 +12435 futex(0x208778c, FUTEX_WAIT_PRIVATE, 6, NULL +12434 futex(0x208778c, FUTEX_WAIT_PRIVATE, 6, NULL +12433 <... futex resumed> ) = 0 +12432 <... futex resumed> ) = 5 +12433 futex(0x208778c, FUTEX_WAIT_PRIVATE, 6, NULL +12432 futex(0x208778c, FUTEX_WAKE_PRIVATE, 2147483647 +12433 <... futex resumed> ) = -1 EAGAIN (Resource temporarily unavailable) +12429 <... futex resumed> ) = 0 +12436 <... futex resumed> ) = 0 +12435 <... futex resumed> ) = 0 +12436 futex(0x1475c070, FUTEX_WAIT_PRIVATE, 0, NULL +12435 futex(0x10292840, FUTEX_WAIT_PRIVATE, 0, NULL +12434 <... futex resumed> ) = 0 +12433 futex(0x78cf9b0, FUTEX_WAIT_PRIVATE, 0, NULL +12434 futex(0xbda97d0, FUTEX_WAIT_PRIVATE, 0, NULL +12432 <... futex resumed> ) = 4 +12432 futex(0x36116f0, FUTEX_WAIT_PRIVATE, 0, NULL +12429 openat(AT_FDCWD, "./ocamlrun", O_RDONLY) = 4 +12429 lseek(4, 0, SEEK_END) = 2291768 +12429 mmap(NULL, 2291768, PROT_READ|PROT_WRITE, MAP_PRIVATE, 4, 0) = 0x7f4c68369000 +12429 lseek(4, 65536, SEEK_SET) = 65536 +12429 brk(0x19472000) = 0x19472000 +12429 read(4, "\0\0\0\340\0\0\0\200\0\30\0\0@\0\270\17\205\2\304\17\305\362\3d\0\264\5\361\0\0\0\0"..., 467144) = 467144 +12429 lseek(4, 540872, SEEK_SET) = 540872 +12429 read(4, "\377\377\377\377\377\377\377\377\0\0\0\0\0\0\0\0\377\377\377\377\377\377\377\377\0\0\0\0\0\0\0\0", 32) = 32 +12429 lseek(4, 540904, SEEK_SET) = 540904 +12429 read(4, "", 0) = 0 +12429 lseek(4, 589824, SEEK_SET) = 589824 +12429 read(4, "", 0) = 0 +12429 lseek(4, 589824, SEEK_SET) = 589824 +12429 read(4, "", 0) = 0 +12429 munmap(0x7f4c68369000, 2291768) = 0 +12429 close(4) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 getcwd("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun", 4096) = 54 +12429 lstat("/home/monniaux/work/Kalray/tests/ocaml-4.07.1/byterun/ocamlrun", {st_mode=S_IFREG|0755, st_size=2291768, ...}) = 0 +12429 futex(0x36116f0, FUTEX_WAKE_PRIVATE, 1) = 1 +12432 <... futex resumed> ) = 0 +12432 futex(0x36116a0, FUTEX_WAKE_PRIVATE, 1) = 0 +12432 mprotect(0x7f4b20021000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20022000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20023000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20025000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20027000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20029000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2002a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2002c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2002e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20030000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20031000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20033000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20035000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20037000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20038000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2003a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2003c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2003e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2003f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20041000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20043000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20045000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20047000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20048000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2004a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2004c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2004e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2004f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20051000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20053000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20055000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20056000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20058000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2005a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2005c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2005d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2005f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20061000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20063000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20064000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20066000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20068000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2006a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2006c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2006d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2006f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20071000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20073000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20074000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20076000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20078000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2007a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2007b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2007d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2007f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20081000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20082000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20084000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20086000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20088000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20089000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2008b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2008d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2008f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20091000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20092000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20094000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20096000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20098000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20099000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2009b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2009d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2009f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12429 futex(0x20877f8, FUTEX_WAIT_PRIVATE, 0, NULL +12432 mprotect(0x7f4b200a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200a6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200af000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200b6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200bd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200c4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200cb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200d4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200db000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200e2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200e9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200f0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200f9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b200fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20100000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20101000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20103000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20105000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20107000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20108000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2010a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2010c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2010e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2010f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20111000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20113000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20115000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20116000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20118000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2011a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2011c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2011e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2011f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20121000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20123000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20125000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20126000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20128000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2012a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2012c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2012d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2012f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20131000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20133000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20134000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20136000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20138000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2013a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2013b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2013d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2013f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20141000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20143000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20144000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20146000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20148000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2014a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2014b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2014d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2014f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20151000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20152000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20154000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20156000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20158000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20159000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2015b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2015d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2015f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20160000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20162000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20164000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20166000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20168000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20169000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2016b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2016d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2016f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20170000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20172000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20174000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20176000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20177000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20179000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2017b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2017d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2017e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20180000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20182000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20184000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20185000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20187000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20189000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2018b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2018d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2018e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20190000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20192000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20194000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20195000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20197000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20199000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2019b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2019c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2019e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201a2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201a9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201b2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201b9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201c0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201c7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ce000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201d7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201de000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201e5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ec000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201f3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201fc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b201ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20201000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20203000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20204000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20206000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20208000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2020a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2020b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2020d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2020f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20211000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20212000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20214000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20216000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20218000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20219000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2021b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2021d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2021f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20220000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20222000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20224000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20226000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20228000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20229000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2022b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2022d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2022f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20230000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20232000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20234000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20236000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20237000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20239000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2023b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2023d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2023e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20240000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20242000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20244000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20245000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20247000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20249000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2024b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2024d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2024e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20250000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20252000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20254000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20255000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20257000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20259000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2025b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2025c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2025e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20260000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20262000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20263000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20265000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20267000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20269000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2026a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2026c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2026e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20270000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20272000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20273000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20275000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20277000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20279000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2027a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2027c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2027e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20280000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20281000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20283000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20285000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20287000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20288000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2028a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2028c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2028e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2028f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20291000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20293000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20295000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20297000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20298000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2029a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2029c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2029e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2029f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202a5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202ac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202b3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202bc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202c3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202ca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202d1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202d8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202e1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202e8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202ef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202f6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202fd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b202fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20300000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20302000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20304000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20306000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20307000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20309000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2030b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2030d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2030e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20310000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20312000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20314000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20315000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20317000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20319000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2031b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2031c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2031e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20320000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20322000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20323000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20325000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20327000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20329000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2032b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2032c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2032e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20330000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20332000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20333000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20335000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20337000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20339000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2033a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2033c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2033e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20340000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20341000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20343000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20345000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20347000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20348000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2034a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2034c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2034e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20350000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20351000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20353000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20355000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20357000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20358000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2035a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2035c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2035e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2035f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20361000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20363000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20365000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20366000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20368000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2036a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2036c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2036d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2036f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20371000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20373000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20375000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20376000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20378000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2037a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2037c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2037d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2037f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20381000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20383000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20384000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20386000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20388000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2038a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2038b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2038d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2038f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20391000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20392000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20394000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20396000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20398000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2039a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2039b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2039d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2039f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203a1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203a8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203af000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203b6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203bf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203c6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203cd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203d4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203db000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203e2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203eb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203f2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203f9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b203fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20400000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20401000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20403000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20405000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20407000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20408000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2040a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2040c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2040e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20410000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20411000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20413000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20415000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20417000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20418000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2041a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2041c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2041e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2041f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20421000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20423000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20425000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20426000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20428000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2042a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2042c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2042d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2042f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20431000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20433000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20435000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20436000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20438000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2043a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2043c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2043d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2043f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20441000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20443000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20444000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20446000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20448000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2044a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2044b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2044d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2044f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20451000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20452000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20454000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20456000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20458000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2045a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2045b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2045d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2045f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20461000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20462000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20464000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20466000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20468000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20469000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2046b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2046d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2046f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20470000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20472000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20474000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20476000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20477000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20479000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2047b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2047d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2047f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20480000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20482000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20484000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20486000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20487000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20489000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2048b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2048d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2048e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20490000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20492000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20494000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20495000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20497000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20499000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2049b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2049c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2049e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204a4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ab000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204b2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204b9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204c0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204c9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204d0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204d7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204de000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204e5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204f5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204fc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b204ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20501000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20503000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20504000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20506000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20508000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2050a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2050b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2050d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2050f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20511000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20513000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20514000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20516000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20518000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2051a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2051b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2051d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2051f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20521000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20522000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20524000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20526000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20528000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20529000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2052b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2052d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2052f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20530000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20532000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20534000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20536000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20538000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20539000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2053b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2053d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2053f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20540000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20542000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20544000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20546000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20547000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20549000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2054b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2054d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2054e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20550000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20552000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20554000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20555000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20557000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20559000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2055b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2055d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2055e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20560000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20562000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20564000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20565000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20567000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20569000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2056b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2056c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2056e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20570000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20572000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20573000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20575000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20577000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20579000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2057a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2057c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2057e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20580000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20582000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20583000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20585000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20587000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20589000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2058a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2058c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2058e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20590000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20591000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20593000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20595000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20597000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20598000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2059a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2059c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2059e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2059f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205a5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205ae000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205b5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205bc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205c3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205ca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205d3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205da000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205e1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205e8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205ef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205f8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b205ff000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20600000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20602000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20604000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20606000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20607000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20609000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2060b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2060d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2060e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20610000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20612000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20614000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20615000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20617000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20619000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2061b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2061d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2061e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20620000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20622000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20624000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20625000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20627000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20629000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2062b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2062c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2062e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20630000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20632000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20633000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20635000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20637000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20639000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2063a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2063c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2063e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20640000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20642000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20643000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20645000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20647000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20649000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2064a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2064c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2064e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20650000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20651000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20653000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20655000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20657000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20658000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2065a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2065c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2065e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2065f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20661000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20663000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20665000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20667000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20668000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2066a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2066c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2066e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2066f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20671000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20673000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20675000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20676000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20678000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2067a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2067c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2067d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2067f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20681000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20683000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20684000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20686000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20688000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2068a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2068c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2068d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2068f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20691000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20693000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20694000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20696000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20698000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2069a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2069b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2069d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2069f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206a1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206a8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206b1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206b8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206bf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206c6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206cd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206dd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206e4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206eb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206f2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b206fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20700000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20702000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20703000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20705000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20707000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20709000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2070a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2070c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2070e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20710000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20711000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20713000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20715000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20717000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20718000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2071a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2071c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2071e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20720000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20721000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20723000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20725000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20727000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20728000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2072a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2072c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2072e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2072f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20731000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20733000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20735000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20736000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20738000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2073a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2073c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2073d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2073f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20741000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20743000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20745000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20746000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20748000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2074a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2074c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2074d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2074f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20751000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20753000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20754000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20756000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20758000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2075a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2075b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2075d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2075f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20761000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20762000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20764000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20766000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20768000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20769000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2076b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2076d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2076f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20771000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20772000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20774000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20776000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20778000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20779000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2077b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2077d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2077f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20780000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20782000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20784000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20786000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20787000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20789000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2078b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2078d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2078e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20790000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20792000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20794000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20796000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20797000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20799000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2079b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2079d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2079e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207a4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ab000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207b2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207bb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207c2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207c9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207d0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207d7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207e0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207e7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207f5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207fc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b207ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20801000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20803000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20805000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20806000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20808000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2080a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2080c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2080d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2080f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20811000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20813000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20814000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20816000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20818000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2081a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2081b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2081d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2081f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20821000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20822000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20824000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20826000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20828000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2082a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2082b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2082d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2082f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20831000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20832000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20834000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20836000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20838000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20839000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2083b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2083d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2083f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20840000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20842000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20844000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20846000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20847000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20849000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2084b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2084d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2084f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20850000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20852000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20854000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20856000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20857000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20859000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2085b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2085d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2085e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20860000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20862000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20864000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20865000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20867000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20869000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2086b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2086c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2086e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20870000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20872000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20874000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20875000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20877000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20879000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2087b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2087c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2087e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20880000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20882000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20883000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20885000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20887000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20889000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2088a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2088c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2088e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20890000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20891000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20893000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20895000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20897000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20899000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2089a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2089c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2089e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208a0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208a7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208ae000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208b5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208be000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208c5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208cc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208d3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208da000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208e3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208ea000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208f1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208f8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b208ff000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20900000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20902000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20904000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20906000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20908000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20909000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2090b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2090d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2090f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20910000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20912000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20914000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20916000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20917000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20919000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2091b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2091d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2091e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20920000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20922000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20924000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20925000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20927000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20929000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2092b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2092c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2092e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20930000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20932000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20934000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20935000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20937000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20939000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2093b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2093c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2093e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20940000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20942000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20943000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20945000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20947000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20949000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2094a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2094c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2094e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20950000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20951000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20953000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20955000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20957000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20959000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2095a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2095c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2095e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20960000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20961000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20963000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20965000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20967000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20968000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2096a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2096c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2096e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2096f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20971000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20973000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20975000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20976000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20978000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2097a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2097c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2097e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2097f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20981000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20983000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20985000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20986000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20988000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2098a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2098c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2098d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2098f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20991000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20993000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20994000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20996000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20998000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2099a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2099b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2099d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2099f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209a3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209aa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209b1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209b8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209bf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209c8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209cf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209dd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209e4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209ed000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209f4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b209fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a00000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a02000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a07000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a09000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a12000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a17000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a19000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a1e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a20000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a25000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a27000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a2c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a2e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a37000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a3c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a3e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a43000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a45000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a4a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a4c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a51000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a53000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a5c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a61000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a63000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a68000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a6a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a6f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a71000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a76000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a78000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a81000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a86000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a88000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a8d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a8f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a94000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a96000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a9b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a9d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20a9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aa0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aa2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aa4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aa6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aa7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aa9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aad000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ab0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ab2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ab4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ab5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ab7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ab9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20abb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20abc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20abe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ac0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ac2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ac3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ac5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ac7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ac9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20acb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20acc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ace000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ad0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ad2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ad3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ad5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ad7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ad9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ada000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20adc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ade000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ae0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ae1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ae3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ae5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ae7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ae8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20af1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20af3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20af5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20af7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20af8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20afa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20afc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20afe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20aff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b05000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b0c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b13000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b1c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b23000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b2a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b31000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b38000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b41000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 openat(AT_FDCWD, "./ocamlrun", O_RDONLY|O_NONBLOCK) = 4 +12432 fcntl(4, F_SETFL, O_RDONLY) = 0 +12432 mprotect(0x7f4b20b42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b48000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b4f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b56000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b5f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b66000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b6d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "\177E", 2) = 2 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b20b74000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b7b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b84000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b8b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b92000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b99000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20b9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ba0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ba1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 lseek(4, -16, SEEK_END) = 2291752 +12432 mprotect(0x7f4b20ba3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ba5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ba7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ba9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20baa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bb0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bb5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bb7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bbc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bbe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "\1\0\0\0\0\0\0\0\0\0\0\0\0\0\0\0", 16) = 16 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b20bc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bc3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bc5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bcc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bce000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bcf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bd1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bd3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bd5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bd8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bda000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bdc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20be1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20be3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20be4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20be6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20be8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bea000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20beb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bf1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bf3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bf4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bf6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bf8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bfa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bfb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bfd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20bff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c01000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c08000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c09000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c0f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c18000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 close(4) = 0 +12432 mprotect(0x7f4b20c1b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c1f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c22000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c26000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c2d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c34000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c3d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c40000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c44000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c47000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c4b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c52000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c59000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c5e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c62000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c65000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c69000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c6c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c70000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c77000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c7e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c83000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c87000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c8a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c8e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c91000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c95000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c9c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20c9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ca1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ca3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ca4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ca6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ca8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20caa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20caf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cb3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cb4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cb6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cba000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 openat(AT_FDCWD, "/home/monniaux/work/caml/quicksort", O_RDONLY|O_NONBLOCK) = 4 +12432 fcntl(4, F_SETFL, O_RDONLY) = 0 +12432 lseek(4, -16, SEEK_END) = 146122 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "\0\0\0\7Caml1999X023", 16) = 16 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b20cbb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cc1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cc2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cc8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cc9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ccb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ccd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ccf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cd1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cd4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cd8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cdb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cdf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ce0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ce2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ce4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ce6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ce7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ce9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ceb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ced000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cf0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cf2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cf4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cf6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cf7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cf9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cfb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20cfd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 lseek(4, -72, SEEK_END) = 146066 +12432 mprotect(0x7f4b20cfe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d00000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d04000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "CODE\0\2\3\230DLPT\0\0\0\0DLLS\0\0\0\0PRIM\0\0\35\177"..., 56) = 56 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b20d05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d07000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d09000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d0b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d14000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d17000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d1b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d1e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d22000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d25000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d29000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d2c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d30000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d37000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d3c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d40000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d43000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d47000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d4a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d4e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d51000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d55000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d5c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d61000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d65000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d68000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d6c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d6f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d73000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d76000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d7a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d81000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d86000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d8a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d8d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d91000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d94000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d98000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d9b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20d9f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20da0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20da2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20da4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20da6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20da7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20da9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20daf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20db0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20db2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20db4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20db6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20db7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20db9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dbb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dbd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dbe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dc0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dc2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dc4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dc5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dc7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dc9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dcb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dcc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dd0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dd4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dd5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dd7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ddb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ddc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dde000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20de0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20de2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20de3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20de5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20de7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20de9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20df0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20df1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20df3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20df5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20df7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20df9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dfa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dfc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20dfe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e00000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e07000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e0e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e15000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e1e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e25000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e2c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e33000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e3a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e43000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e4a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e51000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e58000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e5f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e68000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e6f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e76000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e7d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e84000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e8d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e94000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e9b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20e9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ea0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ea2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ea3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ea5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ea7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ea9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eaa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eb0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eb2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eb5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eb7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eb9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ebc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ebe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ec0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ec1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ec3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ec5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ec7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ec8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ecc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ece000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ecf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ed1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ed3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ed5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ed7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ed8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eda000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20edc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ede000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20edf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ee1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ee3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ee5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ee6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ee8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eec000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ef1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ef3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ef4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ef6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ef8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20efa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20efb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20efd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20eff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f03000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f0a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f11000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f18000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f1b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f1f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f22000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f28000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f2f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f36000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f3d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f40000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f44000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f47000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f4d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f54000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f5b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f5e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f62000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f65000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f69000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f6c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f72000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f79000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f80000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f83000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f87000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f8a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f8e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f91000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f97000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f9e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20f9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fa1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fa3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fa5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fa6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fa8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20faa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20faf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fb3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fb4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fb6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fbc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fc3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fcb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fcd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fcf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fd1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fd4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fd8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fdb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fe1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fe2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fe4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fe6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fe8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fe9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20feb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20fef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ff0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ff2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ff4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ff6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ff7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ff9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ffb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ffd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b20ffe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21000000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21002000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21004000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21006000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21007000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21009000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2100b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2100d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2100e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21010000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21012000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21014000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21015000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21017000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21019000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2101b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2101c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2101e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21020000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21022000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21023000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21025000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21027000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21029000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2102b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2102c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2102e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21030000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21032000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21033000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21035000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21037000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21039000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2103a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2103c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2103e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21040000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21041000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21043000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21045000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21047000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21048000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2104a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2104c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2104e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21050000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21051000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21053000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21055000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21057000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21058000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2105a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2105c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2105e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2105f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21061000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21063000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21065000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21066000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21068000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2106a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2106c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2106d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2106f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21071000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21073000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21075000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21076000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21078000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2107a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2107c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2107d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2107f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21081000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21083000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21084000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21086000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21088000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2108a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2108b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2108d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2108f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21091000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21092000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21094000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21096000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21098000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2109a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2109b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2109d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2109f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210a1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210a8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210af000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210b6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210bd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210c6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210cd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210d4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210db000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210e2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210eb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210f2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210f9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b210fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21100000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21101000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21103000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21105000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21107000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21108000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2110a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2110c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2110e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21110000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21111000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21113000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21115000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21117000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21118000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2111a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2111c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2111e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2111f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21121000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21123000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21125000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21126000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21128000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2112a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2112c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2112d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2112f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21131000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21133000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21135000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21136000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21138000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2113a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2113c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2113d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2113f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21141000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21143000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21144000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21146000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21148000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2114a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2114b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2114d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2114f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21151000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21152000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21154000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21156000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21158000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2115a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2115b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2115d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2115f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21161000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21162000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21164000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21166000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21168000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21169000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2116b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2116d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2116f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21170000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21172000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21174000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21176000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21177000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21179000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2117b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2117d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2117f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21180000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21182000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21184000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21186000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21187000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21189000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2118b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2118d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2118e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21190000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21192000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21194000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21195000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21197000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21199000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2119b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2119c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2119e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211a4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ab000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211b2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211b9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211c0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211c9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211d0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211d7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211de000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211e5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211f5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211fc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b211ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21201000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21203000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21204000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21206000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21208000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2120a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2120b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2120d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2120f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21211000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21213000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21214000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21216000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21218000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2121a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2121b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2121d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2121f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21221000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21222000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21224000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21226000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21228000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21229000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2122b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2122d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2122f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21230000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21232000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21234000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21236000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21238000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21239000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2123b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2123d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2123f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21240000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21242000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21244000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21246000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21247000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21249000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2124b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2124d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2124e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21250000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21252000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21254000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21255000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21257000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21259000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2125b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2125d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2125e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21260000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21262000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21264000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21265000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21267000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21269000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2126b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2126c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2126e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21270000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21272000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21273000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21275000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21277000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21279000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2127a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2127c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2127e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21280000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21281000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21283000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21285000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21287000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21289000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2128a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2128c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2128e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21290000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21291000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21293000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21295000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21297000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21298000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2129a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2129c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2129e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2129f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212a5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212ae000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212b5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212bc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212c3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212ca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212d3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212da000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212e1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212e8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212ef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212f8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b212ff000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21300000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21302000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21304000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21306000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21307000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21309000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2130b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2130d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2130e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21310000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21312000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21314000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21315000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21317000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21319000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2131b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2131d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2131e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21320000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21322000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21324000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21325000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21327000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21329000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2132b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2132c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2132e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21330000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21332000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21333000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21335000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21337000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21339000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2133a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2133c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2133e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21340000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21342000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21343000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21345000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21347000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21349000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2134a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2134c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2134e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21350000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21351000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21353000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21355000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21357000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21358000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2135a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2135c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2135e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2135f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21361000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21363000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21365000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21367000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21368000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2136a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2136c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2136e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2136f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21371000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21373000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21375000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21376000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21378000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2137a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2137c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2137d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2137f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21381000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21383000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21384000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21386000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21388000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2138a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2138c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2138d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2138f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21391000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21393000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21394000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21396000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21398000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2139a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2139b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2139d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2139f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213a1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213a8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213b1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213b8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213bf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213c6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213cd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213dd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213e4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213eb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213f2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b213fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21400000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21402000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21403000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21405000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21407000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21409000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2140a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2140c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2140e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21410000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21411000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21413000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21415000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21417000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21418000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2141a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2141c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2141e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21420000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21421000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21423000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21425000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21427000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21428000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2142a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2142c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2142e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2142f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21431000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21433000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21435000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21436000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21438000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2143a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2143c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2143d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2143f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21441000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21443000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21444000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21446000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21448000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2144a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2144c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2144d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2144f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21451000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21453000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21454000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21456000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21458000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2145a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2145b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2145d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2145f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21461000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21462000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21464000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21466000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21468000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21469000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2146b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2146d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2146f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21471000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21472000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21474000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21476000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21478000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21479000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2147b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2147d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2147f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21480000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21482000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21484000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21486000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21487000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21489000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2148b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2148d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2148e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21490000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21492000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21494000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21496000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21497000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21499000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2149b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2149d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2149e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214a4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ab000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214b2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214bb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214c2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214c9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214d0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214d7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214e0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214e7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214f5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214fc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b214ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21501000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21503000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21505000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21506000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21508000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2150a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2150c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2150d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2150f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21511000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21513000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21514000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21516000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21518000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2151a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2151b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2151d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2151f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21521000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21522000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21524000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21526000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21528000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2152a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2152b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2152d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2152f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21531000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21532000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21534000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21536000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21538000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21539000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2153b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2153d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2153f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21540000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21542000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21544000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21546000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21547000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21549000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2154b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2154d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2154f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21550000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21552000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21554000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21556000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21557000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21559000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2155b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2155d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2155e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21560000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21562000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21564000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21565000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21567000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21569000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2156b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2156c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2156e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21570000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21572000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21574000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21575000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21577000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21579000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2157b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2157c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2157e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21580000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21582000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21583000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21585000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21587000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21589000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2158a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2158c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2158e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21590000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21591000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21593000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21595000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21597000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21599000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2159a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2159c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2159e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215a0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215a7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215ae000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215b5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215be000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215c5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215cc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215d3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215da000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 write(2, "allocating ", 11) = 11 +12432 mprotect(0x7f4b215df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215e3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215ea000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215f1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215f8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b215ff000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21600000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21602000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21604000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21606000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21608000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21609000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2160b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2160d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2160f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21610000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21612000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21614000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21616000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21617000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 write(2, "491520", 6) = 6 +12432 mprotect(0x7f4b21619000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2161b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2161d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2161e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21620000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21622000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21624000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21625000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21627000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21629000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2162b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2162d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2162e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21630000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21632000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21634000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21635000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21637000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21639000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2163b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2163c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2163e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21640000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21642000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21643000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21645000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21647000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21649000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2164a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2164c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2164e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21650000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21652000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21653000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21655000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21657000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21659000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2165a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2165c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2165e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21660000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21661000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21663000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21665000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21667000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21668000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2166a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2166c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2166e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2166f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21671000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21673000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21675000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21677000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21678000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2167a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 write(2, "\n", 1) = 1 +12432 mprotect(0x7f4b2167c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2167e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2167f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21681000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21683000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21685000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21686000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21688000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2168a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2168c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2168d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2168f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21691000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21693000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21694000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21696000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21698000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2169a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2169c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2169d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2169f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216a3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216aa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216b1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216b8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216c1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216c8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216cf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216dd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216e4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216ed000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216f4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b216fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21700000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21702000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21703000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21705000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21707000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21709000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2170a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2170c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2170e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21710000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21712000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21713000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21715000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21717000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21719000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2171a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2171c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2171e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21720000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21721000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21723000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21725000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21727000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21728000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2172a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2172c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2172e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2172f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21731000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21733000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21735000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21737000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21738000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2173a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2173c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2173e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2173f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21741000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21743000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21745000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21746000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21748000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2174a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2174c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2174d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2174f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21751000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21753000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21754000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21756000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21758000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2175a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2175c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2175d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2175f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21761000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21763000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21764000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21766000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21768000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2176a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2176b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2176d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2176f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21771000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21772000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21774000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21776000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21778000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21779000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2177b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2177d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2177f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21781000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21782000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21784000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21786000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21788000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21789000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2178b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2178d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2178f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21790000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21792000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21794000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21796000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21797000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21799000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2179b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2179d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2179e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217a6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217ad000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217b4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217bb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217c2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217cb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217d2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217d9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217e0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217e7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217f0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217f7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217fe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b217ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21801000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21803000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21805000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21806000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21808000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2180a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2180c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2180d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2180f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21811000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21813000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21815000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21816000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21818000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2181a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2181c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2181d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2181f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21821000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21823000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21824000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21826000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21828000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2182a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2182b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2182d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2182f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21831000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21832000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21834000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21836000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21838000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2183a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2183b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2183d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2183f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21841000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21842000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21844000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21846000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21848000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21849000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2184b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2184d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2184f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21850000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21852000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21854000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21856000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21857000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21859000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2185b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2185d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2185f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21860000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21862000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21864000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21866000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21867000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21869000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2186b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2186d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2186e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21870000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21872000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21874000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21875000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21877000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21879000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2187b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2187c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2187e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21880000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21882000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21884000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21885000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21887000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21889000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2188b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2188c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2188e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21890000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21892000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21893000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21895000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21897000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21899000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2189a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2189c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2189e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218a0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218a7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218b0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218b7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218be000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218c5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218cc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218d5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218dc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218e3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218ea000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218f1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218fa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b218ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21901000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21902000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21904000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21906000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21908000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21909000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2190b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2190d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2190f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21910000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21912000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21914000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21916000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21917000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21919000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2191b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2191d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2191f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21920000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21922000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21924000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21926000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21927000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21929000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2192b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2192d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2192e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21930000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21932000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21934000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21935000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21937000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21939000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2193b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2193c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2193e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21940000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21942000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21944000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21945000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21947000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21949000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2194b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2194c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2194e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21950000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21952000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21953000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21955000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21957000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21959000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2195a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2195c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2195e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21960000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21961000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21963000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21965000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21967000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21969000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2196a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2196c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2196e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21970000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21971000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21973000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21975000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21977000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21978000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2197a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2197c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2197e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2197f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21981000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21983000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21985000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21986000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21988000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2198a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2198c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2198e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2198f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21991000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21993000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21995000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21996000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21998000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2199a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2199c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2199d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2199f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219a3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219aa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219b3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219ba000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219c1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219c8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219cf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219d8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219df000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219e6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219ed000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219f4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219fd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b219fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a00000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a04000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a07000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a09000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a0b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a12000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a17000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a19000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a1e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a22000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a25000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a29000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a2c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a30000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a37000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a3c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a3e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a43000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a47000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a4a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a4e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a51000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a55000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a5c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a61000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a63000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a68000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a6a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a6f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a73000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a76000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a7a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 lseek(4, -146099, SEEK_END) = 39 +12432 mprotect(0x7f4b21a81000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a86000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a88000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a8d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a8f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a94000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a98000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a9b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21a9f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aa0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aa2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aa4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aa6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aa7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aa9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aad000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ab0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ab2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ab4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ab5000, 135168, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "T\0\0\0\337\2\0\0\0\0\0\0W\0\0\0\1\0\17\0\20\0\0\0\23\0\0\0\34\0\0\0"..., 131992) = 131992 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b21ad6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ad7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ad9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ada000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21adc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ade000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ae0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ae2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ae3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ae5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ae7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ae9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21af0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21af1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21af3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21af5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21af7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21af8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21afa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21afc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21afe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21aff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b07000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b0e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b15000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b1c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b23000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b2c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b33000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b3a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b41000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b48000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b51000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b58000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b5f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b66000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b6d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b76000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b7d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b84000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b8b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b92000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b9b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21b9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ba0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ba2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ba3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ba5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ba7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ba9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21baa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bb0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bb5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bb7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bbc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bbe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bc0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bc3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bc5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bc7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bcc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bce000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bcf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bd1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bd3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bd5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bd8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bda000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bdc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21be1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21be3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21be5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21be6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21be8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bec000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bf1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bf3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bf4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bf6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bf8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bfa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bfb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bfd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21bff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c01000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c0a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c11000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c18000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c1b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c1f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c22000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c26000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c2d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c36000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c3d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c40000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c44000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c47000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c4b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c52000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c5b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c5e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c62000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c65000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c69000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c6c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c70000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c77000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c80000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c83000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c87000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c8a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c8e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c91000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c95000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c9c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21c9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ca1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ca3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ca5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ca6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ca8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21caa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21caf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cb3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cb4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cb6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cba000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cbb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cc1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cc2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ccb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ccd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ccf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cd1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cd4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cd8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cdb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cdf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ce0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ce2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ce4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ce6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ce7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ce9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ceb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ced000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cf0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cf2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cf4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cf6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cf7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cf9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cfb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cfd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21cfe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d00000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d04000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d07000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d09000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d0b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d14000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d17000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d1b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d1e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d22000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d25000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d29000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d2c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d30000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d39000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d3c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d40000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d43000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d47000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d4a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d4e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d51000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d55000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d5e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d61000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d65000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d68000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d6c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d6f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d73000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d76000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d7a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d83000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d86000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d8a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d8d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d91000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d94000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d98000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d9b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21d9f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21da0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21da2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21da4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21da6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21da8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21da9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21daf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21db0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21db2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21db4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21db6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21db7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21db9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dbb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dbd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dbe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dc0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dc2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dc4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dc5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dc7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dc9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dcb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dcd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dd0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dd4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dd5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dd7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ddb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ddc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dde000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21de0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21de2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21de3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21de5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21de7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21de9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21df0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21df1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21df3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21df5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21df7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21df9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dfa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dfc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21dfe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e00000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e07000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e0e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e15000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e1e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e25000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e2c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e33000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e3a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e43000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e4a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e51000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e58000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e5f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e68000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e6f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e76000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e7d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e84000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e8d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e94000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e9b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21e9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ea0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ea2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ea3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ea5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ea7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ea9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eaa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eb0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eb2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eb5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eb7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eb9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ebc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ebe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ec0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ec1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ec3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ec5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ec7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ec8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ecc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ece000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ecf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ed1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ed3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ed5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ed7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ed8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eda000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21edc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ede000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21edf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ee1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ee3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ee5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ee6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ee8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eec000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ef1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ef3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ef4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ef6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ef8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21efa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21efc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21efd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21eff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f03000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f0a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f11000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f18000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f1b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f21000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f22000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f28000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f2f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f36000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f3d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f40000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f46000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f47000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f4d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f54000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f5b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f5e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f62000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f65000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f6b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f6c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f72000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f79000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f80000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f83000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f87000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f8a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f90000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f91000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f97000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f9e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21f9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fa1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fa3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fa5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fa6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fa8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21faa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21faf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fb3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fb4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fb6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fbc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fc3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fcb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fcd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fcf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fd1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fd4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fd8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fdb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fe1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fe2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fe4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fe6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fe8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fe9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21feb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21fef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ff0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ff2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ff4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ff6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ff7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ff9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ffb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ffd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b21ffe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22000000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22002000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22004000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22006000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22007000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22009000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2200b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2200d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2200e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22010000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22012000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22014000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22015000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22017000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22019000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2201b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2201c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2201e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22020000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22022000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22023000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22025000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22027000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22029000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2202b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2202c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2202e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22030000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22032000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22033000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22035000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22037000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22039000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2203a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2203c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2203e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22040000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22041000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22043000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22045000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22047000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22048000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2204a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2204c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2204e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22050000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22051000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22053000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22055000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22057000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22058000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2205a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2205c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2205e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2205f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22061000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22063000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22065000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22066000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22068000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2206a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2206c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2206d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2206f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22071000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22073000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22075000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22076000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22078000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2207a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2207c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2207d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2207f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22081000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22083000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22084000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22086000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22088000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2208a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2208b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2208d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2208f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22091000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22092000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22094000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22096000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22098000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2209a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2209b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2209d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2209f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220a1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220a8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220af000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220b6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220bf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220c6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220cd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220d4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220db000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220e4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220eb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220f2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220f9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b220fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22100000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22101000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22103000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22105000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22107000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22109000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2210a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2210c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2210e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22110000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22111000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22113000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22115000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22117000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22118000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2211a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2211c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2211e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2211f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22121000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22123000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22125000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22126000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22128000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2212a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2212c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2212e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2212f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22131000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22133000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22135000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22136000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22138000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2213a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2213c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2213d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2213f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22141000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22143000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22144000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22146000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22148000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2214a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2214b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2214d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2214f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22151000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22153000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22154000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22156000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22158000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2215a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2215b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2215d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2215f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22161000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22162000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22164000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 lseek(4, -14107, SEEK_END) = 132031 +12432 mprotect(0x7f4b22166000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22168000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22169000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2216b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2216d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2216f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22170000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22172000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22174000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22176000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22177000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22179000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2217b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2217d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "", 0) = 0 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b2217f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22180000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22182000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22184000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22186000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22187000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22189000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2218b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2218d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2218e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22190000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22192000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22194000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22195000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22197000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22199000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2219b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2219d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2219e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221a4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ab000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221b2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221b9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221c0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221c9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221d0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221d7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221de000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221e5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221f5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221fc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b221ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22201000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22203000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22204000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22206000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22208000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2220a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2220b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2220d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2220f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22211000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22213000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22214000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22216000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22218000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2221a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2221b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2221d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2221f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22221000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22222000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22224000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22226000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22228000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22229000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2222b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2222d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 lseek(4, -14107, SEEK_END) = 132031 +12432 mprotect(0x7f4b2222f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22230000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22232000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22234000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22236000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22238000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22239000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2223b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2223d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2223f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22240000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22242000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22244000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22246000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "", 0) = 0 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b22247000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22249000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2224b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2224d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2224e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22250000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22252000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22254000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22255000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22257000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22259000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2225b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2225d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2225e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22260000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22262000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22264000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22265000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22267000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22269000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2226b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2226c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2226e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22270000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22272000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22273000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22275000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22277000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22279000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2227a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2227c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2227e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22280000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22282000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22283000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22285000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22287000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22289000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2228a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2228c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2228e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22290000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22291000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22293000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22295000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22297000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22298000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2229a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2229c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2229e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2229f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222a7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222ae000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222b5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222bc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222c3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222cc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222d3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222da000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222e1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 lseek(4, -14107, SEEK_END) = 132031 +12432 mprotect(0x7f4b222e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222e8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222f1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222f8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b222ff000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "caml_abs_float\0caml_acos_float\0c"..., 7551) = 7551 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b22300000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22302000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22304000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22306000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22307000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22309000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2230b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2230d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2230e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22310000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22312000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22314000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22316000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22317000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22319000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2231b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2231d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2231e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22320000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22322000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22324000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22325000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22327000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22329000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2232b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2232c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2232e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22330000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22332000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22333000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22335000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22337000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22339000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2233b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2233c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2233e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22340000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22342000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22343000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22345000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22347000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22349000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2234a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2234c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2234e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22350000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22351000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22353000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22355000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22357000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22358000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2235a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2235c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2235e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22360000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22361000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22363000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22365000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22367000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22368000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2236a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2236c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2236e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2236f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22371000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22373000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22375000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22376000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22378000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2237a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2237c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2237d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2237f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22381000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22383000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22384000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22386000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22388000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2238a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2238c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2238d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2238f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22391000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22393000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22394000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22396000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22398000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2239a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2239b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2239d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2239f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223a1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223a8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223b1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223b8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223bf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223c6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223cd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223dd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223e4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223eb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223f2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b223fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22400000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22402000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22403000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22405000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22407000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22409000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2240a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2240c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2240e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22410000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22411000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22413000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22415000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22417000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22418000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2241a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2241c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2241e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22420000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22421000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22423000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22425000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22427000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22428000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2242a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2242c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2242e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2242f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22431000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22433000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22435000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22436000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22438000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2243a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2243c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2243d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2243f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22441000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22443000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22445000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22446000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22448000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2244a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2244c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2244d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2244f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22451000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22453000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22454000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22456000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22458000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2245a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2245b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2245d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2245f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22461000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22462000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22464000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22466000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22468000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2246a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2246b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2246d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2246f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22471000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22472000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22474000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22476000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22478000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22479000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2247b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2247d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2247f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22480000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22482000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22484000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22486000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22487000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22489000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2248b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2248d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 stat("/opt/ccomp/ocaml/4.07.1/lib/ocaml/ld.conf", {st_mode=S_IFREG|0664, st_size=77, ...}) = 0 +12432 mprotect(0x7f4b2248f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22490000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22492000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22494000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22496000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22497000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22499000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2249b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2249d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2249e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224a4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ad000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224b4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224bb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224c2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224c9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224d0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224d9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 openat(AT_FDCWD, "/opt/ccomp/ocaml/4.07.1/lib/ocaml/ld.conf", O_RDONLY|O_NONBLOCK) = 5 +12432 fcntl(5, F_SETFL, O_RDONLY) = 0 +12432 mprotect(0x7f4b224e0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224e7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(5, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(5, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(5, "/opt/ccomp/ocaml/4.07.1/lib/ocam"..., 77) = 77 +12432 fcntl(5, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b224f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224f7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224fe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b224ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22501000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22503000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22505000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22506000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22508000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2250a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2250c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2250d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2250f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22511000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22513000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22515000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22516000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22518000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2251a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2251c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2251d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2251f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22521000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22523000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22524000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22526000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22528000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2252a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2252b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2252d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2252f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22531000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22532000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22534000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22536000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22538000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22539000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 close(5) = 0 +12432 mprotect(0x7f4b2253b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2253d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2253f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22541000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22542000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22544000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22546000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22548000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22549000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2254b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2254d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2254f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22550000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22552000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22554000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22556000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22557000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22559000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2255b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2255d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2255e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22560000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22562000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22564000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22566000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22567000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22569000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2256b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2256d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2256e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22570000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22572000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22574000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22575000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22577000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22579000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2257b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2257c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2257e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22580000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22582000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22583000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22585000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22587000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22589000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2258b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2258c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2258e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22590000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22592000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22593000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22595000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22597000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22599000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2259a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2259c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2259e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225a0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225a7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225b0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225b7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225be000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225c5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225cc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225d5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225dc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225e3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225ea000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225f1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225fa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b225ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22601000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22602000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22604000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22606000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22608000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22609000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2260b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2260d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2260f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22610000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22612000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22614000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22616000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22617000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22619000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2261b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2261d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2261f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22620000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22622000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22624000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22626000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22627000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22629000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2262b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2262d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2262e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22630000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22632000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22634000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22635000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22637000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22639000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2263b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2263c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2263e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22640000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22642000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22644000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22645000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22647000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22649000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2264b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2264c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2264e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22650000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22652000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22653000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22655000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22657000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22659000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2265a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2265c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2265e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22660000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22661000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22663000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22665000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22667000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22669000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2266a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2266c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2266e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22670000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22671000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22673000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22675000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22677000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22678000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2267a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2267c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2267e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 lseek(4, -6556, SEEK_END) = 139582 +12432 mprotect(0x7f4b2267f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22681000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22683000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22685000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22686000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22688000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2268a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2268c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2268e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2268f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22691000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22693000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22695000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22696000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22698000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2269a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2269c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2269d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2269f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226a3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226aa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226b3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226ba000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226c1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226c8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226cf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226d8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226df000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226e6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 lseek(4, 0, SEEK_CUR) = 139582 +12432 mprotect(0x7f4b226eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226ed000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226f4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b226fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22700000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22702000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22704000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22705000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22707000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22709000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2270b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2270c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2270e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22710000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22712000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22713000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22715000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22717000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22719000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2271a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2271c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2271e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22720000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22721000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22723000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22725000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22727000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22729000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2272a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2272c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2272e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22730000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22731000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22733000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22735000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22737000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22738000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2273a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2273c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2273e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2273f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22741000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22743000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22745000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22746000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22748000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2274a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2274c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2274e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2274f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22751000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22753000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22755000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22756000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22758000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2275a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2275c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2275d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2275f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22761000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22763000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22764000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22766000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22768000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2276a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2276b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2276d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2276f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22771000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22773000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22774000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22776000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22778000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2277a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2277b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2277d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2277f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22781000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22782000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22784000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22786000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22788000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22789000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2278b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2278d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2278f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22790000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22792000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22794000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22796000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22798000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22799000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2279b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2279d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2279f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227a6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227ad000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227b4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227bd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227c4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227cb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227d2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227d7000, 65536, PROT_READ|PROT_WRITE) = 0 +12432 fcntl(4, F_GETFL) = 0x8000 (flags O_RDONLY|O_LARGEFILE) +12432 fcntl(4, F_SETFL, O_RDONLY|O_NONBLOCK|O_LARGEFILE) = 0 +12432 read(4, "\204\225\246\276\0\0\22:\0\0\1\346\0\0\t\26\0\0\7K\10\0\4\360\0\10\0\0\10\370-O"..., 65536) = 6556 +12432 fcntl(4, F_SETFL, O_RDONLY|O_LARGEFILE) = 0 +12432 mprotect(0x7f4b227e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227e9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227f0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227f7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227fe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b227ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22801000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22803000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22805000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22807000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22808000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2280a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2280c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2280e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2280f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22811000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22813000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22815000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22816000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22818000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2281a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2281c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2281d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2281f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22821000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22823000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22824000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22826000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22828000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2282a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2282c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2282d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2282f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22831000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22833000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22834000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22836000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22838000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2283a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2283b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2283d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2283f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22841000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22842000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22844000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22846000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22848000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22849000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2284b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2284d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2284f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22851000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22852000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22854000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22856000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22858000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22859000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2285b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2285d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2285f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22860000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22862000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22864000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22866000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22867000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22869000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2286b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2286d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2286e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22870000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22872000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22874000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22876000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22877000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22879000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2287b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2287d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2287e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22880000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22882000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22884000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22885000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22887000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22889000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2288b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2288c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2288e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22890000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22892000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22893000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22895000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22897000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22899000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2289b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2289c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2289e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228a2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228a9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228b0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228b7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228be000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228c7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ce000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228d5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228dc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228e3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ec000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228f3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228fa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b228ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22901000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22902000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22904000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22906000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22908000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22909000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2290b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2290d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2290f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22911000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22912000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22914000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22916000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22918000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22919000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2291b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2291d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2291f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22920000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22922000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22924000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22926000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22927000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22929000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2292b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2292d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2292e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22930000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22932000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22934000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22936000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22937000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22939000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2293b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2293d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2293e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22940000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22942000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22944000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22945000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22947000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22949000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2294b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2294c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2294e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22950000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22952000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22953000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22955000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22957000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22959000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2295b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2295c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2295e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22960000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22962000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22963000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22965000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22967000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22969000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2296a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2296c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2296e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22970000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22971000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22973000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22975000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22977000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22978000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2297a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2297c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2297e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22980000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22981000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22983000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22985000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22987000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22988000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2298a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2298c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2298e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2298f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22991000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22993000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22995000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22996000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22998000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2299a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2299c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2299d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2299f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229a5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229ac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229b3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229ba000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229c1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229ca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229d1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229d8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229df000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229e6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229ef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229f6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229fd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b229fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a00000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a04000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a07000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a09000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a0b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a14000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a17000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a1b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a1e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a22000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a25000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a29000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a2c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a30000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a39000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a3c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a40000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a43000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a47000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a4a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a4e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a51000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a55000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a5e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a61000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a65000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a68000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a6c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a6f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a73000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a76000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a7a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a81000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a86000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a8a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a8d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a91000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a94000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a98000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a9b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22a9f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aa0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aa2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aa4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aa6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aa7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aa9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aaf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ab0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ab2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ab4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ab6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ab7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ab9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22abb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22abd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22abe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ac0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ac2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ac4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ac5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ac7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ac9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22acb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22acc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ace000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ad0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ad2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ad4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ad5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ad7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ad9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22adb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22adc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ade000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ae0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ae2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ae3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ae5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ae7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ae9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22aee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22af0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22af1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22af3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22af5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22af7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22af9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22afa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22afc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22afe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b00000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b07000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b0e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b15000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b1e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b25000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b2c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b33000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b3a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b43000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b4a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b51000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b58000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b5f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b68000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b6f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b76000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b7d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b84000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b8d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b94000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b9b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22b9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ba0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ba2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ba3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ba5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ba7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ba9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22baa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bb0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bb2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bb5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bb7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bb9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bbc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bbe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bc0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bc3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bc5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bc7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bcc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bce000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bcf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bd1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bd3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bd5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bd7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bd8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bda000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bdc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bde000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22be1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22be3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22be5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22be6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22be8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bec000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bf1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bf3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bf4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bf6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bf8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bfa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bfc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bfd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22bff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c03000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c0a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c11000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c18000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c1b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c21000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c22000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c28000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c2f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c36000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c3d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c40000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c44000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c47000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c4d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c54000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c5b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c5e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c62000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c65000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c69000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c6c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c72000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c79000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c80000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c83000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c87000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c8a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c8e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c91000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c97000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c9e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22c9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ca1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ca3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ca5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ca6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ca8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22caa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22caf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cb3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cb4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cb6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cbc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cc3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ccb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ccd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ccf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cd1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cd4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cd8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cdb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ce1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ce2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ce4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ce6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ce8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ce9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ceb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ced000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cf0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cf2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cf4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cf6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cf7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cf9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cfb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cfd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22cfe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d00000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d06000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d07000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d09000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d0d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d14000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 close(4) = 0 +12432 mprotect(0x7f4b22d15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d17000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d1b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d1e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d22000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d25000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d2b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d2c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d32000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d39000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d3c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d40000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d43000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d47000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d4a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d50000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d51000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d57000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d5e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d61000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d65000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d68000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d6c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d6f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d75000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d76000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d7c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d83000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d86000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d8a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d8d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d91000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d94000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d9a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d9b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22d9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22da1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22da2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22da4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22da6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22da8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22da9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22daf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22db0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22db2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22db4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22db6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22db7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22db9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dbb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dbf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dc0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dc2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dc6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dc7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dc9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dcb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dcd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dd0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dd4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dd5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dd7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ddb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ddc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dde000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22de0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22de2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22de4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22de5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22de7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22de9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22deb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22df0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22df2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22df3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22df5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22df7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22df9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dfa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dfc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22dfe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e00000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e07000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e10000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e17000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e1e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e25000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e2c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e35000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e3c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e43000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e4a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e51000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e5a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e61000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e68000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e6f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e76000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e7f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e86000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e8d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e94000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e9b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22e9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ea0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ea2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ea4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ea5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ea7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ea9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eab000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eb0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eb2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eb5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eb7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eb9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ebc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ebe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ec0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ec1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ec3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ec5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ec7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ec9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ecc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ece000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ed0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ed1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ed3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ed5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ed7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ed8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eda000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22edc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ede000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22edf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ee1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ee3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ee5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ee6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ee8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ef1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ef3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ef5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ef6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ef8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22efa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22efc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22efd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22eff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f03000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f0a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f13000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f1a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f1b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f21000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f22000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f28000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f2f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f38000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f3f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f40000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f46000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f47000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f4d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f54000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f5d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f5e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f64000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f65000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f6b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f6c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f72000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f79000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f82000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f83000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f89000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f8a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f90000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f91000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f97000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f9e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22f9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fa1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fa3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fa5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fa7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fa8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22faa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fae000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22faf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fb5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fb6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fbc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fc3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fcb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fcd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fcf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fd1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fd3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fd4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fd8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fda000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fdb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fe1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fe2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fe4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fe6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fe8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fe9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22feb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ff0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ff2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ff4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ff6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ff8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ff9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ffb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22ffd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b22fff000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23000000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23002000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23004000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23006000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23007000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23009000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2300b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2300d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2300e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23010000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23012000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23014000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23015000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23017000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23019000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2301b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2301d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2301e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23020000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23022000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23024000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23025000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23027000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23029000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2302b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2302c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2302e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23030000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23032000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23033000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23035000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23037000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23039000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2303a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2303c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2303e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23040000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23042000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23043000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23045000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23047000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23049000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2304a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2304c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2304e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23050000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23051000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23053000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23055000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23057000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23058000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2305a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2305c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2305e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2305f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23061000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23063000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23065000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23067000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23068000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2306a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2306c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2306e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2306f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23071000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23073000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23075000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23076000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23078000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2307a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2307c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2307d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2307f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23081000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23083000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23084000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23086000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23088000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2308a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2308c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2308d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2308f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23091000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23093000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23094000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23096000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23098000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2309a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2309b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2309d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2309f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230a1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230a8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230b1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230b8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230bf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230c6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230cd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230dd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230e4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230eb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230f2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b230fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23100000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23102000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23103000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23105000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23107000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23109000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2310a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2310c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2310e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23110000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23111000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23113000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23115000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23117000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23118000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2311a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2311c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2311e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23120000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23121000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23123000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23125000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23127000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23128000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2312a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2312c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2312e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2312f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23131000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23133000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23135000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23136000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23138000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2313a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2313c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2313d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2313f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23141000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23143000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23145000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23146000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23148000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2314a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2314c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2314d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2314f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23151000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23153000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23154000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23156000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23158000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2315a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2315b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2315d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2315f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23161000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23162000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23164000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23166000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23168000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2316a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2316b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2316d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2316f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23171000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23172000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23174000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23176000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23178000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23179000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2317b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2317d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2317f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23180000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23182000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23184000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23186000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23187000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23189000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2318b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2318d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2318e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23190000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23192000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23194000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23196000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23197000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23199000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2319b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 lseek(0, 0, SEEK_CUR) = -1 ESPIPE (Illegal seek) +12432 mprotect(0x7f4b2319d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2319e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231a4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ab000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231b2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231bb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231c2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231c9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231d0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231d7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231e0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231e7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231f5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231fc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b231ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23201000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23203000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23205000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23206000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23208000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2320a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2320c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2320d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2320f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23211000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23213000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23214000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23216000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23218000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2321a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2321b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2321d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2321f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23221000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23222000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23224000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23226000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23228000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2322a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2322b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2322d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2322f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23231000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23232000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23234000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23236000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23238000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23239000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2323b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2323d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2323f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23240000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23242000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23244000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23246000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23247000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23249000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2324b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2324d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2324f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23250000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23252000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23254000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23256000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23257000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23259000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2325b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2325d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2325e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23260000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23262000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23264000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23265000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23267000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23269000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2326b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2326c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2326e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23270000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23272000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23274000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23275000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23277000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23279000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2327b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2327c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2327e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23280000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23282000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23283000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23285000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23287000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23289000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2328a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2328c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2328e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23290000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23291000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23293000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23295000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23297000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23299000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2329a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2329c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2329e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232a0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232a7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232ae000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232b5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232be000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232c5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232cc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232d3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232da000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232e3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232ea000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232f1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232f8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b232ff000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23300000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23302000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23304000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23306000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23308000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23309000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2330b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2330d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2330f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23310000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23312000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23314000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23316000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23317000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23319000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2331b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2331d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2331e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23320000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23322000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23324000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23325000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23327000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23329000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2332b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2332d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2332e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23330000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23332000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23334000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23335000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23337000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23339000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2333b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2333c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2333e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23340000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23342000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23343000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23345000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23347000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23349000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2334a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2334c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2334e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23350000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23351000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23353000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23355000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23357000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23359000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2335a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2335c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2335e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23360000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23361000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23363000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23365000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23367000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23368000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2336a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2336c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2336e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2336f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23371000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23373000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23375000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23376000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23378000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2337a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2337c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2337e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2337f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23381000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23383000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23385000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23386000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23388000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2338a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2338c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2338d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2338f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23391000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23393000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23394000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23396000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23398000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2339a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2339b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2339d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2339f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233a3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233aa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233b1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233b8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233bf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233c8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233cf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233dd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233e4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233ed000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233f4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b233fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23400000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23402000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23403000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23405000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23407000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23409000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2340a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2340c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2340e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23410000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23412000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23413000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23415000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23417000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23419000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2341a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2341c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2341e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23420000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23421000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23423000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23425000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23427000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23428000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2342a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2342c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2342e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2342f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23431000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23433000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23435000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23437000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23438000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2343a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2343c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2343e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2343f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23441000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23443000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23445000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23446000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23448000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2344a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2344c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2344d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2344f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23451000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23453000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23454000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23456000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23458000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2345a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2345c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2345d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2345f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23461000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23463000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23464000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23466000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23468000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2346a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2346b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2346d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2346f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23471000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23472000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23474000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23476000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23478000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23479000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2347b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2347d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2347f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23481000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23482000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23484000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23486000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23488000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23489000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2348b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2348d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2348f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23490000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23492000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23494000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23496000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23497000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23499000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2349b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2349d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2349e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234a6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234ad000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234b4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234bb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234c2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234cb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234d2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234d9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234e0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234e7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234f0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234f7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234fe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b234ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23501000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23503000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23505000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23506000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23508000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2350a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 lseek(1, 0, SEEK_CUR) = -1 ESPIPE (Illegal seek) +12432 mprotect(0x7f4b2350c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2350d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2350f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23511000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23513000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23514000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23516000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23518000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2351a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2351c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2351d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2351f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23521000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23523000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23524000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23526000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23528000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2352a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2352b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2352d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2352f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23531000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23532000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23534000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23536000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23538000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23539000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2353b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2353d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2353f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23541000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23542000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23544000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23546000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23548000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23549000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2354b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2354d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2354f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23550000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23552000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23554000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23556000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23557000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23559000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2355b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2355d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2355e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 lseek(2, 0, SEEK_CUR) = -1 ESPIPE (Illegal seek) +12432 mprotect(0x7f4b23560000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23562000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23564000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23566000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23567000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23569000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2356b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2356d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2356e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23570000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23572000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23574000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23575000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23577000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23579000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2357b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2357c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2357e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23580000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23582000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23583000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23585000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23587000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23589000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2358b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2358c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2358e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23590000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23592000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23593000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23595000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23597000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23599000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2359a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2359c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2359e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235a0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235a7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235b0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235b7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235be000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235c5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235cc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235d5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235dc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235e3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235ea000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235f1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235fa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b235ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23601000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23602000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23604000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23606000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23608000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23609000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2360b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2360d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2360f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23610000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23612000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23614000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23616000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23617000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23619000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2361b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2361d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2361f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23620000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23622000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23624000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23626000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23627000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23629000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2362b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2362d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2362e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23630000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23632000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23634000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23635000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23637000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23639000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2363b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2363c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2363e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23640000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23642000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23644000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23645000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23647000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23649000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2364b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2364c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2364e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23650000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23652000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23653000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23655000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23657000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23659000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2365a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2365c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2365e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23660000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23661000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23663000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23665000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23667000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23669000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2366a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2366c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2366e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23670000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23671000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23673000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23675000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23677000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23678000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2367a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2367c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2367e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2367f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23681000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23683000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23685000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23686000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23688000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2368a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2368c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2368e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2368f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23691000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23693000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23695000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23696000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23698000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2369a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2369c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2369d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2369f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236a3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236aa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236b3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236ba000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236c1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236c8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236cf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236df000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236e6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236ed000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236f4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b236fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23700000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23702000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23704000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23705000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23707000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23709000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2370b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2370c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2370e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23710000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23712000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23713000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23715000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23717000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23719000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2371a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2371c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2371e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23720000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23721000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23723000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23725000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23727000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23729000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2372a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2372c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2372e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23730000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23731000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23733000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23735000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23737000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23738000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2373a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2373c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2373e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2373f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23741000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23743000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23745000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23746000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23748000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2374a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2374c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2374e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2374f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23751000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23753000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23755000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23756000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23758000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2375a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2375c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2375d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2375f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23761000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23763000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23764000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23766000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23768000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2376a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2376b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2376d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2376f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23771000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23773000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23774000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23776000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23778000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2377a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2377b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2377d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2377f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23781000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23782000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23784000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23786000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23788000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23789000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2378b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2378d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2378f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23790000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23792000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23794000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23796000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23798000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23799000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2379b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2379d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2379f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237a6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237ad000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237b4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237bd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237c4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237cb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237d2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237d9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237e2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237e9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237f0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237f7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237fe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b237ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23801000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23803000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23805000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23807000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23808000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2380a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2380c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2380e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2380f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23811000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23813000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23815000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23816000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23818000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2381a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2381c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2381d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2381f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23821000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23823000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23824000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23826000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23828000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2382a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2382c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2382d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2382f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23831000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23833000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23834000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23836000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23838000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2383a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2383b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2383d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2383f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23841000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23842000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23844000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23846000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23848000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23849000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2384b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2384d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2384f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23851000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23852000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23854000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23856000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23858000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23859000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2385b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2385d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2385f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23860000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23862000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23864000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23866000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23867000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23869000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2386b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2386d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2386e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23870000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23872000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23874000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23876000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23877000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23879000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2387b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2387d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2387e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23880000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23882000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23884000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23885000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23887000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23889000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2388b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2388c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2388e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23890000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23892000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23893000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23895000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23897000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23899000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2389a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2389c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2389e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238a2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238a9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238b0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238b3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238b7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238be000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238c1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238c3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238c7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ce000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238d1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238d5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238d8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238dc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238df000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238e3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238e6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238e8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ec000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238f3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238f6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238fa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238fd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b238ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23901000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23902000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23904000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23906000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23908000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23909000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2390b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2390d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2390f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23911000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23912000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23914000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23916000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23918000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23919000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2391b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2391d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2391f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23920000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23922000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23924000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23926000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23927000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23929000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2392b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2392d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2392e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23930000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23932000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23934000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23936000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23937000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23939000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2393b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2393d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2393e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23940000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23942000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23944000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23945000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23947000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23949000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2394b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2394c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2394e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23950000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23952000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23953000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23955000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23957000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23959000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2395b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2395c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2395e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23960000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23962000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23963000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23965000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23967000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23969000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2396a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2396c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2396e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23970000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23971000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23973000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23975000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23977000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23978000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2397a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2397c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2397e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23980000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23981000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23983000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23985000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23987000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23988000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2398a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2398c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2398e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2398f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23991000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23993000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23995000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23996000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23998000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2399a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2399c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2399d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2399f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239a5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239ac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239b1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239b3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239ba000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239c1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239c8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239ca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239cf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239d1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239d6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239d8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239df000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239e6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239ed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239ef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239f4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239f6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239fb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239fd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b239fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a00000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a04000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a07000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a09000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a0b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a14000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a17000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a1b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a1e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a22000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a25000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a29000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a2c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a30000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a39000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a3c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a40000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a43000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a47000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a4a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a4e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a51000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a55000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a5c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a61000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a65000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a68000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a6c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a6f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a73000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a76000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a7a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a81000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a86000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a8a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a8d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a91000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a94000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a98000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a9b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23a9f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aa0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aa2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aa4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aa6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aa7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aa9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aaf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ab0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ab2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ab4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ab6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ab7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ab9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23abb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23abd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23abe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ac0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ac2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ac4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ac5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ac7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ac9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23acb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23acc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ace000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ad0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ad2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ad4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ad5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ad7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ad9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23adb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23adc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ade000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ae0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ae2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ae3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ae5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ae7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ae9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23aee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23af0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23af1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23af3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23af5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23af7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23af9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23afa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23afc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23afe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b00000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b07000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b0e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b15000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b1e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b25000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b2c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b33000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b3a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b43000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b4a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b51000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b58000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b5f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b68000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b6f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b76000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b7d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b84000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b8d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b94000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b9b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23b9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ba0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ba2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ba3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ba5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ba7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ba9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23baa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bb0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bb2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bb5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bb7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bb9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bbc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bbe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bc0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bc3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bc5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bc7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bcc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bce000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bcf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bd1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bd3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bd5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bd7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bd8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bda000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bdc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bde000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23be1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23be3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23be5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23be6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23be8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bec000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bf1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bf3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bf4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bf6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bf8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bfa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bfc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bfd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23bff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c03000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c0a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c11000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c18000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c1b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c1f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c22000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c28000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c2f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c36000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c3d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c40000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c44000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c47000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c4d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c54000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c5b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c5e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c62000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c65000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c69000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c6c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c72000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c79000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c80000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c83000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c87000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c8a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c8e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c91000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c97000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c9e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23c9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ca1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ca3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ca5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ca6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ca8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23caa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cac000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23caf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cb3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cb4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cb6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cbc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cc3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ccb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ccd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ccf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cd1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cd4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cd8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cdb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ce1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ce2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ce4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ce6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ce8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ce9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ceb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ced000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cf0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cf2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cf4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cf6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cf7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cf9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cfb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cfd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23cfe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d00000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d02000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d06000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d07000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d09000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d0d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d10000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d12000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d14000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d17000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d19000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d1b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d1e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d20000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d22000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d25000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d27000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d2b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d2c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d2e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d32000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d35000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d37000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d39000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d3c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d3e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d40000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d43000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d45000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d47000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d4a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d4c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d50000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d51000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d53000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d57000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d5a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d5c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d5e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d61000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d63000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d65000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d68000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d6a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d6c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d6f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d71000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d75000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d76000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d78000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d7c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d7f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d81000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d83000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d86000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d88000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d8a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d8d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d8f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d91000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d94000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d96000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d9a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d9b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d9d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23d9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23da1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23da2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23da4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23da6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23da8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23da9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23daf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23db0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23db2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23db4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23db6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23db7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23db9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dbb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dbf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dc0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dc2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dc6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dc7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dc9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dcb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dcd000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dd0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dd2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dd4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dd5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dd7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dd9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ddb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ddc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dde000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23de0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23de2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23de3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23de5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23de7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23de9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23deb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23df0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23df2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23df3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23df5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23df7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23df9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dfa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dfc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23dfe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e00000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e03000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e05000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e07000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e0a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e0c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e0e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e10000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e13000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e15000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e17000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e1a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e1c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e1e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e21000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e23000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e25000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e28000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e2a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e2c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e2f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e31000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e33000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e35000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e38000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e3a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e3c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e3f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e41000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e43000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e46000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e48000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e4a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e4d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e4f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e51000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e54000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e56000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e58000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e5a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e5d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e5f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e61000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e64000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e66000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e68000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e6b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e6d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e6f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e72000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e74000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e76000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e79000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e7b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e7d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e7f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e82000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e84000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e86000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e89000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e8b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e8d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e90000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e92000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e94000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e97000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e99000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e9b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23e9e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ea0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ea2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ea4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ea5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ea7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ea9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eab000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eb0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eb2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eb5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eb7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eb9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ebc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ebe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ec0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ec1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ec3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ec5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ec7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ec9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eca000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ecc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ece000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ed0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ed1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ed3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ed5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ed7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ed8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eda000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23edc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ede000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23edf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ee1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ee3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ee5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ee6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ee8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eee000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eef000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ef1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ef3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ef5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ef6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ef8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23efa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23efc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23efd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23eff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f01000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f03000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f04000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f06000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f08000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f0a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f0b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f0d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f0f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f11000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f13000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f14000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f16000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f18000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f1a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f1b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f1d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f1f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f21000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f22000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f24000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f26000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f28000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f29000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f2b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f2d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f2f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f30000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f32000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f34000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f36000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f38000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f39000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f3b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f3d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f3f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f40000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f42000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f44000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f46000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f47000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f49000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f4b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f4d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f4e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f50000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f52000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f54000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f55000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f57000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f59000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f5b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f5d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f5e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f60000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f62000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f64000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f65000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f67000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f69000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f6b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f6c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f6e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f70000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f72000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f73000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f75000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f77000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f79000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f7a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f7c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f7e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f80000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f82000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f83000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f85000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f87000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f89000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f8a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f8c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f8e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f90000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f91000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f93000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f95000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f97000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f98000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f9a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f9c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f9e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23f9f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fa1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fa3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fa5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fa6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fa8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23faa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fae000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23faf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fb1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fb3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fb5000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fb6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fb8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fba000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fbc000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fbd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fbf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fc1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fc3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fc4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fc6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fc8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fca000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fcb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fcd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fcf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fd1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fd3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fd4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fd6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fd8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fda000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fdb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fdd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fdf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fe1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fe2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fe4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fe6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fe8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fe9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23feb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fed000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fef000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ff0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ff2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ff4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ff6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ff8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ff9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ffb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23ffd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b23fff000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mmap(0x7f4b24000000, 67108864, PROT_NONE, MAP_PRIVATE|MAP_ANONYMOUS|MAP_NORESERVE, -1, 0) = 0x7f4b24000000 +12432 mprotect(0x7f4b24000000, 139264, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24022000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24024000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24026000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24027000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24029000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2402b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2402d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2402e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24030000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24032000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24034000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24035000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24037000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24039000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2403b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2403c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2403e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24040000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24042000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24043000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24045000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24047000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24049000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2404b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2404c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2404e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24050000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24052000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24053000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24055000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24057000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24059000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2405a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2405c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2405e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24060000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24061000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24063000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24065000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24067000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24068000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2406a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2406c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2406e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24070000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24071000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24073000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24075000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24077000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24078000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2407a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2407c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2407e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2407f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24081000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24083000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24085000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24086000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24088000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2408a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2408c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2408d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2408f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24091000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24093000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24094000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24096000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24098000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2409a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2409c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2409d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2409f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240a1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240a3000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 write(1, "-31; 0; 1; 2; 4; 65; 83; 99; 782"..., 35) = 35 +12432 mprotect(0x7f4b240a6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240a8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240aa000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240ad000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240af000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240b1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240b4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240b6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240b8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240ba000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240bb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240bd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240bf000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240c1000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240c4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240c6000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240c8000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240cb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240cd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240cf000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240d2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240d4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240d6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240d9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240db000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240dd000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240df000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240e0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240e2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240e4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240e6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240e9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240eb000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240ed000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240f0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240f2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240f4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240f7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240f9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240fb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b240fe000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24100000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24102000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24103000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24105000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24107000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24109000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2410b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2410c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2410e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24110000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24112000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24113000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24115000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24117000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24119000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2411a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2411c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2411e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24120000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24121000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24123000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24125000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24127000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24128000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2412a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2412c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2412e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24130000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24131000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24133000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24135000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24137000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24138000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2413a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2413c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2413e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2413f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24141000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24143000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24145000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24146000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24148000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2414a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2414c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2414d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2414f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24151000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24153000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24155000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24156000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24158000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2415a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2415c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2415d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2415f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24161000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24163000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24164000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24166000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24168000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2416a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2416b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2416d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2416f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24171000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24172000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24174000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24176000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24178000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2417a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2417b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2417d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2417f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24181000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24182000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24184000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24186000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24188000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24189000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2418b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2418d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2418f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24190000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24192000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24194000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24196000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24197000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24199000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2419b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2419d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2419f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241a2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241a4000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241a6000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241a9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241ab000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241ad000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241b0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241b2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241b4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241b5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241b7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241b9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241bb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241bc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241be000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241c0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241c2000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241c4000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241c5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241c7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241c9000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241cb000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241cc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241ce000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241d0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241d2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241d3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241d5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241d7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241d9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241da000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241dc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241de000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241e0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241e1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241e3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241e5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241e7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241e9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241ea000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241ec000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241ee000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241f0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241f1000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241f3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241f5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241f7000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241f8000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241fa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241fc000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241fe000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b241ff000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24201000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24203000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24205000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24206000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24208000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2420a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2420c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2420e000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2420f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24211000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24213000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24215000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24216000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24218000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2421a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2421c000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2421d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2421f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24221000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24223000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24224000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24226000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24228000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2422a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2422b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2422d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2422f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24231000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24233000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24234000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24236000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24238000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2423a000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2423b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2423d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2423f000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24241000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24242000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24244000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24246000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24248000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24249000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2424b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2424d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2424f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24250000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24252000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24254000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24256000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24258000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24259000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2425b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2425d000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2425f000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24260000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24262000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24264000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24266000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24267000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24269000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2426b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2426d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2426e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24270000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24272000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24274000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24275000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24277000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24279000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2427b000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2427d000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2427e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24280000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24282000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24284000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24285000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24287000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24289000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2428b000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2428c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2428e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24290000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24292000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24293000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24295000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24297000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b24299000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2429a000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2429c000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b2429e000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242a0000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242a2000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242a3000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242a5000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242a7000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242a9000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242aa000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242ac000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242ae000, 8192, PROT_READ|PROT_WRITE) = 0 +12432 mprotect(0x7f4b242b0000, 4096, PROT_READ|PROT_WRITE) = 0 +12432 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, {tv_sec=13, tv_nsec=793296369}) = 0 +12432 futex(0x78cf9b0, FUTEX_WAKE_PRIVATE, 1 +12433 <... futex resumed> ) = 0 +12432 <... futex resumed> ) = 1 +12433 futex(0x78cf960, FUTEX_WAIT_PRIVATE, 2, NULL +12432 futex(0x78cf960, FUTEX_WAKE_PRIVATE, 1 +12433 <... futex resumed> ) = -1 EAGAIN (Resource temporarily unavailable) +12432 <... futex resumed> ) = 0 +12433 futex(0x78cf960, FUTEX_WAKE_PRIVATE, 1 +12432 futex(0xbda97d0, FUTEX_WAKE_PRIVATE, 1 +12433 <... futex resumed> ) = 0 +12434 <... futex resumed> ) = 0 +12433 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, +12434 futex(0xbda9780, FUTEX_WAIT_PRIVATE, 2, NULL +12433 <... clock_gettime resumed> {tv_sec=13, tv_nsec=793373038}) = 0 +12432 <... futex resumed> ) = 1 +12433 futex(0x20877a8, FUTEX_WAIT_PRIVATE, 2, NULL +12432 futex(0xbda9780, FUTEX_WAKE_PRIVATE, 1 +12434 <... futex resumed> ) = 0 +12432 <... futex resumed> ) = 1 +12434 futex(0xbda9780, FUTEX_WAKE_PRIVATE, 1 +12432 futex(0x10292840, FUTEX_WAKE_PRIVATE, 1 +12434 <... futex resumed> ) = 0 +12435 <... futex resumed> ) = 0 +12434 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, +12435 futex(0x102927f0, FUTEX_WAIT_PRIVATE, 2, NULL +12434 <... clock_gettime resumed> {tv_sec=13, tv_nsec=793425864}) = 0 +12432 <... futex resumed> ) = 1 +12434 futex(0x20877a8, FUTEX_WAIT_PRIVATE, 2, NULL +12432 futex(0x102927f0, FUTEX_WAKE_PRIVATE, 1 +12435 <... futex resumed> ) = 0 +12432 <... futex resumed> ) = 1 +12435 futex(0x102927f0, FUTEX_WAKE_PRIVATE, 1 +12432 futex(0x1475c070, FUTEX_WAKE_PRIVATE, 1 +12435 <... futex resumed> ) = 0 +12432 <... futex resumed> ) = 1 +12436 <... futex resumed> ) = 0 +12435 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, +12436 futex(0x1475c020, FUTEX_WAKE_PRIVATE, 1 +12435 <... clock_gettime resumed> {tv_sec=13, tv_nsec=793479454}) = 0 +12436 <... futex resumed> ) = 0 +12435 futex(0x20877a8, FUTEX_WAIT_PRIVATE, 2, NULL +12436 clock_gettime(CLOCK_PROCESS_CPUTIME_ID, +12432 futex(0x20877f8, FUTEX_WAKE_PRIVATE, 1 +12436 <... clock_gettime resumed> {tv_sec=13, tv_nsec=793494482}) = 0 +12429 <... futex resumed> ) = 0 +12436 futex(0x20877a8, FUTEX_WAIT_PRIVATE, 2, NULL +12429 futex(0x20877a8, FUTEX_WAIT_PRIVATE, 2, NULL +12432 <... futex resumed> ) = 1 +12432 futex(0x20877a8, FUTEX_WAKE_PRIVATE, 1 +12433 <... futex resumed> ) = 0 +12432 <... futex resumed> ) = 1 +12433 futex(0x20877a8, FUTEX_WAKE_PRIVATE, 1) = 1 +12434 <... futex resumed> ) = 0 +12432 mprotect(0x7f4b242b1000, 4096, PROT_READ|PROT_WRITE +12434 futex(0x20877a8, FUTEX_WAKE_PRIVATE, 1 +12432 <... mprotect resumed> ) = 0 +12435 <... futex resumed> ) = 0 +12434 <... futex resumed> ) = 1 +12435 futex(0x20877a8, FUTEX_WAKE_PRIVATE, 1 +12434 madvise(0x7f4b4cef7000, 8368128, MADV_DONTNEED +12436 <... futex resumed> ) = 0 +12435 <... futex resumed> ) = 1 +12436 futex(0x20877a8, FUTEX_WAKE_PRIVATE, 1 +12433 madvise(0x7f4b4d6f8000, 8368128, MADV_DONTNEED +12435 madvise(0x7f4b4c6f6000, 8368128, MADV_DONTNEED +12429 <... futex resumed> ) = 0 +12436 <... futex resumed> ) = 1 +12435 <... madvise resumed> ) = 0 +12429 futex(0x7f4c68d999d0, FUTEX_WAIT, 12432, NULL +12436 madvise(0x7f4b4bef5000, 8368128, MADV_DONTNEED +12435 exit(0 +12436 <... madvise resumed> ) = 0 +12435 <... exit resumed>) = ? +12436 exit(0 +12434 <... madvise resumed> ) = 0 +12436 <... exit resumed>) = ? +12435 +++ exited with 0 +++ +12436 +++ exited with 0 +++ +12434 exit(0 +12433 <... madvise resumed> ) = 0 +12434 <... exit resumed>) = ? +12433 exit(0 +12434 +++ exited with 0 +++ +12433 <... exit resumed>) = ? +12432 madvise(0x7f4c68599000, 8368128, MADV_DONTNEED +12433 +++ exited with 0 +++ +12432 <... madvise resumed> ) = 0 +12432 exit(0) = ? +12429 <... futex resumed> ) = 0 +12432 +++ exited with 0 +++ +12429 munmap(0x7f4c68599000, 8392704) = 0 +12429 futex(0x20877a8, FUTEX_WAKE_PRIVATE, 1) = 0 +12429 munmap(0x7f4c706f2000, 34607104) = 0 +12429 munmap(0x7f4c69eff000, 34607104) = 0 +12429 munmap(0x7f4c61eff000, 34607104) = 0 +12429 munmap(0x7f4b61efe000, 4294971392) = 0 +12429 munmap(0x7f4b24000000, 67108864) = 0 +12429 openat(AT_FDCWD, "/proc/sys/vm/overcommit_memory", O_RDONLY|O_CLOEXEC) = 4 +12429 read(4, "0", 1) = 1 +12429 close(4) = 0 +12429 madvise(0x7f4b22ceb000, 20008960, MADV_DONTNEED) = 0 +12429 madvise(0x7f4b20021000, 46964736, MADV_DONTNEED) = 0 +12429 munmap(0x7f4c68f9f000, 6819840) = 0 +12429 munmap(0x7f4b5defd000, 67112960) = 0 +12429 munmap(0x7f4b59efc000, 67112960) = 0 +12429 munmap(0x7f4b55efb000, 67112960) = 0 +12429 munmap(0x7f4b51efa000, 67112960) = 0 +12429 munmap(0x7f4c68d9a000, 2114360) = 0 +12429 munmap(0x7f4c7038d000, 3556520) = 0 +12429 munmap(0x7f4c70173000, 2202280) = 0 +12429 munmap(0x7f4c69a56000, 4885520) = 0 +12429 munmap(0x7f4c69848000, 2151080) = 0 +12429 munmap(0x7f4c69620000, 2259944) = 0 +12429 munmap(0x7f4b4def9000, 67112960) = 0 +12429 openat(AT_FDCWD, "/tmp/mppa-tlm-input-12429", O_WRONLY) = 4 +12431 <... openat resumed> ) = 3 +12429 close(4 +12431 read(3, +12429 <... close resumed> ) = 0 +12429 unlink("/tmp/mppa-tlm-input-12429" +12431 <... read resumed> "", 1024) = 0 +12429 <... unlink resumed> ) = 0 +12431 madvise(0x7f4c727f3000, 8368128, MADV_DONTNEED +12429 unlink("/tmp/mppa-tlm-output-12429" +12431 <... madvise resumed> ) = 0 +12431 exit(0 +12429 <... unlink resumed> ) = 0 +12431 <... exit resumed>) = ? +12429 futex(0x7f4c72ff39d0, FUTEX_WAIT, 12431, NULL) = -1 EAGAIN (Resource temporarily unavailable) +12431 +++ exited with 0 +++ +12429 munmap(0x7f4b4d6f8000, 8392704) = 0 +12429 brk(0x19bfe000) = 0x19bfe000 +12429 exit_group(0) = ? +12430 <... rt_sigtimedwait resumed> ) = ? +12430 +++ exited with 0 +++ +12429 +++ exited with 0 +++ diff --git a/test/monniaux/ocaml/byterun/unix.c b/test/monniaux/ocaml/byterun/unix.c new file mode 100644 index 00000000..15c67642 --- /dev/null +++ b/test/monniaux/ocaml/byterun/unix.c @@ -0,0 +1,459 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 2001 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Unix-specific stuff */ + +#define _GNU_SOURCE + /* Helps finding RTLD_DEFAULT in glibc */ + /* also secure_getenv */ + +#include +#include +#include +#include +#include +#include +#if HAS_IOCTL +#include +#endif +#include +#include "caml/config.h" +#ifdef SUPPORT_DYNAMIC_LINKING +#ifdef __CYGWIN__ +#include "flexdll.h" +#else +#include +#endif +#endif +#ifdef HAS_UNISTD +#include +#endif +#if HAS_DIR +#ifdef HAS_DIRENT +#include +#else +#include +#endif +#endif +#ifdef __APPLE__ +#include +#endif +#include "caml/fail.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/osdeps.h" +#include "caml/signals.h" +#include "caml/sys.h" +#include "caml/io.h" +#include "caml/alloc.h" + +#ifndef S_ISREG +#define S_ISREG(mode) (((mode) & S_IFMT) == S_IFREG) +#endif + +#ifndef EINTR +#define EINTR (-1) +#endif +#ifndef EAGAIN +#define EAGAIN (-1) +#endif +#ifndef EWOULDBLOCK +#define EWOULDBLOCK (-1) +#endif + +int caml_read_fd(int fd, int flags, void * buf, int n) +{ + int retcode; + do { + caml_enter_blocking_section(); + retcode = read(fd, buf, n); + caml_leave_blocking_section(); + } while (retcode == -1 && errno == EINTR); + if (retcode == -1) caml_sys_io_error(NO_ARG); + return retcode; +} + +int caml_write_fd(int fd, int flags, void * buf, int n) +{ + int retcode; + again: +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + if (flags & CHANNEL_FLAG_BLOCKING_WRITE) { + retcode = write(fd, buf, n); + } else { +#endif + caml_enter_blocking_section(); + retcode = write(fd, buf, n); + caml_leave_blocking_section(); +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + } +#endif + if (retcode == -1) { + if (errno == EINTR) goto again; + if ((errno == EAGAIN || errno == EWOULDBLOCK) && n > 1) { + /* We couldn't do a partial write here, probably because + n <= PIPE_BUF and POSIX says that writes of less than + PIPE_BUF characters must be atomic. + We first try again with a partial write of 1 character. + If that fails too, we'll return an error code. */ + n = 1; goto again; + } + } + if (retcode == -1) caml_sys_io_error(NO_ARG); + CAMLassert (retcode > 0); + return retcode; +} + +caml_stat_string caml_decompose_path(struct ext_table * tbl, char * path) +{ + char * p, * q; + size_t n; + + if (path == NULL) return NULL; + p = caml_stat_strdup(path); + q = p; + while (1) { + for (n = 0; q[n] != 0 && q[n] != ':'; n++) /*nothing*/; + caml_ext_table_add(tbl, q); + q = q + n; + if (*q == 0) break; + *q = 0; + q += 1; + } + return p; +} + +caml_stat_string caml_search_in_path(struct ext_table * path, const char * name) +{ + const char * p; + char * dir, * fullname; + int i; + struct stat st; + + for (p = name; *p != 0; p++) { + if (*p == '/') goto not_found; + } + for (i = 0; i < path->size; i++) { + dir = path->contents[i]; + if (dir[0] == 0) dir = "."; /* empty path component = current dir */ + fullname = caml_stat_strconcat(3, dir, "/", name); + if (stat(fullname, &st) == 0 && S_ISREG(st.st_mode)) + return fullname; + caml_stat_free(fullname); + } + not_found: + return caml_stat_strdup(name); +} + +#ifdef __CYGWIN__ + +/* Cygwin needs special treatment because of the implicit ".exe" at the + end of executable file names */ + +static int cygwin_file_exists(const char * name) +{ + int fd, ret; + struct stat st; + /* Cannot use stat() here because it adds ".exe" implicitly */ + fd = open(name, O_RDONLY); + if (fd == -1) return 0; + ret = fstat(fd, &st); + close(fd); + return ret == 0 && S_ISREG(st.st_mode); +} + +static caml_stat_string cygwin_search_exe_in_path(struct ext_table * path, const char * name) +{ + const char * p; + char * dir, * fullname; + int i; + + for (p = name; *p != 0; p++) { + if (*p == '/' || *p == '\\') goto not_found; + } + for (i = 0; i < path->size; i++) { + dir = path->contents[i]; + if (dir[0] == 0) dir = "."; /* empty path component = current dir */ + fullname = caml_stat_strconcat(3, dir, "/", name); + if (cygwin_file_exists(fullname)) return fullname; + caml_stat_free(fullname); + fullname = caml_stat_strconcat(4, dir, "/", name, ".exe"); + if (cygwin_file_exists(fullname)) return fullname; + caml_stat_free(fullname); + } + not_found: + if (cygwin_file_exists(name)) return caml_stat_strdup(name); + fullname = caml_stat_strconcat(2, name, ".exe"); + if (cygwin_file_exists(fullname)) return fullname; + caml_stat_free(fullname); + return caml_stat_strdup(name); +} + +#endif + +caml_stat_string caml_search_exe_in_path(const char * name) +{ + struct ext_table path; + char * tofree; + caml_stat_string res; + + caml_ext_table_init(&path, 8); + tofree = caml_decompose_path(&path, getenv("PATH")); +#ifndef __CYGWIN__ + res = caml_search_in_path(&path, name); +#else + res = cygwin_search_exe_in_path(&path, name); +#endif + caml_stat_free(tofree); + caml_ext_table_free(&path, 0); + return res; +} + +caml_stat_string caml_search_dll_in_path(struct ext_table * path, const char * name) +{ + caml_stat_string dllname; + caml_stat_string res; + + dllname = caml_stat_strconcat(2, name, ".so"); + res = caml_search_in_path(path, dllname); + caml_stat_free(dllname); + return res; +} + +#ifdef SUPPORT_DYNAMIC_LINKING +#ifdef __CYGWIN__ +/* Use flexdll */ + +void * caml_dlopen(char * libname, int for_execution, int global) +{ + int flags = (global ? FLEXDLL_RTLD_GLOBAL : 0); + if (!for_execution) flags |= FLEXDLL_RTLD_NOEXEC; + return flexdll_dlopen(libname, flags); +} + +void caml_dlclose(void * handle) +{ + flexdll_dlclose(handle); +} + +void * caml_dlsym(void * handle, const char * name) +{ + return flexdll_dlsym(handle, name); +} + +void * caml_globalsym(const char * name) +{ + return flexdll_dlsym(flexdll_dlopen(NULL,0), name); +} + +char * caml_dlerror(void) +{ + return flexdll_dlerror(); +} + +#else +/* Use normal dlopen */ + +#ifndef RTLD_GLOBAL +#define RTLD_GLOBAL 0 +#endif +#ifndef RTLD_LOCAL +#define RTLD_LOCAL 0 +#endif + +void * caml_dlopen(char * libname, int for_execution, int global) +{ + return dlopen(libname, RTLD_NOW | (global ? RTLD_GLOBAL : RTLD_LOCAL)); + /* Could use RTLD_LAZY if for_execution == 0, but needs testing */ +} + +void caml_dlclose(void * handle) +{ + dlclose(handle); +} + +void * caml_dlsym(void * handle, const char * name) +{ +#ifdef DL_NEEDS_UNDERSCORE + char _name[1000] = "_"; + strncat (_name, name, 998); + name = _name; +#endif + return dlsym(handle, name); +} + +void * caml_globalsym(const char * name) +{ +#ifdef RTLD_DEFAULT + return caml_dlsym(RTLD_DEFAULT, name); +#else + return NULL; +#endif +} + +char * caml_dlerror(void) +{ + return (char*) dlerror(); +} + +#endif +#else + +void * caml_dlopen(char * libname, int for_execution, int global) +{ + return NULL; +} + +void caml_dlclose(void * handle) +{ +} + +void * caml_dlsym(void * handle, const char * name) +{ + return NULL; +} + +void * caml_globalsym(const char * name) +{ + return NULL; +} + +char * caml_dlerror(void) +{ + return "dynamic loading not supported on this platform"; +} + +#endif + +/* Add to [contents] the (short) names of the files contained in + the directory named [dirname]. No entries are added for [.] and [..]. + Return 0 on success, -1 on error; set errno in the case of error. */ + +CAMLexport int caml_read_directory(char * dirname, struct ext_table * contents) +{ +#if HAS_DIR + DIR * d; +#ifdef HAS_DIRENT + struct dirent * e; +#else + struct direct * e; +#endif + + d = opendir(dirname); + if (d == NULL) return -1; + while (1) { + e = readdir(d); + if (e == NULL) break; + if (strcmp(e->d_name, ".") == 0 || strcmp(e->d_name, "..") == 0) continue; + caml_ext_table_add(contents, caml_stat_strdup(e->d_name)); + } + closedir(d); + return 0; +#else + caml_invalid_argument("Unix.read_directory not implemented"); +#endif +} + +/* Recover executable name from /proc/self/exe if possible */ + +char * caml_executable_name(void) +{ +#if defined(__linux__) + int namelen, retcode; + char * name; + struct stat st; + + /* lstat("/proc/self/exe") returns st_size == 0 so we cannot use it + to determine the size of the buffer. Instead, we guess and adjust. */ + namelen = 256; + while (1) { + name = caml_stat_alloc(namelen); + retcode = readlink("/proc/self/exe", name, namelen); + if (retcode == -1) { caml_stat_free(name); return NULL; } + if (retcode < namelen) break; + caml_stat_free(name); + if (namelen >= 1024*1024) return NULL; /* avoid runaway and overflow */ + namelen *= 2; + } + /* readlink() does not zero-terminate its result. + There is room for a final zero since retcode < namelen. */ + name[retcode] = 0; + /* Make sure that the contents of /proc/self/exe is a regular file. + (Old Linux kernels return an inode number instead.) */ + if (stat(name, &st) == -1 || ! S_ISREG(st.st_mode)) { + caml_stat_free(name); return NULL; + } + return name; + +#elif defined(__APPLE__) + unsigned int namelen; + char * name; + + namelen = 256; + name = caml_stat_alloc(namelen); + if (_NSGetExecutablePath(name, &namelen) == 0) return name; + caml_stat_free(name); + /* Buffer is too small, but namelen now contains the size needed */ + name = caml_stat_alloc(namelen); + if (_NSGetExecutablePath(name, &namelen) == 0) return name; + caml_stat_free(name); + return NULL; + +#else + return NULL; + +#endif +} + +char *caml_secure_getenv (char const *var) +{ +#ifdef HAS_GETENV +#ifdef HAS_SECURE_GETENV + return secure_getenv (var); +#elif defined (HAS___SECURE_GETENV) + return __secure_getenv (var); +#elif defined(HAS_ISSETUGID) + if (!issetugid ()) + return CAML_SYS_GETENV (var); + else + return NULL; +#else + if (geteuid () == getuid () && getegid () == getgid ()) + return CAML_SYS_GETENV (var); + else + return NULL; +#endif +#else + return NULL; +#endif +} + +int caml_num_rows_fd(int fd) +{ +#ifdef TIOCGWINSZ + struct winsize w; + w.ws_row = -1; + if (ioctl(fd, TIOCGWINSZ, &w) == 0) + return w.ws_row; + else + return -1; +#else + return -1; +#endif +} + + diff --git a/test/monniaux/ocaml/byterun/weak.c b/test/monniaux/ocaml/byterun/weak.c new file mode 100644 index 00000000..f430ef8f --- /dev/null +++ b/test/monniaux/ocaml/byterun/weak.c @@ -0,0 +1,427 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Damien Doligez, projet Para, INRIA Rocquencourt */ +/* */ +/* Copyright 1997 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Operations on weak arrays and ephemerons (named ephe here)*/ + +#include + +#include "caml/alloc.h" +#include "caml/fail.h" +#include "caml/major_gc.h" +#include "caml/memory.h" +#include "caml/mlvalues.h" +#include "caml/weak.h" + +value caml_ephe_list_head = 0; + +static value ephe_dummy = 0; +value caml_ephe_none = (value) &ephe_dummy; + +#if defined (NATIVE_CODE) && defined (NO_NAKED_POINTERS) +/** The minor heap is considered alive. + Outside minor and major heap, x must be black. +*/ +static inline int Is_Dead_during_clean(value x){ + CAMLassert (x != caml_ephe_none); + CAMLassert (caml_gc_phase == Phase_clean); + return Is_block (x) && !Is_young (x) && Is_white_val(x); +} +/** The minor heap doesn't have to be marked, outside they should + already be black +*/ +static inline int Must_be_Marked_during_mark(value x){ + CAMLassert (x != caml_ephe_none); + CAMLassert (caml_gc_phase == Phase_mark); + return Is_block (x) && !Is_young (x); +} +#else +static inline int Is_Dead_during_clean(value x){ + CAMLassert (x != caml_ephe_none); + CAMLassert (caml_gc_phase == Phase_clean); + return Is_block (x) && Is_in_heap (x) && Is_white_val(x); +} +static inline int Must_be_Marked_during_mark(value x){ + CAMLassert (x != caml_ephe_none); + CAMLassert (caml_gc_phase == Phase_mark); + return Is_block (x) && Is_in_heap (x); +} +#endif + + +/* [len] is a value that represents a number of words (fields) */ +CAMLprim value caml_ephe_create (value len) +{ + mlsize_t size, i; + value res; + + size = Long_val (len) + CAML_EPHE_FIRST_KEY; + if (size < CAML_EPHE_FIRST_KEY || size > Max_wosize) + caml_invalid_argument ("Weak.create"); + res = caml_alloc_shr (size, Abstract_tag); + for (i = 1; i < size; i++) Field (res, i) = caml_ephe_none; + Field (res, CAML_EPHE_LINK_OFFSET) = caml_ephe_list_head; + caml_ephe_list_head = res; + return res; +} + +CAMLprim value caml_weak_create (value len) +{ + return caml_ephe_create(len); +} + +/** + Specificity of the cleaning phase (Phase_clean): + + The dead keys must be removed from the ephemerons and data removed + when one the keys is dead. Here we call it cleaning the ephemerons. + A specific phase of the GC is dedicated to this, Phase_clean. This + phase is just after the mark phase, so the white values are dead + values. It iterates the function caml_ephe_clean through all the + ephemerons. + + However the GC is incremental and ocaml code can run on the middle + of this cleaning phase. In order to respect the semantic of the + ephemerons concerning dead values, the getter and setter must work + as if the cleaning of all the ephemerons have been done at once. + + - key getter: Even if a dead key have not yet been replaced by + caml_ephe_none, getting it should return none. + - key setter: If we replace a dead key we need to set the data to + caml_ephe_none and clean the ephemeron. + + This two cases are dealt by a call to do_check_key_clean that + trigger the cleaning of the ephemerons when the accessed key is + dead. This test is fast. + + In the case of value getter and value setter, there is no fast + test because the removing of the data depend of the deadliness of the keys. + We must always try to clean the ephemerons. + + */ + +#define None_val (Val_int(0)) +#define Some_tag 0 + +/* If we are in Phase_clean we need to check if the key + that is going to disappear is dead and so should trigger a cleaning + */ +static void do_check_key_clean(value ar, mlsize_t offset){ + CAMLassert (offset >= CAML_EPHE_FIRST_KEY); + if (caml_gc_phase == Phase_clean){ + value elt = Field (ar, offset); + if (elt != caml_ephe_none && Is_Dead_during_clean(elt)){ + Field(ar,offset) = caml_ephe_none; + Field(ar,CAML_EPHE_DATA_OFFSET) = caml_ephe_none; + }; + }; +} + +/* If we are in Phase_clean we need to do as if the key is empty when + it will be cleaned during this phase */ +static inline int is_ephe_key_none(value ar, mlsize_t offset){ + value elt = Field (ar, offset); + if (elt == caml_ephe_none){ + return 1; + }else if (caml_gc_phase == Phase_clean && Is_Dead_during_clean(elt)){ + Field(ar,offset) = caml_ephe_none; + Field(ar,CAML_EPHE_DATA_OFFSET) = caml_ephe_none; + return 1; + } else { + return 0; + } +} + + +static void do_set (value ar, mlsize_t offset, value v) +{ + if (Is_block (v) && Is_young (v)){ + /* modified version of Modify */ + value old = Field (ar, offset); + Field (ar, offset) = v; + if (!(Is_block (old) && Is_young (old))){ + add_to_ephe_ref_table (&caml_ephe_ref_table, ar, offset); + } + }else{ + Field (ar, offset) = v; + } +} + +CAMLprim value caml_ephe_set_key (value ar, value n, value el) +{ + mlsize_t offset = Long_val (n) + CAML_EPHE_FIRST_KEY; + CAMLassert (Is_in_heap (ar)); + if (offset < CAML_EPHE_FIRST_KEY || offset >= Wosize_val (ar)){ + caml_invalid_argument ("Weak.set"); + } + do_check_key_clean(ar,offset); + do_set (ar, offset, el); + return Val_unit; +} + +CAMLprim value caml_ephe_unset_key (value ar, value n) +{ + mlsize_t offset = Long_val (n) + CAML_EPHE_FIRST_KEY; + CAMLassert (Is_in_heap (ar)); + if (offset < CAML_EPHE_FIRST_KEY || offset >= Wosize_val (ar)){ + caml_invalid_argument ("Weak.set"); + } + do_check_key_clean(ar,offset); + Field (ar, offset) = caml_ephe_none; + return Val_unit; +} + +value caml_ephe_set_key_option (value ar, value n, value el) +{ + mlsize_t offset = Long_val (n) + CAML_EPHE_FIRST_KEY; + CAMLassert (Is_in_heap (ar)); + if (offset < CAML_EPHE_FIRST_KEY || offset >= Wosize_val (ar)){ + caml_invalid_argument ("Weak.set"); + } + do_check_key_clean(ar,offset); + if (el != None_val && Is_block (el)){ + CAMLassert (Wosize_val (el) == 1); + do_set (ar, offset, Field (el, 0)); + }else{ + Field (ar, offset) = caml_ephe_none; + } + return Val_unit; +} + +CAMLprim value caml_weak_set (value ar, value n, value el){ + return caml_ephe_set_key_option(ar,n,el); +} + +CAMLprim value caml_ephe_set_data (value ar, value el) +{ + CAMLassert (Is_in_heap (ar)); + if (caml_gc_phase == Phase_clean){ + /* During this phase since we don't know which ephemeron have been + cleaned we always need to check it. */ + caml_ephe_clean(ar); + }; + do_set (ar, CAML_EPHE_DATA_OFFSET, el); + return Val_unit; +} + +CAMLprim value caml_ephe_unset_data (value ar) +{ + CAMLassert (Is_in_heap (ar)); + Field (ar, CAML_EPHE_DATA_OFFSET) = caml_ephe_none; + return Val_unit; +} + +CAMLprim value caml_ephe_get_key (value ar, value n) +{ + CAMLparam2 (ar, n); + mlsize_t offset = Long_val (n) + CAML_EPHE_FIRST_KEY; + CAMLlocal2 (res, elt); + CAMLassert (Is_in_heap (ar)); + if (offset < CAML_EPHE_FIRST_KEY || offset >= Wosize_val (ar)){ + caml_invalid_argument ("Weak.get_key"); + } + if (is_ephe_key_none(ar, offset)){ + res = None_val; + }else{ + elt = Field (ar, offset); + if (caml_gc_phase == Phase_mark && Must_be_Marked_during_mark(elt)){ + caml_darken (elt, NULL); + } + res = caml_alloc_small (1, Some_tag); + Field (res, 0) = elt; + } + CAMLreturn (res); +} + +CAMLprim value caml_weak_get (value ar, value n){ + return caml_ephe_get_key(ar, n); +} + +CAMLprim value caml_ephe_get_data (value ar) +{ + CAMLparam1 (ar); + CAMLlocal2 (res, elt); + CAMLassert (Is_in_heap (ar)); + elt = Field (ar, CAML_EPHE_DATA_OFFSET); + if(caml_gc_phase == Phase_clean) caml_ephe_clean(ar); + if (elt == caml_ephe_none){ + res = None_val; + }else{ + if (caml_gc_phase == Phase_mark && Must_be_Marked_during_mark(elt)){ + caml_darken (elt, NULL); + } + res = caml_alloc_small (1, Some_tag); + Field (res, 0) = elt; + } + CAMLreturn (res); +} + +CAMLprim value caml_ephe_get_key_copy (value ar, value n) +{ + CAMLparam2 (ar, n); + mlsize_t offset = Long_val (n) + CAML_EPHE_FIRST_KEY; + CAMLlocal2 (res, elt); + value v; /* Caution: this is NOT a local root. */ + CAMLassert (Is_in_heap (ar)); + if (offset < CAML_EPHE_FIRST_KEY || offset >= Wosize_val (ar)){ + caml_invalid_argument ("Weak.get_copy"); + } + + if (is_ephe_key_none(ar, offset)) CAMLreturn (None_val); + v = Field (ar, offset); + /** Don't copy custom_block #7279 */ + if (Is_block (v) && Is_in_heap_or_young(v) && Tag_val(v) != Custom_tag ) { + elt = caml_alloc (Wosize_val (v), Tag_val (v)); + /* The GC may erase or move v during this call to caml_alloc. */ + v = Field (ar, offset); + if (is_ephe_key_none(ar, offset)) CAMLreturn (None_val); + if (Tag_val (v) < No_scan_tag){ + mlsize_t i; + for (i = 0; i < Wosize_val (v); i++){ + value f = Field (v, i); + if (caml_gc_phase == Phase_mark && Must_be_Marked_during_mark(f)){ + caml_darken (f, NULL); + } + Modify (&Field (elt, i), f); + } + }else{ + memmove (Bp_val (elt), Bp_val (v), Bosize_val (v)); + } + }else{ + if ( caml_gc_phase == Phase_mark && Must_be_Marked_during_mark(v) ){ + caml_darken (v, NULL); + }; + elt = v; + } + res = caml_alloc_small (1, Some_tag); + Field (res, 0) = elt; + + CAMLreturn (res); +} + +CAMLprim value caml_weak_get_copy (value ar, value n){ + return caml_ephe_get_key_copy(ar,n); +} + +CAMLprim value caml_ephe_get_data_copy (value ar) +{ + CAMLparam1 (ar); + mlsize_t offset = CAML_EPHE_DATA_OFFSET; + CAMLlocal2 (res, elt); + value v; /* Caution: this is NOT a local root. */ + CAMLassert (Is_in_heap (ar)); + + v = Field (ar, offset); + if (caml_gc_phase == Phase_clean) caml_ephe_clean(ar); + if (v == caml_ephe_none) CAMLreturn (None_val); + /** Don't copy custom_block #7279 */ + if (Is_block (v) && Is_in_heap_or_young(v) && Tag_val(v) != Custom_tag ) { + elt = caml_alloc (Wosize_val (v), Tag_val (v)); + /* The GC may erase or move v during this call to caml_alloc. */ + v = Field (ar, offset); + if (caml_gc_phase == Phase_clean) caml_ephe_clean(ar); + if (v == caml_ephe_none) CAMLreturn (None_val); + if (Tag_val (v) < No_scan_tag){ + mlsize_t i; + for (i = 0; i < Wosize_val (v); i++){ + value f = Field (v, i); + if (caml_gc_phase == Phase_mark && Must_be_Marked_during_mark(f)){ + caml_darken (f, NULL); + } + Modify (&Field (elt, i), f); + } + }else{ + memmove (Bp_val (elt), Bp_val (v), Bosize_val (v)); + } + }else{ + if ( caml_gc_phase == Phase_mark && Must_be_Marked_during_mark(v) ){ + caml_darken (v, NULL); + }; + elt = v; + } + res = caml_alloc_small (1, Some_tag); + Field (res, 0) = elt; + + CAMLreturn (res); +} + +CAMLprim value caml_ephe_check_key (value ar, value n) +{ + mlsize_t offset = Long_val (n) + CAML_EPHE_FIRST_KEY; + CAMLassert (Is_in_heap (ar)); + if (offset < CAML_EPHE_FIRST_KEY || offset >= Wosize_val (ar)){ + caml_invalid_argument ("Weak.check"); + } + return Val_bool (!is_ephe_key_none(ar, offset)); +} + +CAMLprim value caml_weak_check (value ar, value n) +{ + return caml_ephe_check_key(ar,n); +} + +CAMLprim value caml_ephe_check_data (value ar) +{ + if(caml_gc_phase == Phase_clean) caml_ephe_clean(ar); + return Val_bool (Field (ar, CAML_EPHE_DATA_OFFSET) != caml_ephe_none); +} + +CAMLprim value caml_ephe_blit_key (value ars, value ofs, + value ard, value ofd, value len) +{ + mlsize_t offset_s = Long_val (ofs) + CAML_EPHE_FIRST_KEY; + mlsize_t offset_d = Long_val (ofd) + CAML_EPHE_FIRST_KEY; + mlsize_t length = Long_val (len); + long i; + CAMLassert (Is_in_heap (ars)); + CAMLassert (Is_in_heap (ard)); + if (offset_s < CAML_EPHE_FIRST_KEY || offset_s + length > Wosize_val (ars)){ + caml_invalid_argument ("Weak.blit"); + } + if (offset_d < CAML_EPHE_FIRST_KEY || offset_d + length > Wosize_val (ard)){ + caml_invalid_argument ("Weak.blit"); + } + if (caml_gc_phase == Phase_clean){ + caml_ephe_clean(ars); + caml_ephe_clean(ard); + } + if (offset_d < offset_s){ + for (i = 0; i < length; i++){ + do_set (ard, offset_d + i, Field (ars, offset_s + i)); + } + }else{ + for (i = length - 1; i >= 0; i--){ + do_set (ard, offset_d + i, Field (ars, offset_s + i)); + } + } + return Val_unit; +} + +CAMLprim value caml_ephe_blit_data (value ars, value ard) +{ + if(caml_gc_phase == Phase_clean) { + caml_ephe_clean(ars); + caml_ephe_clean(ard); + }; + do_set (ard, CAML_EPHE_DATA_OFFSET, Field (ars, CAML_EPHE_DATA_OFFSET)); + return Val_unit; +} + +CAMLprim value caml_weak_blit (value ars, value ofs, + value ard, value ofd, value len) +{ + return caml_ephe_blit_key (ars, ofs, ard, ofd, len); +} diff --git a/test/monniaux/ocaml/byterun/win32.c b/test/monniaux/ocaml/byterun/win32.c new file mode 100644 index 00000000..1ce8ad5e --- /dev/null +++ b/test/monniaux/ocaml/byterun/win32.c @@ -0,0 +1,1019 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +#define CAML_INTERNALS + +/* Win32-specific stuff */ + +/* FILE_INFO_BY_HANDLE_CLASS and FILE_NAME_INFO are only available from Windows + Vista onwards */ +#undef _WIN32_WINNT +#define _WIN32_WINNT 0x0600 + +#define WIN32_LEAN_AND_MEAN +#include +#include +#include +#include +#include +#include +#include +#include +#include +#include +#include +#include +#include +#include +#include "caml/alloc.h" +#include "caml/address_class.h" +#include "caml/fail.h" +#include "caml/io.h" +#include "caml/memory.h" +#include "caml/misc.h" +#include "caml/osdeps.h" +#include "caml/signals.h" +#include "caml/sys.h" + +#include "caml/config.h" +#ifdef SUPPORT_DYNAMIC_LINKING +#include +#endif + +#ifndef S_ISREG +#define S_ISREG(mode) (((mode) & S_IFMT) == S_IFREG) +#endif + +unsigned short caml_win32_major = 0; +unsigned short caml_win32_minor = 0; +unsigned short caml_win32_build = 0; +unsigned short caml_win32_revision = 0; + +CAMLnoreturn_start +static void caml_win32_sys_error (int errnum) +CAMLnoreturn_end; + +static void caml_win32_sys_error(int errnum) +{ + wchar_t buffer[512]; + value msg; + if (FormatMessage(FORMAT_MESSAGE_FROM_SYSTEM | FORMAT_MESSAGE_IGNORE_INSERTS, + NULL, + errnum, + 0, + buffer, + sizeof(buffer)/sizeof(wchar_t), + NULL)) { + msg = caml_copy_string_of_utf16(buffer); + } else { + msg = caml_alloc_sprintf("unknown error #%d", errnum); + } + caml_raise_sys_error(msg); +} + +int caml_read_fd(int fd, int flags, void * buf, int n) +{ + int retcode; + if ((flags & CHANNEL_FLAG_FROM_SOCKET) == 0) { + caml_enter_blocking_section(); + retcode = read(fd, buf, n); + /* Large reads from console can fail with ENOMEM. Reduce requested size + and try again. */ + if (retcode == -1 && errno == ENOMEM && n > 16384) { + retcode = read(fd, buf, 16384); + } + caml_leave_blocking_section(); + if (retcode == -1) caml_sys_io_error(NO_ARG); + } else { + caml_enter_blocking_section(); + retcode = recv((SOCKET) _get_osfhandle(fd), buf, n, 0); + caml_leave_blocking_section(); + if (retcode == -1) caml_win32_sys_error(WSAGetLastError()); + } + return retcode; +} + +int caml_write_fd(int fd, int flags, void * buf, int n) +{ + int retcode; + if ((flags & CHANNEL_FLAG_FROM_SOCKET) == 0) { +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + if (flags & CHANNEL_FLAG_BLOCKING_WRITE) { + retcode = write(fd, buf, n); + } else { +#endif + caml_enter_blocking_section(); + retcode = write(fd, buf, n); + caml_leave_blocking_section(); +#if defined(NATIVE_CODE) && defined(WITH_SPACETIME) + } +#endif + if (retcode == -1) caml_sys_io_error(NO_ARG); + } else { + caml_enter_blocking_section(); + retcode = send((SOCKET) _get_osfhandle(fd), buf, n, 0); + caml_leave_blocking_section(); + if (retcode == -1) caml_win32_sys_error(WSAGetLastError()); + } + CAMLassert (retcode > 0); + return retcode; +} + +wchar_t * caml_decompose_path(struct ext_table * tbl, wchar_t * path) +{ + wchar_t * p, * q; + int n; + + if (path == NULL) return NULL; + p = caml_stat_wcsdup(path); + q = p; + while (1) { + for (n = 0; q[n] != 0 && q[n] != L';'; n++) /*nothing*/; + caml_ext_table_add(tbl, q); + q = q + n; + if (*q == 0) break; + *q = 0; + q += 1; + } + return p; +} + +wchar_t * caml_search_in_path(struct ext_table * path, const wchar_t * name) +{ + wchar_t * dir, * fullname; + char * u8; + const wchar_t * p; + int i; + struct _stati64 st; + + for (p = name; *p != 0; p++) { + if (*p == '/' || *p == '\\') goto not_found; + } + for (i = 0; i < path->size; i++) { + dir = path->contents[i]; + if (dir[0] == 0) continue; + /* not sure what empty path components mean under Windows */ + fullname = caml_stat_wcsconcat(3, dir, L"\\", name); + u8 = caml_stat_strdup_of_utf16(fullname); + caml_gc_message(0x100, "Searching %s\n", u8); + caml_stat_free(u8); + if (_wstati64(fullname, &st) == 0 && S_ISREG(st.st_mode)) + return fullname; + caml_stat_free(fullname); + } + not_found: + u8 = caml_stat_strdup_of_utf16(name); + caml_gc_message(0x100, "%s not found in search path\n", u8); + caml_stat_free(u8); + return caml_stat_wcsdup(name); +} + +CAMLexport wchar_t * caml_search_exe_in_path(const wchar_t * name) +{ + wchar_t * fullname, * filepart; + char * u8; + size_t fullnamelen; + DWORD retcode; + + fullnamelen = wcslen(name) + 1; + if (fullnamelen < 256) fullnamelen = 256; + while (1) { + fullname = caml_stat_alloc(fullnamelen*sizeof(wchar_t)); + retcode = SearchPath(NULL, /* use system search path */ + name, + L".exe", /* add .exe extension if needed */ + fullnamelen, + fullname, + &filepart); + if (retcode == 0) { + u8 = caml_stat_strdup_of_utf16(name); + caml_gc_message(0x100, "%s not found in search path\n", u8); + caml_stat_free(u8); + caml_stat_free(fullname); + return caml_stat_strdup_os(name); + } + if (retcode < fullnamelen) + return fullname; + caml_stat_free(fullname); + fullnamelen = retcode + 1; + } +} + +wchar_t * caml_search_dll_in_path(struct ext_table * path, const wchar_t * name) +{ + wchar_t * dllname; + wchar_t * res; + + dllname = caml_stat_wcsconcat(2, name, L".dll"); + res = caml_search_in_path(path, dllname); + caml_stat_free(dllname); + return res; +} + +#ifdef SUPPORT_DYNAMIC_LINKING + +void * caml_dlopen(wchar_t * libname, int for_execution, int global) +{ + void *handle; + int flags = (global ? FLEXDLL_RTLD_GLOBAL : 0); + if (!for_execution) flags |= FLEXDLL_RTLD_NOEXEC; + handle = flexdll_wdlopen(libname, flags); + if ((handle != NULL) && ((caml_verb_gc & 0x100) != 0)) { + flexdll_dump_exports(handle); + fflush(stdout); + } + return handle; +} + +void caml_dlclose(void * handle) +{ + flexdll_dlclose(handle); +} + +void * caml_dlsym(void * handle, const char * name) +{ + return flexdll_dlsym(handle, name); +} + +void * caml_globalsym(const char * name) +{ + return flexdll_dlsym(flexdll_dlopen(NULL,0), name); +} + +char * caml_dlerror(void) +{ + return flexdll_dlerror(); +} + +#else + +void * caml_dlopen(wchar_t * libname, int for_execution, int global) +{ + return NULL; +} + +void caml_dlclose(void * handle) +{ +} + +void * caml_dlsym(void * handle, const char * name) +{ + return NULL; +} + +void * caml_globalsym(const char * name) +{ + return NULL; +} + +char * caml_dlerror(void) +{ + return "dynamic loading not supported on this platform"; +} + +#endif + +/* Proper emulation of signal(), including ctrl-C and ctrl-break */ + +typedef void (*sighandler)(int sig); +static int ctrl_handler_installed = 0; +static volatile sighandler ctrl_handler_action = SIG_DFL; + +static BOOL WINAPI ctrl_handler(DWORD event) +{ + /* Only ctrl-C and ctrl-Break are handled */ + if (event != CTRL_C_EVENT && event != CTRL_BREAK_EVENT) return FALSE; + /* Default behavior is to exit, which we get by not handling the event */ + if (ctrl_handler_action == SIG_DFL) return FALSE; + /* Ignore behavior is to do nothing, which we get by claiming that we + have handled the event */ + if (ctrl_handler_action == SIG_IGN) return TRUE; + /* Win32 doesn't like it when we do a longjmp() at this point + (it looks like we're running in a different thread than + the main program!). So, just record the signal. */ + caml_record_signal(SIGINT); + /* We have handled the event */ + return TRUE; +} + +sighandler caml_win32_signal(int sig, sighandler action) +{ + sighandler oldaction; + + if (sig != SIGINT) return signal(sig, action); + if (! ctrl_handler_installed) { + SetConsoleCtrlHandler(ctrl_handler, TRUE); + ctrl_handler_installed = 1; + } + oldaction = ctrl_handler_action; + ctrl_handler_action = action; + return oldaction; +} + +/* Expansion of @responsefile and *? file patterns in the command line */ + +static int argc; +static wchar_t ** argv; +static int argvsize; + +static void store_argument(wchar_t * arg); +static void expand_argument(wchar_t * arg); +static void expand_pattern(wchar_t * arg); + +static void out_of_memory(void) +{ + fprintf(stderr, "Out of memory while expanding command line\n"); + exit(2); +} + +static void store_argument(wchar_t * arg) +{ + if (argc + 1 >= argvsize) { + argvsize *= 2; + argv = (wchar_t **) caml_stat_resize_noexc(argv, argvsize * sizeof(wchar_t *)); + if (argv == NULL) out_of_memory(); + } + argv[argc++] = arg; +} + +static void expand_argument(wchar_t * arg) +{ + wchar_t * p; + + for (p = arg; *p != 0; p++) { + if (*p == L'*' || *p == L'?') { + expand_pattern(arg); + return; + } + } + store_argument(arg); +} + +static void expand_pattern(wchar_t * pat) +{ + wchar_t * prefix, * p, * name; + intptr_t handle; + struct _wfinddata_t ffblk; + size_t i; + + handle = _wfindfirst(pat, &ffblk); + if (handle == -1) { + store_argument(pat); /* a la Bourne shell */ + return; + } + prefix = caml_stat_wcsdup(pat); + /* We need to stop at the first directory or drive boundary, because the + * _findata_t structure contains the filename, not the leading directory. */ + for (i = wcslen(prefix); i > 0; i--) { + wchar_t c = prefix[i - 1]; + if (c == L'\\' || c == L'/' || c == L':') { prefix[i] = 0; break; } + } + /* No separator was found, it's a filename pattern without a leading directory. */ + if (i == 0) + prefix[0] = 0; + do { + name = caml_stat_wcsconcat(2, prefix, ffblk.name); + store_argument(name); + } while (_wfindnext(handle, &ffblk) != -1); + _findclose(handle); + caml_stat_free(prefix); +} + + +CAMLexport void caml_expand_command_line(int * argcp, wchar_t *** argvp) +{ + int i; + argc = 0; + argvsize = 16; + argv = (wchar_t **) caml_stat_alloc_noexc(argvsize * sizeof(wchar_t *)); + if (argv == NULL) out_of_memory(); + for (i = 0; i < *argcp; i++) expand_argument((*argvp)[i]); + argv[argc] = NULL; + *argcp = argc; + *argvp = argv; +} + +/* Add to [contents] the (short) names of the files contained in + the directory named [dirname]. No entries are added for [.] and [..]. + Return 0 on success, -1 on error; set errno in the case of error. */ + +int caml_read_directory(wchar_t * dirname, struct ext_table * contents) +{ + size_t dirnamelen; + wchar_t * template; + intptr_t h; + struct _wfinddata_t fileinfo; + + dirnamelen = wcslen(dirname); + if (dirnamelen > 0 && + (dirname[dirnamelen - 1] == L'/' + || dirname[dirnamelen - 1] == L'\\' + || dirname[dirnamelen - 1] == L':')) + template = caml_stat_wcsconcat(2, dirname, L"*.*"); + else + template = caml_stat_wcsconcat(2, dirname, L"\\*.*"); + h = _wfindfirst(template, &fileinfo); + if (h == -1) { + caml_stat_free(template); + return errno == ENOENT ? 0 : -1; + } + do { + if (wcscmp(fileinfo.name, L".") != 0 && wcscmp(fileinfo.name, L"..") != 0) { + caml_ext_table_add(contents, caml_stat_strdup_of_utf16(fileinfo.name)); + } + } while (_wfindnext(h, &fileinfo) == 0); + _findclose(h); + caml_stat_free(template); + return 0; +} + +#ifndef NATIVE_CODE + +/* Set up a new thread for control-C emulation and termination */ + +void caml_signal_thread(void * lpParam) +{ + wchar_t *endptr; + HANDLE h; + /* Get an hexa-code raw handle through the environment */ + h = (HANDLE) (uintptr_t) + wcstol(caml_secure_getenv(_T("CAMLSIGPIPE")), &endptr, 16); + while (1) { + DWORD numread; + BOOL ret; + char iobuf[2]; + /* This shall always return a single character */ + ret = ReadFile(h, iobuf, 1, &numread, NULL); + if (!ret || numread != 1) caml_sys_exit(Val_int(2)); + switch (iobuf[0]) { + case 'C': + caml_record_signal(SIGINT); + break; + case 'T': + raise(SIGTERM); + return; + } + } +} + +#endif /* NATIVE_CODE */ + +#if defined(NATIVE_CODE) + +/* Handling of system stack overflow. + * Based on code provided by Olivier Andrieu. + + * An EXCEPTION_STACK_OVERFLOW is signaled when the guard page at the + * end of the stack has been accessed. Windows clears the PAGE_GUARD + * protection (making it a regular PAGE_READWRITE) and then calls our + * exception handler. This means that although we're handling an "out + * of stack" condition, there is a bit of stack available to call + * functions and allocate temporaries. + * + * PAGE_GUARD is a one-shot access protection mechanism: we need to + * restore the PAGE_GUARD protection on this page otherwise the next + * stack overflow won't be detected and the program will abruptly exit + * with STATUS_ACCESS_VIOLATION. + * + * Visual Studio 2003 and later (_MSC_VER >= 1300) have a + * _resetstkoflw() function that resets this protection. + * Unfortunately, it cannot work when called directly from the + * exception handler because at this point we are using the page that + * is to be protected. + * + * A solution is to use an alternate stack when restoring the + * protection. However it's not possible to use _resetstkoflw() then + * since it determines the stack pointer by calling alloca(): it would + * try to protect the alternate stack. + * + * Finally, we call caml_raise_stack_overflow; it will either call + * caml_raise_exception which switches back to the normal stack, or + * call caml_fatal_uncaught_exception which terminates the program + * quickly. + */ + +static uintnat win32_alt_stack[0x100]; + +static void caml_reset_stack (void *faulting_address) +{ + SYSTEM_INFO si; + DWORD page_size; + MEMORY_BASIC_INFORMATION mbi; + DWORD oldprot; + + /* get the system's page size. */ + GetSystemInfo (&si); + page_size = si.dwPageSize; + + /* get some information on the page the fault occurred */ + if (! VirtualQuery (faulting_address, &mbi, sizeof mbi)) + goto failed; + + VirtualProtect (mbi.BaseAddress, page_size, + mbi.Protect | PAGE_GUARD, &oldprot); + + failed: + caml_raise_stack_overflow(); +} + + +#ifndef _WIN64 +static LONG CALLBACK + caml_stack_overflow_VEH (EXCEPTION_POINTERS* exn_info) +{ + DWORD code = exn_info->ExceptionRecord->ExceptionCode; + CONTEXT *ctx = exn_info->ContextRecord; + DWORD *ctx_ip = &(ctx->Eip); + DWORD *ctx_sp = &(ctx->Esp); + + if (code == EXCEPTION_STACK_OVERFLOW && Is_in_code_area (*ctx_ip)) + { + uintnat faulting_address; + uintnat * alt_esp; + + /* grab the address that caused the fault */ + faulting_address = exn_info->ExceptionRecord->ExceptionInformation[1]; + + /* call caml_reset_stack(faulting_address) using the alternate stack */ + alt_esp = win32_alt_stack + sizeof(win32_alt_stack) / sizeof(uintnat); + *--alt_esp = faulting_address; + *ctx_sp = (uintnat) (alt_esp - 1); + *ctx_ip = (uintnat) &caml_reset_stack; + + return EXCEPTION_CONTINUE_EXECUTION; + } + + return EXCEPTION_CONTINUE_SEARCH; +} + +#else +extern char *caml_exception_pointer; +extern value *caml_young_ptr; + +/* Do not use the macro from address_class.h here. */ +#undef Is_in_code_area +#define Is_in_code_area(pc) \ + ( ((char *)(pc) >= caml_code_area_start && \ + (char *)(pc) <= caml_code_area_end) \ +|| ((char *)(pc) >= &caml_system__code_begin && \ + (char *)(pc) <= &caml_system__code_end) \ +|| (Classify_addr(pc) & In_code_area) ) +extern char caml_system__code_begin, caml_system__code_end; + + +static LONG CALLBACK + caml_stack_overflow_VEH (EXCEPTION_POINTERS* exn_info) +{ + DWORD code = exn_info->ExceptionRecord->ExceptionCode; + CONTEXT *ctx = exn_info->ContextRecord; + + if (code == EXCEPTION_STACK_OVERFLOW && Is_in_code_area (ctx->Rip)) + { + uintnat faulting_address; + uintnat * alt_rsp; + + /* grab the address that caused the fault */ + faulting_address = exn_info->ExceptionRecord->ExceptionInformation[1]; + + /* refresh runtime parameters from registers */ + caml_exception_pointer = (char *) ctx->R14; + caml_young_ptr = (value *) ctx->R15; + + /* call caml_reset_stack(faulting_address) using the alternate stack */ + alt_rsp = win32_alt_stack + sizeof(win32_alt_stack) / sizeof(uintnat); + ctx->Rcx = faulting_address; + ctx->Rsp = (uintnat) (alt_rsp - 4 - 1); + ctx->Rip = (uintnat) &caml_reset_stack; + + return EXCEPTION_CONTINUE_EXECUTION; + } + + return EXCEPTION_CONTINUE_SEARCH; +} +#endif /* _WIN64 */ + +void caml_win32_overflow_detection(void) +{ + AddVectoredExceptionHandler(1, caml_stack_overflow_VEH); +} + +#endif /* NATIVE_CODE */ + +/* Seeding of pseudo-random number generators */ + +int caml_win32_random_seed (intnat data[16]) +{ + /* For better randomness, consider: + http://msdn.microsoft.com/library/en-us/seccrypto/security/rtlgenrandom.asp + http://blogs.msdn.com/b/michael_howard/archive/2005/01/14/353379.aspx + */ + FILETIME t; + LARGE_INTEGER pc; + GetSystemTimeAsFileTime(&t); + QueryPerformanceCounter(&pc); /* PR#6032 */ + data[0] = t.dwLowDateTime; + data[1] = t.dwHighDateTime; + data[2] = GetCurrentProcessId(); + data[3] = pc.LowPart; + data[4] = pc.HighPart; + return 5; +} + + +#if defined(_MSC_VER) && __STDC_SECURE_LIB__ >= 200411L + +static void invalid_parameter_handler(const wchar_t* expression, + const wchar_t* function, + const wchar_t* file, + unsigned int line, + uintptr_t pReserved) +{ + /* no crash box */ +} + + +void caml_install_invalid_parameter_handler() +{ + _set_invalid_parameter_handler(invalid_parameter_handler); +} + +#endif + + +/* Recover executable name */ + +wchar_t * caml_executable_name(void) +{ + wchar_t * name; + DWORD namelen, ret; + + namelen = 256; + while (1) { + name = caml_stat_alloc(namelen*sizeof(wchar_t)); + ret = GetModuleFileName(NULL, name, namelen); + if (ret == 0) { caml_stat_free(name); return NULL; } + if (ret < namelen) break; + caml_stat_free(name); + if (namelen >= 1024*1024) return NULL; /* avoid runaway and overflow */ + namelen *= 2; + } + return name; +} + +/* snprintf emulation */ + +#ifdef LACKS_VSCPRINTF +/* No _vscprintf until Visual Studio .NET 2002 and sadly no version number + in the CRT headers until Visual Studio 2005 so forced to predicate this + on the compiler version instead */ +int _vscprintf(const char * format, va_list args) +{ + int n; + int sz = 5; + char* buf = (char*)malloc(sz); + n = _vsnprintf(buf, sz, format, args); + while (n < 0 || n > sz) { + sz += 512; + buf = (char*)realloc(buf, sz); + n = _vsnprintf(buf, sz, format, args); + } + free(buf); + return n; +} +#endif + +#if defined(_WIN32) && !defined(_UCRT) +int caml_snprintf(char * buf, size_t size, const char * format, ...) +{ + int len; + va_list args; + + if (size > 0) { + va_start(args, format); + len = _vsnprintf(buf, size, format, args); + va_end(args); + if (len >= 0 && len < size) { + /* [len] characters were stored in [buf], + a null-terminator was appended. */ + return len; + } + /* [size] characters were stored in [buf], without null termination. + Put a null terminator, truncating the output. */ + buf[size - 1] = 0; + } + /* Compute the actual length of output, excluding null terminator */ + va_start(args, format); + len = _vscprintf(format, args); + va_end(args); + return len; +} +#endif + +wchar_t *caml_secure_getenv (wchar_t const *var) +{ + /* Win32 doesn't have a notion of setuid bit, so getenv is safe. */ + return _wgetenv(var); +} + +/* caml_win32_getenv is used to implement Sys.getenv and Unix.getenv in such a + way that they get direct access to the Win32 environment rather than to the + copy that is cached by the C runtime system. The result of caml_win32_getenv + is dynamically allocated and must be explicitly deallocated. + + In contrast, the OCaml runtime system still calls _wgetenv from the C runtime + system, via caml_secure_getenv. The result is statically allocated and needs + no deallocation. */ +CAMLexport wchar_t *caml_win32_getenv(wchar_t const *lpName) +{ + wchar_t * lpBuffer; + DWORD nSize = 256, res; + + lpBuffer = caml_stat_alloc_noexc(nSize * sizeof(wchar_t)); + + if (lpBuffer == NULL) + return NULL; + + res = GetEnvironmentVariable(lpName, lpBuffer, nSize); + + if (res == 0) { + caml_stat_free(lpBuffer); + return NULL; + } + + if (res < nSize) + return lpBuffer; + + nSize = res; + lpBuffer = caml_stat_resize_noexc(lpBuffer, nSize * sizeof(wchar_t)); + + if (lpBuffer == NULL) + return NULL; + + res = GetEnvironmentVariable(lpName, lpBuffer, nSize); + + if (res == 0 || res >= nSize) { + caml_stat_free(lpBuffer); + return NULL; + } + + return lpBuffer; +} + +/* The rename() implementation in MSVC's CRT is based on MoveFile() + and therefore fails if the new name exists. This is inconsistent + with POSIX and a problem in practice. Here we reimplement + rename() using MoveFileEx() to make it more POSIX-like. + There are no official guarantee that the rename operation is atomic, + but it is widely believed to be atomic on NTFS. */ + +int caml_win32_rename(const wchar_t * oldpath, const wchar_t * newpath) +{ + /* MOVEFILE_REPLACE_EXISTING: to be closer to POSIX + MOVEFILE_COPY_ALLOWED: MoveFile performs a copy if old and new + paths are on different devices, so we do the same here for + compatibility with the old rename()-based implementation. + MOVEFILE_WRITE_THROUGH: not sure it's useful; affects only + the case where a copy is done. */ + if (MoveFileEx(oldpath, newpath, + MOVEFILE_REPLACE_EXISTING | MOVEFILE_WRITE_THROUGH | + MOVEFILE_COPY_ALLOWED)) { + return 0; + } + /* Modest attempt at mapping Win32 error codes to POSIX error codes. + The __dosmaperr() function from the CRT does a better job but is + generally not accessible. */ + switch (GetLastError()) { + case ERROR_FILE_NOT_FOUND: case ERROR_PATH_NOT_FOUND: + errno = ENOENT; break; + case ERROR_ACCESS_DENIED: case ERROR_WRITE_PROTECT: case ERROR_CANNOT_MAKE: + errno = EACCES; break; + case ERROR_CURRENT_DIRECTORY: case ERROR_BUSY: + errno = EBUSY; break; + case ERROR_NOT_SAME_DEVICE: + errno = EXDEV; break; + case ERROR_ALREADY_EXISTS: + errno = EEXIST; break; + default: + errno = EINVAL; + } + return -1; +} + +/* Windows Unicode support */ +static uintnat windows_unicode_enabled = WINDOWS_UNICODE; + +/* If [windows_unicode_strict] is non-zero, then illegal UTF-8 characters (on + the OCaml side) or illegal UTF-16 characters (on the Windows side) cause an + error to be signaled. What happens then depends on the variable + [windows_unicode_fallback]. + + If [windows_unicode_strict] is zero, then illegal characters are silently + dropped. */ +static uintnat windows_unicode_strict = 1; + +/* If [windows_unicode_fallback] is non-zero, then if an error is signaled when + translating to UTF-16, the translation is re-done under the assumption that + the argument string is encoded in the local codepage. */ +static uintnat windows_unicode_fallback = 1; + +CAMLexport int win_multi_byte_to_wide_char(const char *s, int slen, wchar_t *out, int outlen) +{ + int retcode; + + CAMLassert (s != NULL); + + if (slen == 0) + return 0; + + if (windows_unicode_enabled != 0) { + retcode = MultiByteToWideChar(CP_UTF8, windows_unicode_strict ? MB_ERR_INVALID_CHARS : 0, s, slen, out, outlen); + if (retcode == 0 && windows_unicode_fallback != 0) + retcode = MultiByteToWideChar(CP_THREAD_ACP, 0, s, slen, out, outlen); + } else { + retcode = MultiByteToWideChar(CP_THREAD_ACP, 0, s, slen, out, outlen); + } + + if (retcode == 0) + caml_win32_sys_error(GetLastError()); + + return retcode; +} + +#ifndef WC_ERR_INVALID_CHARS /* For old versions of Windows we simply ignore the flag */ +#define WC_ERR_INVALID_CHARS 0 +#endif + +CAMLexport int win_wide_char_to_multi_byte(const wchar_t *s, int slen, char *out, int outlen) +{ + int retcode; + + CAMLassert(s != NULL); + + if (slen == 0) + return 0; + + if (windows_unicode_enabled != 0) + retcode = WideCharToMultiByte(CP_UTF8, windows_unicode_strict ? WC_ERR_INVALID_CHARS : 0, s, slen, out, outlen, NULL, NULL); + else + retcode = WideCharToMultiByte(CP_THREAD_ACP, 0, s, slen, out, outlen, NULL, NULL); + + if (retcode == 0) + caml_win32_sys_error(GetLastError()); + + return retcode; +} + +CAMLexport value caml_copy_string_of_utf16(const wchar_t *s) +{ + int retcode, slen; + value v; + + slen = wcslen(s); + retcode = win_wide_char_to_multi_byte(s, slen, NULL, 0); /* Do not include final NULL */ + v = caml_alloc_string(retcode); + win_wide_char_to_multi_byte(s, slen, String_val(v), retcode); + + return v; +} + +CAMLexport inline wchar_t* caml_stat_strdup_to_utf16(const char *s) +{ + wchar_t * ws; + int retcode; + + retcode = win_multi_byte_to_wide_char(s, -1, NULL, 0); + ws = malloc(retcode * sizeof(*ws)); + win_multi_byte_to_wide_char(s, -1, ws, retcode); + + return ws; +} + +CAMLexport caml_stat_string caml_stat_strdup_of_utf16(const wchar_t *s) +{ + caml_stat_string out; + int retcode; + + retcode = win_wide_char_to_multi_byte(s, -1, NULL, 0); + out = caml_stat_alloc(retcode); + win_wide_char_to_multi_byte(s, -1, out, retcode); + + return out; +} + +void caml_probe_win32_version(void) +{ + /* Determine the version of Windows we're running, and cache it */ + WCHAR fileName[MAX_PATH]; + DWORD size = + GetModuleFileName(GetModuleHandle(L"kernel32"), fileName, MAX_PATH); + DWORD dwHandle = 0; + BYTE* versionInfo; + fileName[size] = 0; + size = GetFileVersionInfoSize(fileName, &dwHandle); + versionInfo = (BYTE*)malloc(size * sizeof(BYTE)); + if (GetFileVersionInfo(fileName, 0, size, versionInfo)) { + UINT len = 0; + VS_FIXEDFILEINFO* vsfi = NULL; + VerQueryValue(versionInfo, L"\\", (void**)&vsfi, &len); + caml_win32_major = HIWORD(vsfi->dwProductVersionMS); + caml_win32_minor = LOWORD(vsfi->dwProductVersionMS); + caml_win32_build = HIWORD(vsfi->dwProductVersionLS); + caml_win32_revision = LOWORD(vsfi->dwProductVersionLS); + } + free(versionInfo); +} + +static UINT startup_codepage = 0; + +void caml_setup_win32_terminal(void) +{ + if (caml_win32_major >= 10) { + startup_codepage = GetConsoleOutputCP(); + if (startup_codepage != CP_UTF8) + SetConsoleOutputCP(CP_UTF8); + } +} + +void caml_restore_win32_terminal(void) +{ + if (startup_codepage != 0) + SetConsoleOutputCP(startup_codepage); +} + +/* Detect if a named pipe corresponds to a Cygwin/MSYS pty: see + https://github.com/mirror/newlib-cygwin/blob/00e9bf2/winsup/cygwin/dtable.cc#L932 +*/ +typedef +BOOL (WINAPI *tGetFileInformationByHandleEx)(HANDLE, FILE_INFO_BY_HANDLE_CLASS, + LPVOID, DWORD); + +static int caml_win32_is_cygwin_pty(HANDLE hFile) +{ + char buffer[1024]; + FILE_NAME_INFO * nameinfo = (FILE_NAME_INFO *) buffer; + static tGetFileInformationByHandleEx pGetFileInformationByHandleEx = INVALID_HANDLE_VALUE; + + if (pGetFileInformationByHandleEx == INVALID_HANDLE_VALUE) + pGetFileInformationByHandleEx = + (tGetFileInformationByHandleEx)GetProcAddress(GetModuleHandle(L"KERNEL32.DLL"), + "GetFileInformationByHandleEx"); + + if (pGetFileInformationByHandleEx == NULL) + return 0; + + /* Get pipe name. GetFileInformationByHandleEx does not NULL-terminate the string, so reduce + the buffer size to allow for adding one. */ + if (! pGetFileInformationByHandleEx(hFile, FileNameInfo, buffer, sizeof(buffer) - sizeof(WCHAR))) + return 0; + + nameinfo->FileName[nameinfo->FileNameLength / sizeof(WCHAR)] = L'\0'; + + /* check if this could be a msys pty pipe ('msys-XXXX-ptyN-XX') + or a cygwin pty pipe ('cygwin-XXXX-ptyN-XX') */ + if ((wcsstr(nameinfo->FileName, L"msys-") || + wcsstr(nameinfo->FileName, L"cygwin-")) && wcsstr(nameinfo->FileName, L"-pty")) + return 1; + + return 0; +} + +CAMLexport int caml_win32_isatty(int fd) +{ + DWORD lpMode; + HANDLE hFile = (HANDLE)_get_osfhandle(fd); + + if (hFile == INVALID_HANDLE_VALUE) + return 0; + + switch (GetFileType(hFile)) { + case FILE_TYPE_CHAR: + /* Both console handles and the NUL device are FILE_TYPE_CHAR. The NUL + device returns FALSE for a GetConsoleMode call. _isatty incorrectly + only uses GetFileType (see GPR#1321). */ + return GetConsoleMode(hFile, &lpMode); + case FILE_TYPE_PIPE: + /* Cygwin PTYs are implemented using named pipes */ + return caml_win32_is_cygwin_pty(hFile); + default: + break; + } + + return 0; +} + +int caml_num_rows_fd(int fd) +{ + return -1; +} diff --git a/test/monniaux/ocaml/config/Makefile b/test/monniaux/ocaml/config/Makefile new file mode 100644 index 00000000..8fa72626 --- /dev/null +++ b/test/monniaux/ocaml/config/Makefile @@ -0,0 +1,109 @@ +# generated by ./configure -prefix /opt/ccomp/ocaml/4.07.1 -cc /opt/CompCert/3.4/x86_64-linux/bin/ccomp -D_DEFAULT_SOURCE --no-shared-libs +CONFIGURE_ARGS=-prefix /opt/ccomp/ocaml/4.07.1 -cc /opt/CompCert/3.4/x86_64-linux/bin/ccomp -D_DEFAULT_SOURCE --no-shared-libs +PREFIX=/opt/ccomp/ocaml/4.07.1 +BINDIR=$(PREFIX)/bin +BYTERUN=$(BINDIR)/ocamlrun +LIBDIR=$(PREFIX)/lib/ocaml +STUBLIBDIR=$(LIBDIR)/stublibs +MANDIR=$(PREFIX)/man +PROGRAMS_MAN_SECTION=1 +LIBRARIES_MAN_SECTION=3 +RANLIB=ranlib +RANLIBCMD=ranlib +ARCMD=ar +HASHBANGSCRIPTS=true +UNIX_OR_WIN32=unix +UNIXLIB=unix +GRAPHLIB=graph +PTHREAD_LINK=-lpthread +PTHREAD_CAML_LINK=-cclib -lpthread +X11_INCLUDES= +X11_LINK=-lX11 +LIBBFD_LINK= +LIBBFD_INCLUDE= +# DM CC=k1-mbr-gcc +CC=/home/monniaux/work/Kalray/CompCert/ccomp +CPP=$(CC) -E +CFLAGS=-O -Wall -fall +# DM CFLAGS=-O3 -Wall +CPPFLAGS= -D_FILE_OFFSET_BITS=64 -D_REENTRANT -DCAML_NAME_SPACE +OCAMLC_CFLAGS=-O +OCAMLC_CPPFLAGS= -D_FILE_OFFSET_BITS=64 -D_REENTRANT +LDFLAGS= +BYTECCLIBS= -lm +RPATH= +EXE= +EMPTY= +OUTPUTEXE=-o $(EMPTY) +SUPPORTS_SHARED_LIBRARIES=false +SHAREDCCCOMPOPTS= +MKSHAREDLIBRPATH= +NATDYNLINKOPTS= +SYSLIB=-l$(1) +#ml let syslib x = "-l"^x;; + +### How to build a static library +MKLIB=rm -f $(1) && ar rc $(1) $(2) && ranlib $(1) +#ml let mklib out files opts = (* "" *) +#ml Printf.sprintf "ar rc %s %s %s && ranlib %s" +#ml out opts files out;; +ARCH=amd64 +MODEL=default +SYSTEM=linux +OCAMLOPT_CFLAGS=-O +OCAMLOPT_CPPFLAGS= -D_FILE_OFFSET_BITS=64 -D_REENTRANT +NATIVECCPROFOPTS= +NATIVECCLIBS= -lm -ldl +ASM=as +ASPP=gcc -c +ASPPPROFFLAGS=-DPROFILING +PROFILING=true +DYNLINKOPTS= -ldl +OTHERLIBRARIES=unix str dynlink bigarray raw_spacetime_lib systhreads threads graph +CC_PROFILE=-pg +SYSTHREAD_SUPPORT=true +PACKLD=ld -r -o\ $(EMPTY) +IFLEXDIR= +O=o +A=a +SO=so +EXT_OBJ=.o +OUTPUTOBJ=-o $(EMPTY) +EXT_ASM=.s +EXT_LIB=.a +EXT_DLL=.so +EXTRALIBS= +CCOMPTYPE=cc +TOOLCHAIN=cc +NATDYNLINK=false +CMXS=cmxa +MKEXE=$(CC) $(CFLAGS) $(CPPFLAGS) $(LDFLAGS) +MKEXEDEBUGFLAG=-g +MKDLL=shared-libs-not-available +MKMAINDLL=shared-libs-not-available +RUNTIMED=true +RUNTIMEI=true +WITH_DEBUGGER=ocamldebugger +WITH_OCAMLDOC=ocamldoc +ASM_CFI_SUPPORTED=true +WITH_FRAME_POINTERS=false +WITH_SPACETIME=false +ENABLE_CALL_COUNTS=true +WITH_PROFINFO=false +LIBUNWIND_AVAILABLE=false +LIBUNWIND_INCLUDE_FLAGS= +LIBUNWIND_LINK_FLAGS= +PROFINFO_WIDTH=0 +WITH_CPLUGINS=false +WITH_FPIC=false +TARGET=x86_64-unknown-linux-gnu +HOST=x86_64-unknown-linux-gnu +FLAMBDA=false +WITH_FLAMBDA_INVARIANTS=false +FORCE_SAFE_STRING=false +DEFAULT_SAFE_STRING=true +WINDOWS_UNICODE=0 +AFL_INSTRUMENT=false +MAX_TESTSUITE_DIR_RETRIES=0 +FLAT_FLOAT_ARRAY=true +AWK=awk diff --git a/test/monniaux/ocaml/config/Makefile-templ b/test/monniaux/ocaml/config/Makefile-templ new file mode 100644 index 00000000..34af691e --- /dev/null +++ b/test/monniaux/ocaml/config/Makefile-templ @@ -0,0 +1,202 @@ +#************************************************************************** +#* * +#* OCaml * +#* * +#* Xavier Leroy, projet Cristal, INRIA Rocquencourt * +#* * +#* Copyright 1999 Institut National de Recherche en Informatique et * +#* en Automatique. * +#* * +#* All rights reserved. This file is distributed under the terms of * +#* the GNU Lesser General Public License version 2.1, with the * +#* special exception on linking described in the file LICENSE. * +#* * +#************************************************************************** + +### Compile-time configuration + +########## General configuration + +### Where to install the binaries +BINDIR=/usr/local/bin + +### Where to install the standard library +LIBDIR=/usr/local/lib/ocaml +STUBLIBDIR=$(LIBDIR)/stublibs + +### Where to install the man pages +# Man pages for commands go in $(MANDIR)/man$(MANEXT) +# Man pages for the library go in $(MANDIR)/mano +MANDIR=/usr/local/man +MANEXT=1 + +### Do #! scripts work on your system? +### Beware: on some systems (e.g. SunOS 4), this will work only if +### the string "#!$(BINDIR)/ocamlrun" is less than 32 characters long. +### In doubt, set HASHBANGSCRIPTS to false. +HASHBANGSCRIPTS=true +#HASHBANGSCRIPTS=false + +########## Configuration for the bytecode compiler + +### Which C compiler to use for the bytecode interpreter. +### Performance of the bytecode interpreter is *much* improved +### if Gnu CC version 2 is used. +#CC=gcc +#BYTECFLAGS= + +### Additional compile-time options for $(BYTECC). +# If using gcc on Intel x86: +# (the -fno-defer-pop option circumvents a bug in certain versions of gcc) +#BYTECCCOMPOPTS=-fno-defer-pop -Wall +# If using gcc and being cautious: +#BYTECCCOMPOPTS=-Wall +# Otherwise: +#BYTECCCOMPOPTS= + +### Additional link-time options for $(BYTECC) +# To support dynamic loading of shared libraries (they need to look at +# our own symbols): +#LDFLAGS=-Wl,-E +# Otherwise: +#LDFLAGS= + +### Libraries needed +# On most platforms: +#CCLIBS=-lcurses -ltermcap -lm + +### How to invoke the C preprocessor +# This is not needed anymore. Leave these lines commented out. +# On most machines: +#CPP=/lib/cpp -P +# Under Solaris: +#CPP=/usr/ccs/lib/cpp -P +# Under FreeBSD: +#CPP=cpp -P + +### How to invoke ranlib +RANLIB=ranlib +RANLIBCMD=ranlib + +# If ranlib is not needed: +#RANLIB=ar rs +#RANLIBCMD= + +### How to invoke ar +#ARCMD=ar + +### Shared library support +# Extension for shared libraries: so if supported, a if not supported +#SO=so +#SO=a +# Set to nothing if shared libraries supported, and to -custom if not supported +#CUSTOM_IF_NOT_SHARED= +#CUSTOM_IF_NOT_SHARED=-custom +# Options to $(BYTECC) to produce shared objects (e.g. PIC) +#SHAREDCCCOMPOPTS=-fPIC +# How to build a shared library, invoked with output .so as first arg +# and object files as remaining args +#MKSHAREDLIB=gcc -shared -o +# Compile-time option to $(BYTECC) to add a directory to be searched +# at run-time for shared libraries +#RPATH=-Wl,-rpath + +############# Configuration for the native-code compiler + +### Name of architecture for the native-code compiler +### Currently supported: +### +### i386 Intel Pentium PCs under Linux, *BSD*, NextStep +### power Macintosh under Mac OS X and Linux +### arm ARM under Linux +### +### Set ARCH=none if your machine is not supported +#ARCH=i386 +#ARCH=power +#ARCH=arm +#ARCH=none + +### Name of architecture model for the native-code compiler. +### Some architectures come in several slightly different flavors +### that share a common code generator. This variable tailors the +### behavior of the code generator to the particular flavor used. +### Currently needed only if ARCH=power; leave MODEL=default for +### other architectures. +### If ARCH=power: set MODEL=ppc +### For other architectures: leave MODEL=default +### +#MODEL=ppc +#MODEL=default + +### Name of operating system family for the native-code compiler. +#SYSTEM=solaris +#SYSTEM=linux +#SYSTEM=linux_elf +#SYSTEM=bsd +#SYSTEM=unknown + +#NATIVECFLAGS= + +# For gcc if cautious: +#NATIVECCCOMPOPTS=-Wall + +# Compile-time option to $(NATIVECC) to add a directory to be searched +# at run-time for shared libraries +#RPATH=-Wl,-rpath + +### Command and flags to use for assembling ocamlopt-generated code +#ASM=as + +### Command and flags to use for assembling .S files (often with preprocessing) +# If gcc is available: +#ASPP=gcc -c +# On Solaris: +#ASPP=as -P + +### Extra flags to use for assembling .S files in profiling mode +#ASPPPROFFLAGS=-DPROFILING + +### Whether profiling with gprof is supported +# If yes: (e.g. x86/Linux): +#PROFILING=true +# If no: +#PROFILING=false + +### Option to give to the C compiler for profiling +#CC_PROFILE=-pg +#CC_PROFILE=-xpg + +############# Configuration for the contributed libraries + +### Which libraries to compile and install +# Currently available: +# unix Unix system calls +# str Regular expressions and high-level string processing +# threads Lightweight concurrent processes +# systhreads Same as threads, requires POSIX threads +# graph Portable drawing primitives for X11 +# dynlink Dynamic linking of bytecode +# bigarray Large, multidimensional numerical arrays + +OTHERLIBRARIES=unix str threads graph dynlink bigarray + +### Link-time options to ocamlc or ocamlopt for linking with POSIX threads +# Needed for the "systhreads" package +# Usually: +#PTHREAD_LINK=-cclib -lpthread +# For Solaris: +#PTHREAD_LINK=-cclib -lpthread -cclib -lposix4 + +### -I options for finding the X11/*.h includes +# Needed for the "graph" package +# Usually: +#X11_INCLUDES=-I/usr/X11R6/include +# For SunOS with OpenLook: +#X11_INCLUDES=/usr/openwin/include + +### Link-time options to ocamlc or ocamlopt for linking with X11 libraries +# Needed for the "graph" package +# Usually: +#X11_LINK=-lX11 +# For SunOS with OpenLook: +#X11_LINK=-L$(X11_LIB) -lX11 diff --git a/test/monniaux/ocaml/config/m-nt.h b/test/monniaux/ocaml/config/m-nt.h new file mode 100644 index 00000000..97ae9cf0 --- /dev/null +++ b/test/monniaux/ocaml/config/m-nt.h @@ -0,0 +1,65 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Machine configuration, Intel x86 processors, Win32, + Visual C++ or Mingw compiler */ + +#ifdef _WIN64 +#define ARCH_SIXTYFOUR +#else +#undef ARCH_SIXTYFOUR +#endif +#undef ARCH_BIG_ENDIAN +#undef ARCH_ALIGN_DOUBLE + +#define SIZEOF_INT 4 +#define SIZEOF_LONG 4 +#ifdef _WIN64 +#define SIZEOF_PTR 8 +#else +#define SIZEOF_PTR 4 +#endif +#define SIZEOF_SHORT 2 + +#ifdef __MINGW32__ +#define ARCH_INT64_TYPE long long +#define ARCH_UINT64_TYPE unsigned long long +#else +#define ARCH_INT64_TYPE __int64 +#define ARCH_UINT64_TYPE unsigned __int64 +#endif +#define ARCH_INT64_PRINTF_FORMAT "I64" +#if _MSC_VER >= 1800 +#define ARCH_SIZET_PRINTF_FORMAT "z" +#else +#define ARCH_SIZET_PRINTF_FORMAT "I" +#endif + +#if defined(_MSC_VER) && !defined(__cplusplus) +#define inline __inline +#endif + +#undef NONSTANDARD_DIV_MOD + +#define PROFINFO_WIDTH 0 + +/* Microsoft introduced the LL integer literal suffix in Visual C++ .NET 2003 */ +#if defined(_MSC_VER) && _MSC_VER < 1400 +#define INT64_LITERAL(s) s ## i64 +#else +#define INT64_LITERAL(s) s ## LL +#endif + +#define FLAT_FLOAT_ARRAY diff --git a/test/monniaux/ocaml/config/m-templ.h b/test/monniaux/ocaml/config/m-templ.h new file mode 100644 index 00000000..74ce64e1 --- /dev/null +++ b/test/monniaux/ocaml/config/m-templ.h @@ -0,0 +1,83 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Processor dependencies */ + +#define ARCH_SIXTYFOUR + +/* Define ARCH_SIXTYFOUR if the processor has a natural word size of 64 bits. + That is, sizeof(char *) = 8. + Otherwise, leave ARCH_SIXTYFOUR undefined. + This assumes sizeof(char *) = 4. */ + +#define ARCH_BIG_ENDIAN + +/* Define ARCH_BIG_ENDIAN if the processor is big endian (the most + significant byte of an integer stored in memory comes first). + Leave ARCH_BIG_ENDIAN undefined if the processor is little-endian + (the least significant byte comes first). +*/ + +#define ARCH_ALIGN_DOUBLE + +/* Define ARCH_ALIGN_DOUBLE if the processor requires doubles to be + doubleword-aligned. Leave ARCH_ALIGN_DOUBLE undefined if the processor + supports word-aligned doubles. */ + +#undef ARCH_CODE32 + +/* Define ARCH_CODE32 if, on a 64-bit machine, code pointers fit in 32 bits, + i.e. the code segment resides in the low 4G of the addressing space. + ARCH_CODE32 is ignored on 32-bit machines. */ + +#define SIZEOF_INT 4 +#define SIZEOF_LONG 4 +#define SIZEOF_PTR 4 +#define SIZEOF_SHORT 2 + +/* Define SIZEOF_INT, SIZEOF_LONG, SIZEOF_PTR and SIZEOF_SHORT + to the sizes in bytes of the C types "int", "long", "char *" and "short", + respectively. */ + +#define ARCH_INT64_TYPE long long +#define ARCH_UINT64_TYPE unsigned long long + +/* Define ARCH_INT64_TYPE and ARCH_UINT64_TYPE to 64-bit integer types, + typically "long long" and "unsigned long long" on 32-bit platforms, + and "long" and "unsigned long" on 64-bit platforms. + If the C compiler doesn't support any 64-bit integer type, + leave both ARCH_INT64_TYPE and ARCH_UINT64_TYPE undefined. */ + +#define ARCH_INT64_PRINTF_FORMAT "ll" + +/* Define ARCH_INT64_PRINTF_FORMAT to the printf format used for formatting + values of type ARCH_INT64_TYPE. This is usually "ll" on 32-bit + platforms and "l" on 64-bit platforms. + Leave undefined if ARCH_INT64_TYPE is undefined. */ + +#define ARCH_ALIGN_INT64 + +/* Define ARCH_ALIGN_INT64 if the processor requires 64-bit integers to be + doubleword-aligned. Leave ARCH_ALIGN_INT64 undefined if the processor + supports word-aligned 64-bit integers. Leave undefined if + 64-bit integers are not supported. */ + +#undef NONSTANDARD_DIV_MOD + +/* Leave NONSTANDARD_DIV_MOD undefined if the C operators / and % implement + round-towards-zero semantics, as specified by ISO C 9x and implemented + by most contemporary processors. Otherwise, or if you don't know, + define NONSTANDARD_DIV_MOD: this will select a slower but correct + software emulation of division and modulus. */ diff --git a/test/monniaux/ocaml/config/s-nt.h b/test/monniaux/ocaml/config/s-nt.h new file mode 100644 index 00000000..ab4046b3 --- /dev/null +++ b/test/monniaux/ocaml/config/s-nt.h @@ -0,0 +1,41 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Operating system dependencies, Intel x86 processors, Windows NT */ + +#define OCAML_OS_TYPE "Win32" + +#if defined(__MINGW32__) || _MSC_VER >= 1600 +#define HAS_STDINT_H +#endif +#undef BSD_SIGNALS +#define HAS_STRERROR +#define HAS_SOCKETS +#define HAS_GETCWD +#define HAS_UTIME +#define HAS_DUP2 +#define HAS_GETHOSTNAME +#define HAS_MKTIME +#define HAS_PUTENV +#define HAS_LOCALE +#define HAS_BROKEN_PRINTF +#define HAS_IPV6 +#define HAS_NICE +#define SUPPORT_DYNAMIC_LINKING +#define HAS_EXECVPE +#if defined(_MSC_VER) && _MSC_VER < 1300 +#define LACKS_SANE_NAN +#define LACKS_VSCPRINTF +#endif diff --git a/test/monniaux/ocaml/config/s-templ.h b/test/monniaux/ocaml/config/s-templ.h new file mode 100644 index 00000000..1df4eb31 --- /dev/null +++ b/test/monniaux/ocaml/config/s-templ.h @@ -0,0 +1,214 @@ +/**************************************************************************/ +/* */ +/* OCaml */ +/* */ +/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */ +/* */ +/* Copyright 1996 Institut National de Recherche en Informatique et */ +/* en Automatique. */ +/* */ +/* All rights reserved. This file is distributed under the terms of */ +/* the GNU Lesser General Public License version 2.1, with the */ +/* special exception on linking described in the file LICENSE. */ +/* */ +/**************************************************************************/ + +/* Operating system and standard library dependencies. */ + +/* 0. Operating system type string. */ + +#define OCAML_OS_TYPE "Unix" +/* #define OCAML_OS_TYPE "Win32" */ +/* #define OCAML_OS_TYPE "MacOS" */ + +/* 1. For the runtime system. */ + +#define POSIX_SIGNALS + +/* Define POSIX_SIGNALS if signal handling is POSIX-compliant. + In particular, sigaction(), sigprocmask() and the operations on + sigset_t are provided. */ + +#define BSD_SIGNALS + +/* Define BSD_SIGNALS if signal handlers have the BSD semantics: the handler + remains attached to the signal when the signal is received. Leave it + undefined if signal handlers have the System V semantics: the signal + resets the behavior to default. */ + +#define HAS_SIGSETMASK + +/* Define HAS_SIGSETMASK if you have sigsetmask(), as in BSD. */ + +#define SUPPORT_DYNAMIC_LINKING + +/* Define SUPPORT_DYNAMIC_LINKING if dynamic loading of C stub code + via dlopen() is available. */ + +#define HAS_C99_FLOAT_OPS + +/* Define HAS_C99_FLOAT_OPS if conforms to ISO C99. + In particular, it should provide expm1(), log1p(), hypot(), copysign(). */ + +/* 2. For the Unix library. */ + +#define HAS_SOCKETS + +/* Define HAS_SOCKETS if you have BSD sockets. */ + +#define HAS_SOCKLEN_T + +/* Define HAS_SOCKLEN_T if the type socklen_t is defined in + /usr/include/sys/socket.h. */ + +#define HAS_UNISTD + +/* Define HAS_UNISTD if you have /usr/include/unistd.h. */ + +#define HAS_DIRENT + +/* Define HAS_DIRENT if you have /usr/include/dirent.h and the result of + readdir() is of type struct dirent *. + Otherwise, we'll load /usr/include/sys/dir.h, and readdir() is expected to + return a struct direct *. */ + +#define HAS_REWINDDIR + +/* Define HAS_REWINDDIR if you have rewinddir(). */ + +#define HAS_LOCKF + +/* Define HAS_LOCKF if the library provides the lockf() function. */ + +#define HAS_MKFIFO + +/* Define HAS_MKFIFO if the library provides the mkfifo() function. */ + +#define HAS_GETCWD + +/* Define HAS_GETCWD if the library provides the getcwd() function. */ + +#define HAS_GETPRIORITY + +/* Define HAS_GETPRIORITY if the library provides getpriority() and + setpriority(). Otherwise, we'll use nice(). */ + +#define HAS_UTIME +#define HAS_UTIMES + +/* Define HAS_UTIME if you have /usr/include/utime.h and the library + provides utime(). Define HAS_UTIMES if the library provides utimes(). */ + +#define HAS_DUP2 + +/* Define HAS_DUP2 if you have dup2(). */ + +#define HAS_FCHMOD + +/* Define HAS_FCHMOD if you have fchmod() and fchown(). */ + +#define HAS_TRUNCATE + +/* Define HAS_TRUNCATE if you have truncate() and + ftruncate(). */ + +#define HAS_SELECT + +/* Define HAS_SELECT if you have select(). */ + +#define HAS_SYS_SELECT_H + +/* Define HAS_SYS_SELECT_H if /usr/include/sys/select.h exists + and should be included before using select(). */ + +#define HAS_NANOSLEEP +/* Define HAS_NANOSLEEP if you have nanosleep(). */ + +#define HAS_SYMLINK + +/* Define HAS_SYMLINK if you have symlink() and readlink() and lstat(). */ + +#define HAS_WAIT4 +#define HAS_WAITPID + +/* Define HAS_WAIT4 if you have wait4(). + Define HAS_WAITPID if you have waitpid(). */ + +#define HAS_GETGROUPS + +/* Define HAS_GETGROUPS if you have getgroups(). */ + +#define HAS_SETGROUPS + +/* Define HAS_SETGROUPS if you have setgroups(). */ + +#define HAS_INITGROUPS + +/* Define HAS_INITGROUPS if you have initgroups(). */ + +#define HAS_TERMIOS + +/* Define HAS_TERMIOS if you have /usr/include/termios.h and it is + Posix-compliant. */ + +#define HAS_ASYNC_IO + +/* Define HAS_ASYNC_IO if BSD-style asynchronous I/O are supported + (the process can request to be sent a SIGIO signal when a descriptor + is ready for reading). */ + +#define HAS_SETITIMER + +/* Define HAS_SETITIMER if you have setitimer(). */ + +#define HAS_GETHOSTNAME + +/* Define HAS_GETHOSTNAME if you have gethostname(). */ + +#define HAS_UNAME + +/* Define HAS_UNAME if you have uname(). */ + +#define HAS_GETTIMEOFDAY + +/* Define HAS_GETTIMEOFDAY if you have gettimeofday(). */ + +#define HAS_MKTIME + +/* Define HAS_MKTIME if you have mktime(). */ + +#define HAS_SETSID + +/* Define HAS_SETSID if you have setsid(). */ + +#define HAS_PUTENV + +/* Define HAS_PUTENV if you have putenv(). */ + +#define HAS_LOCALE + +/* Define HAS_LOCALE if you have the include file and the + setlocale() function. */ + +#define HAS_MMAP + +/* Define HAS_MMAP if you have the include file and the + functions mmap() and munmap(). */ + +#define HAS_GETHOSTBYNAME_R 6 + +/* Define HAS_GETHOSTBYNAME_R if gethostbyname_r() is available. + The value of this symbol is the number of arguments of + gethostbyname_r(): either 5 or 6 depending on prototype. + (5 is the Solaris version, 6 is the Linux version). */ + +#define HAS_GETHOSTBYADDR_R 8 + +/* Define HAS_GETHOSTBYADDR_R if gethostbyname_r() is available. + The value of this symbol is the number of arguments of + gethostbyaddr_r(): either 7 or 8 depending on prototype. + (7 is the Solaris version, 8 is the Linux version). */ + +#define HAS_NICE + +/* Define HAS_NICE if you have nice(). */ diff --git a/test/monniaux/ocaml/examples/quicksort b/test/monniaux/ocaml/examples/quicksort new file mode 100755 index 00000000..0b820167 Binary files /dev/null and b/test/monniaux/ocaml/examples/quicksort differ diff --git a/test/monniaux/ocaml/examples/quicksort.ml b/test/monniaux/ocaml/examples/quicksort.ml new file mode 100644 index 00000000..57ca5a03 --- /dev/null +++ b/test/monniaux/ocaml/examples/quicksort.ml @@ -0,0 +1,11 @@ +let rec quicksort gt = function + | [] -> [] + | x::xs -> + let ys, zs = List.partition (gt x) xs in + (quicksort gt ys) @ (x :: (quicksort gt zs));; + +let l = + quicksort ( > ) [4; 65; 2; -31; 0; 99; 83; 782; 1] +in +List.iter (fun x -> Printf.printf "%d; " x) l; +print_newline(); diff --git a/test/monniaux/ocaml/tools/make-version-header.sh b/test/monniaux/ocaml/tools/make-version-header.sh new file mode 100755 index 00000000..707d04fa --- /dev/null +++ b/test/monniaux/ocaml/tools/make-version-header.sh @@ -0,0 +1,55 @@ +#!/bin/sh + +#************************************************************************** +#* * +#* OCaml * +#* * +#* Damien Doligez, projet Gallium, INRIA Rocquencourt * +#* * +#* Copyright 2003 Institut National de Recherche en Informatique et * +#* en Automatique. * +#* * +#* All rights reserved. As an exception to the licensing rules of * +#* OCaml, this file is freely redistributable, modified or not, * +#* without constraints. * +#* * +#************************************************************************** + +# This script extracts the components from an OCaml version number +# and provides them as C defines: +# OCAML_VERSION_MAJOR: the major version number +# OCAML_VERSION_MAJOR: the minor version number +# OCAML_VERSION_PATCHLEVEL: the patchlevel number if present, or 0 if absent +# OCAML_VERSION_ADDITIONAL: this is defined only if the additional-info +# field is present, and is a string that contains that field. +# Note that additional-info is always absent in officially-released +# versions of OCaml. + +# usage: +# make-version-header.sh [version-file] +# The argument is the VERSION file from the OCaml sources. +# If the argument is not given, the version number from "ocamlc -v" will +# be used. + +case $# in + 0) version="`ocamlc -v | tr -d '\r' | sed -n -e 's/.*version //p'`";; + 1) version="`sed -e 1q $1 | tr -d '\r'`";; + *) echo "usage: make-version-header.sh [version-file]" >&2 + exit 2;; +esac + +major="`echo "$version" | sed -n -e '1s/^\([0-9]*\)\..*/\1/p'`" +minor="`echo "$version" | sed -n -e '1s/^[0-9]*\.0*\([0-9]*\).*/\1/p'`" +patchlvl="`echo "$version" | sed -n -e '1s/^[0-9]*\.[0-9]*\.\([0-9]*\).*/\1/p'`" +suffix="`echo "$version" | sed -n -e '1s/^[^+]*+\(.*\)/\1/p'`" + +echo "#define OCAML_VERSION_MAJOR $major" +printf "#define OCAML_VERSION_MINOR %d\n" $minor +case $patchlvl in "") patchlvl=0;; esac +echo "#define OCAML_VERSION_PATCHLEVEL $patchlvl" +case "$suffix" in + "") echo "#undef OCAML_VERSION_ADDITIONAL";; + *) echo "#define OCAML_VERSION_ADDITIONAL \"$suffix\"";; +esac +printf "#define OCAML_VERSION %d%02d%02d\n" $major $minor $patchlvl +echo "#define OCAML_VERSION_STRING \"$version\"" -- cgit From 4f17ea42b77c02f8632906cfc33f29fc62a43123 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 17:41:01 +0100 Subject: Makefile for ocamlrun testing --- test/monniaux/ocaml/Makefile | 7 +++++++ 1 file changed, 7 insertions(+) create mode 100644 test/monniaux/ocaml/Makefile diff --git a/test/monniaux/ocaml/Makefile b/test/monniaux/ocaml/Makefile new file mode 100644 index 00000000..46ce8994 --- /dev/null +++ b/test/monniaux/ocaml/Makefile @@ -0,0 +1,7 @@ +test: byterun/ocamlrun + k1-cluster --syscall=libstd_scalls.so -- byterun/ocamlrun examples/quicksort + +byterun/ocamlrun: + (cd byterun ; $(MAKE)) + +.PHONY: test -- cgit From c4661a0181e6db3aa4a36939bdbea5948eadb6c9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 18:26:04 +0100 Subject: csmith for testing --- test/monniaux/csmith/Makefile | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) create mode 100644 test/monniaux/csmith/Makefile diff --git a/test/monniaux/csmith/Makefile b/test/monniaux/csmith/Makefile new file mode 100644 index 00000000..b094e425 --- /dev/null +++ b/test/monniaux/csmith/Makefile @@ -0,0 +1,23 @@ +CSMITH=csmith +MAX=100 + +include ../rules.mk +K1C_CCOMPFLAGS+=-I/usr/include/csmith -fstruct-passing -fbitfields + +TARGETS_S=$(shell seq --format src%06.f.ccomp.k1c.s 0 $(MAX)) +TARGETS_C=$(shell seq --format src%06.f.c 0 $(MAX)) +TARGETS_O=$(shell seq --format src%06.f.ccomp.k1c.o 0 $(MAX)) + +all: c s o + +s: $(TARGETS_S) +c: $(TARGETS_C) +o: $(TARGETS_O) + +src%.c : + $(CSMITH) --output $@ --seed $* + +clean: + -rm -f $(TARGETS_C) $(TARGETS_S) $(TARGETS_O) + +.PHONY: s c o clean -- cgit From 0b393f7e37315e055eddc5cfba8333639ca265be Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 20 Mar 2019 18:35:03 +0100 Subject: rm tests inherited from Risc-V --- mppa_k1c/Asmexpand.ml | 35 +++++------------------------------ 1 file changed, 5 insertions(+), 30 deletions(-) diff --git a/mppa_k1c/Asmexpand.ml b/mppa_k1c/Asmexpand.ml index 1c9e4e4c..f1528389 100644 --- a/mppa_k1c/Asmexpand.ml +++ b/mppa_k1c/Asmexpand.ml @@ -124,9 +124,6 @@ let expand_annot_val kind txt targ args res = assert false (* Handling of memcpy *) -let offset_in_range ofs = - let ofs = Z.to_int64 ofs in -2048L <= ofs && ofs < 2048L - let emit_move dst r = if dst <> r then emit (Paddil(dst, r, Z.zero));; @@ -175,7 +172,7 @@ let expand_builtin_memcpy sz al args = | _ -> assert false;; (* Handling of volatile reads and writes *) - +(* FIXME probably need to check for size of displacement *) let expand_builtin_vload_common chunk base ofs res = match chunk, res with | Mint8unsigned, BR(Asmblock.IR res) -> @@ -211,21 +208,9 @@ let expand_builtin_vload chunk args res = | [BA(Asmblock.IR addr)] -> expand_builtin_vload_common chunk addr _0 res | [BA_addrstack ofs] -> - if offset_in_range (Z.add ofs (Memdata.size_chunk chunk)) then - expand_builtin_vload_common chunk stack_pointer ofs res - else begin - assert false (* FIXME - expand_addptrofs Asmblock.GPR32 stack_pointer ofs; (* X31 <- sp + ofs *) - expand_builtin_vload_common chunk GPR32 _0 res *) - end + expand_builtin_vload_common chunk stack_pointer ofs res | [BA_addptr(BA(Asmblock.IR addr), (BA_int ofs | BA_long ofs))] -> - if offset_in_range (Z.add ofs (Memdata.size_chunk chunk)) then - expand_builtin_vload_common chunk addr ofs res - else begin - assert false (* FIXME - expand_addptrofs Asmblock.GPR32 addr ofs; (* X31 <- addr + ofs *) - expand_builtin_vload_common chunk Asmblock.GPR32 _0 res *) - end + expand_builtin_vload_common chunk addr ofs res | _ -> assert false @@ -256,19 +241,9 @@ let expand_builtin_vstore chunk args = | [BA(Asmblock.IR addr); src] -> expand_builtin_vstore_common chunk addr _0 src | [BA_addrstack ofs; src] -> - if offset_in_range (Z.add ofs (Memdata.size_chunk chunk)) then - expand_builtin_vstore_common chunk stack_pointer ofs src - else begin (* FIXME - expand_addptrofs X31 X2 ofs; (* X31 <- sp + ofs *) - expand_builtin_vstore_common chunk X31 _0 src *) - end + expand_builtin_vstore_common chunk stack_pointer ofs src | [BA_addptr(BA(Asmblock.IR addr), (BA_int ofs | BA_long ofs)); src] -> - if offset_in_range (Z.add ofs (Memdata.size_chunk chunk)) then - expand_builtin_vstore_common chunk addr ofs src - else begin (* FIXME - expand_addptrofs X31 addr ofs; (* X31 <- addr + ofs *) - expand_builtin_vstore_common chunk X31 _0 src *) - end + expand_builtin_vstore_common chunk addr ofs src | _ -> assert false -- cgit From 23fa2a18e015b9d330ad6f1f08cf50adf90bd80b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 21 Mar 2019 22:39:27 +0100 Subject: try to be portable across archs --- arm/Machregsaux.ml | 5 +++++ arm/Machregsaux.mli | 2 ++ backend/IRC.ml | 9 +++------ backend/IRC.mli | 1 - backend/Regalloc.ml | 2 +- mppa_k1c/Machregsaux.ml | 5 +++++ mppa_k1c/Machregsaux.mli | 2 ++ powerpc/Machregsaux.ml | 5 +++++ powerpc/Machregsaux.mli | 2 ++ riscV/Machregsaux.ml | 5 +++++ riscV/Machregsaux.mli | 2 ++ x86/Conventions1.v | 3 ++- x86/Machregsaux.ml | 9 +++++++-- x86/Machregsaux.mli | 2 ++ 14 files changed, 43 insertions(+), 11 deletions(-) diff --git a/arm/Machregsaux.ml b/arm/Machregsaux.ml index ce5c67f6..14c75155 100644 --- a/arm/Machregsaux.ml +++ b/arm/Machregsaux.ml @@ -33,3 +33,8 @@ let register_by_name s = let can_reserve_register r = List.mem r Conventions1.int_callee_save_regs || List.mem r Conventions1.float_callee_save_regs + +let class_of_type = function + | AST.Tint | AST.Tlong -> 0 + | AST.Tfloat | AST.Tsingle -> 1 + | AST.Tany32 | AST.Tany64 -> assert false diff --git a/arm/Machregsaux.mli b/arm/Machregsaux.mli index 9404568d..d7117c21 100644 --- a/arm/Machregsaux.mli +++ b/arm/Machregsaux.mli @@ -16,3 +16,5 @@ val name_of_register: Machregs.mreg -> string option val register_by_name: string -> Machregs.mreg option val is_scratch_register: string -> bool val can_reserve_register: Machregs.mreg -> bool + +val class_of_type: AST.typ -> int diff --git a/backend/IRC.ml b/backend/IRC.ml index c7b1bf04..67da47da 100644 --- a/backend/IRC.ml +++ b/backend/IRC.ml @@ -15,6 +15,7 @@ open Camlcoq open AST open Registers open Machregs +open Machregsaux open Locations open Conventions1 open XTL @@ -237,13 +238,9 @@ type graph = { according to their types. A variable can be forced into class 2 by giving it a negative spill cost. *) -let class_of_type = function - | Tint | Tlong -> 0 - | Tfloat | Tsingle -> 0 (* normal: 1 *) - | Tany32 | Tany64 -> assert false -let class_of_reg r = 0 - (* normal: if Conventions1.is_float_reg r then 1 else 0 *) +let class_of_reg r = + if Conventions1.is_float_reg r then 1 else 0 let class_of_loc = function | R r -> class_of_reg r diff --git a/backend/IRC.mli b/backend/IRC.mli index 30b6d5c1..f7bbf9c5 100644 --- a/backend/IRC.mli +++ b/backend/IRC.mli @@ -43,5 +43,4 @@ val coloring: graph -> (var -> loc) val reserved_registers: mreg list ref (* Auxiliaries to deal with register classes *) -val class_of_type: AST.typ -> int val class_of_loc: loc -> int diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml index 19aba4f6..7db8a866 100644 --- a/backend/Regalloc.ml +++ b/backend/Regalloc.ml @@ -1067,7 +1067,7 @@ let make_parmove srcs dsts itmp ftmp k = | Locations.S(sl, ofs, ty), R rd -> code := LTL.Lgetstack(sl, ofs, ty, rd) :: !code | Locations.S(sls, ofss, tys), Locations.S(sld, ofsd, tyd) -> - let tmp = temp_for (class_of_type tys) in + let tmp = temp_for (Machregsaux.class_of_type tys) in (* code will be reversed at the end *) code := LTL.Lsetstack(tmp, sld, ofsd, tyd) :: LTL.Lgetstack(sls, ofss, tys, tmp) :: !code diff --git a/mppa_k1c/Machregsaux.ml b/mppa_k1c/Machregsaux.ml index 473e0602..9c4175ed 100644 --- a/mppa_k1c/Machregsaux.ml +++ b/mppa_k1c/Machregsaux.ml @@ -31,3 +31,8 @@ let register_by_name s = Machregs.register_by_name (coqstring_uppercase_ascii_of_camlstring s) let can_reserve_register r = Conventions1.is_callee_save r + +let class_of_type = function + | AST.Tint | AST.Tlong + | AST.Tfloat | AST.Tsingle -> 0 + | AST.Tany32 | AST.Tany64 -> assert false diff --git a/mppa_k1c/Machregsaux.mli b/mppa_k1c/Machregsaux.mli index 9404568d..d7117c21 100644 --- a/mppa_k1c/Machregsaux.mli +++ b/mppa_k1c/Machregsaux.mli @@ -16,3 +16,5 @@ val name_of_register: Machregs.mreg -> string option val register_by_name: string -> Machregs.mreg option val is_scratch_register: string -> bool val can_reserve_register: Machregs.mreg -> bool + +val class_of_type: AST.typ -> int diff --git a/powerpc/Machregsaux.ml b/powerpc/Machregsaux.ml index 664f71a0..0b0d4548 100644 --- a/powerpc/Machregsaux.ml +++ b/powerpc/Machregsaux.ml @@ -33,3 +33,8 @@ let register_by_name s = let can_reserve_register r = List.mem r Conventions1.int_callee_save_regs || List.mem r Conventions1.float_callee_save_regs + +let class_of_type = function + | AST.Tint | AST.Tlong -> 0 + | AST.Tfloat | AST.Tsingle -> 1 + | AST.Tany32 | AST.Tany64 -> assert false diff --git a/powerpc/Machregsaux.mli b/powerpc/Machregsaux.mli index 9404568d..d7117c21 100644 --- a/powerpc/Machregsaux.mli +++ b/powerpc/Machregsaux.mli @@ -16,3 +16,5 @@ val name_of_register: Machregs.mreg -> string option val register_by_name: string -> Machregs.mreg option val is_scratch_register: string -> bool val can_reserve_register: Machregs.mreg -> bool + +val class_of_type: AST.typ -> int diff --git a/riscV/Machregsaux.ml b/riscV/Machregsaux.ml index 473e0602..07097eaf 100644 --- a/riscV/Machregsaux.ml +++ b/riscV/Machregsaux.ml @@ -31,3 +31,8 @@ let register_by_name s = Machregs.register_by_name (coqstring_uppercase_ascii_of_camlstring s) let can_reserve_register r = Conventions1.is_callee_save r + +let class_of_type = function + | AST.Tint | AST.Tlong -> 0 + | AST.Tfloat | AST.Tsingle -> 1 + | AST.Tany32 | AST.Tany64 -> assert false diff --git a/riscV/Machregsaux.mli b/riscV/Machregsaux.mli index 9404568d..d7117c21 100644 --- a/riscV/Machregsaux.mli +++ b/riscV/Machregsaux.mli @@ -16,3 +16,5 @@ val name_of_register: Machregs.mreg -> string option val register_by_name: string -> Machregs.mreg option val is_scratch_register: string -> bool val can_reserve_register: Machregs.mreg -> bool + +val class_of_type: AST.typ -> int diff --git a/x86/Conventions1.v b/x86/Conventions1.v index 646c4afb..35d555f9 100644 --- a/x86/Conventions1.v +++ b/x86/Conventions1.v @@ -15,6 +15,7 @@ Require Import Coqlib Decidableplus. Require Import AST Machregs Locations. +Require Import Errors. (** * Classification of machine registers *) @@ -26,7 +27,7 @@ Require Import AST Machregs Locations. We follow the x86-32 and x86-64 application binary interfaces (ABI) in our choice of callee- and caller-save registers. *) - + Definition is_callee_save (r: mreg) : bool := match r with | AX | CX | DX => false diff --git a/x86/Machregsaux.ml b/x86/Machregsaux.ml index 473e0602..80066b00 100644 --- a/x86/Machregsaux.ml +++ b/x86/Machregsaux.ml @@ -14,9 +14,9 @@ open Camlcoq open Machregs - + let register_names : (mreg, string) Hashtbl.t = Hashtbl.create 31 - + let _ = List.iter (fun (s, r) -> Hashtbl.add register_names r (camlstring_of_coqstring s)) @@ -31,3 +31,8 @@ let register_by_name s = Machregs.register_by_name (coqstring_uppercase_ascii_of_camlstring s) let can_reserve_register r = Conventions1.is_callee_save r + +let class_of_type = function + | AST.Tint | AST.Tlong -> 0 + | AST.Tfloat | AST.Tsingle -> 1 + | AST.Tany32 | AST.Tany64 -> assert false diff --git a/x86/Machregsaux.mli b/x86/Machregsaux.mli index 9404568d..d7117c21 100644 --- a/x86/Machregsaux.mli +++ b/x86/Machregsaux.mli @@ -16,3 +16,5 @@ val name_of_register: Machregs.mreg -> string option val register_by_name: string -> Machregs.mreg option val is_scratch_register: string -> bool val can_reserve_register: Machregs.mreg -> bool + +val class_of_type: AST.typ -> int -- cgit From 49ea6f7d4b9e18f8aa740d068bb3fb9e49596e00 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 00:28:46 +0100 Subject: I think it should now compile for all architectures. --- mppa_k1c/ConstpropOp.v | 613 ------------------------------------------------- 1 file changed, 613 deletions(-) delete mode 100644 mppa_k1c/ConstpropOp.v diff --git a/mppa_k1c/ConstpropOp.v b/mppa_k1c/ConstpropOp.v deleted file mode 100644 index e7391ab5..00000000 --- a/mppa_k1c/ConstpropOp.v +++ /dev/null @@ -1,613 +0,0 @@ -(* *********************************************************************) -(* *) -(* The Compcert verified compiler *) -(* *) -(* Xavier Leroy, INRIA Paris *) -(* *) -(* Copyright Institut National de Recherche en Informatique et en *) -(* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) -(* *) -(* *********************************************************************) - -(** Strength reduction for operators and conditions. - This is the machine-dependent part of [Constprop]. *) - -Require Archi. -Require Import Coqlib Compopts. -Require Import AST Integers Floats. -Require Import Op Registers. -Require Import ValueDomain. - -(** * Converting known values to constants *) - -Definition const_for_result (a: aval) : option operation := - match a with - | I n => Some(Ointconst n) - | L n => if Archi.ptr64 then Some(Olongconst n) else None - | F n => if Compopts.generate_float_constants tt then Some(Ofloatconst n) else None - | FS n => if Compopts.generate_float_constants tt then Some(Osingleconst n) else None - | Ptr(Gl id ofs) => Some(Oaddrsymbol id ofs) - | Ptr(Stk ofs) => Some(Oaddrstack ofs) - | _ => None - end. - -(** * Operator strength reduction *) - -(** We now define auxiliary functions for strength reduction of - operators and addressing modes: replacing an operator with a cheaper - one if some of its arguments are statically known. These are again - large pattern-matchings expressed in indirect style. *) - -(** Original definition: -<< -Nondetfunction cond_strength_reduction - (cond: condition) (args: list reg) (vl: list aval) := - match cond, args, vl with - | Ccomp c, r1 :: r2 :: nil, I n1 :: v2 :: nil => - (Ccompimm (swap_comparison c) n1, r2 :: nil) - | Ccomp c, r1 :: r2 :: nil, v1 :: I n2 :: nil => - (Ccompimm c n2, r1 :: nil) - | Ccompu c, r1 :: r2 :: nil, I n1 :: v2 :: nil => - (Ccompuimm (swap_comparison c) n1, r2 :: nil) - | Ccompu c, r1 :: r2 :: nil, v1 :: I n2 :: nil => - (Ccompuimm c n2, r1 :: nil) - | Ccompl c, r1 :: r2 :: nil, L n1 :: v2 :: nil => - (Ccomplimm (swap_comparison c) n1, r2 :: nil) - | Ccompl c, r1 :: r2 :: nil, v1 :: L n2 :: nil => - (Ccomplimm c n2, r1 :: nil) - | Ccomplu c, r1 :: r2 :: nil, L n1 :: v2 :: nil => - (Ccompluimm (swap_comparison c) n1, r2 :: nil) - | Ccomplu c, r1 :: r2 :: nil, v1 :: L n2 :: nil => - (Ccompluimm c n2, r1 :: nil) - | _, _, _ => - (cond, args) - end. ->> -*) - -Inductive cond_strength_reduction_cases: forall (cond: condition) (args: list reg) (vl: list aval), Type := - | cond_strength_reduction_case1: forall c r1 r2 n1 v2, cond_strength_reduction_cases (Ccomp c) (r1 :: r2 :: nil) (I n1 :: v2 :: nil) - | cond_strength_reduction_case2: forall c r1 r2 v1 n2, cond_strength_reduction_cases (Ccomp c) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | cond_strength_reduction_case3: forall c r1 r2 n1 v2, cond_strength_reduction_cases (Ccompu c) (r1 :: r2 :: nil) (I n1 :: v2 :: nil) - | cond_strength_reduction_case4: forall c r1 r2 v1 n2, cond_strength_reduction_cases (Ccompu c) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | cond_strength_reduction_case5: forall c r1 r2 n1 v2, cond_strength_reduction_cases (Ccompl c) (r1 :: r2 :: nil) (L n1 :: v2 :: nil) - | cond_strength_reduction_case6: forall c r1 r2 v1 n2, cond_strength_reduction_cases (Ccompl c) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | cond_strength_reduction_case7: forall c r1 r2 n1 v2, cond_strength_reduction_cases (Ccomplu c) (r1 :: r2 :: nil) (L n1 :: v2 :: nil) - | cond_strength_reduction_case8: forall c r1 r2 v1 n2, cond_strength_reduction_cases (Ccomplu c) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | cond_strength_reduction_default: forall (cond: condition) (args: list reg) (vl: list aval), cond_strength_reduction_cases cond args vl. - -Definition cond_strength_reduction_match (cond: condition) (args: list reg) (vl: list aval) := - match cond as zz1, args as zz2, vl as zz3 return cond_strength_reduction_cases zz1 zz2 zz3 with - | Ccomp c, r1 :: r2 :: nil, I n1 :: v2 :: nil => cond_strength_reduction_case1 c r1 r2 n1 v2 - | Ccomp c, r1 :: r2 :: nil, v1 :: I n2 :: nil => cond_strength_reduction_case2 c r1 r2 v1 n2 - | Ccompu c, r1 :: r2 :: nil, I n1 :: v2 :: nil => cond_strength_reduction_case3 c r1 r2 n1 v2 - | Ccompu c, r1 :: r2 :: nil, v1 :: I n2 :: nil => cond_strength_reduction_case4 c r1 r2 v1 n2 - | Ccompl c, r1 :: r2 :: nil, L n1 :: v2 :: nil => cond_strength_reduction_case5 c r1 r2 n1 v2 - | Ccompl c, r1 :: r2 :: nil, v1 :: L n2 :: nil => cond_strength_reduction_case6 c r1 r2 v1 n2 - | Ccomplu c, r1 :: r2 :: nil, L n1 :: v2 :: nil => cond_strength_reduction_case7 c r1 r2 n1 v2 - | Ccomplu c, r1 :: r2 :: nil, v1 :: L n2 :: nil => cond_strength_reduction_case8 c r1 r2 v1 n2 - | cond, args, vl => cond_strength_reduction_default cond args vl - end. - -Definition cond_strength_reduction (cond: condition) (args: list reg) (vl: list aval) := - match cond_strength_reduction_match cond args vl with - | cond_strength_reduction_case1 c r1 r2 n1 v2 => (* Ccomp c, r1 :: r2 :: nil, I n1 :: v2 :: nil *) - (Ccompimm (swap_comparison c) n1, r2 :: nil) - | cond_strength_reduction_case2 c r1 r2 v1 n2 => (* Ccomp c, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - (Ccompimm c n2, r1 :: nil) - | cond_strength_reduction_case3 c r1 r2 n1 v2 => (* Ccompu c, r1 :: r2 :: nil, I n1 :: v2 :: nil *) - (Ccompuimm (swap_comparison c) n1, r2 :: nil) - | cond_strength_reduction_case4 c r1 r2 v1 n2 => (* Ccompu c, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - (Ccompuimm c n2, r1 :: nil) - | cond_strength_reduction_case5 c r1 r2 n1 v2 => (* Ccompl c, r1 :: r2 :: nil, L n1 :: v2 :: nil *) - (Ccomplimm (swap_comparison c) n1, r2 :: nil) - | cond_strength_reduction_case6 c r1 r2 v1 n2 => (* Ccompl c, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - (Ccomplimm c n2, r1 :: nil) - | cond_strength_reduction_case7 c r1 r2 n1 v2 => (* Ccomplu c, r1 :: r2 :: nil, L n1 :: v2 :: nil *) - (Ccompluimm (swap_comparison c) n1, r2 :: nil) - | cond_strength_reduction_case8 c r1 r2 v1 n2 => (* Ccomplu c, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - (Ccompluimm c n2, r1 :: nil) - | cond_strength_reduction_default cond args vl => - (cond, args) - end. - - -Definition make_cmp_base (c: condition) (args: list reg) (vl: list aval) := - let (c', args') := cond_strength_reduction c args vl in (Ocmp c', args'). - -Definition make_cmp_imm_eq (c: condition) (args: list reg) (vl: list aval) - (n: int) (r1: reg) (v1: aval) := - if Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1) then (Omove, r1 :: nil) - else if Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1) then (Oxorimm Int.one, r1 :: nil) - else make_cmp_base c args vl. - -Definition make_cmp_imm_ne (c: condition) (args: list reg) (vl: list aval) - (n: int) (r1: reg) (v1: aval) := - if Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1) then (Omove, r1 :: nil) - else if Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1) then (Oxorimm Int.one, r1 :: nil) - else make_cmp_base c args vl. - -(** Original definition: -<< -Nondetfunction make_cmp (c: condition) (args: list reg) (vl: list aval) := - match c, args, vl with - | Ccompimm Ceq n, r1 :: nil, v1 :: nil => - make_cmp_imm_eq c args vl n r1 v1 - | Ccompimm Cne n, r1 :: nil, v1 :: nil => - make_cmp_imm_ne c args vl n r1 v1 - | Ccompuimm Ceq n, r1 :: nil, v1 :: nil => - make_cmp_imm_eq c args vl n r1 v1 - | Ccompuimm Cne n, r1 :: nil, v1 :: nil => - make_cmp_imm_ne c args vl n r1 v1 - | _, _, _ => - make_cmp_base c args vl - end. ->> -*) - -Inductive make_cmp_cases: forall (c: condition) (args: list reg) (vl: list aval), Type := - | make_cmp_case1: forall n r1 v1, make_cmp_cases (Ccompimm Ceq n) (r1 :: nil) (v1 :: nil) - | make_cmp_case2: forall n r1 v1, make_cmp_cases (Ccompimm Cne n) (r1 :: nil) (v1 :: nil) - | make_cmp_case3: forall n r1 v1, make_cmp_cases (Ccompuimm Ceq n) (r1 :: nil) (v1 :: nil) - | make_cmp_case4: forall n r1 v1, make_cmp_cases (Ccompuimm Cne n) (r1 :: nil) (v1 :: nil) - | make_cmp_default: forall (c: condition) (args: list reg) (vl: list aval), make_cmp_cases c args vl. - -Definition make_cmp_match (c: condition) (args: list reg) (vl: list aval) := - match c as zz1, args as zz2, vl as zz3 return make_cmp_cases zz1 zz2 zz3 with - | Ccompimm Ceq n, r1 :: nil, v1 :: nil => make_cmp_case1 n r1 v1 - | Ccompimm Cne n, r1 :: nil, v1 :: nil => make_cmp_case2 n r1 v1 - | Ccompuimm Ceq n, r1 :: nil, v1 :: nil => make_cmp_case3 n r1 v1 - | Ccompuimm Cne n, r1 :: nil, v1 :: nil => make_cmp_case4 n r1 v1 - | c, args, vl => make_cmp_default c args vl - end. - -Definition make_cmp (c: condition) (args: list reg) (vl: list aval) := - match make_cmp_match c args vl with - | make_cmp_case1 n r1 v1 => (* Ccompimm Ceq n, r1 :: nil, v1 :: nil *) - make_cmp_imm_eq c args vl n r1 v1 - | make_cmp_case2 n r1 v1 => (* Ccompimm Cne n, r1 :: nil, v1 :: nil *) - make_cmp_imm_ne c args vl n r1 v1 - | make_cmp_case3 n r1 v1 => (* Ccompuimm Ceq n, r1 :: nil, v1 :: nil *) - make_cmp_imm_eq c args vl n r1 v1 - | make_cmp_case4 n r1 v1 => (* Ccompuimm Cne n, r1 :: nil, v1 :: nil *) - make_cmp_imm_ne c args vl n r1 v1 - | make_cmp_default c args vl => - make_cmp_base c args vl - end. - - -Definition make_addimm (n: int) (r: reg) := - if Int.eq n Int.zero - then (Omove, r :: nil) - else (Oaddimm n, r :: nil). - -Definition make_shlimm (n: int) (r1 r2: reg) := - if Int.eq n Int.zero then (Omove, r1 :: nil) - else if Int.ltu n Int.iwordsize then (Oshlimm n, r1 :: nil) - else (Oshl, r1 :: r2 :: nil). - -Definition make_shrimm (n: int) (r1 r2: reg) := - if Int.eq n Int.zero then (Omove, r1 :: nil) - else if Int.ltu n Int.iwordsize then (Oshrimm n, r1 :: nil) - else (Oshr, r1 :: r2 :: nil). - -Definition make_shruimm (n: int) (r1 r2: reg) := - if Int.eq n Int.zero then (Omove, r1 :: nil) - else if Int.ltu n Int.iwordsize then (Oshruimm n, r1 :: nil) - else (Oshru, r1 :: r2 :: nil). - -Definition make_mulimm (n: int) (r1 r2: reg) := - if Int.eq n Int.zero then - (Ointconst Int.zero, nil) - else if Int.eq n Int.one then - (Omove, r1 :: nil) - else - match Int.is_power2 n with - | Some l => (Oshlimm l, r1 :: nil) - | None => (Omul, r1 :: r2 :: nil) - end. - -Definition make_andimm (n: int) (r: reg) (a: aval) := - if Int.eq n Int.zero then (Ointconst Int.zero, nil) - else if Int.eq n Int.mone then (Omove, r :: nil) - else if match a with Uns _ m => Int.eq (Int.zero_ext m (Int.not n)) Int.zero - | _ => false end - then (Omove, r :: nil) - else (Oandimm n, r :: nil). - -Definition make_orimm (n: int) (r: reg) := - if Int.eq n Int.zero then (Omove, r :: nil) - else if Int.eq n Int.mone then (Ointconst Int.mone, nil) - else (Oorimm n, r :: nil). - -Definition make_xorimm (n: int) (r: reg) := - if Int.eq n Int.zero then (Omove, r :: nil) - else (Oxorimm n, r :: nil). - -Definition make_divimm n (r1 r2: reg) := - if Int.eq n Int.one then - (Omove, r1 :: nil) - else - match Int.is_power2 n with - | Some l => if Int.ltu l (Int.repr 31) - then (Oshrximm l, r1 :: nil) - else (Odiv, r1 :: r2 :: nil) - | None => (Odiv, r1 :: r2 :: nil) - end. - -Definition make_divuimm n (r1 r2: reg) := - if Int.eq n Int.one then - (Omove, r1 :: nil) - else - match Int.is_power2 n with - | Some l => (Oshruimm l, r1 :: nil) - | None => (Odivu, r1 :: r2 :: nil) - end. - -Definition make_moduimm n (r1 r2: reg) := - match Int.is_power2 n with - | Some l => (Oandimm (Int.sub n Int.one), r1 :: nil) - | None => (Omodu, r1 :: r2 :: nil) - end. - -Definition make_addlimm (n: int64) (r: reg) := - if Int64.eq n Int64.zero - then (Omove, r :: nil) - else (Oaddlimm n, r :: nil). - -Definition make_shllimm (n: int) (r1 r2: reg) := - if Int.eq n Int.zero then (Omove, r1 :: nil) - else if Int.ltu n Int64.iwordsize' then (Oshllimm n, r1 :: nil) - else (Oshll, r1 :: r2 :: nil). - -Definition make_shrlimm (n: int) (r1 r2: reg) := - if Int.eq n Int.zero then (Omove, r1 :: nil) - else if Int.ltu n Int64.iwordsize' then (Oshrlimm n, r1 :: nil) - else (Oshrl, r1 :: r2 :: nil). - -Definition make_shrluimm (n: int) (r1 r2: reg) := - if Int.eq n Int.zero then (Omove, r1 :: nil) - else if Int.ltu n Int64.iwordsize' then (Oshrluimm n, r1 :: nil) - else (Oshrlu, r1 :: r2 :: nil). - -Definition make_mullimm (n: int64) (r1 r2: reg) := - if Int64.eq n Int64.zero then - (Olongconst Int64.zero, nil) - else if Int64.eq n Int64.one then - (Omove, r1 :: nil) - else - match Int64.is_power2' n with - | Some l => (Oshllimm l, r1 :: nil) - | None => (Omull, r1 :: r2 :: nil) - end. - -Definition make_andlimm (n: int64) (r: reg) (a: aval) := - if Int64.eq n Int64.zero then (Olongconst Int64.zero, nil) - else if Int64.eq n Int64.mone then (Omove, r :: nil) - else (Oandlimm n, r :: nil). - -Definition make_orlimm (n: int64) (r: reg) := - if Int64.eq n Int64.zero then (Omove, r :: nil) - else if Int64.eq n Int64.mone then (Olongconst Int64.mone, nil) - else (Oorlimm n, r :: nil). - -Definition make_xorlimm (n: int64) (r: reg) := - if Int64.eq n Int64.zero then (Omove, r :: nil) - else (Oxorlimm n, r :: nil). - -Definition make_divlimm n (r1 r2: reg) := - match Int64.is_power2' n with - | Some l => if Int.ltu l (Int.repr 63) - then (Oshrxlimm l, r1 :: nil) - else (Odivl, r1 :: r2 :: nil) - | None => (Odivl, r1 :: r2 :: nil) - end. - -Definition make_divluimm n (r1 r2: reg) := - match Int64.is_power2' n with - | Some l => (Oshrluimm l, r1 :: nil) - | None => (Odivlu, r1 :: r2 :: nil) - end. - -Definition make_modluimm n (r1 r2: reg) := - match Int64.is_power2 n with - | Some l => (Oandlimm (Int64.sub n Int64.one), r1 :: nil) - | None => (Omodlu, r1 :: r2 :: nil) - end. - -Definition make_mulfimm (n: float) (r r1 r2: reg) := - if Float.eq_dec n (Float.of_int (Int.repr 2)) - then (Oaddf, r :: r :: nil) - else (Omulf, r1 :: r2 :: nil). - -Definition make_mulfsimm (n: float32) (r r1 r2: reg) := - if Float32.eq_dec n (Float32.of_int (Int.repr 2)) - then (Oaddfs, r :: r :: nil) - else (Omulfs, r1 :: r2 :: nil). - -Definition make_cast8signed (r: reg) (a: aval) := - if vincl a (Sgn Ptop 8) then (Omove, r :: nil) else (Ocast8signed, r :: nil). -Definition make_cast16signed (r: reg) (a: aval) := - if vincl a (Sgn Ptop 16) then (Omove, r :: nil) else (Ocast16signed, r :: nil). - -(** Original definition: -<< -Nondetfunction op_strength_reduction - (op: operation) (args: list reg) (vl: list aval) := - match op, args, vl with - | Ocast8signed, r1 :: nil, v1 :: nil => make_cast8signed r1 v1 - | Ocast16signed, r1 :: nil, v1 :: nil => make_cast16signed r1 v1 - | Oadd, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_addimm n1 r2 - | Oadd, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_addimm n2 r1 - | Osub, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_addimm (Int.neg n2) r1 - | Omul, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_mulimm n1 r2 r1 - | Omul, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_mulimm n2 r1 r2 - | Odiv, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_divimm n2 r1 r2 - | Odivu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_divuimm n2 r1 r2 - | Omodu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_moduimm n2 r1 r2 - | Oand, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_andimm n1 r2 v2 - | Oand, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm n2 r1 v1 - | Oandimm n, r1 :: nil, v1 :: nil => make_andimm n r1 v1 - | Oor, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_orimm n1 r2 - | Oor, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_orimm n2 r1 - | Oxor, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_xorimm n1 r2 - | Oxor, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_xorimm n2 r1 - | Oshl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shlimm n2 r1 r2 - | Oshr, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrimm n2 r1 r2 - | Oshru, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shruimm n2 r1 r2 - | Oaddl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_addlimm n1 r2 - | Oaddl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_addlimm n2 r1 - | Osubl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_addlimm (Int64.neg n2) r1 - | Omull, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_mullimm n1 r2 r1 - | Omull, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_mullimm n2 r1 r2 - | Odivl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_divlimm n2 r1 r2 - | Odivlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_divluimm n2 r1 r2 - | Omodlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_modluimm n2 r1 r2 - | Oandl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_andlimm n1 r2 v2 - | Oandl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_andlimm n2 r1 v1 - | Oandlimm n, r1 :: nil, v1 :: nil => make_andlimm n r1 v1 - | Oorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_orlimm n1 r2 - | Oorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_orlimm n2 r1 - | Oxorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_xorlimm n1 r2 - | Oxorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_xorlimm n2 r1 - | Oshll, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shllimm n2 r1 r2 - | Oshrl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrlimm n2 r1 r2 - | Oshrlu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrluimm n2 r1 r2 - | Ocmp c, args, vl => make_cmp c args vl - | Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => make_mulfimm n2 r1 r1 r2 - | Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil => make_mulfimm n1 r2 r1 r2 - | Omulfs, r1 :: r2 :: nil, v1 :: FS n2 :: nil => make_mulfsimm n2 r1 r1 r2 - | Omulfs, r1 :: r2 :: nil, FS n1 :: v2 :: nil => make_mulfsimm n1 r2 r1 r2 - | _, _, _ => (op, args) - end. ->> -*) - -Inductive op_strength_reduction_cases: forall (op: operation) (args: list reg) (vl: list aval), Type := - | op_strength_reduction_case1: forall r1 v1, op_strength_reduction_cases (Ocast8signed) (r1 :: nil) (v1 :: nil) - | op_strength_reduction_case2: forall r1 v1, op_strength_reduction_cases (Ocast16signed) (r1 :: nil) (v1 :: nil) - | op_strength_reduction_case3: forall r1 r2 n1 v2, op_strength_reduction_cases (Oadd) (r1 :: r2 :: nil) (I n1 :: v2 :: nil) - | op_strength_reduction_case4: forall r1 r2 v1 n2, op_strength_reduction_cases (Oadd) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case5: forall r1 r2 v1 n2, op_strength_reduction_cases (Osub) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case6: forall r1 r2 n1 v2, op_strength_reduction_cases (Omul) (r1 :: r2 :: nil) (I n1 :: v2 :: nil) - | op_strength_reduction_case7: forall r1 r2 v1 n2, op_strength_reduction_cases (Omul) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case8: forall r1 r2 v1 n2, op_strength_reduction_cases (Odiv) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case9: forall r1 r2 v1 n2, op_strength_reduction_cases (Odivu) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case10: forall r1 r2 v1 n2, op_strength_reduction_cases (Omodu) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case11: forall r1 r2 n1 v2, op_strength_reduction_cases (Oand) (r1 :: r2 :: nil) (I n1 :: v2 :: nil) - | op_strength_reduction_case12: forall r1 r2 v1 n2, op_strength_reduction_cases (Oand) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case13: forall n r1 v1, op_strength_reduction_cases (Oandimm n) (r1 :: nil) (v1 :: nil) - | op_strength_reduction_case14: forall r1 r2 n1 v2, op_strength_reduction_cases (Oor) (r1 :: r2 :: nil) (I n1 :: v2 :: nil) - | op_strength_reduction_case15: forall r1 r2 v1 n2, op_strength_reduction_cases (Oor) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case16: forall r1 r2 n1 v2, op_strength_reduction_cases (Oxor) (r1 :: r2 :: nil) (I n1 :: v2 :: nil) - | op_strength_reduction_case17: forall r1 r2 v1 n2, op_strength_reduction_cases (Oxor) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case18: forall r1 r2 v1 n2, op_strength_reduction_cases (Oshl) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case19: forall r1 r2 v1 n2, op_strength_reduction_cases (Oshr) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case20: forall r1 r2 v1 n2, op_strength_reduction_cases (Oshru) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case21: forall r1 r2 n1 v2, op_strength_reduction_cases (Oaddl) (r1 :: r2 :: nil) (L n1 :: v2 :: nil) - | op_strength_reduction_case22: forall r1 r2 v1 n2, op_strength_reduction_cases (Oaddl) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case23: forall r1 r2 v1 n2, op_strength_reduction_cases (Osubl) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case24: forall r1 r2 n1 v2, op_strength_reduction_cases (Omull) (r1 :: r2 :: nil) (L n1 :: v2 :: nil) - | op_strength_reduction_case25: forall r1 r2 v1 n2, op_strength_reduction_cases (Omull) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case26: forall r1 r2 v1 n2, op_strength_reduction_cases (Odivl) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case27: forall r1 r2 v1 n2, op_strength_reduction_cases (Odivlu) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case28: forall r1 r2 v1 n2, op_strength_reduction_cases (Omodlu) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case29: forall r1 r2 n1 v2, op_strength_reduction_cases (Oandl) (r1 :: r2 :: nil) (L n1 :: v2 :: nil) - | op_strength_reduction_case30: forall r1 r2 v1 n2, op_strength_reduction_cases (Oandl) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case31: forall n r1 v1, op_strength_reduction_cases (Oandlimm n) (r1 :: nil) (v1 :: nil) - | op_strength_reduction_case32: forall r1 r2 n1 v2, op_strength_reduction_cases (Oorl) (r1 :: r2 :: nil) (L n1 :: v2 :: nil) - | op_strength_reduction_case33: forall r1 r2 v1 n2, op_strength_reduction_cases (Oorl) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case34: forall r1 r2 n1 v2, op_strength_reduction_cases (Oxorl) (r1 :: r2 :: nil) (L n1 :: v2 :: nil) - | op_strength_reduction_case35: forall r1 r2 v1 n2, op_strength_reduction_cases (Oxorl) (r1 :: r2 :: nil) (v1 :: L n2 :: nil) - | op_strength_reduction_case36: forall r1 r2 v1 n2, op_strength_reduction_cases (Oshll) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case37: forall r1 r2 v1 n2, op_strength_reduction_cases (Oshrl) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case38: forall r1 r2 v1 n2, op_strength_reduction_cases (Oshrlu) (r1 :: r2 :: nil) (v1 :: I n2 :: nil) - | op_strength_reduction_case39: forall c args vl, op_strength_reduction_cases (Ocmp c) (args) (vl) - | op_strength_reduction_case40: forall r1 r2 v1 n2, op_strength_reduction_cases (Omulf) (r1 :: r2 :: nil) (v1 :: F n2 :: nil) - | op_strength_reduction_case41: forall r1 r2 n1 v2, op_strength_reduction_cases (Omulf) (r1 :: r2 :: nil) (F n1 :: v2 :: nil) - | op_strength_reduction_case42: forall r1 r2 v1 n2, op_strength_reduction_cases (Omulfs) (r1 :: r2 :: nil) (v1 :: FS n2 :: nil) - | op_strength_reduction_case43: forall r1 r2 n1 v2, op_strength_reduction_cases (Omulfs) (r1 :: r2 :: nil) (FS n1 :: v2 :: nil) - | op_strength_reduction_default: forall (op: operation) (args: list reg) (vl: list aval), op_strength_reduction_cases op args vl. - -Definition op_strength_reduction_match (op: operation) (args: list reg) (vl: list aval) := - match op as zz1, args as zz2, vl as zz3 return op_strength_reduction_cases zz1 zz2 zz3 with - | Ocast8signed, r1 :: nil, v1 :: nil => op_strength_reduction_case1 r1 v1 - | Ocast16signed, r1 :: nil, v1 :: nil => op_strength_reduction_case2 r1 v1 - | Oadd, r1 :: r2 :: nil, I n1 :: v2 :: nil => op_strength_reduction_case3 r1 r2 n1 v2 - | Oadd, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case4 r1 r2 v1 n2 - | Osub, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case5 r1 r2 v1 n2 - | Omul, r1 :: r2 :: nil, I n1 :: v2 :: nil => op_strength_reduction_case6 r1 r2 n1 v2 - | Omul, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case7 r1 r2 v1 n2 - | Odiv, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case8 r1 r2 v1 n2 - | Odivu, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case9 r1 r2 v1 n2 - | Omodu, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case10 r1 r2 v1 n2 - | Oand, r1 :: r2 :: nil, I n1 :: v2 :: nil => op_strength_reduction_case11 r1 r2 n1 v2 - | Oand, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case12 r1 r2 v1 n2 - | Oandimm n, r1 :: nil, v1 :: nil => op_strength_reduction_case13 n r1 v1 - | Oor, r1 :: r2 :: nil, I n1 :: v2 :: nil => op_strength_reduction_case14 r1 r2 n1 v2 - | Oor, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case15 r1 r2 v1 n2 - | Oxor, r1 :: r2 :: nil, I n1 :: v2 :: nil => op_strength_reduction_case16 r1 r2 n1 v2 - | Oxor, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case17 r1 r2 v1 n2 - | Oshl, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case18 r1 r2 v1 n2 - | Oshr, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case19 r1 r2 v1 n2 - | Oshru, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case20 r1 r2 v1 n2 - | Oaddl, r1 :: r2 :: nil, L n1 :: v2 :: nil => op_strength_reduction_case21 r1 r2 n1 v2 - | Oaddl, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case22 r1 r2 v1 n2 - | Osubl, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case23 r1 r2 v1 n2 - | Omull, r1 :: r2 :: nil, L n1 :: v2 :: nil => op_strength_reduction_case24 r1 r2 n1 v2 - | Omull, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case25 r1 r2 v1 n2 - | Odivl, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case26 r1 r2 v1 n2 - | Odivlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case27 r1 r2 v1 n2 - | Omodlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case28 r1 r2 v1 n2 - | Oandl, r1 :: r2 :: nil, L n1 :: v2 :: nil => op_strength_reduction_case29 r1 r2 n1 v2 - | Oandl, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case30 r1 r2 v1 n2 - | Oandlimm n, r1 :: nil, v1 :: nil => op_strength_reduction_case31 n r1 v1 - | Oorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => op_strength_reduction_case32 r1 r2 n1 v2 - | Oorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case33 r1 r2 v1 n2 - | Oxorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => op_strength_reduction_case34 r1 r2 n1 v2 - | Oxorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => op_strength_reduction_case35 r1 r2 v1 n2 - | Oshll, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case36 r1 r2 v1 n2 - | Oshrl, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case37 r1 r2 v1 n2 - | Oshrlu, r1 :: r2 :: nil, v1 :: I n2 :: nil => op_strength_reduction_case38 r1 r2 v1 n2 - | Ocmp c, args, vl => op_strength_reduction_case39 c args vl - | Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => op_strength_reduction_case40 r1 r2 v1 n2 - | Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil => op_strength_reduction_case41 r1 r2 n1 v2 - | Omulfs, r1 :: r2 :: nil, v1 :: FS n2 :: nil => op_strength_reduction_case42 r1 r2 v1 n2 - | Omulfs, r1 :: r2 :: nil, FS n1 :: v2 :: nil => op_strength_reduction_case43 r1 r2 n1 v2 - | op, args, vl => op_strength_reduction_default op args vl - end. - -Definition op_strength_reduction (op: operation) (args: list reg) (vl: list aval) := - match op_strength_reduction_match op args vl with - | op_strength_reduction_case1 r1 v1 => (* Ocast8signed, r1 :: nil, v1 :: nil *) - make_cast8signed r1 v1 - | op_strength_reduction_case2 r1 v1 => (* Ocast16signed, r1 :: nil, v1 :: nil *) - make_cast16signed r1 v1 - | op_strength_reduction_case3 r1 r2 n1 v2 => (* Oadd, r1 :: r2 :: nil, I n1 :: v2 :: nil *) - make_addimm n1 r2 - | op_strength_reduction_case4 r1 r2 v1 n2 => (* Oadd, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_addimm n2 r1 - | op_strength_reduction_case5 r1 r2 v1 n2 => (* Osub, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_addimm (Int.neg n2) r1 - | op_strength_reduction_case6 r1 r2 n1 v2 => (* Omul, r1 :: r2 :: nil, I n1 :: v2 :: nil *) - make_mulimm n1 r2 r1 - | op_strength_reduction_case7 r1 r2 v1 n2 => (* Omul, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_mulimm n2 r1 r2 - | op_strength_reduction_case8 r1 r2 v1 n2 => (* Odiv, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_divimm n2 r1 r2 - | op_strength_reduction_case9 r1 r2 v1 n2 => (* Odivu, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_divuimm n2 r1 r2 - | op_strength_reduction_case10 r1 r2 v1 n2 => (* Omodu, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_moduimm n2 r1 r2 - | op_strength_reduction_case11 r1 r2 n1 v2 => (* Oand, r1 :: r2 :: nil, I n1 :: v2 :: nil *) - make_andimm n1 r2 v2 - | op_strength_reduction_case12 r1 r2 v1 n2 => (* Oand, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_andimm n2 r1 v1 - | op_strength_reduction_case13 n r1 v1 => (* Oandimm n, r1 :: nil, v1 :: nil *) - make_andimm n r1 v1 - | op_strength_reduction_case14 r1 r2 n1 v2 => (* Oor, r1 :: r2 :: nil, I n1 :: v2 :: nil *) - make_orimm n1 r2 - | op_strength_reduction_case15 r1 r2 v1 n2 => (* Oor, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_orimm n2 r1 - | op_strength_reduction_case16 r1 r2 n1 v2 => (* Oxor, r1 :: r2 :: nil, I n1 :: v2 :: nil *) - make_xorimm n1 r2 - | op_strength_reduction_case17 r1 r2 v1 n2 => (* Oxor, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_xorimm n2 r1 - | op_strength_reduction_case18 r1 r2 v1 n2 => (* Oshl, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_shlimm n2 r1 r2 - | op_strength_reduction_case19 r1 r2 v1 n2 => (* Oshr, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_shrimm n2 r1 r2 - | op_strength_reduction_case20 r1 r2 v1 n2 => (* Oshru, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_shruimm n2 r1 r2 - | op_strength_reduction_case21 r1 r2 n1 v2 => (* Oaddl, r1 :: r2 :: nil, L n1 :: v2 :: nil *) - make_addlimm n1 r2 - | op_strength_reduction_case22 r1 r2 v1 n2 => (* Oaddl, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_addlimm n2 r1 - | op_strength_reduction_case23 r1 r2 v1 n2 => (* Osubl, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_addlimm (Int64.neg n2) r1 - | op_strength_reduction_case24 r1 r2 n1 v2 => (* Omull, r1 :: r2 :: nil, L n1 :: v2 :: nil *) - make_mullimm n1 r2 r1 - | op_strength_reduction_case25 r1 r2 v1 n2 => (* Omull, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_mullimm n2 r1 r2 - | op_strength_reduction_case26 r1 r2 v1 n2 => (* Odivl, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_divlimm n2 r1 r2 - | op_strength_reduction_case27 r1 r2 v1 n2 => (* Odivlu, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_divluimm n2 r1 r2 - | op_strength_reduction_case28 r1 r2 v1 n2 => (* Omodlu, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_modluimm n2 r1 r2 - | op_strength_reduction_case29 r1 r2 n1 v2 => (* Oandl, r1 :: r2 :: nil, L n1 :: v2 :: nil *) - make_andlimm n1 r2 v2 - | op_strength_reduction_case30 r1 r2 v1 n2 => (* Oandl, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_andlimm n2 r1 v1 - | op_strength_reduction_case31 n r1 v1 => (* Oandlimm n, r1 :: nil, v1 :: nil *) - make_andlimm n r1 v1 - | op_strength_reduction_case32 r1 r2 n1 v2 => (* Oorl, r1 :: r2 :: nil, L n1 :: v2 :: nil *) - make_orlimm n1 r2 - | op_strength_reduction_case33 r1 r2 v1 n2 => (* Oorl, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_orlimm n2 r1 - | op_strength_reduction_case34 r1 r2 n1 v2 => (* Oxorl, r1 :: r2 :: nil, L n1 :: v2 :: nil *) - make_xorlimm n1 r2 - | op_strength_reduction_case35 r1 r2 v1 n2 => (* Oxorl, r1 :: r2 :: nil, v1 :: L n2 :: nil *) - make_xorlimm n2 r1 - | op_strength_reduction_case36 r1 r2 v1 n2 => (* Oshll, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_shllimm n2 r1 r2 - | op_strength_reduction_case37 r1 r2 v1 n2 => (* Oshrl, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_shrlimm n2 r1 r2 - | op_strength_reduction_case38 r1 r2 v1 n2 => (* Oshrlu, r1 :: r2 :: nil, v1 :: I n2 :: nil *) - make_shrluimm n2 r1 r2 - | op_strength_reduction_case39 c args vl => (* Ocmp c, args, vl *) - make_cmp c args vl - | op_strength_reduction_case40 r1 r2 v1 n2 => (* Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil *) - make_mulfimm n2 r1 r1 r2 - | op_strength_reduction_case41 r1 r2 n1 v2 => (* Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil *) - make_mulfimm n1 r2 r1 r2 - | op_strength_reduction_case42 r1 r2 v1 n2 => (* Omulfs, r1 :: r2 :: nil, v1 :: FS n2 :: nil *) - make_mulfsimm n2 r1 r1 r2 - | op_strength_reduction_case43 r1 r2 n1 v2 => (* Omulfs, r1 :: r2 :: nil, FS n1 :: v2 :: nil *) - make_mulfsimm n1 r2 r1 r2 - | op_strength_reduction_default op args vl => - (op, args) - end. - - -(** Original definition: -<< -Nondetfunction addr_strength_reduction - (addr: addressing) (args: list reg) (vl: list aval) := - match addr, args, vl with - | Aindexed n, r1 :: nil, Ptr(Gl symb n1) :: nil => - if Archi.pic_code tt - then (addr, args) - else (Aglobal symb (Ptrofs.add n1 n), nil) - | Aindexed n, r1 :: nil, Ptr(Stk n1) :: nil => - (Ainstack (Ptrofs.add n1 n), nil) - | _, _, _ => - (addr, args) - end. ->> -*) - -Inductive addr_strength_reduction_cases: forall (addr: addressing) (args: list reg) (vl: list aval), Type := - | addr_strength_reduction_case1: forall n r1 symb n1, addr_strength_reduction_cases (Aindexed n) (r1 :: nil) (Ptr(Gl symb n1) :: nil) - | addr_strength_reduction_case2: forall n r1 n1, addr_strength_reduction_cases (Aindexed n) (r1 :: nil) (Ptr(Stk n1) :: nil) - | addr_strength_reduction_default: forall (addr: addressing) (args: list reg) (vl: list aval), addr_strength_reduction_cases addr args vl. - -Definition addr_strength_reduction_match (addr: addressing) (args: list reg) (vl: list aval) := - match addr as zz1, args as zz2, vl as zz3 return addr_strength_reduction_cases zz1 zz2 zz3 with - | Aindexed n, r1 :: nil, Ptr(Gl symb n1) :: nil => addr_strength_reduction_case1 n r1 symb n1 - | Aindexed n, r1 :: nil, Ptr(Stk n1) :: nil => addr_strength_reduction_case2 n r1 n1 - | addr, args, vl => addr_strength_reduction_default addr args vl - end. - -Definition addr_strength_reduction (addr: addressing) (args: list reg) (vl: list aval) := - match addr_strength_reduction_match addr args vl with - | addr_strength_reduction_case1 n r1 symb n1 => (* Aindexed n, r1 :: nil, Ptr(Gl symb n1) :: nil *) - if Archi.pic_code tt then (addr, args) else (Aglobal symb (Ptrofs.add n1 n), nil) - | addr_strength_reduction_case2 n r1 n1 => (* Aindexed n, r1 :: nil, Ptr(Stk n1) :: nil *) - (Ainstack (Ptrofs.add n1 n), nil) - | addr_strength_reduction_default addr args vl => - (addr, args) - end. - - -- cgit From 5a425ba34f3f2520de752ea95a4636a9437c312a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 08:12:41 +0100 Subject: ça recompile sur x86 MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Makefile | 2 +- backend/OpHelpers.v | 42 ++++++++++++++++++++++++++ backend/OpHelpersproof.v | 78 ++++++++++++++++++++++++++++++++++++++++++++++++ backend/SelectDiv.vp | 2 +- backend/SelectDivproof.v | 1 + backend/Selection.v | 2 +- backend/Selectionproof.v | 1 + backend/SplitLong.vp | 1 + backend/SplitLongproof.v | 3 +- mppa_k1c/SelectOp.vp | 39 +----------------------- mppa_k1c/SelectOpproof.v | 65 ---------------------------------------- x86/SelectLong.vp | 2 +- x86/SelectLongproof.v | 1 + x86/SelectOp.vp | 8 +++++ x86/SelectOpproof.v | 29 ++++++++++++++++-- 15 files changed, 166 insertions(+), 110 deletions(-) create mode 100644 backend/OpHelpers.v create mode 100644 backend/OpHelpersproof.v diff --git a/Makefile b/Makefile index ed9059db..d8cd428a 100644 --- a/Makefile +++ b/Makefile @@ -72,7 +72,7 @@ COMMON=Errors.v AST.v Linking.v \ # Back-end modules (in backend/, $(ARCH)/) BACKEND=\ - Cminor.v Op.v CminorSel.v \ + Cminor.v Op.v CminorSel.v OpHelpers.v OpHelpersproof.v \ SelectOp.v SelectDiv.v SplitLong.v SelectLong.v Selection.v \ SelectOpproof.v SelectDivproof.v SplitLongproof.v \ SelectLongproof.v Selectionproof.v \ diff --git a/backend/OpHelpers.v b/backend/OpHelpers.v new file mode 100644 index 00000000..53414dab --- /dev/null +++ b/backend/OpHelpers.v @@ -0,0 +1,42 @@ +Require Import Coqlib. +Require Import AST Integers Floats. +Require Import Op CminorSel. + +(** Some arithmetic operations are transformed into calls to + runtime library functions. The following type class collects + the names of these functions. *) + +Definition sig_l_l := mksignature (Tlong :: nil) (Some Tlong) cc_default. +Definition sig_l_f := mksignature (Tlong :: nil) (Some Tfloat) cc_default. +Definition sig_l_s := mksignature (Tlong :: nil) (Some Tsingle) cc_default. +Definition sig_f_l := mksignature (Tfloat :: nil) (Some Tlong) cc_default. +Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_default. +Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default. +Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default. +Definition sig_ii_i := mksignature (Tint :: Tint :: nil) (Some Tint) cc_default. +Definition sig_ff_f := mksignature (Tfloat :: Tfloat :: nil) (Some Tfloat) cc_default. +Definition sig_ss_s := mksignature (Tsingle :: Tsingle :: nil) (Some Tsingle) cc_default. + +Class helper_functions := mk_helper_functions { + i64_dtos: ident; (**r float64 -> signed long *) + i64_dtou: ident; (**r float64 -> unsigned long *) + i64_stod: ident; (**r signed long -> float64 *) + i64_utod: ident; (**r unsigned long -> float64 *) + i64_stof: ident; (**r signed long -> float32 *) + i64_utof: ident; (**r unsigned long -> float32 *) + i64_sdiv: ident; (**r signed division *) + i64_udiv: ident; (**r unsigned division *) + i64_smod: ident; (**r signed remainder *) + i64_umod: ident; (**r unsigned remainder *) + i64_shl: ident; (**r shift left *) + i64_shr: ident; (**r shift right unsigned *) + i64_sar: ident; (**r shift right signed *) + i64_umulh: ident; (**r unsigned multiply high *) + i64_smulh: ident; (**r signed multiply high *) + i32_sdiv: ident; (**r signed division *) + i32_udiv: ident; (**r unsigned division *) + i32_smod: ident; (**r signed remainder *) + i32_umod: ident; (**r unsigned remainder *) + f64_div: ident; (**float division*) + f32_div: ident; (**float division*) +}. diff --git a/backend/OpHelpersproof.v b/backend/OpHelpersproof.v new file mode 100644 index 00000000..63040c5f --- /dev/null +++ b/backend/OpHelpersproof.v @@ -0,0 +1,78 @@ +Require Import Coqlib. +Require Import Maps. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Import Memory. +Require Import Globalenvs. +Require Import Cminor. +Require Import Op. +Require Import CminorSel. +Require Import Events. +Require Import OpHelpers. + +(** * Axiomatization of the helper functions *) + +Definition external_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop := + forall F V (ge: Genv.t F V) m, + external_call (EF_runtime name sg) ge vargs m E0 vres m. + +Definition builtin_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop := + forall F V (ge: Genv.t F V) m, + external_call (EF_builtin name sg) ge vargs m E0 vres m. + +Axiom arith_helpers_correct : + (forall x z, Val.longoffloat x = Some z -> external_implements "__compcert_i64_dtos" sig_f_l (x::nil) z) + /\ (forall x z, Val.longuoffloat x = Some z -> external_implements "__compcert_i64_dtou" sig_f_l (x::nil) z) + /\ (forall x z, Val.floatoflong x = Some z -> external_implements "__compcert_i64_stod" sig_l_f (x::nil) z) + /\ (forall x z, Val.floatoflongu x = Some z -> external_implements "__compcert_i64_utod" sig_l_f (x::nil) z) + /\ (forall x z, Val.singleoflong x = Some z -> external_implements "__compcert_i64_stof" sig_l_s (x::nil) z) + /\ (forall x z, Val.singleoflongu x = Some z -> external_implements "__compcert_i64_utof" sig_l_s (x::nil) z) + /\ (forall x, builtin_implements "__builtin_negl" sig_l_l (x::nil) (Val.negl x)) + /\ (forall x y, builtin_implements "__builtin_addl" sig_ll_l (x::y::nil) (Val.addl x y)) + /\ (forall x y, builtin_implements "__builtin_subl" sig_ll_l (x::y::nil) (Val.subl x y)) + /\ (forall x y, builtin_implements "__builtin_mull" sig_ii_l (x::y::nil) (Val.mull' x y)) + /\ (forall x y z, Val.divls x y = Some z -> external_implements "__compcert_i64_sdiv" sig_ll_l (x::y::nil) z) + /\ (forall x y z, Val.divlu x y = Some z -> external_implements "__compcert_i64_udiv" sig_ll_l (x::y::nil) z) + /\ (forall x y z, Val.modls x y = Some z -> external_implements "__compcert_i64_smod" sig_ll_l (x::y::nil) z) + /\ (forall x y z, Val.modlu x y = Some z -> external_implements "__compcert_i64_umod" sig_ll_l (x::y::nil) z) + /\ (forall x y, external_implements "__compcert_i64_shl" sig_li_l (x::y::nil) (Val.shll x y)) + /\ (forall x y, external_implements "__compcert_i64_shr" sig_li_l (x::y::nil) (Val.shrlu x y)) + /\ (forall x y, external_implements "__compcert_i64_sar" sig_li_l (x::y::nil) (Val.shrl x y)) + /\ (forall x y, external_implements "__compcert_i64_umulh" sig_ll_l (x::y::nil) (Val.mullhu x y)) + /\ (forall x y, external_implements "__compcert_i64_smulh" sig_ll_l (x::y::nil) (Val.mullhs x y)) + /\ (forall x y z, Val.divs x y = Some z -> external_implements "__compcert_i32_sdiv" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.divu x y = Some z -> external_implements "__compcert_i32_udiv" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.mods x y = Some z -> external_implements "__compcert_i32_smod" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.modu x y = Some z -> external_implements "__compcert_i32_umod" sig_ii_i (x::y::nil) z) + /\ (forall x y z, Val.divfs x y = z -> external_implements "__compcert_f32_div" sig_ss_s (x::y::nil) z) + /\ (forall x y z, Val.divf x y = z -> external_implements "__compcert_f64_div" sig_ff_f (x::y::nil) z) +. + +Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop := + (prog_defmap p)!id = Some (Gfun (External (EF_runtime name sg))). + +Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F) V) (hf: helper_functions) : Prop := + helper_declared p i64_dtos "__compcert_i64_dtos" sig_f_l + /\ helper_declared p i64_dtou "__compcert_i64_dtou" sig_f_l + /\ helper_declared p i64_stod "__compcert_i64_stod" sig_l_f + /\ helper_declared p i64_utod "__compcert_i64_utod" sig_l_f + /\ helper_declared p i64_stof "__compcert_i64_stof" sig_l_s + /\ helper_declared p i64_utof "__compcert_i64_utof" sig_l_s + /\ helper_declared p i64_sdiv "__compcert_i64_sdiv" sig_ll_l + /\ helper_declared p i64_udiv "__compcert_i64_udiv" sig_ll_l + /\ helper_declared p i64_smod "__compcert_i64_smod" sig_ll_l + /\ helper_declared p i64_umod "__compcert_i64_umod" sig_ll_l + /\ helper_declared p i64_shl "__compcert_i64_shl" sig_li_l + /\ helper_declared p i64_shr "__compcert_i64_shr" sig_li_l + /\ helper_declared p i64_sar "__compcert_i64_sar" sig_li_l + /\ helper_declared p i64_umulh "__compcert_i64_umulh" sig_ll_l + /\ helper_declared p i64_smulh "__compcert_i64_smulh" sig_ll_l + /\ helper_declared p i32_sdiv "__compcert_i32_sdiv" sig_ii_i + /\ helper_declared p i32_udiv "__compcert_i32_udiv" sig_ii_i + /\ helper_declared p i32_smod "__compcert_i32_smod" sig_ii_i + /\ helper_declared p i32_umod "__compcert_i32_umod" sig_ii_i + /\ helper_declared p f32_div "__compcert_f32_div" sig_ss_s + /\ helper_declared p f64_div "__compcert_f64_div" sig_ff_f +. diff --git a/backend/SelectDiv.vp b/backend/SelectDiv.vp index 7241d028..9f852616 100644 --- a/backend/SelectDiv.vp +++ b/backend/SelectDiv.vp @@ -15,7 +15,7 @@ Require Import Coqlib. Require Import Compopts. Require Import AST Integers Floats. -Require Import Op CminorSel SelectOp SplitLong SelectLong. +Require Import Op CminorSel OpHelpers SelectOp SplitLong SelectLong. Local Open Scope cminorsel_scope. diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v index dbd3f58d..e2249ddb 100644 --- a/backend/SelectDivproof.v +++ b/backend/SelectDivproof.v @@ -15,6 +15,7 @@ Require Import Zquot Coqlib. Require Import AST Integers Floats Values Memory Globalenvs Events. Require Import Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. Require Import SelectOp SelectOpproof SplitLong SplitLongproof SelectLong SelectLongproof SelectDiv. Local Open Scope cminorsel_scope. diff --git a/backend/Selection.v b/backend/Selection.v index aba84049..05a06abf 100644 --- a/backend/Selection.v +++ b/backend/Selection.v @@ -26,7 +26,7 @@ Require String. Require Import Coqlib Maps. Require Import AST Errors Integers Globalenvs Switch. Require Cminor. -Require Import Op CminorSel. +Require Import Op CminorSel OpHelpers. Require Import SelectOp SplitLong SelectLong SelectDiv. Require Machregs. diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v index 5a0fafaf..50ff377a 100644 --- a/backend/Selectionproof.v +++ b/backend/Selectionproof.v @@ -16,6 +16,7 @@ Require Import FunInd. Require Import Coqlib Maps. Require Import AST Linking Errors Integers Values Memory Events Globalenvs Smallstep. Require Import Switch Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. Require Import SelectOp SelectDiv SplitLong SelectLong Selection. Require Import SelectOpproof SelectDivproof SplitLongproof SelectLongproof. diff --git a/backend/SplitLong.vp b/backend/SplitLong.vp index 6a7d4f5c..dfe42df0 100644 --- a/backend/SplitLong.vp +++ b/backend/SplitLong.vp @@ -16,6 +16,7 @@ Require String. Require Import Coqlib. Require Import AST Integers Floats. Require Import Op CminorSel. +Require Import OpHelpers. Require Import SelectOp. Local Open Scope cminorsel_scope. diff --git a/backend/SplitLongproof.v b/backend/SplitLongproof.v index a74276c7..6718ba5b 100644 --- a/backend/SplitLongproof.v +++ b/backend/SplitLongproof.v @@ -16,6 +16,7 @@ Require Import String. Require Import Coqlib Maps. Require Import AST Errors Integers Floats. Require Import Values Memory Globalenvs Events Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. Require Import SelectOp SelectOpproof SplitLong. Local Open Scope cminorsel_scope. @@ -33,7 +34,7 @@ Variable sp: val. Variable e: env. Variable m: mem. -Ltac UseHelper := decompose [Logic.and] i64_helpers_correct; eauto. +Ltac UseHelper := decompose [Logic.and] arith_helpers_correct; eauto. Ltac DeclHelper := red in HELPERS; decompose [Logic.and] HELPERS; eauto. Lemma eval_helper: diff --git a/mppa_k1c/SelectOp.vp b/mppa_k1c/SelectOp.vp index 19712d2f..f6605c11 100644 --- a/mppa_k1c/SelectOp.vp +++ b/mppa_k1c/SelectOp.vp @@ -50,6 +50,7 @@ Require Import Integers. Require Import Floats. Require Import Op. Require Import CminorSel. +Require Import OpHelpers. Local Open Scope cminorsel_scope. @@ -57,44 +58,6 @@ Local Open Scope string_scope. Local Open Scope error_monad_scope. Section SELECT. -(** Some operations on 64-bit integers are transformed into calls to - runtime library functions. The following type class collects - the names of these functions. *) - -Definition sig_l_l := mksignature (Tlong :: nil) (Some Tlong) cc_default. -Definition sig_l_f := mksignature (Tlong :: nil) (Some Tfloat) cc_default. -Definition sig_l_s := mksignature (Tlong :: nil) (Some Tsingle) cc_default. -Definition sig_f_l := mksignature (Tfloat :: nil) (Some Tlong) cc_default. -Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_default. -Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default. -Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default. -Definition sig_ii_i := mksignature (Tint :: Tint :: nil) (Some Tint) cc_default. -Definition sig_ff_f := mksignature (Tfloat :: Tfloat :: nil) (Some Tfloat) cc_default. -Definition sig_ss_s := mksignature (Tsingle :: Tsingle :: nil) (Some Tsingle) cc_default. - -Class helper_functions := mk_helper_functions { - i64_dtos: ident; (**r float64 -> signed long *) - i64_dtou: ident; (**r float64 -> unsigned long *) - i64_stod: ident; (**r signed long -> float64 *) - i64_utod: ident; (**r unsigned long -> float64 *) - i64_stof: ident; (**r signed long -> float32 *) - i64_utof: ident; (**r unsigned long -> float32 *) - i64_sdiv: ident; (**r signed division *) - i64_udiv: ident; (**r unsigned division *) - i64_smod: ident; (**r signed remainder *) - i64_umod: ident; (**r unsigned remainder *) - i64_shl: ident; (**r shift left *) - i64_shr: ident; (**r shift right unsigned *) - i64_sar: ident; (**r shift right signed *) - i64_umulh: ident; (**r unsigned multiply high *) - i64_smulh: ident; (**r signed multiply high *) - i32_sdiv: ident; (**r signed division *) - i32_udiv: ident; (**r unsigned division *) - i32_smod: ident; (**r signed remainder *) - i32_umod: ident; (**r unsigned remainder *) - f64_div: ident; (**float division*) - f32_div: ident; (**float division*) -}. Context {hf: helper_functions}. diff --git a/mppa_k1c/SelectOpproof.v b/mppa_k1c/SelectOpproof.v index a4a5f72b..111c08e4 100644 --- a/mppa_k1c/SelectOpproof.v +++ b/mppa_k1c/SelectOpproof.v @@ -35,71 +35,6 @@ Local Open Scope cminorsel_scope. Local Open Scope string_scope. -(** * Axiomatization of the helper functions *) - -Definition external_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop := - forall F V (ge: Genv.t F V) m, - external_call (EF_runtime name sg) ge vargs m E0 vres m. - -Definition builtin_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop := - forall F V (ge: Genv.t F V) m, - external_call (EF_builtin name sg) ge vargs m E0 vres m. - -Axiom i64_helpers_correct : - (forall x z, Val.longoffloat x = Some z -> external_implements "__compcert_i64_dtos" sig_f_l (x::nil) z) - /\ (forall x z, Val.longuoffloat x = Some z -> external_implements "__compcert_i64_dtou" sig_f_l (x::nil) z) - /\ (forall x z, Val.floatoflong x = Some z -> external_implements "__compcert_i64_stod" sig_l_f (x::nil) z) - /\ (forall x z, Val.floatoflongu x = Some z -> external_implements "__compcert_i64_utod" sig_l_f (x::nil) z) - /\ (forall x z, Val.singleoflong x = Some z -> external_implements "__compcert_i64_stof" sig_l_s (x::nil) z) - /\ (forall x z, Val.singleoflongu x = Some z -> external_implements "__compcert_i64_utof" sig_l_s (x::nil) z) - /\ (forall x, builtin_implements "__builtin_negl" sig_l_l (x::nil) (Val.negl x)) - /\ (forall x y, builtin_implements "__builtin_addl" sig_ll_l (x::y::nil) (Val.addl x y)) - /\ (forall x y, builtin_implements "__builtin_subl" sig_ll_l (x::y::nil) (Val.subl x y)) - /\ (forall x y, builtin_implements "__builtin_mull" sig_ii_l (x::y::nil) (Val.mull' x y)) - /\ (forall x y z, Val.divls x y = Some z -> external_implements "__compcert_i64_sdiv" sig_ll_l (x::y::nil) z) - /\ (forall x y z, Val.divlu x y = Some z -> external_implements "__compcert_i64_udiv" sig_ll_l (x::y::nil) z) - /\ (forall x y z, Val.modls x y = Some z -> external_implements "__compcert_i64_smod" sig_ll_l (x::y::nil) z) - /\ (forall x y z, Val.modlu x y = Some z -> external_implements "__compcert_i64_umod" sig_ll_l (x::y::nil) z) - /\ (forall x y, external_implements "__compcert_i64_shl" sig_li_l (x::y::nil) (Val.shll x y)) - /\ (forall x y, external_implements "__compcert_i64_shr" sig_li_l (x::y::nil) (Val.shrlu x y)) - /\ (forall x y, external_implements "__compcert_i64_sar" sig_li_l (x::y::nil) (Val.shrl x y)) - /\ (forall x y, external_implements "__compcert_i64_umulh" sig_ll_l (x::y::nil) (Val.mullhu x y)) - /\ (forall x y, external_implements "__compcert_i64_smulh" sig_ll_l (x::y::nil) (Val.mullhs x y)) - /\ (forall x y z, Val.divs x y = Some z -> external_implements "__compcert_i32_sdiv" sig_ii_i (x::y::nil) z) - /\ (forall x y z, Val.divu x y = Some z -> external_implements "__compcert_i32_udiv" sig_ii_i (x::y::nil) z) - /\ (forall x y z, Val.mods x y = Some z -> external_implements "__compcert_i32_smod" sig_ii_i (x::y::nil) z) - /\ (forall x y z, Val.modu x y = Some z -> external_implements "__compcert_i32_umod" sig_ii_i (x::y::nil) z) - /\ (forall x y z, Val.divfs x y = z -> external_implements "__compcert_f32_div" sig_ss_s (x::y::nil) z) - /\ (forall x y z, Val.divf x y = z -> external_implements "__compcert_f64_div" sig_ff_f (x::y::nil) z) -. - -Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop := - (prog_defmap p)!id = Some (Gfun (External (EF_runtime name sg))). - -Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F) V) (hf: helper_functions) : Prop := - helper_declared p i64_dtos "__compcert_i64_dtos" sig_f_l - /\ helper_declared p i64_dtou "__compcert_i64_dtou" sig_f_l - /\ helper_declared p i64_stod "__compcert_i64_stod" sig_l_f - /\ helper_declared p i64_utod "__compcert_i64_utod" sig_l_f - /\ helper_declared p i64_stof "__compcert_i64_stof" sig_l_s - /\ helper_declared p i64_utof "__compcert_i64_utof" sig_l_s - /\ helper_declared p i64_sdiv "__compcert_i64_sdiv" sig_ll_l - /\ helper_declared p i64_udiv "__compcert_i64_udiv" sig_ll_l - /\ helper_declared p i64_smod "__compcert_i64_smod" sig_ll_l - /\ helper_declared p i64_umod "__compcert_i64_umod" sig_ll_l - /\ helper_declared p i64_shl "__compcert_i64_shl" sig_li_l - /\ helper_declared p i64_shr "__compcert_i64_shr" sig_li_l - /\ helper_declared p i64_sar "__compcert_i64_sar" sig_li_l - /\ helper_declared p i64_umulh "__compcert_i64_umulh" sig_ll_l - /\ helper_declared p i64_smulh "__compcert_i64_smulh" sig_ll_l - /\ helper_declared p i32_sdiv "__compcert_i32_sdiv" sig_ii_i - /\ helper_declared p i32_udiv "__compcert_i32_udiv" sig_ii_i - /\ helper_declared p i32_smod "__compcert_i32_smod" sig_ii_i - /\ helper_declared p i32_umod "__compcert_i32_umod" sig_ii_i - /\ helper_declared p f32_div "__compcert_f32_div" sig_ss_s - /\ helper_declared p f64_div "__compcert_f64_div" sig_ff_f -. - (** * Useful lemmas and tactics *) (** The following are trivial lemmas and custom tactics that help diff --git a/x86/SelectLong.vp b/x86/SelectLong.vp index b213e23f..4f9fb518 100644 --- a/x86/SelectLong.vp +++ b/x86/SelectLong.vp @@ -16,7 +16,7 @@ Require Import Coqlib. Require Import Compopts. Require Import AST Integers Floats. Require Import Op CminorSel. -Require Import SelectOp SplitLong. +Require Import OpHelpers SelectOp SplitLong. Local Open Scope cminorsel_scope. Local Open Scope string_scope. diff --git a/x86/SelectLongproof.v b/x86/SelectLongproof.v index 3bef632d..f008f39e 100644 --- a/x86/SelectLongproof.v +++ b/x86/SelectLongproof.v @@ -16,6 +16,7 @@ Require Import String Coqlib Maps Integers Floats Errors. Require Archi. Require Import AST Values Memory Globalenvs Events. Require Import Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. Require Import SelectOp SelectOpproof SplitLong SplitLongproof. Require Import SelectLong. diff --git a/x86/SelectOp.vp b/x86/SelectOp.vp index a1583600..eadda093 100644 --- a/x86/SelectOp.vp +++ b/x86/SelectOp.vp @@ -40,6 +40,7 @@ Require Import Coqlib. Require Import Compopts. Require Import AST Integers Floats. Require Import Op CminorSel. +Require Import OpHelpers. Local Open Scope cminorsel_scope. @@ -540,3 +541,10 @@ Nondetfunction builtin_arg (e: expr) := | Eload chunk (Ainstack ofs) Enil => BA_loadstack chunk ofs | _ => BA e end. + +(* floats *) +Definition divf_base (e1: expr) (e2: expr) := + Eop Odivf (e1 ::: e2 ::: Enil). + +Definition divfs_base (e1: expr) (e2: expr) := + Eop Odivfs (e1 ::: e2 ::: Enil). diff --git a/x86/SelectOpproof.v b/x86/SelectOpproof.v index fdbadaa8..1eeb5906 100644 --- a/x86/SelectOpproof.v +++ b/x86/SelectOpproof.v @@ -23,6 +23,8 @@ Require Import Cminor. Require Import Op. Require Import CminorSel. Require Import SelectOp. +Require Import OpHelpers. +Require Import OpHelpersproof. Local Open Scope cminorsel_scope. @@ -74,8 +76,10 @@ Ltac TrivialExists := (** * Correctness of the smart constructors *) Section CMCONSTR. - -Variable ge: genv. +Variable prog: program. +Variable hf: helper_functions. +Hypothesis HELPERS: helper_functions_declared prog hf. +Let ge := Genv.globalenv prog. Variable sp: val. Variable e: env. Variable m: mem. @@ -984,4 +988,25 @@ Proof. - constructor; auto. Qed. + +(* floating-point division without HELPERS *) +Theorem eval_divf_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v. +Proof. + intros; unfold divf_base. + TrivialExists. +Qed. + +Theorem eval_divfs_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v. +Proof. + intros; unfold divfs_base. + TrivialExists. +Qed. End CMCONSTR. -- cgit From 1fd8a013a4fa57fcfa6c9e295638871153e596ee Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 09:27:40 +0100 Subject: restructuration pour compilation toute archi --- mppa_k1c/SelectLong.vp | 1 + mppa_k1c/SelectLongproof.v | 1 + mppa_k1c/SelectOpproof.v | 4 +++- 3 files changed, 5 insertions(+), 1 deletion(-) diff --git a/mppa_k1c/SelectLong.vp b/mppa_k1c/SelectLong.vp index 3a17ab17..0c3618d7 100644 --- a/mppa_k1c/SelectLong.vp +++ b/mppa_k1c/SelectLong.vp @@ -21,6 +21,7 @@ Require Import Coqlib. Require Import Compopts. Require Import AST Integers Floats. Require Import Op CminorSel. +Require Import OpHelpers. Require Import SelectOp SplitLong. Local Open Scope cminorsel_scope. diff --git a/mppa_k1c/SelectLongproof.v b/mppa_k1c/SelectLongproof.v index 4723278a..79187338 100644 --- a/mppa_k1c/SelectLongproof.v +++ b/mppa_k1c/SelectLongproof.v @@ -21,6 +21,7 @@ Require Import String Coqlib Maps Integers Floats Errors. Require Archi. Require Import AST Values Memory Globalenvs Events. Require Import Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. Require Import SelectOp SelectOpproof SplitLong SplitLongproof. Require Import SelectLong. diff --git a/mppa_k1c/SelectOpproof.v b/mppa_k1c/SelectOpproof.v index 111c08e4..89af39ee 100644 --- a/mppa_k1c/SelectOpproof.v +++ b/mppa_k1c/SelectOpproof.v @@ -30,6 +30,8 @@ Require Import Op. Require Import CminorSel. Require Import SelectOp. Require Import Events. +Require Import OpHelpers. +Require Import OpHelpersproof. Local Open Scope cminorsel_scope. Local Open Scope string_scope. @@ -93,7 +95,7 @@ Variable m: mem. (* Helper lemmas - from SplitLongproof.v *) -Ltac UseHelper := decompose [Logic.and] i64_helpers_correct; eauto. +Ltac UseHelper := decompose [Logic.and] arith_helpers_correct; eauto. Ltac DeclHelper := red in HELPERS; decompose [Logic.and] HELPERS; eauto. Lemma eval_helper: -- cgit From 5281141d38a492eea3e080e087b91a314b0be020 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 09:51:56 +0100 Subject: ARM repasse --- arm/SelectLong.vp | 2 +- arm/SelectLongproof.v | 1 + arm/SelectOp.vp | 8 ++++++++ arm/SelectOpproof.v | 28 ++++++++++++++++++++++++++-- 4 files changed, 36 insertions(+), 3 deletions(-) diff --git a/arm/SelectLong.vp b/arm/SelectLong.vp index cc7a38f6..b4cdd0e3 100644 --- a/arm/SelectLong.vp +++ b/arm/SelectLong.vp @@ -16,6 +16,6 @@ Require Import Coqlib. Require Import Compopts. Require Import AST Integers Floats. Require Import Op CminorSel. -Require Import SelectOp SplitLong. +Require Import OpHelpers SelectOp SplitLong. (** This file is empty because we use the default implementation provided in [SplitLong]. *) diff --git a/arm/SelectLongproof.v b/arm/SelectLongproof.v index a82c082c..a65a38d4 100644 --- a/arm/SelectLongproof.v +++ b/arm/SelectLongproof.v @@ -16,6 +16,7 @@ Require Import String Coqlib Maps Integers Floats Errors. Require Archi. Require Import AST Values Memory Globalenvs Events. Require Import Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. Require Import SelectOp SelectOpproof SplitLong SplitLongproof. Require Import SelectLong. diff --git a/arm/SelectOp.vp b/arm/SelectOp.vp index c361df65..f3f01730 100644 --- a/arm/SelectOp.vp +++ b/arm/SelectOp.vp @@ -42,6 +42,7 @@ Require Import AST. Require Import Integers. Require Import Floats. Require Import Op. +Require Import OpHelpers. Require Import CminorSel. Local Open Scope cminorsel_scope. @@ -508,3 +509,10 @@ Nondetfunction builtin_arg (e: expr) := | Eop (Oaddimm n) (e1:::Enil) => BA_addptr (BA e1) (BA_int n) | _ => BA e end. + +(* floats *) +Definition divf_base (e1: expr) (e2: expr) := + Eop Odivf (e1 ::: e2 ::: Enil). + +Definition divfs_base (e1: expr) (e2: expr) := + Eop Odivfs (e1 ::: e2 ::: Enil). diff --git a/arm/SelectOpproof.v b/arm/SelectOpproof.v index d4aac9b6..212bcfd7 100644 --- a/arm/SelectOpproof.v +++ b/arm/SelectOpproof.v @@ -24,6 +24,7 @@ Require Import Cminor. Require Import Op. Require Import CminorSel. Require Import SelectOp. +Require Import OpHelpers OpHelpersproof. Local Open Scope cminorsel_scope. Local Transparent Archi.ptr64. @@ -76,8 +77,10 @@ Ltac TrivialExists := (** * Correctness of the smart constructors *) Section CMCONSTR. - -Variable ge: genv. +Variable prog: program. +Variable hf: helper_functions. +Hypothesis HELPERS: helper_functions_declared prog hf. +Let ge := Genv.globalenv prog. Variable sp: val. Variable e: env. Variable m: mem. @@ -893,4 +896,25 @@ Proof. - constructor; auto. Qed. + +(* floating-point division without HELPERS *) +Theorem eval_divf_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v. +Proof. + intros; unfold divf_base. + TrivialExists. +Qed. + +Theorem eval_divfs_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v. +Proof. + intros; unfold divfs_base. + TrivialExists. +Qed. End CMCONSTR. -- cgit From 4d53251d4e9a07da1170f4be9b04b5d44381c9ed Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 10:16:16 +0100 Subject: Risc-V works again (32/64). --- riscV/SelectLong.vp | 2 +- riscV/SelectLongproof.v | 1 + riscV/SelectOp.vp | 7 +++++++ riscV/SelectOpproof.v | 29 +++++++++++++++++++++++++++-- 4 files changed, 36 insertions(+), 3 deletions(-) diff --git a/riscV/SelectLong.vp b/riscV/SelectLong.vp index b3e07bf5..0ccc4725 100644 --- a/riscV/SelectLong.vp +++ b/riscV/SelectLong.vp @@ -21,7 +21,7 @@ Require Import Coqlib. Require Import Compopts. Require Import AST Integers Floats. Require Import Op CminorSel. -Require Import SelectOp SplitLong. +Require Import OpHelpers SelectOp SplitLong. Local Open Scope cminorsel_scope. Local Open Scope string_scope. diff --git a/riscV/SelectLongproof.v b/riscV/SelectLongproof.v index 78a1935d..d47b6d64 100644 --- a/riscV/SelectLongproof.v +++ b/riscV/SelectLongproof.v @@ -21,6 +21,7 @@ Require Import String Coqlib Maps Integers Floats Errors. Require Archi. Require Import AST Values Memory Globalenvs Events. Require Import Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. Require Import SelectOp SelectOpproof SplitLong SplitLongproof. Require Import SelectLong. diff --git a/riscV/SelectOp.vp b/riscV/SelectOp.vp index bb8af2ed..181b9d05 100644 --- a/riscV/SelectOp.vp +++ b/riscV/SelectOp.vp @@ -448,3 +448,10 @@ Nondetfunction builtin_arg (e: expr) := if Archi.ptr64 then BA_addptr (BA e1) (BA_long n) else BA e | _ => BA e end. + +(* floats *) +Definition divf_base (e1: expr) (e2: expr) := + Eop Odivf (e1 ::: e2 ::: Enil). + +Definition divfs_base (e1: expr) (e2: expr) := + Eop Odivfs (e1 ::: e2 ::: Enil). diff --git a/riscV/SelectOpproof.v b/riscV/SelectOpproof.v index 90f077db..9966305c 100644 --- a/riscV/SelectOpproof.v +++ b/riscV/SelectOpproof.v @@ -29,6 +29,8 @@ Require Import Cminor. Require Import Op. Require Import CminorSel. Require Import SelectOp. +Require Import OpHelpers. +Require Import OpHelpersproof. Local Open Scope cminorsel_scope. @@ -80,8 +82,10 @@ Ltac TrivialExists := (** * Correctness of the smart constructors *) Section CMCONSTR. - -Variable ge: genv. +Variable prog: program. +Variable hf: helper_functions. +Hypothesis HELPERS: helper_functions_declared prog hf. +Let ge := Genv.globalenv prog. Variable sp: val. Variable e: env. Variable m: mem. @@ -922,4 +926,25 @@ Proof. - constructor; auto. Qed. + +(* floating-point division without HELPERS *) +Theorem eval_divf_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v. +Proof. + intros; unfold divf_base. + TrivialExists. +Qed. + +Theorem eval_divfs_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v. +Proof. + intros; unfold divfs_base. + TrivialExists. +Qed. End CMCONSTR. -- cgit From 88c46de03f37dd9edb78e68734a9976ab2ccc056 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 10:28:54 +0100 Subject: seems like powerpc runs but the result segfaults --- powerpc/SelectLong.vp | 2 +- powerpc/SelectLongproof.v | 1 + powerpc/SelectOp.vp | 8 ++++++++ powerpc/SelectOpproof.v | 28 ++++++++++++++++++++++++++-- 4 files changed, 36 insertions(+), 3 deletions(-) diff --git a/powerpc/SelectLong.vp b/powerpc/SelectLong.vp index 5f13774b..e4274ba5 100644 --- a/powerpc/SelectLong.vp +++ b/powerpc/SelectLong.vp @@ -16,7 +16,7 @@ Require Import Coqlib. Require Import Compopts. Require Import AST Integers Floats. Require Import Op CminorSel. -Require Import SelectOp SplitLong. +Require Import OpHelpers SelectOp SplitLong. Local Open Scope cminorsel_scope. Local Open Scope string_scope. diff --git a/powerpc/SelectLongproof.v b/powerpc/SelectLongproof.v index a214d131..b4e48596 100644 --- a/powerpc/SelectLongproof.v +++ b/powerpc/SelectLongproof.v @@ -16,6 +16,7 @@ Require Import String Coqlib Maps Integers Floats Errors. Require Archi. Require Import AST Values Memory Globalenvs Events. Require Import Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. Require Import SelectOp SelectOpproof SplitLong SplitLongproof. Require Import SelectLong. diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp index d2ca408a..478ce251 100644 --- a/powerpc/SelectOp.vp +++ b/powerpc/SelectOp.vp @@ -43,6 +43,7 @@ Require Import Integers. Require Import Floats. Require Import Op. Require Import CminorSel. +Require Import OpHelpers. Local Open Scope cminorsel_scope. @@ -552,3 +553,10 @@ Nondetfunction builtin_arg (e: expr) := | Eop Oadd (e1:::e2:::Enil) => BA_addptr (BA e1) (BA e2) | _ => BA e end. + +(* floats *) +Definition divf_base (e1: expr) (e2: expr) := + Eop Odivf (e1 ::: e2 ::: Enil). + +Definition divfs_base (e1: expr) (e2: expr) := + Eop Odivfs (e1 ::: e2 ::: Enil). diff --git a/powerpc/SelectOpproof.v b/powerpc/SelectOpproof.v index 5f87d9b9..00b91e70 100644 --- a/powerpc/SelectOpproof.v +++ b/powerpc/SelectOpproof.v @@ -25,6 +25,8 @@ Require Import Cminor. Require Import Op. Require Import CminorSel. Require Import SelectOp. +Require Import OpHelpers. +Require Import OpHelpersproof. Local Open Scope cminorsel_scope. Local Transparent Archi.ptr64. @@ -77,8 +79,10 @@ Ltac TrivialExists := (** * Correctness of the smart constructors *) Section CMCONSTR. - -Variable ge: genv. +Variable prog: program. +Variable hf: helper_functions. +Hypothesis HELPERS: helper_functions_declared prog hf. +Let ge := Genv.globalenv prog. Variable sp: val. Variable e: env. Variable m: mem. @@ -1044,4 +1048,24 @@ Proof. - constructor; auto. Qed. +(* floating-point division without HELPERS *) +Theorem eval_divf_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v. +Proof. + intros; unfold divf_base. + TrivialExists. +Qed. + +Theorem eval_divfs_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v. +Proof. + intros; unfold divfs_base. + TrivialExists. +Qed. End CMCONSTR. -- cgit From 1bc0ce716ff90a5384a70b5f9426108bb6380549 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 10:49:09 +0100 Subject: rm Pdiv / Pdivu --- mppa_k1c/Asm.v | 5 ----- mppa_k1c/Asmblock.v | 16 +--------------- mppa_k1c/Asmblockdeps.v | 29 ----------------------------- mppa_k1c/Asmblockgen.v | 12 ------------ mppa_k1c/PostpassSchedulingOracle.ml | 1 - mppa_k1c/TargetPrinter.ml | 4 ---- 6 files changed, 1 insertion(+), 66 deletions(-) diff --git a/mppa_k1c/Asm.v b/mppa_k1c/Asm.v index 493f4a05..0ca554ab 100644 --- a/mppa_k1c/Asm.v +++ b/mppa_k1c/Asm.v @@ -50,9 +50,6 @@ Inductive instruction : Type := | Psemi (**r semi colon separating bundles *) | Pnop (**r instruction that does nothing *) - | Pdiv (**r 32 bits integer division *) - | Pdivu (**r 32 bits integer division *) - (** builtins *) | Pclzll (rd rs: ireg) | Pstsud (rd rs1 rs2: ireg) @@ -218,8 +215,6 @@ Inductive instruction : Type := Definition control_to_instruction (c: control) := match c with | PExpand (Asmblock.Pbuiltin ef args res) => Pbuiltin ef args res - | PExpand (Asmblock.Pdiv) => Pdiv - | PExpand (Asmblock.Pdivu) => Pdivu | PCtlFlow Asmblock.Pret => Pret | PCtlFlow (Asmblock.Pcall l) => Pcall l | PCtlFlow (Asmblock.Picall r) => Picall r diff --git a/mppa_k1c/Asmblock.v b/mppa_k1c/Asmblock.v index d335801e..fdec8ed2 100644 --- a/mppa_k1c/Asmblock.v +++ b/mppa_k1c/Asmblock.v @@ -178,10 +178,7 @@ Inductive ex_instruction : Type := | Pbtbl (r: ireg) (tbl: list label) (**r N-way branch through a jump table *) *) | Pbuiltin: external_function -> list (builtin_arg preg) - -> builtin_res preg -> ex_instruction (**r built-in function (pseudo) *) - | Pdiv (**r 32 bits integer division, call to __divdi3 *) - | Pdivu (**r 32 bits integer division, call to __udivdi3 *) -. + -> builtin_res preg -> ex_instruction (**r built-in function (pseudo) *). (** FIXME: comment not up to date ! @@ -1486,17 +1483,6 @@ Definition exec_control (f: function) (oc: option control) (rs: regset) (m: mem) (** Pseudo-instructions *) | Pbuiltin ef args res => Stuck (**r treated specially below *) - | Pdiv => - match Val.divs (rs GPR0) (rs GPR1) with - | Some v => Next (rs # GPR0 <- v # RA <- (rs RA)) m - | None => Stuck - end - - | Pdivu => - match Val.divu (rs GPR0) (rs GPR1) with - | Some v => Next (rs # GPR0 <- v # RA <- (rs RA)) m - | None => Stuck - end end | None => Next rs m end. diff --git a/mppa_k1c/Asmblockdeps.v b/mppa_k1c/Asmblockdeps.v index 501ec212..ad96ae87 100644 --- a/mppa_k1c/Asmblockdeps.v +++ b/mppa_k1c/Asmblockdeps.v @@ -540,8 +540,6 @@ Definition trans_control (ctl: control) : macro := | Pj_l l => [(#PC, Op (Control (Oj_l l)) (Name (#PC) @ Enil))] | Pcb bt r l => [(#PC, Op (Control (Ocb bt l)) (Name (#r) @ Name (#PC) @ Enil))] | Pcbu bt r l => [(#PC, Op (Control (Ocbu bt l)) (Name (#r) @ Name (#PC) @ Enil))] - | Pdiv => [(#GPR0, Op (Control Odiv) (Name (#GPR0) @ Name (#GPR1) @ Enil)); (#RA, Name (#RA))] - | Pdivu => [(#GPR0, Op (Control Odivu) (Name (#GPR0) @ Name (#GPR1) @ Enil)); (#RA, Name (#RA))] | Pbuiltin ef args res => [(#PC, Op (Control (OError)) Enil)] end. @@ -863,22 +861,6 @@ Proof. intros. destruct ex. - simpl in *. inv H1. destruct c; destruct i; try discriminate. all: try (inv H0; eexists; split; try split; [ simpl control_eval; pose (H3 PC); simpl in e; rewrite e; reflexivity | Simpl | intros rr; destruct rr; Simpl]). - (* Pdiv *) - + destruct (Val.divs _ _) eqn:DIVS; try discriminate. inv H0. unfold nextblock in DIVS. repeat (rewrite Pregmap.gso in DIVS; try discriminate). - eexists; split; try split. - * simpl control_eval. pose (H3 PC); simpl in e; rewrite e; clear e. simpl. - Simpl. pose (H3 GPR0); rewrite e; clear e. pose (H3 GPR1); rewrite e; clear e. rewrite DIVS. - Simpl. - * Simpl. - * intros rr; destruct rr; Simpl. destruct (preg_eq GPR0 g); Simpl. rewrite e. Simpl. - (* Pdivu *) - + destruct (Val.divu _ _) eqn:DIVU; try discriminate. inv H0. unfold nextblock in DIVU. repeat (rewrite Pregmap.gso in DIVU; try discriminate). - eexists; split; try split. - * simpl control_eval. pose (H3 PC); simpl in e; rewrite e; clear e. simpl. - Simpl. pose (H3 GPR0); rewrite e; clear e. pose (H3 GPR1); rewrite e; clear e. rewrite DIVU. - Simpl. - * Simpl. - * intros rr; destruct rr; Simpl. destruct (preg_eq GPR0 g); Simpl. rewrite e. Simpl. (* Pj_l *) + unfold goto_label in H0. destruct (label_pos _ _ _) eqn:LPOS; try discriminate. destruct (nextblock _ _ _) eqn:NB; try discriminate. inv H0. eexists; split; try split. @@ -1070,12 +1052,6 @@ Lemma exec_exit_none: Proof. intros. inv H0. destruct ex as [ctl|]; try discriminate. destruct ctl; destruct i; try reflexivity; try discriminate. -(* Pdiv *) - - simpl in *. pose (H3 GPR0); rewrite e in H1; clear e. pose (H3 GPR1); rewrite e in H1; clear e. - destruct (Val.divs _ _); try discriminate; auto. -(* Pdivu *) - - simpl in *. pose (H3 GPR0); rewrite e in H1; clear e. pose (H3 GPR1); rewrite e in H1; clear e. - destruct (Val.divu _ _); try discriminate; auto. (* Pj_l *) - simpl in *. pose (H3 PC); simpl in e; rewrite e in H1. clear e. unfold goto_label_deps in H1. unfold goto_label. @@ -1187,11 +1163,6 @@ Lemma forward_simu_exit_stuck: Proof. intros. inv H1. destruct ex as [ctl|]; try discriminate. destruct ctl; destruct i; try discriminate; try (simpl; reflexivity). -(* Pdiv *) - - simpl in *. pose (H3 GPR0); simpl in e; rewrite e; clear e. pose (H3 GPR1); simpl in e; rewrite e; clear e. - destruct (Val.divs _ _); try discriminate; auto. - - simpl in *. pose (H3 GPR0); simpl in e; rewrite e; clear e. pose (H3 GPR1); simpl in e; rewrite e; clear e. - destruct (Val.divu _ _); try discriminate; auto. (* Pj_l *) - simpl in *. pose (H3 PC); simpl in e; rewrite e. unfold goto_label_deps. unfold goto_label in H0. destruct (label_pos _ _ _); auto. clear e. destruct (rs PC); auto. discriminate. diff --git a/mppa_k1c/Asmblockgen.v b/mppa_k1c/Asmblockgen.v index 5b00a64f..c03e319c 100644 --- a/mppa_k1c/Asmblockgen.v +++ b/mppa_k1c/Asmblockgen.v @@ -425,18 +425,6 @@ Definition transl_op OK (mulimm32 rd rs1 n ::i k) | Omulhs, _ => Error(msg "Asmblockgen.transl_op: Omulhs") (* Normalement pas émis *) | Omulhu, _ => Error(msg "Asmblockgen.transl_op: Omulhu") (* Normalement pas émis *) - | Odiv, a1 :: a2 :: nil => Error(msg "Asmblockgen.transl_op: Odiv: 32-bits division not supported yet. Please use 64-bits.") - (* do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; - OK (Pdivw rd rs1 rs2 :: k) *) - | Odivu, a1 :: a2 :: nil => Error(msg "Asmblockgen.transl_op: Odivu: 32-bits division not supported yet. Please use 64-bits.") - (* do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; - OK (Pdivuw rd rs1 rs2 :: k) *) - | Omod, a1 :: a2 :: nil => Error(msg "Asmblockgen.transl_op: Omod: 32-bits modulo not supported yet. Please use 64-bits.") - (* do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; - OK (Premw rd rs1 rs2 :: k) *) - | Omodu, a1 :: a2 :: nil => Error(msg "Asmblockgen.transl_op: Omodu: 32-bits modulo not supported yet. Please use 64-bits.") - (* do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; - OK (Premuw rd rs1 rs2 :: k) *) | Oand, a1 :: a2 :: nil => do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2; OK (Pandw rd rs1 rs2 ::i k) diff --git a/mppa_k1c/PostpassSchedulingOracle.ml b/mppa_k1c/PostpassSchedulingOracle.ml index 2c39e342..56b00c7e 100644 --- a/mppa_k1c/PostpassSchedulingOracle.ml +++ b/mppa_k1c/PostpassSchedulingOracle.ml @@ -222,7 +222,6 @@ let basic_rec i = | Pnop -> { inst = "nop"; write_locs = []; read_locs = []; imm = None ; is_control = false} let expand_rec = function - | Pdiv | Pdivu -> { inst = "Pdiv"; write_locs = [Reg (IR GPR0)]; read_locs = [Reg (IR GPR0); Reg (IR GPR1)]; imm = None; is_control = true } | Pbuiltin _ -> raise OpaqueInstruction let ctl_flow_rec = function diff --git a/mppa_k1c/TargetPrinter.ml b/mppa_k1c/TargetPrinter.ml index d4e2afc9..69824852 100644 --- a/mppa_k1c/TargetPrinter.ml +++ b/mppa_k1c/TargetPrinter.ml @@ -258,10 +258,6 @@ module Target (*: TARGET*) = fprintf oc " goto %a\n" symbol s | Pigoto(rs) -> fprintf oc " igoto %a\n" ireg rs - | Pdiv -> - fprintf oc " call __divdi3\n" - | Pdivu -> - fprintf oc " call __udivdi3\n" | Pj_l(s) -> fprintf oc " goto %a\n" print_label s | Pcb (bt, r, lbl) | Pcbu (bt, r, lbl) -> -- cgit From 4f45e2926fa9d0dba400ec6b6a1506c898cad13d Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 22 Mar 2019 11:10:11 +0100 Subject: Reorganized the test/mppa/ tests to have fewer of them --- test/mppa/instr/addw.c | 5 -- test/mppa/instr/andd.c | 5 -- test/mppa/instr/andw.c | 5 -- test/mppa/instr/branch.c | 10 --- test/mppa/instr/branchz.c | 10 --- test/mppa/instr/branchzu.c | 11 --- test/mppa/instr/call.c | 16 ---- test/mppa/instr/cast_S32_S64.c | 7 -- test/mppa/instr/cast_S64_U32.c | 7 -- test/mppa/instr/cast_U32_S64.c | 7 -- test/mppa/instr/cb.deqz.c | 10 --- test/mppa/instr/cb.dgez.c | 10 --- test/mppa/instr/cb.dgtz.c | 10 --- test/mppa/instr/cb.dlez.c | 10 --- test/mppa/instr/cb.dltz.c | 10 --- test/mppa/instr/cb.dnez.c | 10 --- test/mppa/instr/cb.wgez.c | 10 --- test/mppa/instr/cb.wgtz.c | 10 --- test/mppa/instr/cb.wlez.c | 10 --- test/mppa/instr/cb.wltz.c | 10 --- test/mppa/instr/compd.eq.c | 7 -- test/mppa/instr/compd.geu.c | 7 -- test/mppa/instr/compd.gt.c | 7 -- test/mppa/instr/compd.gtu.c | 7 -- test/mppa/instr/compd.le.c | 7 -- test/mppa/instr/compd.leu.c | 7 -- test/mppa/instr/compd.lt.c | 7 -- test/mppa/instr/compd.ltu.c | 7 -- test/mppa/instr/compd.ne.c | 7 -- test/mppa/instr/compw.eq.c | 7 -- test/mppa/instr/compw.geu.c | 7 -- test/mppa/instr/compw.gt.c | 7 -- test/mppa/instr/compw.gtu.c | 7 -- test/mppa/instr/compw.le.c | 7 -- test/mppa/instr/compw.leu.c | 7 -- test/mppa/instr/compw.lt.c | 7 -- test/mppa/instr/compw.ltu.c | 7 -- test/mppa/instr/compw.ne.c | 7 -- test/mppa/instr/div2.c | 7 -- test/mppa/instr/doubleconv.c | 9 --- test/mppa/instr/f32.c | 8 ++ test/mppa/instr/f64.c | 8 ++ test/mppa/instr/faddd.c | 5 -- test/mppa/instr/faddw.c | 5 -- test/mppa/instr/floatconv.c | 9 --- test/mppa/instr/fmuld.c | 7 -- test/mppa/instr/fmulw.c | 7 -- test/mppa/instr/fnegd.c | 7 -- test/mppa/instr/fnegw.c | 7 -- test/mppa/instr/for.c | 9 --- test/mppa/instr/forvar.c | 9 --- test/mppa/instr/forvarl.c | 10 --- test/mppa/instr/fsbfd.c | 7 -- test/mppa/instr/fsbfw.c | 7 -- test/mppa/instr/i32.c | 87 ++++++++++++++++++++ test/mppa/instr/i64.c | 107 +++++++++++++++++++++++++ test/mppa/instr/indirect_call.c | 33 -------- test/mppa/instr/indirect_tailcall.c | 33 -------- test/mppa/instr/individual/andw.c | 5 ++ test/mppa/instr/individual/branch.c | 10 +++ test/mppa/instr/individual/branchz.c | 10 +++ test/mppa/instr/individual/branchzu.c | 11 +++ test/mppa/instr/individual/call.c | 16 ++++ test/mppa/instr/individual/cast_S32_S64.c | 7 ++ test/mppa/instr/individual/cast_S64_U32.c | 7 ++ test/mppa/instr/individual/cb.deqz.c | 10 +++ test/mppa/instr/individual/cb.dgez.c | 10 +++ test/mppa/instr/individual/cb.dgtz.c | 10 +++ test/mppa/instr/individual/cb.dlez.c | 10 +++ test/mppa/instr/individual/cb.dltz.c | 10 +++ test/mppa/instr/individual/cb.dnez.c | 10 +++ test/mppa/instr/individual/cb.wgez.c | 10 +++ test/mppa/instr/individual/cb.wgtz.c | 10 +++ test/mppa/instr/individual/cb.wlez.c | 10 +++ test/mppa/instr/individual/cb.wltz.c | 10 +++ test/mppa/instr/individual/compd.eq.c | 7 ++ test/mppa/instr/individual/compd.geu.c | 7 ++ test/mppa/instr/individual/compd.gt.c | 7 ++ test/mppa/instr/individual/compd.le.c | 7 ++ test/mppa/instr/individual/compd.leu.c | 7 ++ test/mppa/instr/individual/compd.lt.c | 7 ++ test/mppa/instr/individual/compd.ltu.c | 7 ++ test/mppa/instr/individual/compd.ne.c | 7 ++ test/mppa/instr/individual/compw.eq.c | 7 ++ test/mppa/instr/individual/compw.geu.c | 7 ++ test/mppa/instr/individual/compw.gt.c | 7 ++ test/mppa/instr/individual/compw.gtu.c | 7 ++ test/mppa/instr/individual/compw.le.c | 7 ++ test/mppa/instr/individual/compw.leu.c | 7 ++ test/mppa/instr/individual/compw.lt.c | 7 ++ test/mppa/instr/individual/compw.ltu.c | 7 ++ test/mppa/instr/individual/compw.ne.c | 7 ++ test/mppa/instr/individual/div2.c | 7 ++ test/mppa/instr/individual/doubleconv.c | 9 +++ test/mppa/instr/individual/floatconv.c | 9 +++ test/mppa/instr/individual/fmuld.c | 7 ++ test/mppa/instr/individual/fmulw.c | 7 ++ test/mppa/instr/individual/fnegd.c | 7 ++ test/mppa/instr/individual/fnegw.c | 7 ++ test/mppa/instr/individual/for.c | 9 +++ test/mppa/instr/individual/forvar.c | 9 +++ test/mppa/instr/individual/forvarl.c | 10 +++ test/mppa/instr/individual/fsbfd.c | 7 ++ test/mppa/instr/individual/fsbfw.c | 7 ++ test/mppa/instr/individual/indirect_call.c | 33 ++++++++ test/mppa/instr/individual/indirect_tailcall.c | 33 ++++++++ test/mppa/instr/individual/lbs.c | 9 +++ test/mppa/instr/individual/lbz.c | 9 +++ test/mppa/instr/individual/muld.c | 7 ++ test/mppa/instr/individual/mulw.c | 7 ++ test/mppa/instr/individual/negd.c | 7 ++ test/mppa/instr/individual/ord.c | 7 ++ test/mppa/instr/individual/sbfd.c | 7 ++ test/mppa/instr/individual/sbfw.c | 7 ++ test/mppa/instr/individual/simple.c | 7 ++ test/mppa/instr/individual/sllw.c | 7 ++ test/mppa/instr/individual/srad.c | 7 ++ test/mppa/instr/individual/srld.c | 7 ++ test/mppa/instr/individual/tailcall.c | 16 ++++ test/mppa/instr/individual/udivd.c | 7 ++ test/mppa/instr/individual/umodd.c | 7 ++ test/mppa/instr/individual/xord.c | 7 ++ test/mppa/instr/lbs.c | 9 --- test/mppa/instr/lbz.c | 9 --- test/mppa/instr/muld.c | 7 -- test/mppa/instr/mulw.c | 7 -- test/mppa/instr/negd.c | 7 -- test/mppa/instr/ord.c | 7 -- test/mppa/instr/sbfd.c | 7 -- test/mppa/instr/sbfw.c | 7 -- test/mppa/instr/simple.c | 7 -- test/mppa/instr/sllw.c | 7 -- test/mppa/instr/srad.c | 7 -- test/mppa/instr/srld.c | 7 -- test/mppa/instr/tailcall.c | 16 ---- test/mppa/instr/udivd.c | 7 -- test/mppa/instr/ui32.c | 12 +++ test/mppa/instr/ui64.c | 10 +++ test/mppa/instr/umodd.c | 7 -- test/mppa/instr/xord.c | 7 -- test/mppa/interop/i32.c | 13 +++ test/mppa/interop/i64.c | 14 ++++ test/mppa/interop/i_manyiargs.c | 9 --- test/mppa/interop/i_multiiargs.c | 6 -- test/mppa/interop/i_oneiarg.c | 6 -- test/mppa/interop/individual/i_multiiargs.c | 6 ++ test/mppa/interop/individual/i_oneiarg.c | 6 ++ test/mppa/interop/individual/ll_multillargs.c | 7 ++ test/mppa/interop/individual/ll_onellarg.c | 7 ++ test/mppa/interop/individual/ll_void.c | 7 ++ test/mppa/interop/individual/void_void.c | 7 ++ test/mppa/interop/ll_manyllargs.c | 8 -- test/mppa/interop/ll_multillargs.c | 7 -- test/mppa/interop/ll_onellarg.c | 7 -- test/mppa/interop/ll_void.c | 7 -- test/mppa/interop/void_void.c | 7 -- 156 files changed, 870 insertions(+), 662 deletions(-) delete mode 100644 test/mppa/instr/addw.c delete mode 100644 test/mppa/instr/andd.c delete mode 100644 test/mppa/instr/andw.c delete mode 100644 test/mppa/instr/branch.c delete mode 100644 test/mppa/instr/branchz.c delete mode 100644 test/mppa/instr/branchzu.c delete mode 100644 test/mppa/instr/call.c delete mode 100644 test/mppa/instr/cast_S32_S64.c delete mode 100644 test/mppa/instr/cast_S64_U32.c delete mode 100644 test/mppa/instr/cast_U32_S64.c delete mode 100644 test/mppa/instr/cb.deqz.c delete mode 100644 test/mppa/instr/cb.dgez.c delete mode 100644 test/mppa/instr/cb.dgtz.c delete mode 100644 test/mppa/instr/cb.dlez.c delete mode 100644 test/mppa/instr/cb.dltz.c delete mode 100644 test/mppa/instr/cb.dnez.c delete mode 100644 test/mppa/instr/cb.wgez.c delete mode 100644 test/mppa/instr/cb.wgtz.c delete mode 100644 test/mppa/instr/cb.wlez.c delete mode 100644 test/mppa/instr/cb.wltz.c delete mode 100644 test/mppa/instr/compd.eq.c delete mode 100644 test/mppa/instr/compd.geu.c delete mode 100644 test/mppa/instr/compd.gt.c delete mode 100644 test/mppa/instr/compd.gtu.c delete mode 100644 test/mppa/instr/compd.le.c delete mode 100644 test/mppa/instr/compd.leu.c delete mode 100644 test/mppa/instr/compd.lt.c delete mode 100644 test/mppa/instr/compd.ltu.c delete mode 100644 test/mppa/instr/compd.ne.c delete mode 100644 test/mppa/instr/compw.eq.c delete mode 100644 test/mppa/instr/compw.geu.c delete mode 100644 test/mppa/instr/compw.gt.c delete mode 100644 test/mppa/instr/compw.gtu.c delete mode 100644 test/mppa/instr/compw.le.c delete mode 100644 test/mppa/instr/compw.leu.c delete mode 100644 test/mppa/instr/compw.lt.c delete mode 100644 test/mppa/instr/compw.ltu.c delete mode 100644 test/mppa/instr/compw.ne.c delete mode 100644 test/mppa/instr/div2.c delete mode 100644 test/mppa/instr/doubleconv.c create mode 100644 test/mppa/instr/f32.c create mode 100644 test/mppa/instr/f64.c delete mode 100644 test/mppa/instr/faddd.c delete mode 100644 test/mppa/instr/faddw.c delete mode 100644 test/mppa/instr/floatconv.c delete mode 100644 test/mppa/instr/fmuld.c delete mode 100644 test/mppa/instr/fmulw.c delete mode 100644 test/mppa/instr/fnegd.c delete mode 100644 test/mppa/instr/fnegw.c delete mode 100644 test/mppa/instr/for.c delete mode 100644 test/mppa/instr/forvar.c delete mode 100644 test/mppa/instr/forvarl.c delete mode 100644 test/mppa/instr/fsbfd.c delete mode 100644 test/mppa/instr/fsbfw.c create mode 100644 test/mppa/instr/i32.c create mode 100644 test/mppa/instr/i64.c delete mode 100644 test/mppa/instr/indirect_call.c delete mode 100644 test/mppa/instr/indirect_tailcall.c create mode 100644 test/mppa/instr/individual/andw.c create mode 100644 test/mppa/instr/individual/branch.c create mode 100644 test/mppa/instr/individual/branchz.c create mode 100644 test/mppa/instr/individual/branchzu.c create mode 100644 test/mppa/instr/individual/call.c create mode 100644 test/mppa/instr/individual/cast_S32_S64.c create mode 100644 test/mppa/instr/individual/cast_S64_U32.c create mode 100644 test/mppa/instr/individual/cb.deqz.c create mode 100644 test/mppa/instr/individual/cb.dgez.c create mode 100644 test/mppa/instr/individual/cb.dgtz.c create mode 100644 test/mppa/instr/individual/cb.dlez.c create mode 100644 test/mppa/instr/individual/cb.dltz.c create mode 100644 test/mppa/instr/individual/cb.dnez.c create mode 100644 test/mppa/instr/individual/cb.wgez.c create mode 100644 test/mppa/instr/individual/cb.wgtz.c create mode 100644 test/mppa/instr/individual/cb.wlez.c create mode 100644 test/mppa/instr/individual/cb.wltz.c create mode 100644 test/mppa/instr/individual/compd.eq.c create mode 100644 test/mppa/instr/individual/compd.geu.c create mode 100644 test/mppa/instr/individual/compd.gt.c create mode 100644 test/mppa/instr/individual/compd.le.c create mode 100644 test/mppa/instr/individual/compd.leu.c create mode 100644 test/mppa/instr/individual/compd.lt.c create mode 100644 test/mppa/instr/individual/compd.ltu.c create mode 100644 test/mppa/instr/individual/compd.ne.c create mode 100644 test/mppa/instr/individual/compw.eq.c create mode 100644 test/mppa/instr/individual/compw.geu.c create mode 100644 test/mppa/instr/individual/compw.gt.c create mode 100644 test/mppa/instr/individual/compw.gtu.c create mode 100644 test/mppa/instr/individual/compw.le.c create mode 100644 test/mppa/instr/individual/compw.leu.c create mode 100644 test/mppa/instr/individual/compw.lt.c create mode 100644 test/mppa/instr/individual/compw.ltu.c create mode 100644 test/mppa/instr/individual/compw.ne.c create mode 100644 test/mppa/instr/individual/div2.c create mode 100644 test/mppa/instr/individual/doubleconv.c create mode 100644 test/mppa/instr/individual/floatconv.c create mode 100644 test/mppa/instr/individual/fmuld.c create mode 100644 test/mppa/instr/individual/fmulw.c create mode 100644 test/mppa/instr/individual/fnegd.c create mode 100644 test/mppa/instr/individual/fnegw.c create mode 100644 test/mppa/instr/individual/for.c create mode 100644 test/mppa/instr/individual/forvar.c create mode 100644 test/mppa/instr/individual/forvarl.c create mode 100644 test/mppa/instr/individual/fsbfd.c create mode 100644 test/mppa/instr/individual/fsbfw.c create mode 100644 test/mppa/instr/individual/indirect_call.c create mode 100644 test/mppa/instr/individual/indirect_tailcall.c create mode 100644 test/mppa/instr/individual/lbs.c create mode 100644 test/mppa/instr/individual/lbz.c create mode 100644 test/mppa/instr/individual/muld.c create mode 100644 test/mppa/instr/individual/mulw.c create mode 100644 test/mppa/instr/individual/negd.c create mode 100644 test/mppa/instr/individual/ord.c create mode 100644 test/mppa/instr/individual/sbfd.c create mode 100644 test/mppa/instr/individual/sbfw.c create mode 100644 test/mppa/instr/individual/simple.c create mode 100644 test/mppa/instr/individual/sllw.c create mode 100644 test/mppa/instr/individual/srad.c create mode 100644 test/mppa/instr/individual/srld.c create mode 100644 test/mppa/instr/individual/tailcall.c create mode 100644 test/mppa/instr/individual/udivd.c create mode 100644 test/mppa/instr/individual/umodd.c create mode 100644 test/mppa/instr/individual/xord.c delete mode 100644 test/mppa/instr/lbs.c delete mode 100644 test/mppa/instr/lbz.c delete mode 100644 test/mppa/instr/muld.c delete mode 100644 test/mppa/instr/mulw.c delete mode 100644 test/mppa/instr/negd.c delete mode 100644 test/mppa/instr/ord.c delete mode 100644 test/mppa/instr/sbfd.c delete mode 100644 test/mppa/instr/sbfw.c delete mode 100644 test/mppa/instr/simple.c delete mode 100644 test/mppa/instr/sllw.c delete mode 100644 test/mppa/instr/srad.c delete mode 100644 test/mppa/instr/srld.c delete mode 100644 test/mppa/instr/tailcall.c delete mode 100644 test/mppa/instr/udivd.c create mode 100644 test/mppa/instr/ui32.c create mode 100644 test/mppa/instr/ui64.c delete mode 100644 test/mppa/instr/umodd.c delete mode 100644 test/mppa/instr/xord.c create mode 100644 test/mppa/interop/i32.c create mode 100644 test/mppa/interop/i64.c delete mode 100644 test/mppa/interop/i_manyiargs.c delete mode 100644 test/mppa/interop/i_multiiargs.c delete mode 100644 test/mppa/interop/i_oneiarg.c create mode 100644 test/mppa/interop/individual/i_multiiargs.c create mode 100644 test/mppa/interop/individual/i_oneiarg.c create mode 100644 test/mppa/interop/individual/ll_multillargs.c create mode 100644 test/mppa/interop/individual/ll_onellarg.c create mode 100644 test/mppa/interop/individual/ll_void.c create mode 100644 test/mppa/interop/individual/void_void.c delete mode 100644 test/mppa/interop/ll_manyllargs.c delete mode 100644 test/mppa/interop/ll_multillargs.c delete mode 100644 test/mppa/interop/ll_onellarg.c delete mode 100644 test/mppa/interop/ll_void.c delete mode 100644 test/mppa/interop/void_void.c diff --git a/test/mppa/instr/addw.c b/test/mppa/instr/addw.c deleted file mode 100644 index e22024cf..00000000 --- a/test/mppa/instr/addw.c +++ /dev/null @@ -1,5 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) - c = a+b; -END_TEST32() diff --git a/test/mppa/instr/andd.c b/test/mppa/instr/andd.c deleted file mode 100644 index e3221bd7..00000000 --- a/test/mppa/instr/andd.c +++ /dev/null @@ -1,5 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) - c = a&b; -END_TEST64() diff --git a/test/mppa/instr/andw.c b/test/mppa/instr/andw.c deleted file mode 100644 index 799dc7fb..00000000 --- a/test/mppa/instr/andw.c +++ /dev/null @@ -1,5 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) - c = a&b; -END_TEST32() diff --git a/test/mppa/instr/branch.c b/test/mppa/instr/branch.c deleted file mode 100644 index c9937e31..00000000 --- a/test/mppa/instr/branch.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - if ((a & 0x1) == 1) - c = 0; - else - c = 1; -} -END_TEST32() diff --git a/test/mppa/instr/branchz.c b/test/mppa/instr/branchz.c deleted file mode 100644 index d3e021b5..00000000 --- a/test/mppa/instr/branchz.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - if (a & 0x1 == 0) - c = 0; - else - c = 1; -} -END_TEST32() diff --git a/test/mppa/instr/branchzu.c b/test/mppa/instr/branchzu.c deleted file mode 100644 index d0169174..00000000 --- a/test/mppa/instr/branchzu.c +++ /dev/null @@ -1,11 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - b = !(a & 0x01); - if (!b) - c = 0; - else - c = 1; -} -END_TEST32() diff --git a/test/mppa/instr/call.c b/test/mppa/instr/call.c deleted file mode 100644 index ba2ec323..00000000 --- a/test/mppa/instr/call.c +++ /dev/null @@ -1,16 +0,0 @@ -#include "framework.h" - -int sum(int a, int b){ - return a+b; -} - -int make(int a){ - return a; -} - -BEGIN_TEST(int) -{ - c = sum(make(a), make(b)); -} -END_TEST32() -/* RETURN VALUE: 60 */ diff --git a/test/mppa/instr/cast_S32_S64.c b/test/mppa/instr/cast_S32_S64.c deleted file mode 100644 index 09c97e00..00000000 --- a/test/mppa/instr/cast_S32_S64.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = (long long) a; -} -END_TEST32() diff --git a/test/mppa/instr/cast_S64_U32.c b/test/mppa/instr/cast_S64_U32.c deleted file mode 100644 index 2d9dc723..00000000 --- a/test/mppa/instr/cast_S64_U32.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = (unsigned int) a; -} -END_TEST64() diff --git a/test/mppa/instr/cast_U32_S64.c b/test/mppa/instr/cast_U32_S64.c deleted file mode 100644 index 6f9cd059..00000000 --- a/test/mppa/instr/cast_U32_S64.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned int) -{ - c = (long long) a; -} -END_TEST32() diff --git a/test/mppa/instr/cb.deqz.c b/test/mppa/instr/cb.deqz.c deleted file mode 100644 index 6da2ab07..00000000 --- a/test/mppa/instr/cb.deqz.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - if (0 != (a & 0x1LL)) - c = 1; - else - c = 0; -} -END_TEST64() diff --git a/test/mppa/instr/cb.dgez.c b/test/mppa/instr/cb.dgez.c deleted file mode 100644 index 7bef25ad..00000000 --- a/test/mppa/instr/cb.dgez.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - if (0 > (a & 0x1LL)) - c = 1; - else - c = 0; -} -END_TEST64() diff --git a/test/mppa/instr/cb.dgtz.c b/test/mppa/instr/cb.dgtz.c deleted file mode 100644 index 1a43fb1f..00000000 --- a/test/mppa/instr/cb.dgtz.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - if (0 >= (a & 0x1LL) - 1) - c = 1; - else - c = 0; -} -END_TEST64() diff --git a/test/mppa/instr/cb.dlez.c b/test/mppa/instr/cb.dlez.c deleted file mode 100644 index 2fb97939..00000000 --- a/test/mppa/instr/cb.dlez.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - if (a & 0x1LL > 0) - c = 1; - else - c = 0; -} -END_TEST64() diff --git a/test/mppa/instr/cb.dltz.c b/test/mppa/instr/cb.dltz.c deleted file mode 100644 index a431d5d0..00000000 --- a/test/mppa/instr/cb.dltz.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - if ((a & 0x1LL) - 1 >= 0) - c = 1; - else - c = 0; -} -END_TEST64() diff --git a/test/mppa/instr/cb.dnez.c b/test/mppa/instr/cb.dnez.c deleted file mode 100644 index 44516cbe..00000000 --- a/test/mppa/instr/cb.dnez.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - if (0 == (a & 0x1LL)) - c = 1; - else - c = 0; -} -END_TEST64() diff --git a/test/mppa/instr/cb.wgez.c b/test/mppa/instr/cb.wgez.c deleted file mode 100644 index 5779ad92..00000000 --- a/test/mppa/instr/cb.wgez.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - if (0 > (a & 0x1) - 1) - c = 1; - else - c = 0; -} -END_TEST32() diff --git a/test/mppa/instr/cb.wgtz.c b/test/mppa/instr/cb.wgtz.c deleted file mode 100644 index abb695bd..00000000 --- a/test/mppa/instr/cb.wgtz.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - if (0 >= (a & 0x1)) - c = 1; - else - c = 0; -} -END_TEST32() diff --git a/test/mppa/instr/cb.wlez.c b/test/mppa/instr/cb.wlez.c deleted file mode 100644 index 3a2e08c1..00000000 --- a/test/mppa/instr/cb.wlez.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - if ((a & 0x1) > 0) - c = 1; - else - c = 0; -} -END_TEST32() diff --git a/test/mppa/instr/cb.wltz.c b/test/mppa/instr/cb.wltz.c deleted file mode 100644 index 5d52c72a..00000000 --- a/test/mppa/instr/cb.wltz.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - if ((a & 0x1) - 1 >= 0) - c = 1; - else - c = 0; -} -END_TEST32() diff --git a/test/mppa/instr/compd.eq.c b/test/mppa/instr/compd.eq.c deleted file mode 100644 index 4fe8de2a..00000000 --- a/test/mppa/instr/compd.eq.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = ((a & 0x1LL) == (b & 0x1LL)); -} -END_TEST64() diff --git a/test/mppa/instr/compd.geu.c b/test/mppa/instr/compd.geu.c deleted file mode 100644 index fccf0804..00000000 --- a/test/mppa/instr/compd.geu.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned long long) -{ - c = (a >= b); -} -END_TEST64() diff --git a/test/mppa/instr/compd.gt.c b/test/mppa/instr/compd.gt.c deleted file mode 100644 index b9901436..00000000 --- a/test/mppa/instr/compd.gt.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = (a > b); -} -END_TEST64() diff --git a/test/mppa/instr/compd.gtu.c b/test/mppa/instr/compd.gtu.c deleted file mode 100644 index 7b2b96a6..00000000 --- a/test/mppa/instr/compd.gtu.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned long long) -{ - c = (a > b); -} -END_TEST64() diff --git a/test/mppa/instr/compd.le.c b/test/mppa/instr/compd.le.c deleted file mode 100644 index 6fa0f103..00000000 --- a/test/mppa/instr/compd.le.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = (a <= b); -} -END_TEST64() diff --git a/test/mppa/instr/compd.leu.c b/test/mppa/instr/compd.leu.c deleted file mode 100644 index 1ad18281..00000000 --- a/test/mppa/instr/compd.leu.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned long long) -{ - c = (a <= b); -} -END_TEST64() diff --git a/test/mppa/instr/compd.lt.c b/test/mppa/instr/compd.lt.c deleted file mode 100644 index c42cda56..00000000 --- a/test/mppa/instr/compd.lt.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = (a < b); -} -END_TEST64() diff --git a/test/mppa/instr/compd.ltu.c b/test/mppa/instr/compd.ltu.c deleted file mode 100644 index b03d4d53..00000000 --- a/test/mppa/instr/compd.ltu.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned long long) -{ - c = (a < b); -} -END_TEST64() diff --git a/test/mppa/instr/compd.ne.c b/test/mppa/instr/compd.ne.c deleted file mode 100644 index fd9d0b28..00000000 --- a/test/mppa/instr/compd.ne.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned long long) -{ - c = ((a & 0x1ULL) != (b & 0x1ULL)); -} -END_TEST64() diff --git a/test/mppa/instr/compw.eq.c b/test/mppa/instr/compw.eq.c deleted file mode 100644 index cd93f365..00000000 --- a/test/mppa/instr/compw.eq.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = ((a & 0x1) == (b & 0x1)); -} -END_TEST32() diff --git a/test/mppa/instr/compw.geu.c b/test/mppa/instr/compw.geu.c deleted file mode 100644 index b8fb1adf..00000000 --- a/test/mppa/instr/compw.geu.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned int) -{ - c = (a >= b); -} -END_TEST32() diff --git a/test/mppa/instr/compw.gt.c b/test/mppa/instr/compw.gt.c deleted file mode 100644 index 5f6bc907..00000000 --- a/test/mppa/instr/compw.gt.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = (a > b); -} -END_TEST32() diff --git a/test/mppa/instr/compw.gtu.c b/test/mppa/instr/compw.gtu.c deleted file mode 100644 index 947f6a14..00000000 --- a/test/mppa/instr/compw.gtu.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned int) -{ - c = (a > b); -} -END_TEST32() diff --git a/test/mppa/instr/compw.le.c b/test/mppa/instr/compw.le.c deleted file mode 100644 index 35ec6b7d..00000000 --- a/test/mppa/instr/compw.le.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = (a <= b); -} -END_TEST32() diff --git a/test/mppa/instr/compw.leu.c b/test/mppa/instr/compw.leu.c deleted file mode 100644 index 74ebfb42..00000000 --- a/test/mppa/instr/compw.leu.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned int) -{ - c = (a <= b); -} -END_TEST32() diff --git a/test/mppa/instr/compw.lt.c b/test/mppa/instr/compw.lt.c deleted file mode 100644 index cb1f30bd..00000000 --- a/test/mppa/instr/compw.lt.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = (a < b); -} -END_TEST32() diff --git a/test/mppa/instr/compw.ltu.c b/test/mppa/instr/compw.ltu.c deleted file mode 100644 index 6a0c5af1..00000000 --- a/test/mppa/instr/compw.ltu.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned int) -{ - c = (a < b); -} -END_TEST32() diff --git a/test/mppa/instr/compw.ne.c b/test/mppa/instr/compw.ne.c deleted file mode 100644 index 7035e2c7..00000000 --- a/test/mppa/instr/compw.ne.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned int) -{ - c = ((a & 0x1U) != (b & 0x1U)); -} -END_TEST32() diff --git a/test/mppa/instr/div2.c b/test/mppa/instr/div2.c deleted file mode 100644 index b5dfe63a..00000000 --- a/test/mppa/instr/div2.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = (a + b) / 2; -} -END_TEST32() diff --git a/test/mppa/instr/doubleconv.c b/test/mppa/instr/doubleconv.c deleted file mode 100644 index 55b1ddab..00000000 --- a/test/mppa/instr/doubleconv.c +++ /dev/null @@ -1,9 +0,0 @@ -#include "framework.h" - -double long2double(long v){ - return v; -} - -BEGIN_TEST(long) - c = (long) long2double(a) + (long) long2double(b) + (long) long2double(42.3); -END_TEST64() diff --git a/test/mppa/instr/f32.c b/test/mppa/instr/f32.c new file mode 100644 index 00000000..7e304aeb --- /dev/null +++ b/test/mppa/instr/f32.c @@ -0,0 +1,8 @@ +#include "framework.h" + +BEGIN_TEST(float) + c = ((float)a + (float)b); + c += ((float)a * (float)b); + c += (-(float)a); + c += ((float)a - (float)b); +END_TESTF32() diff --git a/test/mppa/instr/f64.c b/test/mppa/instr/f64.c new file mode 100644 index 00000000..be8094c9 --- /dev/null +++ b/test/mppa/instr/f64.c @@ -0,0 +1,8 @@ +#include "framework.h" + +BEGIN_TEST(double) + c = ((double)a + (double)b); + c += ((double)a * (double)b); + c += (-(double)a); + c += ((double)a - (double)b); +END_TESTF64() diff --git a/test/mppa/instr/faddd.c b/test/mppa/instr/faddd.c deleted file mode 100644 index 35b7fc92..00000000 --- a/test/mppa/instr/faddd.c +++ /dev/null @@ -1,5 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(double) - c = ((double)a + (double)b); -END_TESTF64() diff --git a/test/mppa/instr/faddw.c b/test/mppa/instr/faddw.c deleted file mode 100644 index e0e635ae..00000000 --- a/test/mppa/instr/faddw.c +++ /dev/null @@ -1,5 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(float) - c = ((float)a + (float)b); -END_TESTF32() diff --git a/test/mppa/instr/floatconv.c b/test/mppa/instr/floatconv.c deleted file mode 100644 index 32b798e1..00000000 --- a/test/mppa/instr/floatconv.c +++ /dev/null @@ -1,9 +0,0 @@ -#include "framework.h" - -float int2float(int v){ - return v; -} - -BEGIN_TEST(int) - c = (int) int2float(a) + (int) int2float(b) + (int) int2float(42.3); -END_TEST32() diff --git a/test/mppa/instr/fmuld.c b/test/mppa/instr/fmuld.c deleted file mode 100644 index 03c990fa..00000000 --- a/test/mppa/instr/fmuld.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(double) -{ - c = ((double)a * (double)b); -} -END_TESTF64() diff --git a/test/mppa/instr/fmulw.c b/test/mppa/instr/fmulw.c deleted file mode 100644 index f85eba64..00000000 --- a/test/mppa/instr/fmulw.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(float) -{ - c = ((float)a * (float)b); -} -END_TESTF32() diff --git a/test/mppa/instr/fnegd.c b/test/mppa/instr/fnegd.c deleted file mode 100644 index 974eb7e8..00000000 --- a/test/mppa/instr/fnegd.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(double) -{ - c = (-(double)a); -} -END_TESTF64() diff --git a/test/mppa/instr/fnegw.c b/test/mppa/instr/fnegw.c deleted file mode 100644 index fbeaab8e..00000000 --- a/test/mppa/instr/fnegw.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(float) -{ - c = (-(float)a); -} -END_TESTF64() diff --git a/test/mppa/instr/for.c b/test/mppa/instr/for.c deleted file mode 100644 index 373ab6bd..00000000 --- a/test/mppa/instr/for.c +++ /dev/null @@ -1,9 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - int j; - for (j = 0 ; j < 10 ; j++) - c += a; -} -END_TEST32() diff --git a/test/mppa/instr/forvar.c b/test/mppa/instr/forvar.c deleted file mode 100644 index 9e43c198..00000000 --- a/test/mppa/instr/forvar.c +++ /dev/null @@ -1,9 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - int k; - for (k = 0 ; k < (b & 0x8) ; k++) - c += a; -} -END_TEST32() diff --git a/test/mppa/instr/forvarl.c b/test/mppa/instr/forvarl.c deleted file mode 100644 index c1fe90fd..00000000 --- a/test/mppa/instr/forvarl.c +++ /dev/null @@ -1,10 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long int) -{ - int j; - - for (j = 0 ; j < (b & 0x8LL) ; j++) - c += a; -} -END_TEST64() diff --git a/test/mppa/instr/fsbfd.c b/test/mppa/instr/fsbfd.c deleted file mode 100644 index f80c1efe..00000000 --- a/test/mppa/instr/fsbfd.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(double) -{ - c = ((double)a - (double)b); -} -END_TESTF64() diff --git a/test/mppa/instr/fsbfw.c b/test/mppa/instr/fsbfw.c deleted file mode 100644 index 067c40b5..00000000 --- a/test/mppa/instr/fsbfw.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(float) -{ - c = ((float)a - (float)b); -} -END_TESTF64() diff --git a/test/mppa/instr/i32.c b/test/mppa/instr/i32.c new file mode 100644 index 00000000..c48531b1 --- /dev/null +++ b/test/mppa/instr/i32.c @@ -0,0 +1,87 @@ +#include "framework.h" + +int sum(int a, int b){ + return a+b; +} + +int make(int a){ + return a; +} + +int tailsum(int a, int b){ + return make(a+b); +} + +float int2float(int v){ + return v; +} + +BEGIN_TEST(int) + c = a+b; + c += a&b; + + if ((a & 0x1) == 1) + c += 1; + else + c += 2; + + if (a & 0x1 == 0) + c += 4; + else + c += 8; + + b = !(a & 0x01); + if (!b) + c += 16; + else + c += 32; + + c += sum(make(a), make(b)); + c += (long long) a; + + if (0 > (a & 0x1) - 1) + c += 64; + else + c += 128; + + if (0 >= (a & 0x1)) + c += 256; + else + c += 512; + + if ((a & 0x1) > 0) + c += 1024; + else + c += 2048; + + if ((a & 0x1) - 1 >= 0) + c += 4096; + else + c += 8192; + + c += ((a & 0x1) == (b & 0x1)); + c += (a > b); + c += (a <= b); + c += (a < b); + c += (a + b) / 2; + c += (int) int2float(a) + (int) int2float(b) + (int) int2float(42.3); + + int j; + for (j = 0 ; j < 10 ; j++) + c += a; + int k; + for (k = 0 ; k < (b & 0x8) ; k++) + c += a; + + char s[] = "Tome and Cherry at the playa\n"; + c += s[(a & (sizeof(s)-1))]; + + unsigned char s2[] = "Tim is sorry at the playa\n"; + c += s2[a & (sizeof(s) - 1)]; + + c += a*b; + c += a-b; + c += a << (b & 0x8); + + c += sum(a, b); +END_TEST32() diff --git a/test/mppa/instr/i64.c b/test/mppa/instr/i64.c new file mode 100644 index 00000000..00eb159d --- /dev/null +++ b/test/mppa/instr/i64.c @@ -0,0 +1,107 @@ +#include "framework.h" + +long long sum(long long a, long long b){ + return a+b; +} + +long long diff(long long a, long long b){ + return a-b; +} + +long long mul(long long a, long long b){ + return a*b; +} + +long long make(long long a){ + return a; +} + +long long random_op(long long a, long long b){ + long long d = 3; + long long (*op)(long long, long long); + + if (a % d == 0) + op = sum; + else if (a % d == 1) + op = diff; + else + op = mul; + + return op(a, b); +} + +double long2double(long v){ + return v; +} + +BEGIN_TEST(long long) + c = a&b; + c += a*b; + c += -a; + c += a | b; + c += a-b; + c += a >> (b & 0x8LL); + c += a >> (b & 0x8ULL); + c += a % b; + + long long d = 3; + long long (*op)(long long, long long); + + if (a % d == 0) + op = sum; + else if (a % d == 1) + op = diff; + else + op = mul; + + c += op(make(a), make(b)); + c += random_op(a, b); + c += a/b; + c += a^b; + c += (unsigned int) a; + + if (0 != (a & 0x1LL)) + c += 1; + else + c += 2; + + if (0 > (a & 0x1LL)) + c += 4; + else + c += 8; + + if (0 >= (a & 0x1LL) - 1) + c += 16; + else + c += 32; + + if (a & 0x1LL > 0) + c += 64; + else + c += 128; + + if ((a & 0x1LL) - 1 >= 0) + c += 256; + else + c += 512; + + if (0 == (a & 0x1LL)) + c += 1024; + else + c += 2048; + + c += ((a & 0x1LL) == (b & 0x1LL)); + c += (a >= b); + c += (a > b); + c += (a <= b); + c += (a < b); + c += (long) long2double(a) + (long) long2double(b) + (long) long2double(42.3); + + int j; + + for (j = 0 ; j < (b & 0x8LL) ; j++) + c += a; + + c += ((a & 0x1LL) == (b & 0x1LL)); + +END_TEST64() diff --git a/test/mppa/instr/indirect_call.c b/test/mppa/instr/indirect_call.c deleted file mode 100644 index f376c00a..00000000 --- a/test/mppa/instr/indirect_call.c +++ /dev/null @@ -1,33 +0,0 @@ -#include "framework.h" - -long long sum(long long a, long long b){ - return a+b; -} - -long long diff(long long a, long long b){ - return a-b; -} - -long long mul(long long a, long long b){ - return a*b; -} - -long long make(long long a){ - return a; -} - -BEGIN_TEST(long long) -{ - long long d = 3; - long long (*op)(long long, long long); - - if (a % d == 0) - op = sum; - else if (a % d == 1) - op = diff; - else - op = mul; - - c += op(make(a), make(b)); -} -END_TEST64() diff --git a/test/mppa/instr/indirect_tailcall.c b/test/mppa/instr/indirect_tailcall.c deleted file mode 100644 index e6c16ea1..00000000 --- a/test/mppa/instr/indirect_tailcall.c +++ /dev/null @@ -1,33 +0,0 @@ -#include "framework.h" - -long long sum(long long a, long long b){ - return a+b; -} - -long long diff(long long a, long long b){ - return a-b; -} - -long long mul(long long a, long long b){ - return a*b; -} - -long long random_op(long long a, long long b){ - long long d = 3; - long long (*op)(long long, long long); - - if (a % d == 0) - op = sum; - else if (a % d == 1) - op = diff; - else - op = mul; - - return op(a, b); -} - -BEGIN_TEST(long long) -{ - c += random_op(a, b); -} -END_TEST64() diff --git a/test/mppa/instr/individual/andw.c b/test/mppa/instr/individual/andw.c new file mode 100644 index 00000000..799dc7fb --- /dev/null +++ b/test/mppa/instr/individual/andw.c @@ -0,0 +1,5 @@ +#include "framework.h" + +BEGIN_TEST(int) + c = a&b; +END_TEST32() diff --git a/test/mppa/instr/individual/branch.c b/test/mppa/instr/individual/branch.c new file mode 100644 index 00000000..c9937e31 --- /dev/null +++ b/test/mppa/instr/individual/branch.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + if ((a & 0x1) == 1) + c = 0; + else + c = 1; +} +END_TEST32() diff --git a/test/mppa/instr/individual/branchz.c b/test/mppa/instr/individual/branchz.c new file mode 100644 index 00000000..d3e021b5 --- /dev/null +++ b/test/mppa/instr/individual/branchz.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + if (a & 0x1 == 0) + c = 0; + else + c = 1; +} +END_TEST32() diff --git a/test/mppa/instr/individual/branchzu.c b/test/mppa/instr/individual/branchzu.c new file mode 100644 index 00000000..d0169174 --- /dev/null +++ b/test/mppa/instr/individual/branchzu.c @@ -0,0 +1,11 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + b = !(a & 0x01); + if (!b) + c = 0; + else + c = 1; +} +END_TEST32() diff --git a/test/mppa/instr/individual/call.c b/test/mppa/instr/individual/call.c new file mode 100644 index 00000000..ba2ec323 --- /dev/null +++ b/test/mppa/instr/individual/call.c @@ -0,0 +1,16 @@ +#include "framework.h" + +int sum(int a, int b){ + return a+b; +} + +int make(int a){ + return a; +} + +BEGIN_TEST(int) +{ + c = sum(make(a), make(b)); +} +END_TEST32() +/* RETURN VALUE: 60 */ diff --git a/test/mppa/instr/individual/cast_S32_S64.c b/test/mppa/instr/individual/cast_S32_S64.c new file mode 100644 index 00000000..09c97e00 --- /dev/null +++ b/test/mppa/instr/individual/cast_S32_S64.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = (long long) a; +} +END_TEST32() diff --git a/test/mppa/instr/individual/cast_S64_U32.c b/test/mppa/instr/individual/cast_S64_U32.c new file mode 100644 index 00000000..2d9dc723 --- /dev/null +++ b/test/mppa/instr/individual/cast_S64_U32.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = (unsigned int) a; +} +END_TEST64() diff --git a/test/mppa/instr/individual/cb.deqz.c b/test/mppa/instr/individual/cb.deqz.c new file mode 100644 index 00000000..6da2ab07 --- /dev/null +++ b/test/mppa/instr/individual/cb.deqz.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + if (0 != (a & 0x1LL)) + c = 1; + else + c = 0; +} +END_TEST64() diff --git a/test/mppa/instr/individual/cb.dgez.c b/test/mppa/instr/individual/cb.dgez.c new file mode 100644 index 00000000..7bef25ad --- /dev/null +++ b/test/mppa/instr/individual/cb.dgez.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + if (0 > (a & 0x1LL)) + c = 1; + else + c = 0; +} +END_TEST64() diff --git a/test/mppa/instr/individual/cb.dgtz.c b/test/mppa/instr/individual/cb.dgtz.c new file mode 100644 index 00000000..1a43fb1f --- /dev/null +++ b/test/mppa/instr/individual/cb.dgtz.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + if (0 >= (a & 0x1LL) - 1) + c = 1; + else + c = 0; +} +END_TEST64() diff --git a/test/mppa/instr/individual/cb.dlez.c b/test/mppa/instr/individual/cb.dlez.c new file mode 100644 index 00000000..2fb97939 --- /dev/null +++ b/test/mppa/instr/individual/cb.dlez.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + if (a & 0x1LL > 0) + c = 1; + else + c = 0; +} +END_TEST64() diff --git a/test/mppa/instr/individual/cb.dltz.c b/test/mppa/instr/individual/cb.dltz.c new file mode 100644 index 00000000..a431d5d0 --- /dev/null +++ b/test/mppa/instr/individual/cb.dltz.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + if ((a & 0x1LL) - 1 >= 0) + c = 1; + else + c = 0; +} +END_TEST64() diff --git a/test/mppa/instr/individual/cb.dnez.c b/test/mppa/instr/individual/cb.dnez.c new file mode 100644 index 00000000..44516cbe --- /dev/null +++ b/test/mppa/instr/individual/cb.dnez.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + if (0 == (a & 0x1LL)) + c = 1; + else + c = 0; +} +END_TEST64() diff --git a/test/mppa/instr/individual/cb.wgez.c b/test/mppa/instr/individual/cb.wgez.c new file mode 100644 index 00000000..5779ad92 --- /dev/null +++ b/test/mppa/instr/individual/cb.wgez.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + if (0 > (a & 0x1) - 1) + c = 1; + else + c = 0; +} +END_TEST32() diff --git a/test/mppa/instr/individual/cb.wgtz.c b/test/mppa/instr/individual/cb.wgtz.c new file mode 100644 index 00000000..abb695bd --- /dev/null +++ b/test/mppa/instr/individual/cb.wgtz.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + if (0 >= (a & 0x1)) + c = 1; + else + c = 0; +} +END_TEST32() diff --git a/test/mppa/instr/individual/cb.wlez.c b/test/mppa/instr/individual/cb.wlez.c new file mode 100644 index 00000000..3a2e08c1 --- /dev/null +++ b/test/mppa/instr/individual/cb.wlez.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + if ((a & 0x1) > 0) + c = 1; + else + c = 0; +} +END_TEST32() diff --git a/test/mppa/instr/individual/cb.wltz.c b/test/mppa/instr/individual/cb.wltz.c new file mode 100644 index 00000000..5d52c72a --- /dev/null +++ b/test/mppa/instr/individual/cb.wltz.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + if ((a & 0x1) - 1 >= 0) + c = 1; + else + c = 0; +} +END_TEST32() diff --git a/test/mppa/instr/individual/compd.eq.c b/test/mppa/instr/individual/compd.eq.c new file mode 100644 index 00000000..4fe8de2a --- /dev/null +++ b/test/mppa/instr/individual/compd.eq.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = ((a & 0x1LL) == (b & 0x1LL)); +} +END_TEST64() diff --git a/test/mppa/instr/individual/compd.geu.c b/test/mppa/instr/individual/compd.geu.c new file mode 100644 index 00000000..fccf0804 --- /dev/null +++ b/test/mppa/instr/individual/compd.geu.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned long long) +{ + c = (a >= b); +} +END_TEST64() diff --git a/test/mppa/instr/individual/compd.gt.c b/test/mppa/instr/individual/compd.gt.c new file mode 100644 index 00000000..b9901436 --- /dev/null +++ b/test/mppa/instr/individual/compd.gt.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = (a > b); +} +END_TEST64() diff --git a/test/mppa/instr/individual/compd.le.c b/test/mppa/instr/individual/compd.le.c new file mode 100644 index 00000000..6fa0f103 --- /dev/null +++ b/test/mppa/instr/individual/compd.le.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = (a <= b); +} +END_TEST64() diff --git a/test/mppa/instr/individual/compd.leu.c b/test/mppa/instr/individual/compd.leu.c new file mode 100644 index 00000000..1ad18281 --- /dev/null +++ b/test/mppa/instr/individual/compd.leu.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned long long) +{ + c = (a <= b); +} +END_TEST64() diff --git a/test/mppa/instr/individual/compd.lt.c b/test/mppa/instr/individual/compd.lt.c new file mode 100644 index 00000000..c42cda56 --- /dev/null +++ b/test/mppa/instr/individual/compd.lt.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = (a < b); +} +END_TEST64() diff --git a/test/mppa/instr/individual/compd.ltu.c b/test/mppa/instr/individual/compd.ltu.c new file mode 100644 index 00000000..b03d4d53 --- /dev/null +++ b/test/mppa/instr/individual/compd.ltu.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned long long) +{ + c = (a < b); +} +END_TEST64() diff --git a/test/mppa/instr/individual/compd.ne.c b/test/mppa/instr/individual/compd.ne.c new file mode 100644 index 00000000..fd9d0b28 --- /dev/null +++ b/test/mppa/instr/individual/compd.ne.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned long long) +{ + c = ((a & 0x1ULL) != (b & 0x1ULL)); +} +END_TEST64() diff --git a/test/mppa/instr/individual/compw.eq.c b/test/mppa/instr/individual/compw.eq.c new file mode 100644 index 00000000..cd93f365 --- /dev/null +++ b/test/mppa/instr/individual/compw.eq.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = ((a & 0x1) == (b & 0x1)); +} +END_TEST32() diff --git a/test/mppa/instr/individual/compw.geu.c b/test/mppa/instr/individual/compw.geu.c new file mode 100644 index 00000000..b8fb1adf --- /dev/null +++ b/test/mppa/instr/individual/compw.geu.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned int) +{ + c = (a >= b); +} +END_TEST32() diff --git a/test/mppa/instr/individual/compw.gt.c b/test/mppa/instr/individual/compw.gt.c new file mode 100644 index 00000000..5f6bc907 --- /dev/null +++ b/test/mppa/instr/individual/compw.gt.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = (a > b); +} +END_TEST32() diff --git a/test/mppa/instr/individual/compw.gtu.c b/test/mppa/instr/individual/compw.gtu.c new file mode 100644 index 00000000..947f6a14 --- /dev/null +++ b/test/mppa/instr/individual/compw.gtu.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned int) +{ + c = (a > b); +} +END_TEST32() diff --git a/test/mppa/instr/individual/compw.le.c b/test/mppa/instr/individual/compw.le.c new file mode 100644 index 00000000..35ec6b7d --- /dev/null +++ b/test/mppa/instr/individual/compw.le.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = (a <= b); +} +END_TEST32() diff --git a/test/mppa/instr/individual/compw.leu.c b/test/mppa/instr/individual/compw.leu.c new file mode 100644 index 00000000..74ebfb42 --- /dev/null +++ b/test/mppa/instr/individual/compw.leu.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned int) +{ + c = (a <= b); +} +END_TEST32() diff --git a/test/mppa/instr/individual/compw.lt.c b/test/mppa/instr/individual/compw.lt.c new file mode 100644 index 00000000..cb1f30bd --- /dev/null +++ b/test/mppa/instr/individual/compw.lt.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = (a < b); +} +END_TEST32() diff --git a/test/mppa/instr/individual/compw.ltu.c b/test/mppa/instr/individual/compw.ltu.c new file mode 100644 index 00000000..6a0c5af1 --- /dev/null +++ b/test/mppa/instr/individual/compw.ltu.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned int) +{ + c = (a < b); +} +END_TEST32() diff --git a/test/mppa/instr/individual/compw.ne.c b/test/mppa/instr/individual/compw.ne.c new file mode 100644 index 00000000..7035e2c7 --- /dev/null +++ b/test/mppa/instr/individual/compw.ne.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned int) +{ + c = ((a & 0x1U) != (b & 0x1U)); +} +END_TEST32() diff --git a/test/mppa/instr/individual/div2.c b/test/mppa/instr/individual/div2.c new file mode 100644 index 00000000..b5dfe63a --- /dev/null +++ b/test/mppa/instr/individual/div2.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = (a + b) / 2; +} +END_TEST32() diff --git a/test/mppa/instr/individual/doubleconv.c b/test/mppa/instr/individual/doubleconv.c new file mode 100644 index 00000000..55b1ddab --- /dev/null +++ b/test/mppa/instr/individual/doubleconv.c @@ -0,0 +1,9 @@ +#include "framework.h" + +double long2double(long v){ + return v; +} + +BEGIN_TEST(long) + c = (long) long2double(a) + (long) long2double(b) + (long) long2double(42.3); +END_TEST64() diff --git a/test/mppa/instr/individual/floatconv.c b/test/mppa/instr/individual/floatconv.c new file mode 100644 index 00000000..32b798e1 --- /dev/null +++ b/test/mppa/instr/individual/floatconv.c @@ -0,0 +1,9 @@ +#include "framework.h" + +float int2float(int v){ + return v; +} + +BEGIN_TEST(int) + c = (int) int2float(a) + (int) int2float(b) + (int) int2float(42.3); +END_TEST32() diff --git a/test/mppa/instr/individual/fmuld.c b/test/mppa/instr/individual/fmuld.c new file mode 100644 index 00000000..03c990fa --- /dev/null +++ b/test/mppa/instr/individual/fmuld.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(double) +{ + c = ((double)a * (double)b); +} +END_TESTF64() diff --git a/test/mppa/instr/individual/fmulw.c b/test/mppa/instr/individual/fmulw.c new file mode 100644 index 00000000..f85eba64 --- /dev/null +++ b/test/mppa/instr/individual/fmulw.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(float) +{ + c = ((float)a * (float)b); +} +END_TESTF32() diff --git a/test/mppa/instr/individual/fnegd.c b/test/mppa/instr/individual/fnegd.c new file mode 100644 index 00000000..974eb7e8 --- /dev/null +++ b/test/mppa/instr/individual/fnegd.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(double) +{ + c = (-(double)a); +} +END_TESTF64() diff --git a/test/mppa/instr/individual/fnegw.c b/test/mppa/instr/individual/fnegw.c new file mode 100644 index 00000000..fbeaab8e --- /dev/null +++ b/test/mppa/instr/individual/fnegw.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(float) +{ + c = (-(float)a); +} +END_TESTF64() diff --git a/test/mppa/instr/individual/for.c b/test/mppa/instr/individual/for.c new file mode 100644 index 00000000..373ab6bd --- /dev/null +++ b/test/mppa/instr/individual/for.c @@ -0,0 +1,9 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + int j; + for (j = 0 ; j < 10 ; j++) + c += a; +} +END_TEST32() diff --git a/test/mppa/instr/individual/forvar.c b/test/mppa/instr/individual/forvar.c new file mode 100644 index 00000000..9e43c198 --- /dev/null +++ b/test/mppa/instr/individual/forvar.c @@ -0,0 +1,9 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + int k; + for (k = 0 ; k < (b & 0x8) ; k++) + c += a; +} +END_TEST32() diff --git a/test/mppa/instr/individual/forvarl.c b/test/mppa/instr/individual/forvarl.c new file mode 100644 index 00000000..c1fe90fd --- /dev/null +++ b/test/mppa/instr/individual/forvarl.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(long long int) +{ + int j; + + for (j = 0 ; j < (b & 0x8LL) ; j++) + c += a; +} +END_TEST64() diff --git a/test/mppa/instr/individual/fsbfd.c b/test/mppa/instr/individual/fsbfd.c new file mode 100644 index 00000000..f80c1efe --- /dev/null +++ b/test/mppa/instr/individual/fsbfd.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(double) +{ + c = ((double)a - (double)b); +} +END_TESTF64() diff --git a/test/mppa/instr/individual/fsbfw.c b/test/mppa/instr/individual/fsbfw.c new file mode 100644 index 00000000..067c40b5 --- /dev/null +++ b/test/mppa/instr/individual/fsbfw.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(float) +{ + c = ((float)a - (float)b); +} +END_TESTF64() diff --git a/test/mppa/instr/individual/indirect_call.c b/test/mppa/instr/individual/indirect_call.c new file mode 100644 index 00000000..f376c00a --- /dev/null +++ b/test/mppa/instr/individual/indirect_call.c @@ -0,0 +1,33 @@ +#include "framework.h" + +long long sum(long long a, long long b){ + return a+b; +} + +long long diff(long long a, long long b){ + return a-b; +} + +long long mul(long long a, long long b){ + return a*b; +} + +long long make(long long a){ + return a; +} + +BEGIN_TEST(long long) +{ + long long d = 3; + long long (*op)(long long, long long); + + if (a % d == 0) + op = sum; + else if (a % d == 1) + op = diff; + else + op = mul; + + c += op(make(a), make(b)); +} +END_TEST64() diff --git a/test/mppa/instr/individual/indirect_tailcall.c b/test/mppa/instr/individual/indirect_tailcall.c new file mode 100644 index 00000000..e6c16ea1 --- /dev/null +++ b/test/mppa/instr/individual/indirect_tailcall.c @@ -0,0 +1,33 @@ +#include "framework.h" + +long long sum(long long a, long long b){ + return a+b; +} + +long long diff(long long a, long long b){ + return a-b; +} + +long long mul(long long a, long long b){ + return a*b; +} + +long long random_op(long long a, long long b){ + long long d = 3; + long long (*op)(long long, long long); + + if (a % d == 0) + op = sum; + else if (a % d == 1) + op = diff; + else + op = mul; + + return op(a, b); +} + +BEGIN_TEST(long long) +{ + c += random_op(a, b); +} +END_TEST64() diff --git a/test/mppa/instr/individual/lbs.c b/test/mppa/instr/individual/lbs.c new file mode 100644 index 00000000..22a50632 --- /dev/null +++ b/test/mppa/instr/individual/lbs.c @@ -0,0 +1,9 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + char s[] = "Tome and Cherry at the playa\n"; + + c = s[(a & (sizeof(s)-1))]; +} +END_TEST32() diff --git a/test/mppa/instr/individual/lbz.c b/test/mppa/instr/individual/lbz.c new file mode 100644 index 00000000..04ba098d --- /dev/null +++ b/test/mppa/instr/individual/lbz.c @@ -0,0 +1,9 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + unsigned char s[] = "Tim is sorry at the playa\n"; + + c = s[a & (sizeof(s) - 1)]; +} +END_TEST32() diff --git a/test/mppa/instr/individual/muld.c b/test/mppa/instr/individual/muld.c new file mode 100644 index 00000000..f7e23850 --- /dev/null +++ b/test/mppa/instr/individual/muld.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = a*b; +} +END_TEST64() diff --git a/test/mppa/instr/individual/mulw.c b/test/mppa/instr/individual/mulw.c new file mode 100644 index 00000000..a91d966e --- /dev/null +++ b/test/mppa/instr/individual/mulw.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = a * b; +} +END_TEST32() diff --git a/test/mppa/instr/individual/negd.c b/test/mppa/instr/individual/negd.c new file mode 100644 index 00000000..837b9828 --- /dev/null +++ b/test/mppa/instr/individual/negd.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = -a; +} +END_TEST64() diff --git a/test/mppa/instr/individual/ord.c b/test/mppa/instr/individual/ord.c new file mode 100644 index 00000000..cae1ae8b --- /dev/null +++ b/test/mppa/instr/individual/ord.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = a | b; +} +END_TEST64() diff --git a/test/mppa/instr/individual/sbfd.c b/test/mppa/instr/individual/sbfd.c new file mode 100644 index 00000000..77c28c77 --- /dev/null +++ b/test/mppa/instr/individual/sbfd.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = a-b; +} +END_TEST64() diff --git a/test/mppa/instr/individual/sbfw.c b/test/mppa/instr/individual/sbfw.c new file mode 100644 index 00000000..e38a1fff --- /dev/null +++ b/test/mppa/instr/individual/sbfw.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = a-b; +} +END_TEST32() diff --git a/test/mppa/instr/individual/simple.c b/test/mppa/instr/individual/simple.c new file mode 100644 index 00000000..944f09c9 --- /dev/null +++ b/test/mppa/instr/individual/simple.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = a+b; +} +END_TEST32() diff --git a/test/mppa/instr/individual/sllw.c b/test/mppa/instr/individual/sllw.c new file mode 100644 index 00000000..6dd41a6c --- /dev/null +++ b/test/mppa/instr/individual/sllw.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(int) +{ + c = a << (b & 0x8); +} +END_TEST32() diff --git a/test/mppa/instr/individual/srad.c b/test/mppa/instr/individual/srad.c new file mode 100644 index 00000000..00be9d0c --- /dev/null +++ b/test/mppa/instr/individual/srad.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = a >> (b & 0x8LL); +} +END_TEST64() diff --git a/test/mppa/instr/individual/srld.c b/test/mppa/instr/individual/srld.c new file mode 100644 index 00000000..14970efd --- /dev/null +++ b/test/mppa/instr/individual/srld.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned long long) +{ + c = a >> (b & 0x8ULL); +} +END_TEST64() diff --git a/test/mppa/instr/individual/tailcall.c b/test/mppa/instr/individual/tailcall.c new file mode 100644 index 00000000..6c659a01 --- /dev/null +++ b/test/mppa/instr/individual/tailcall.c @@ -0,0 +1,16 @@ +#include "framework.h" + +int make(int a){ + return a; +} + +int sum(int a, int b){ + return make(a+b); +} + +BEGIN_TEST(int) +{ + c = sum(a, b); +} +END_TEST32() +/* RETURN VALUE: 60 */ diff --git a/test/mppa/instr/individual/udivd.c b/test/mppa/instr/individual/udivd.c new file mode 100644 index 00000000..cfb31881 --- /dev/null +++ b/test/mppa/instr/individual/udivd.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned long long) +{ + c = a/b; +} +END_TEST64() diff --git a/test/mppa/instr/individual/umodd.c b/test/mppa/instr/individual/umodd.c new file mode 100644 index 00000000..a7f25f1c --- /dev/null +++ b/test/mppa/instr/individual/umodd.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned long long) +{ + c = a%b; +} +END_TEST64() diff --git a/test/mppa/instr/individual/xord.c b/test/mppa/instr/individual/xord.c new file mode 100644 index 00000000..b6a90cb0 --- /dev/null +++ b/test/mppa/instr/individual/xord.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(long long) +{ + c = a^b; +} +END_TEST64() diff --git a/test/mppa/instr/lbs.c b/test/mppa/instr/lbs.c deleted file mode 100644 index 22a50632..00000000 --- a/test/mppa/instr/lbs.c +++ /dev/null @@ -1,9 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - char s[] = "Tome and Cherry at the playa\n"; - - c = s[(a & (sizeof(s)-1))]; -} -END_TEST32() diff --git a/test/mppa/instr/lbz.c b/test/mppa/instr/lbz.c deleted file mode 100644 index 04ba098d..00000000 --- a/test/mppa/instr/lbz.c +++ /dev/null @@ -1,9 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - unsigned char s[] = "Tim is sorry at the playa\n"; - - c = s[a & (sizeof(s) - 1)]; -} -END_TEST32() diff --git a/test/mppa/instr/muld.c b/test/mppa/instr/muld.c deleted file mode 100644 index f7e23850..00000000 --- a/test/mppa/instr/muld.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = a*b; -} -END_TEST64() diff --git a/test/mppa/instr/mulw.c b/test/mppa/instr/mulw.c deleted file mode 100644 index a91d966e..00000000 --- a/test/mppa/instr/mulw.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = a * b; -} -END_TEST32() diff --git a/test/mppa/instr/negd.c b/test/mppa/instr/negd.c deleted file mode 100644 index 837b9828..00000000 --- a/test/mppa/instr/negd.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = -a; -} -END_TEST64() diff --git a/test/mppa/instr/ord.c b/test/mppa/instr/ord.c deleted file mode 100644 index cae1ae8b..00000000 --- a/test/mppa/instr/ord.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = a | b; -} -END_TEST64() diff --git a/test/mppa/instr/sbfd.c b/test/mppa/instr/sbfd.c deleted file mode 100644 index 77c28c77..00000000 --- a/test/mppa/instr/sbfd.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = a-b; -} -END_TEST64() diff --git a/test/mppa/instr/sbfw.c b/test/mppa/instr/sbfw.c deleted file mode 100644 index e38a1fff..00000000 --- a/test/mppa/instr/sbfw.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = a-b; -} -END_TEST32() diff --git a/test/mppa/instr/simple.c b/test/mppa/instr/simple.c deleted file mode 100644 index 944f09c9..00000000 --- a/test/mppa/instr/simple.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = a+b; -} -END_TEST32() diff --git a/test/mppa/instr/sllw.c b/test/mppa/instr/sllw.c deleted file mode 100644 index 6dd41a6c..00000000 --- a/test/mppa/instr/sllw.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(int) -{ - c = a << (b & 0x8); -} -END_TEST32() diff --git a/test/mppa/instr/srad.c b/test/mppa/instr/srad.c deleted file mode 100644 index 00be9d0c..00000000 --- a/test/mppa/instr/srad.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = a >> (b & 0x8LL); -} -END_TEST64() diff --git a/test/mppa/instr/srld.c b/test/mppa/instr/srld.c deleted file mode 100644 index 14970efd..00000000 --- a/test/mppa/instr/srld.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned long long) -{ - c = a >> (b & 0x8ULL); -} -END_TEST64() diff --git a/test/mppa/instr/tailcall.c b/test/mppa/instr/tailcall.c deleted file mode 100644 index 6c659a01..00000000 --- a/test/mppa/instr/tailcall.c +++ /dev/null @@ -1,16 +0,0 @@ -#include "framework.h" - -int make(int a){ - return a; -} - -int sum(int a, int b){ - return make(a+b); -} - -BEGIN_TEST(int) -{ - c = sum(a, b); -} -END_TEST32() -/* RETURN VALUE: 60 */ diff --git a/test/mppa/instr/udivd.c b/test/mppa/instr/udivd.c deleted file mode 100644 index cfb31881..00000000 --- a/test/mppa/instr/udivd.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned long long) -{ - c = a/b; -} -END_TEST64() diff --git a/test/mppa/instr/ui32.c b/test/mppa/instr/ui32.c new file mode 100644 index 00000000..f56a9b95 --- /dev/null +++ b/test/mppa/instr/ui32.c @@ -0,0 +1,12 @@ +#include "framework.h" + +BEGIN_TEST(unsigned int) +{ + c = (long long) a; + c += (a >= b); + c += (a > b); + c += (a <= b); + c += (a < b); + c += ((a & 0x1U) != (b & 0x1U)); +} +END_TEST32() diff --git a/test/mppa/instr/ui64.c b/test/mppa/instr/ui64.c new file mode 100644 index 00000000..908dec3c --- /dev/null +++ b/test/mppa/instr/ui64.c @@ -0,0 +1,10 @@ +#include "framework.h" + +BEGIN_TEST(unsigned long long) +{ + c = (a > b); + c += (a <= b); + c += (a < b); + c += ((a & 0x1ULL) != (b & 0x1ULL)); +} +END_TEST64() diff --git a/test/mppa/instr/umodd.c b/test/mppa/instr/umodd.c deleted file mode 100644 index a7f25f1c..00000000 --- a/test/mppa/instr/umodd.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(unsigned long long) -{ - c = a%b; -} -END_TEST64() diff --git a/test/mppa/instr/xord.c b/test/mppa/instr/xord.c deleted file mode 100644 index b6a90cb0..00000000 --- a/test/mppa/instr/xord.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" - -BEGIN_TEST(long long) -{ - c = a^b; -} -END_TEST64() diff --git a/test/mppa/interop/i32.c b/test/mppa/interop/i32.c new file mode 100644 index 00000000..6bc2705c --- /dev/null +++ b/test/mppa/interop/i32.c @@ -0,0 +1,13 @@ +#include "framework.h" +#include "common.h" + +BEGIN_TEST(int) + c = i_manyiargs(a, b, a-b, a+b, a*2, b*2, a*2-b, a+b*2, (a-b)*2, (a+b)*2, + -2*a, -2*b, a-b, a+b, a*3, b*3, a*3-b, a+b*3, (a-b)*3, (a+b)*3, + -3*a, -3*b, a-b, a+b, a*4, b*4, a*4-b, a+b*4, (a-b)*4, (a+b)*4); + c += i_multiiargs(a, b, a-b, a+b); + c += i_oneiarg(a); + void_void(); + c += a; +END_TEST32() + diff --git a/test/mppa/interop/i64.c b/test/mppa/interop/i64.c new file mode 100644 index 00000000..3e7240f7 --- /dev/null +++ b/test/mppa/interop/i64.c @@ -0,0 +1,14 @@ +#include "framework.h" +#include "common.h" + +BEGIN_TEST(long long) + c = ll_manyllargs(a, b, a-b, a+b, a*2, b*2, a*2-b, a+b*2, (a-b)*2, (a+b)*2, + -2*a, -2*b, a-b, a+b, a*3, b*3, a*3-b, a+b*3, (a-b)*3, (a+b)*3, + -3*a, -3*b, a-b, a+b, a*4, b*4, a*4-b, a+b*4, (a-b)*4, (a+b)*4); + c += ll_multillargs(a, b, a-b, a+b); + c += ll_onellarg(a); + c = ll_void(); + c += a; + void_void(); + c += a; +END_TEST64() diff --git a/test/mppa/interop/i_manyiargs.c b/test/mppa/interop/i_manyiargs.c deleted file mode 100644 index d674c26f..00000000 --- a/test/mppa/interop/i_manyiargs.c +++ /dev/null @@ -1,9 +0,0 @@ -#include "framework.h" -#include "common.h" - -BEGIN_TEST(int) - c = i_manyiargs(a, b, a-b, a+b, a*2, b*2, a*2-b, a+b*2, (a-b)*2, (a+b)*2, - -2*a, -2*b, a-b, a+b, a*3, b*3, a*3-b, a+b*3, (a-b)*3, (a+b)*3, - -3*a, -3*b, a-b, a+b, a*4, b*4, a*4-b, a+b*4, (a-b)*4, (a+b)*4); -END_TEST() - diff --git a/test/mppa/interop/i_multiiargs.c b/test/mppa/interop/i_multiiargs.c deleted file mode 100644 index 0e8c8936..00000000 --- a/test/mppa/interop/i_multiiargs.c +++ /dev/null @@ -1,6 +0,0 @@ -#include "framework.h" -#include "common.h" - -BEGIN_TEST(int) - c = i_multiiargs(a, b, a-b, a+b); -END_TEST() diff --git a/test/mppa/interop/i_oneiarg.c b/test/mppa/interop/i_oneiarg.c deleted file mode 100644 index 42cd1540..00000000 --- a/test/mppa/interop/i_oneiarg.c +++ /dev/null @@ -1,6 +0,0 @@ -#include "framework.h" -#include "common.h" - -BEGIN_TEST(int) - c = i_oneiarg(a); -END_TEST() diff --git a/test/mppa/interop/individual/i_multiiargs.c b/test/mppa/interop/individual/i_multiiargs.c new file mode 100644 index 00000000..888742b5 --- /dev/null +++ b/test/mppa/interop/individual/i_multiiargs.c @@ -0,0 +1,6 @@ +#include "framework.h" +#include "common.h" + +BEGIN_TEST(int) + c = i_multiiargs(a, b, a-b, a+b); +END_TEST32() diff --git a/test/mppa/interop/individual/i_oneiarg.c b/test/mppa/interop/individual/i_oneiarg.c new file mode 100644 index 00000000..9c969fb8 --- /dev/null +++ b/test/mppa/interop/individual/i_oneiarg.c @@ -0,0 +1,6 @@ +#include "framework.h" +#include "common.h" + +BEGIN_TEST(int) + c = i_oneiarg(a); +END_TEST32() diff --git a/test/mppa/interop/individual/ll_multillargs.c b/test/mppa/interop/individual/ll_multillargs.c new file mode 100644 index 00000000..34b422eb --- /dev/null +++ b/test/mppa/interop/individual/ll_multillargs.c @@ -0,0 +1,7 @@ +#include "framework.h" +#include "common.h" + +BEGIN_TEST(long long) + c = ll_multillargs(a, b, a-b, a+b); +END_TEST64() + diff --git a/test/mppa/interop/individual/ll_onellarg.c b/test/mppa/interop/individual/ll_onellarg.c new file mode 100644 index 00000000..a2fbbbe9 --- /dev/null +++ b/test/mppa/interop/individual/ll_onellarg.c @@ -0,0 +1,7 @@ +#include "framework.h" +#include "common.h" + +BEGIN_TEST(long long) + c = ll_onellarg(a); +END_TEST64() + diff --git a/test/mppa/interop/individual/ll_void.c b/test/mppa/interop/individual/ll_void.c new file mode 100644 index 00000000..da128fdd --- /dev/null +++ b/test/mppa/interop/individual/ll_void.c @@ -0,0 +1,7 @@ +#include "framework.h" +#include "common.h" + +BEGIN_TEST(long long) + c = ll_void(); + c += a; +END_TEST64() diff --git a/test/mppa/interop/individual/void_void.c b/test/mppa/interop/individual/void_void.c new file mode 100644 index 00000000..976a721b --- /dev/null +++ b/test/mppa/interop/individual/void_void.c @@ -0,0 +1,7 @@ +#include "framework.h" +#include "common.h" + +BEGIN_TEST(long long) + void_void(); + c = a; +END_TEST64() diff --git a/test/mppa/interop/ll_manyllargs.c b/test/mppa/interop/ll_manyllargs.c deleted file mode 100644 index 6e0b3b36..00000000 --- a/test/mppa/interop/ll_manyllargs.c +++ /dev/null @@ -1,8 +0,0 @@ -#include "framework.h" -#include "common.h" - -BEGIN_TEST(long long) - c = ll_manyllargs(a, b, a-b, a+b, a*2, b*2, a*2-b, a+b*2, (a-b)*2, (a+b)*2, - -2*a, -2*b, a-b, a+b, a*3, b*3, a*3-b, a+b*3, (a-b)*3, (a+b)*3, - -3*a, -3*b, a-b, a+b, a*4, b*4, a*4-b, a+b*4, (a-b)*4, (a+b)*4); -END_TEST() diff --git a/test/mppa/interop/ll_multillargs.c b/test/mppa/interop/ll_multillargs.c deleted file mode 100644 index edb03b12..00000000 --- a/test/mppa/interop/ll_multillargs.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" -#include "common.h" - -BEGIN_TEST(long long) - c = ll_multillargs(a, b, a-b, a+b); -END_TEST() - diff --git a/test/mppa/interop/ll_onellarg.c b/test/mppa/interop/ll_onellarg.c deleted file mode 100644 index 0d182166..00000000 --- a/test/mppa/interop/ll_onellarg.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" -#include "common.h" - -BEGIN_TEST(long long) - c = ll_onellarg(a); -END_TEST() - diff --git a/test/mppa/interop/ll_void.c b/test/mppa/interop/ll_void.c deleted file mode 100644 index fa350c9b..00000000 --- a/test/mppa/interop/ll_void.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" -#include "common.h" - -BEGIN_TEST(long long) - c = ll_void(); - c += a; -END_TEST() diff --git a/test/mppa/interop/void_void.c b/test/mppa/interop/void_void.c deleted file mode 100644 index e729edb2..00000000 --- a/test/mppa/interop/void_void.c +++ /dev/null @@ -1,7 +0,0 @@ -#include "framework.h" -#include "common.h" - -BEGIN_TEST(long long) - void_void(); - c = a; -END_TEST() -- cgit From 75ca8b4dcb1dbc75feccd2b53ebf5acdee51f3d3 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 11:10:22 +0100 Subject: Now we have division, simplify the patch. --- test/monniaux/mbedtls/mbedtls_Kalray.patch | 50 ------------------------------ 1 file changed, 50 deletions(-) diff --git a/test/monniaux/mbedtls/mbedtls_Kalray.patch b/test/monniaux/mbedtls/mbedtls_Kalray.patch index 25cd5922..e35fb68f 100644 --- a/test/monniaux/mbedtls/mbedtls_Kalray.patch +++ b/test/monniaux/mbedtls/mbedtls_Kalray.patch @@ -39,56 +39,6 @@ index f89bf9604..9f8a4faeb 100644 /** * \def MBEDTLS_VERSION_C -diff --git a/library/bignum.c b/library/bignum.c -index 87015af0c..a2bad7b84 100644 ---- a/library/bignum.c -+++ b/library/bignum.c -@@ -1696,6 +1696,8 @@ cleanup: - return( ret ); - } - -+#define INT_DIV(x,y) ((x) / (y)) -+ - /* - * Modulo: r = A mod b - */ -@@ -1734,12 +1736,12 @@ int mbedtls_mpi_mod_int( mbedtls_mpi_uint *r, const mbedtls_mpi *A, mbedtls_mpi_ - { - x = A->p[i - 1]; - y = ( y << biH ) | ( x >> biH ); -- z = y / b; -+ z = INT_DIV(y, b); - y -= z * b; - - x <<= biH; - y = ( y << biH ) | ( x >> biH ); -- z = y / b; -+ z = INT_DIV(y, b); - y -= z * b; - } - -diff --git a/library/ssl_tls.c b/library/ssl_tls.c -index 8710a5076..e47c859fe 100644 ---- a/library/ssl_tls.c -+++ b/library/ssl_tls.c -@@ -2482,6 +2482,8 @@ static int ssl_decompress_buf( mbedtls_ssl_context *ssl ) - } - #endif /* MBEDTLS_ZLIB_SUPPORT */ - -+#define INT_DIV(x, y) (((long)(x))/(y)) -+ - #if defined(MBEDTLS_SSL_SRV_C) && defined(MBEDTLS_SSL_RENEGOTIATION) - static int ssl_write_hello_request( mbedtls_ssl_context *ssl ); - -@@ -2492,7 +2494,7 @@ static int ssl_resend_hello_request( mbedtls_ssl_context *ssl ) - * timeout if we were using the usual handshake doubling scheme */ - if( ssl->conf->renego_max_records < 0 ) - { -- uint32_t ratio = ssl->conf->hs_timeout_max / ssl->conf->hs_timeout_min + 1; -+ uint32_t ratio = INT_DIV(ssl->conf->hs_timeout_max, ssl->conf->hs_timeout_min + 1); - unsigned char doublings = 1; - - while( ratio != 0 ) diff --git a/tests/suites/helpers.function b/tests/suites/helpers.function index 1255ff4be..103abe9a8 100644 --- a/tests/suites/helpers.function -- cgit From 0edb9804b295e2cb332f86c252de70fa8e760710 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 22 Mar 2019 11:32:19 +0100 Subject: Rajout de tests test/mppa pour division/modulo --- test/mppa/instr/div32.c | 5 +++++ test/mppa/instr/divf32.c | 5 +++++ test/mppa/instr/divf64.c | 5 +++++ test/mppa/instr/divu32.c | 7 +++++++ test/mppa/instr/modi32.c | 5 +++++ test/mppa/instr/modui32.c | 7 +++++++ 6 files changed, 34 insertions(+) create mode 100644 test/mppa/instr/div32.c create mode 100644 test/mppa/instr/divf32.c create mode 100644 test/mppa/instr/divf64.c create mode 100644 test/mppa/instr/divu32.c create mode 100644 test/mppa/instr/modi32.c create mode 100644 test/mppa/instr/modui32.c diff --git a/test/mppa/instr/div32.c b/test/mppa/instr/div32.c new file mode 100644 index 00000000..83c3a0e3 --- /dev/null +++ b/test/mppa/instr/div32.c @@ -0,0 +1,5 @@ +#include "framework.h" + +BEGIN_TEST(int) + c = a/b; +END_TEST32() diff --git a/test/mppa/instr/divf32.c b/test/mppa/instr/divf32.c new file mode 100644 index 00000000..513a3293 --- /dev/null +++ b/test/mppa/instr/divf32.c @@ -0,0 +1,5 @@ +#include "framework.h" + +BEGIN_TEST(float) + c = a / b; +END_TESTF32() diff --git a/test/mppa/instr/divf64.c b/test/mppa/instr/divf64.c new file mode 100644 index 00000000..0dd23826 --- /dev/null +++ b/test/mppa/instr/divf64.c @@ -0,0 +1,5 @@ +#include "framework.h" + +BEGIN_TEST(double) + c = a / b; +END_TESTF64() diff --git a/test/mppa/instr/divu32.c b/test/mppa/instr/divu32.c new file mode 100644 index 00000000..1fe196c4 --- /dev/null +++ b/test/mppa/instr/divu32.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned int) +{ + c = a/b; +} +END_TEST32() diff --git a/test/mppa/instr/modi32.c b/test/mppa/instr/modi32.c new file mode 100644 index 00000000..958ae920 --- /dev/null +++ b/test/mppa/instr/modi32.c @@ -0,0 +1,5 @@ +#include "framework.h" + +BEGIN_TEST(int) + c = a%b; +END_TEST32() diff --git a/test/mppa/instr/modui32.c b/test/mppa/instr/modui32.c new file mode 100644 index 00000000..a39034a8 --- /dev/null +++ b/test/mppa/instr/modui32.c @@ -0,0 +1,7 @@ +#include "framework.h" + +BEGIN_TEST(unsigned int) +{ + c = a%b; +} +END_TEST32() -- cgit From e6b57a54a4449e0c6239ef56984b1ecbf2b3c792 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 11:49:03 +0100 Subject: better patch for running the test suite --- test/monniaux/mbedtls/mbedtls_Kalray.patch | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/test/monniaux/mbedtls/mbedtls_Kalray.patch b/test/monniaux/mbedtls/mbedtls_Kalray.patch index e35fb68f..594ee7fa 100644 --- a/test/monniaux/mbedtls/mbedtls_Kalray.patch +++ b/test/monniaux/mbedtls/mbedtls_Kalray.patch @@ -55,3 +55,16 @@ index 1255ff4be..103abe9a8 100644 #endif /* Type for Hex parameters */ +diff --git a/tests/scripts/run-test-suites.pl b/tests/scripts/run-test-suites.pl +index d0d4046..00cff26 100755 +--- a/tests/scripts/run-test-suites.pl ++++ b/tests/scripts/run-test-suites.pl +@@ -59,7 +59,7 @@ for my $suite (@suites) + if( $verbose ) { + $command .= ' -v'; + } +- my $result = `$command`; ++ my $result = `k1-cluster -- $command`; + + $suite_cases_passed = () = $result =~ /.. PASS/g; + $suite_cases_failed = () = $result =~ /.. FAILED/g; -- cgit