(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* 0 *) (* outputs: q % 2^63, r s.t. x = q * y + r, r < y *) let div21 xh xl y = let maskh = ref zero in let maskl = ref one in let dh = ref zero in let dl = ref y in let cmp = ref true in (* n = ref 0 *) (* loop invariant: mask = 2^n, d = mask * y, (2 * d <= x -> cmp), n >= 0 *) while Int64.equal (l_sr !dh (of_int (uint_size - 1))) zero && !cmp do (* We don't use addmuldiv below to avoid checks on 1 *) dh := l_or (l_sl !dh one) (l_sr !dl (of_int (uint_size - 1))); dl := l_sl !dl one; maskh := l_or (l_sl !maskh one) (l_sr !maskl (of_int (uint_size - 1))); maskl := l_sl !maskl one; (* incr n *) cmp := lt128 !dh !dl xh xl; done; (* mask = 2^n, d = 2^n * d, 2 * d > x *) let remh = ref xh in let reml = ref xl in (* quotienth = ref 0 *) let quotientl = ref zero in (* loop invariant: x = quotient * y + rem, y * 2^(n+1) > r, mask = floor(2^n), d = mask * y, n >= -1 *) while not (Int64.equal (l_or !maskh !maskl) zero) do if le128 !dh !dl !remh !reml then begin (* if rem >= d, add one bit and subtract d *) (* quotienth := !quotienth lor !maskh *) quotientl := l_or !quotientl !maskl; remh := if lt !reml !dl then sub (sub !remh !dh) one else sub !remh !dh; reml := sub !reml !dl end; maskl := l_or (l_sr !maskl one) (l_sl !maskh (of_int (uint_size - 1))); maskh := l_sr !maskh one; dl := l_or (l_sr !dl one) (l_sl !dh (of_int (uint_size - 1))); dh := l_sr !dh one (* decr n *) done; !quotientl, !reml let div21 xh xl y = if Int64.equal y zero then zero, zero else div21 xh xl y (* exact multiplication *) let mulc x y = let lx = ref (Int64.logand x maxuint31) in let ly = ref (Int64.logand y maxuint31) in let hx = Int64.shift_right x 31 in let hy = Int64.shift_right y 31 in let hr = ref (Int64.mul hx hy) in let lr = ref (Int64.logor (Int64.mul !lx !ly) (Int64.shift_left !hr 62)) in hr := (Int64.shift_right_logical !hr 1); lx := Int64.mul !lx hy; ly := Int64.mul hx !ly; hr := Int64.logor !hr (Int64.add (Int64.shift_right !lx 32) (Int64.shift_right !ly 32)); lr := Int64.add !lr (Int64.shift_left !lx 31); hr := Int64.add !hr (Int64.shift_right_logical !lr 63); lr := Int64.add (Int64.shift_left !ly 31) (mask63 !lr); hr := Int64.add !hr (Int64.shift_right_logical !lr 63); if Int64.logand !lr Int64.min_int <> 0L then Int64.(sub !hr one, mask63 !lr) else (!hr, !lr) let equal x y = mask63 x = mask63 y let compare x y = Int64.compare x y (* Number of leading zeroes *) let head0 x = let r = ref 0 in let x = ref x in if Int64.logand !x 0x7FFFFFFF00000000L = 0L then r := !r + 31 else x := Int64.shift_right !x 31; if Int64.logand !x 0xFFFF0000L = 0L then (x := Int64.shift_left !x 16; r := !r + 16); if Int64.logand !x 0xFF000000L = 0L then (x := Int64.shift_left !x 8; r := !r + 8); if Int64.logand !x 0xF0000000L = 0L then (x := Int64.shift_left !x 4; r := !r + 4); if Int64.logand !x 0xC0000000L = 0L then (x := Int64.shift_left !x 2; r := !r + 2); if Int64.logand !x 0x80000000L = 0L then (x := Int64.shift_left !x 1; r := !r + 1); if Int64.logand !x 0x80000000L = 0L then (r := !r + 1); Int64.of_int !r (* Number of trailing zeroes *) let tail0 x = let r = ref 0 in let x = ref x in if Int64.logand !x 0xFFFFFFFFL = 0L then (x := Int64.shift_right !x 32; r := !r + 32); if Int64.logand !x 0xFFFFL = 0L then (x := Int64.shift_right !x 16; r := !r + 16); if Int64.logand !x 0xFFL = 0L then (x := Int64.shift_right !x 8; r := !r + 8); if Int64.logand !x 0xFL = 0L then (x := Int64.shift_right !x 4; r := !r + 4); if Int64.logand !x 0x3L = 0L then (x := Int64.shift_right !x 2; r := !r + 2); if Int64.logand !x 0x1L = 0L then (r := !r + 1); Int64.of_int !r (* May an object be safely cast into an Uint63.t ? *) let is_uint63 t = Obj.is_block t && Int.equal (Obj.tag t) Obj.custom_tag && le (Obj.magic t) maxuint63 (* Register all exported functions so that they can be called from C code *) let () = Callback.register "uint63 add" add; Callback.register "uint63 addcarry" addcarry; Callback.register "uint63 addmuldiv" addmuldiv; Callback.register "uint63 div" div; Callback.register "uint63 div21_ml" div21; Callback.register "uint63 eq" equal; Callback.register "uint63 eq0" (equal Int64.zero); Callback.register "uint63 head0" head0; Callback.register "uint63 land" l_and; Callback.register "uint63 leq" le; Callback.register "uint63 lor" l_or; Callback.register "uint63 lsl" l_sl; Callback.register "uint63 lsl1" (fun x -> l_sl x Int64.one); Callback.register "uint63 lsr" l_sr; Callback.register "uint63 lsr1" (fun x -> l_sr x Int64.one); Callback.register "uint63 lt" lt; Callback.register "uint63 lxor" l_xor; Callback.register "uint63 mod" rem; Callback.register "uint63 mul" mul; Callback.register "uint63 mulc_ml" mulc; Callback.register "uint63 one" one; Callback.register "uint63 sub" sub; Callback.register "uint63 subcarry" subcarry; Callback.register "uint63 tail0" tail0