2020-05-18 22:24:48 +03:00
|
|
|
(define (blodwen-os)
|
|
|
|
(case (machine-type)
|
2020-06-16 13:03:38 +03:00
|
|
|
[(i3le ti3le a6le ta6le) "unix"] ; GNU/Linux
|
|
|
|
[(i3ob ti3ob a6ob ta6ob) "unix"] ; OpenBSD
|
|
|
|
[(i3fb ti3fb a6fb ta6fb) "unix"] ; FreeBSD
|
|
|
|
[(i3nb ti3nb a6nb ta6nb) "unix"] ; NetBSD
|
2020-05-18 22:24:48 +03:00
|
|
|
[(i3osx ti3osx a6osx ta6osx) "darwin"]
|
|
|
|
[(i3nt ti3nt a6nt ta6nt) "windows"]
|
|
|
|
[else "unknown"]))
|
|
|
|
|
2020-05-18 16:55:43 +03:00
|
|
|
(define blodwen-read-args (lambda (desc)
|
|
|
|
(case (vector-ref desc 0)
|
|
|
|
((0) '())
|
|
|
|
((1) (cons (vector-ref desc 2)
|
|
|
|
(blodwen-read-args (vector-ref desc 3)))))))
|
2020-06-01 03:45:15 +03:00
|
|
|
(define b+ (lambda (x y bits) (remainder (+ x y) (ash 1 bits))))
|
|
|
|
(define b- (lambda (x y bits) (remainder (- x y) (ash 1 bits))))
|
|
|
|
(define b* (lambda (x y bits) (remainder (* x y) (ash 1 bits))))
|
|
|
|
(define b/ (lambda (x y bits) (remainder (exact-floor (/ x y)) (ash 1 bits))))
|
2020-05-18 16:55:43 +03:00
|
|
|
|
2021-04-28 11:32:46 +03:00
|
|
|
(define blodwen-toSignedInt (lambda (x y) (modulo x (expt 2 y))))
|
|
|
|
(define blodwen-toUnsignedInt (lambda (x y) (modulo x (expt 2 y))))
|
|
|
|
|
2020-06-01 03:45:15 +03:00
|
|
|
(define integer->bits8 (lambda (x) (modulo x (expt 2 8))))
|
|
|
|
(define integer->bits16 (lambda (x) (modulo x (expt 2 16))))
|
|
|
|
(define integer->bits32 (lambda (x) (modulo x (expt 2 32))))
|
|
|
|
(define integer->bits64 (lambda (x) (modulo x (expt 2 64))))
|
|
|
|
|
2020-08-20 17:01:09 +03:00
|
|
|
(define bits16->bits8 (lambda (x) (modulo x (expt 2 8))))
|
|
|
|
(define bits32->bits8 (lambda (x) (modulo x (expt 2 8))))
|
|
|
|
(define bits32->bits16 (lambda (x) (modulo x (expt 2 16))))
|
|
|
|
(define bits64->bits8 (lambda (x) (modulo x (expt 2 8))))
|
|
|
|
(define bits64->bits16 (lambda (x) (modulo x (expt 2 16))))
|
|
|
|
(define bits64->bits32 (lambda (x) (modulo x (expt 2 32))))
|
|
|
|
|
2021-03-04 23:59:56 +03:00
|
|
|
(define truncate-bits
|
|
|
|
(lambda (x bits)
|
|
|
|
(if (logbit? bits x)
|
|
|
|
(logor x (ash (- 1) bits))
|
|
|
|
(logand x (- (ash 1 bits) 1)))))
|
|
|
|
|
|
|
|
(define blodwen-bits-shl-signed (lambda (x y bits) (truncate-bits (ash x y) bits)))
|
|
|
|
|
2020-06-01 03:45:15 +03:00
|
|
|
(define blodwen-bits-shl (lambda (x y bits) (remainder (ash x y) (ash 1 bits))))
|
2020-05-18 16:55:43 +03:00
|
|
|
(define blodwen-shl (lambda (x y) (ash x y)))
|
|
|
|
(define blodwen-shr (lambda (x y) (ash x (- y))))
|
|
|
|
(define blodwen-and (lambda (x y) (logand x y)))
|
|
|
|
(define blodwen-or (lambda (x y) (logor x y)))
|
|
|
|
(define blodwen-xor (lambda (x y) (logxor x y)))
|
|
|
|
|
|
|
|
(define cast-num
|
|
|
|
(lambda (x)
|
|
|
|
(if (number? x) x 0)))
|
|
|
|
(define destroy-prefix
|
|
|
|
(lambda (x)
|
|
|
|
(cond
|
|
|
|
((equal? x "") "")
|
|
|
|
((equal? (string-ref x 0) #\#) "")
|
|
|
|
(else x))))
|
2020-09-20 17:48:05 +03:00
|
|
|
(define exact-floor
|
|
|
|
(lambda (x)
|
|
|
|
(inexact->exact (floor x))))
|
2020-05-18 16:55:43 +03:00
|
|
|
(define cast-string-int
|
|
|
|
(lambda (x)
|
2020-09-20 17:48:05 +03:00
|
|
|
(exact-floor (cast-num (string->number (destroy-prefix x))))))
|
2020-05-24 00:19:10 +03:00
|
|
|
(define cast-int-char
|
|
|
|
(lambda (x)
|
|
|
|
(if (and (>= x 0)
|
|
|
|
(<= x #x10ffff))
|
|
|
|
(integer->char x)
|
|
|
|
0)))
|
2020-05-18 16:55:43 +03:00
|
|
|
(define cast-string-double
|
|
|
|
(lambda (x)
|
|
|
|
(cast-num (string->number (destroy-prefix x)))))
|
2020-06-11 22:06:56 +03:00
|
|
|
|
|
|
|
(define (from-idris-list xs)
|
|
|
|
(if (= (vector-ref xs 0) 0)
|
|
|
|
'()
|
|
|
|
(cons (vector-ref xs 1) (from-idris-list (vector-ref xs 2)))))
|
2020-08-25 14:30:57 +03:00
|
|
|
(define (string-pack xs) (apply string (from-idris-list xs)))
|
2020-09-19 15:36:57 +03:00
|
|
|
(define (to-idris-list-rev acc xs)
|
|
|
|
(if (null? xs)
|
|
|
|
acc
|
|
|
|
(to-idris-list-rev (vector 1 (car xs) acc) (cdr xs))))
|
|
|
|
(define (string-unpack s) (to-idris-list-rev (vector 0) (reverse (string->list s))))
|
|
|
|
(define (string-concat xs) (apply string-append (from-idris-list xs)))
|
2020-05-18 16:55:43 +03:00
|
|
|
(define string-cons (lambda (x y) (string-append (string x) y)))
|
|
|
|
(define get-tag (lambda (x) (vector-ref x 0)))
|
|
|
|
(define string-reverse (lambda (x)
|
|
|
|
(list->string (reverse (string->list x)))))
|
|
|
|
(define (string-substr off len s)
|
|
|
|
(let* ((l (string-length s))
|
|
|
|
(b (max 0 off))
|
|
|
|
(x (max 0 len))
|
|
|
|
(end (min l (+ b x))))
|
|
|
|
(if (> b l)
|
|
|
|
""
|
|
|
|
(substring s b end))))
|
|
|
|
|
2020-09-19 22:54:34 +03:00
|
|
|
(define (blodwen-string-iterator-new s)
|
2020-09-20 10:56:58 +03:00
|
|
|
0)
|
2020-09-19 22:54:34 +03:00
|
|
|
|
2020-12-31 20:36:07 +03:00
|
|
|
(define (blodwen-string-iterator-to-string _ s ofs f)
|
|
|
|
(f (substring s ofs (string-length s))))
|
|
|
|
|
2020-09-20 10:56:58 +03:00
|
|
|
(define (blodwen-string-iterator-next s ofs)
|
|
|
|
(if (>= ofs (string-length s))
|
|
|
|
(vector 0) ; EOF
|
|
|
|
(vector 1 (string-ref s ofs) (+ ofs 1))))
|
2020-05-18 16:55:43 +03:00
|
|
|
|
|
|
|
(define either-left
|
|
|
|
(lambda (x)
|
|
|
|
(vector 0 x)))
|
|
|
|
|
|
|
|
(define either-right
|
|
|
|
(lambda (x)
|
|
|
|
(vector 1 x)))
|
|
|
|
|
|
|
|
(define blodwen-error-quit
|
|
|
|
(lambda (msg)
|
|
|
|
(display msg)
|
|
|
|
(newline)
|
|
|
|
(exit 1)))
|
|
|
|
|
|
|
|
(define (blodwen-get-line p)
|
|
|
|
(if (port? p)
|
|
|
|
(let ((str (get-line p)))
|
|
|
|
(if (eof-object? str)
|
|
|
|
""
|
|
|
|
str))
|
|
|
|
void))
|
|
|
|
|
|
|
|
(define (blodwen-get-char p)
|
|
|
|
(if (port? p)
|
|
|
|
(let ((chr (get-char p)))
|
|
|
|
(if (eof-object? chr)
|
|
|
|
#\nul
|
|
|
|
chr))
|
|
|
|
void))
|
|
|
|
|
|
|
|
;; Buffers
|
|
|
|
|
|
|
|
(define (blodwen-new-buffer size)
|
|
|
|
(make-bytevector size 0))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-size buf)
|
|
|
|
(bytevector-length buf))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-setbyte buf loc val)
|
|
|
|
(bytevector-u8-set! buf loc val))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-getbyte buf loc)
|
|
|
|
(bytevector-u8-ref buf loc))
|
|
|
|
|
2020-06-01 03:45:15 +03:00
|
|
|
(define (blodwen-buffer-setbits16 buf loc val)
|
|
|
|
(bytevector-u16-set! buf loc val (native-endianness)))
|
2020-05-18 16:55:43 +03:00
|
|
|
|
2020-06-01 03:45:15 +03:00
|
|
|
(define (blodwen-buffer-getbits16 buf loc)
|
|
|
|
(bytevector-u16-ref buf loc (native-endianness)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-setbits32 buf loc val)
|
|
|
|
(bytevector-u32-set! buf loc val (native-endianness)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-getbits32 buf loc)
|
|
|
|
(bytevector-u32-ref buf loc (native-endianness)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-setbits64 buf loc val)
|
|
|
|
(bytevector-u64-set! buf loc val (native-endianness)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-getbits64 buf loc)
|
|
|
|
(bytevector-u64-ref buf loc (native-endianness)))
|
2020-05-19 18:24:23 +03:00
|
|
|
|
|
|
|
(define (blodwen-buffer-setint32 buf loc val)
|
|
|
|
(bytevector-s32-set! buf loc val (native-endianness)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-getint32 buf loc)
|
2020-05-18 16:55:43 +03:00
|
|
|
(bytevector-s32-ref buf loc (native-endianness)))
|
|
|
|
|
2020-06-01 03:45:15 +03:00
|
|
|
(define (blodwen-buffer-setint buf loc val)
|
|
|
|
(bytevector-s64-set! buf loc val (native-endianness)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-getint buf loc)
|
|
|
|
(bytevector-s64-ref buf loc (native-endianness)))
|
|
|
|
|
2020-05-18 16:55:43 +03:00
|
|
|
(define (blodwen-buffer-setdouble buf loc val)
|
|
|
|
(bytevector-ieee-double-set! buf loc val (native-endianness)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-getdouble buf loc)
|
|
|
|
(bytevector-ieee-double-ref buf loc (native-endianness)))
|
|
|
|
|
|
|
|
(define (blodwen-stringbytelen str)
|
|
|
|
(bytevector-length (string->utf8 str)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-setstring buf loc val)
|
|
|
|
(let* [(strvec (string->utf8 val))
|
|
|
|
(len (bytevector-length strvec))]
|
|
|
|
(bytevector-copy! strvec 0 buf loc len)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-getstring buf loc len)
|
|
|
|
(let [(newvec (make-bytevector len))]
|
|
|
|
(bytevector-copy! buf loc newvec 0 len)
|
|
|
|
(utf8->string newvec)))
|
|
|
|
|
|
|
|
(define (blodwen-buffer-copydata buf start len dest loc)
|
|
|
|
(bytevector-copy! buf start dest loc len))
|
|
|
|
|
|
|
|
;; Threads
|
|
|
|
|
2021-02-05 19:16:20 +03:00
|
|
|
(define-record thread-handle (semaphore))
|
2020-05-18 16:55:43 +03:00
|
|
|
|
2021-02-05 19:16:20 +03:00
|
|
|
(define (blodwen-thread proc)
|
|
|
|
(let [(sema (blodwen-make-semaphore 0))]
|
|
|
|
(fork-thread (lambda () (proc (vector 0)) (blodwen-semaphore-post sema)))
|
|
|
|
(make-thread-handle sema)
|
|
|
|
))
|
|
|
|
|
|
|
|
(define (blodwen-thread-wait handle)
|
|
|
|
(blodwen-semaphore-wait (thread-handle-semaphore handle)))
|
|
|
|
|
|
|
|
;; Thread mailboxes
|
|
|
|
|
|
|
|
(define blodwen-thread-data
|
|
|
|
(make-thread-parameter #f))
|
2020-05-18 16:55:43 +03:00
|
|
|
|
|
|
|
(define (blodwen-get-thread-data ty)
|
|
|
|
(blodwen-thread-data))
|
|
|
|
|
|
|
|
(define (blodwen-set-thread-data a)
|
|
|
|
(blodwen-thread-data a))
|
|
|
|
|
2021-02-05 19:16:20 +03:00
|
|
|
;; Semaphore
|
|
|
|
|
|
|
|
(define-record semaphore (box mutex condition))
|
|
|
|
|
|
|
|
(define (blodwen-make-semaphore init)
|
|
|
|
(make-semaphore (box init) (make-mutex) (make-condition)))
|
|
|
|
|
|
|
|
(define (blodwen-semaphore-post sema)
|
|
|
|
(with-mutex (semaphore-mutex sema)
|
|
|
|
(let [(sema-box (semaphore-box sema))]
|
|
|
|
(set-box! sema-box (+ (unbox sema-box) 1))
|
|
|
|
(condition-signal (semaphore-condition sema))
|
|
|
|
)))
|
|
|
|
|
|
|
|
(define (blodwen-semaphore-wait sema)
|
|
|
|
(with-mutex (semaphore-mutex sema)
|
|
|
|
(let [(sema-box (semaphore-box sema))]
|
|
|
|
(when (= (unbox sema-box) 0)
|
|
|
|
(condition-wait (semaphore-condition sema) (semaphore-mutex sema)))
|
|
|
|
(set-box! sema-box (- (unbox sema-box) 1))
|
|
|
|
)))
|
|
|
|
|
|
|
|
;; Barrier
|
|
|
|
|
|
|
|
(define-record barrier (count-box num-threads mutex cond))
|
|
|
|
|
|
|
|
(define (blodwen-make-barrier num-threads)
|
|
|
|
(make-barrier (box 0) num-threads (make-mutex) (make-condition)))
|
|
|
|
|
|
|
|
(define (blodwen-barrier-wait barrier)
|
|
|
|
(let [(count-box (barrier-count-box barrier))
|
|
|
|
(num-threads (barrier-num-threads barrier))
|
|
|
|
(mutex (barrier-mutex barrier))
|
|
|
|
(condition (barrier-cond barrier))]
|
|
|
|
(with-mutex mutex
|
|
|
|
(let* [(count-old (unbox count-box))
|
|
|
|
(count-new (+ count-old 1))]
|
|
|
|
(set-box! count-box count-new)
|
|
|
|
(if (= count-new num-threads)
|
|
|
|
(condition-broadcast condition)
|
|
|
|
(condition-wait condition mutex))
|
|
|
|
))))
|
|
|
|
|
|
|
|
;; Channel
|
|
|
|
|
|
|
|
(define-record channel (box mutex semaphore-get semaphore-put))
|
|
|
|
|
|
|
|
(define (blodwen-make-channel ty)
|
|
|
|
(make-channel
|
|
|
|
(box '())
|
|
|
|
(make-mutex)
|
|
|
|
(blodwen-make-semaphore 0)
|
|
|
|
(blodwen-make-semaphore 0)))
|
|
|
|
|
|
|
|
(define (blodwen-channel-get ty chan)
|
|
|
|
(blodwen-semaphore-post (channel-semaphore-get chan))
|
|
|
|
(blodwen-semaphore-wait (channel-semaphore-put chan))
|
|
|
|
(with-mutex (channel-mutex chan)
|
|
|
|
(let* [(chan-box (channel-box chan))
|
|
|
|
(chan-msg-queue (unbox chan-box))]
|
|
|
|
(set-box! chan-box (cdr chan-msg-queue))
|
|
|
|
(car chan-msg-queue)
|
|
|
|
)))
|
|
|
|
|
|
|
|
(define (blodwen-channel-put ty chan val)
|
|
|
|
(with-mutex (channel-mutex chan)
|
|
|
|
(let* [(chan-box (channel-box chan))
|
|
|
|
(chan-msg-queue (unbox chan-box))]
|
|
|
|
(set-box! chan-box (append chan-msg-queue (list val)))))
|
|
|
|
(blodwen-semaphore-post (channel-semaphore-put chan))
|
|
|
|
(blodwen-semaphore-wait (channel-semaphore-get chan)))
|
|
|
|
|
|
|
|
;; Mutex
|
|
|
|
|
|
|
|
(define (blodwen-make-mutex)
|
|
|
|
(make-mutex))
|
|
|
|
(define (blodwen-mutex-acquire mutex)
|
|
|
|
(mutex-acquire mutex))
|
|
|
|
(define (blodwen-mutex-release mutex)
|
|
|
|
(mutex-release mutex))
|
|
|
|
|
|
|
|
;; Condition variable
|
|
|
|
|
|
|
|
(define (blodwen-make-condition)
|
|
|
|
(make-condition))
|
|
|
|
(define (blodwen-condition-wait condition mutex)
|
|
|
|
(condition-wait condition mutex))
|
|
|
|
(define (blodwen-condition-wait-timeout condition mutex timeout)
|
|
|
|
(let* [(sec (div timeout 1000000))
|
|
|
|
(micro (mod timeout 1000000))]
|
|
|
|
(condition-wait condition mutex (make-time 'time-duration (* 1000 micro) sec))))
|
|
|
|
(define (blodwen-condition-signal condition)
|
|
|
|
(condition-signal condition))
|
|
|
|
(define (blodwen-condition-broadcast condition)
|
|
|
|
(condition-broadcast condition))
|
|
|
|
|
|
|
|
;; Future
|
2020-05-18 16:55:43 +03:00
|
|
|
|
2020-12-04 13:58:26 +03:00
|
|
|
(define-record future-internal (result ready mutex signal))
|
|
|
|
(define (blodwen-make-future work)
|
|
|
|
(let ([future (make-future-internal #f #f (make-mutex) (make-condition))])
|
|
|
|
(fork-thread (lambda ()
|
|
|
|
(let ([result (work)])
|
|
|
|
(with-mutex (future-internal-mutex future)
|
|
|
|
(set-future-internal-result! future result)
|
|
|
|
(set-future-internal-ready! future #t)
|
|
|
|
(condition-broadcast (future-internal-signal future))))))
|
|
|
|
future))
|
|
|
|
(define (blodwen-await-future ty future)
|
|
|
|
(let ([mutex (future-internal-mutex future)])
|
|
|
|
(with-mutex mutex
|
|
|
|
(if (not (future-internal-ready future))
|
|
|
|
(condition-wait (future-internal-signal future) mutex))
|
|
|
|
(future-internal-result future))))
|
|
|
|
|
2020-05-18 16:55:43 +03:00
|
|
|
(define (blodwen-sleep s) (sleep (make-time 'time-duration 0 s)))
|
|
|
|
(define (blodwen-usleep s)
|
|
|
|
(let ((sec (div s 1000000))
|
|
|
|
(micro (mod s 1000000)))
|
|
|
|
(sleep (make-time 'time-duration (* 1000 micro) sec))))
|
|
|
|
|
|
|
|
(define (blodwen-time) (time-second (current-time)))
|
|
|
|
(define (blodwen-clock-time-utc) (current-time 'time-utc))
|
|
|
|
(define (blodwen-clock-time-monotonic) (current-time 'time-monotonic))
|
|
|
|
(define (blodwen-clock-time-duration) (current-time 'time-duration))
|
|
|
|
(define (blodwen-clock-time-process) (current-time 'time-process))
|
|
|
|
(define (blodwen-clock-time-thread) (current-time 'time-thread))
|
|
|
|
(define (blodwen-clock-time-gccpu) (current-time 'time-collector-cpu))
|
|
|
|
(define (blodwen-clock-time-gcreal) (current-time 'time-collector-real))
|
|
|
|
(define (blodwen-is-time? clk) (if (time? clk) 1 0))
|
|
|
|
(define (blodwen-clock-second time) (time-second time))
|
|
|
|
(define (blodwen-clock-nanosecond time) (time-nanosecond time))
|
|
|
|
|
2021-04-03 11:34:33 +03:00
|
|
|
|
|
|
|
(define (blodwen-arg-count)
|
|
|
|
(length (command-line)))
|
|
|
|
|
|
|
|
(define (blodwen-arg n)
|
|
|
|
(if (< n (length (command-line))) (list-ref (command-line) n) ""))
|
2020-05-18 16:55:43 +03:00
|
|
|
|
|
|
|
(define (blodwen-hasenv var)
|
|
|
|
(if (eq? (getenv var) #f) 0 1))
|
|
|
|
|
|
|
|
(define (blodwen-system cmd)
|
|
|
|
(system cmd))
|
|
|
|
|
|
|
|
;; Randoms
|
|
|
|
(define random-seed-register 0)
|
|
|
|
(define (initialize-random-seed-once)
|
|
|
|
(if (= (virtual-register random-seed-register) 0)
|
|
|
|
(let ([seed (time-nanosecond (current-time))])
|
|
|
|
(set-virtual-register! random-seed-register seed)
|
|
|
|
(random-seed seed))))
|
|
|
|
|
|
|
|
(define (blodwen-random-seed seed)
|
|
|
|
(set-virtual-register! random-seed-register seed)
|
|
|
|
(random-seed seed))
|
|
|
|
(define blodwen-random
|
|
|
|
(case-lambda
|
|
|
|
;; no argument, pick a real value from [0, 1.0)
|
|
|
|
[() (begin
|
|
|
|
(initialize-random-seed-once)
|
|
|
|
(random 1.0))]
|
|
|
|
;; single argument k, pick an integral value from [0, k)
|
|
|
|
[(k)
|
|
|
|
(begin
|
|
|
|
(initialize-random-seed-once)
|
|
|
|
(if (> k 0)
|
|
|
|
(random k)
|
|
|
|
(assertion-violationf 'blodwen-random "invalid range argument ~a" k)))]))
|
2020-06-08 22:30:34 +03:00
|
|
|
|
|
|
|
;; For finalisers
|
|
|
|
|
|
|
|
(define blodwen-finaliser (make-guardian))
|
|
|
|
(define (blodwen-register-object obj proc)
|
|
|
|
(let [(x (cons obj proc))]
|
|
|
|
(blodwen-finaliser x)
|
|
|
|
x))
|
|
|
|
(define blodwen-run-finalisers
|
|
|
|
(lambda ()
|
|
|
|
(let run ()
|
|
|
|
(let ([x (blodwen-finaliser)])
|
|
|
|
(when x
|
|
|
|
(((cdr x) (car x)) 'erased)
|
|
|
|
(run))))))
|