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
|
2021-11-03 05:06:16 +03:00
|
|
|
[(i3osx ti3osx a6osx ta6osx tarm64osx) "darwin"]
|
2020-05-18 22:24:48 +03:00
|
|
|
[(i3nt ti3nt a6nt ta6nt) "windows"]
|
|
|
|
[else "unknown"]))
|
|
|
|
|
2021-09-08 18:46:19 +03:00
|
|
|
(define blodwen-lazy
|
|
|
|
(lambda (f)
|
|
|
|
(let ([evaluated #f] [res void])
|
|
|
|
(lambda ()
|
|
|
|
(if (not evaluated)
|
|
|
|
(begin (set! evaluated #t)
|
|
|
|
(set! res (f))
|
|
|
|
(set! f void))
|
|
|
|
(void))
|
|
|
|
res))))
|
|
|
|
|
2021-11-17 14:54:19 +03:00
|
|
|
(define (blodwen-toSignedInt x bits)
|
|
|
|
(if (logbit? bits x)
|
|
|
|
(logor x (ash -1 bits))
|
|
|
|
(logand x (sub1 (ash 1 bits)))))
|
2021-05-04 10:22:06 +03:00
|
|
|
|
2021-11-17 14:54:19 +03:00
|
|
|
(define (blodwen-toUnsignedInt x bits)
|
|
|
|
(logand x (sub1 (ash 1 bits))))
|
2021-05-04 10:22:06 +03:00
|
|
|
|
2022-03-03 07:55:22 +03:00
|
|
|
(define (blodwen-euclidDiv a b)
|
|
|
|
(let ((q (quotient a b))
|
|
|
|
(r (remainder a b)))
|
|
|
|
(if (< r 0)
|
|
|
|
(if (> b 0) (- q 1) (+ q 1))
|
|
|
|
q)))
|
|
|
|
|
|
|
|
(define (blodwen-euclidMod a b)
|
|
|
|
(let ((r (remainder a b)))
|
|
|
|
(if (< r 0)
|
|
|
|
(if (> b 0) (+ r b) (- r b))
|
|
|
|
r)))
|
2021-09-08 18:46:19 +03:00
|
|
|
|
2021-05-04 10:22:06 +03:00
|
|
|
(define bu+ (lambda (x y bits) (blodwen-toUnsignedInt (+ x y) bits)))
|
|
|
|
(define bu- (lambda (x y bits) (blodwen-toUnsignedInt (- x y) bits)))
|
|
|
|
(define bu* (lambda (x y bits) (blodwen-toUnsignedInt (* x y) bits)))
|
|
|
|
(define bu/ (lambda (x y bits) (blodwen-toUnsignedInt (quotient x y) bits)))
|
|
|
|
|
|
|
|
(define bs+ (lambda (x y bits) (blodwen-toSignedInt (+ x y) bits)))
|
|
|
|
(define bs- (lambda (x y bits) (blodwen-toSignedInt (- x y) bits)))
|
|
|
|
(define bs* (lambda (x y bits) (blodwen-toSignedInt (* x y) bits)))
|
2022-03-03 07:55:22 +03:00
|
|
|
(define bs/ (lambda (x y bits) (blodwen-toSignedInt (blodwen-euclidDiv x y) bits)))
|
2021-05-04 10:22:06 +03:00
|
|
|
|
2021-11-17 14:54:19 +03:00
|
|
|
(define (integer->bits8 x) (logand x (sub1 (ash 1 8))))
|
|
|
|
(define (integer->bits16 x) (logand x (sub1 (ash 1 16))))
|
|
|
|
(define (integer->bits32 x) (logand x (sub1 (ash 1 32))))
|
|
|
|
(define (integer->bits64 x) (logand x (sub1 (ash 1 64))))
|
2020-05-18 16:55:43 +03:00
|
|
|
|
2021-11-17 14:54:19 +03:00
|
|
|
(define (bits16->bits8 x) (logand x (sub1 (ash 1 8))))
|
|
|
|
(define (bits32->bits8 x) (logand x (sub1 (ash 1 8))))
|
|
|
|
(define (bits64->bits8 x) (logand x (sub1 (ash 1 8))))
|
|
|
|
(define (bits32->bits16 x) (logand x (sub1 (ash 1 16))))
|
|
|
|
(define (bits64->bits16 x) (logand x (sub1 (ash 1 16))))
|
|
|
|
(define (bits64->bits32 x) (logand x (sub1 (ash 1 32))))
|
2020-06-01 03:45:15 +03:00
|
|
|
|
2021-11-17 14:54:19 +03:00
|
|
|
(define (blodwen-bits-shl-signed x y bits) (blodwen-toSignedInt (ash x y) bits))
|
2020-08-20 17:01:09 +03:00
|
|
|
|
2021-11-17 14:54:19 +03:00
|
|
|
(define (blodwen-bits-shl x y bits) (logand (ash x y) (sub1 (ash 1 bits))))
|
2021-05-04 10:22:06 +03:00
|
|
|
|
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))))
|
2021-05-04 10:22:06 +03:00
|
|
|
|
2020-09-20 17:48:05 +03:00
|
|
|
(define exact-floor
|
|
|
|
(lambda (x)
|
|
|
|
(inexact->exact (floor x))))
|
2021-05-04 10:22:06 +03:00
|
|
|
|
|
|
|
(define exact-truncate
|
|
|
|
(lambda (x)
|
|
|
|
(inexact->exact (truncate x))))
|
|
|
|
|
|
|
|
(define exact-truncate-boundedInt
|
|
|
|
(lambda (x y)
|
|
|
|
(blodwen-toSignedInt (exact-truncate x) y)))
|
|
|
|
|
|
|
|
(define exact-truncate-boundedUInt
|
|
|
|
(lambda (x y)
|
|
|
|
(blodwen-toUnsignedInt (exact-truncate x) y)))
|
|
|
|
|
|
|
|
(define cast-char-boundedInt
|
|
|
|
(lambda (x y)
|
|
|
|
(blodwen-toSignedInt (char->integer x) y)))
|
|
|
|
|
|
|
|
(define cast-char-boundedUInt
|
|
|
|
(lambda (x y)
|
|
|
|
(blodwen-toUnsignedInt (char->integer x) y)))
|
|
|
|
|
2020-05-18 16:55:43 +03:00
|
|
|
(define cast-string-int
|
|
|
|
(lambda (x)
|
2021-05-04 10:22:06 +03:00
|
|
|
(exact-truncate (cast-num (string->number (destroy-prefix x))))))
|
|
|
|
|
|
|
|
(define cast-string-boundedInt
|
|
|
|
(lambda (x y)
|
|
|
|
(blodwen-toSignedInt (cast-string-int x) y)))
|
|
|
|
|
|
|
|
(define cast-string-boundedUInt
|
|
|
|
(lambda (x y)
|
|
|
|
(blodwen-toUnsignedInt (cast-string-int x) y)))
|
|
|
|
|
2020-05-24 00:19:10 +03:00
|
|
|
(define cast-int-char
|
|
|
|
(lambda (x)
|
2021-05-04 10:22:06 +03:00
|
|
|
(if (or
|
|
|
|
(and (>= x 0) (<= x #xd7ff))
|
|
|
|
(and (>= x #xe000) (<= x #x10ffff)))
|
2020-05-24 00:19:10 +03:00
|
|
|
(integer->char x)
|
2021-05-04 10:22:06 +03:00
|
|
|
(integer->char 0))))
|
|
|
|
|
2020-05-18 16:55:43 +03:00
|
|
|
(define cast-string-double
|
|
|
|
(lambda (x)
|
2022-05-08 02:07:28 +03:00
|
|
|
(exact->inexact (cast-num (string->number (destroy-prefix x))))))
|
|
|
|
|
2020-06-11 22:06:56 +03:00
|
|
|
|
2021-05-08 17:42:51 +03:00
|
|
|
(define (string-concat xs) (apply string-append xs))
|
|
|
|
(define (string-unpack s) (string->list s))
|
|
|
|
(define (string-pack xs) (list->string xs))
|
|
|
|
|
2020-05-18 16:55:43 +03:00
|
|
|
(define string-cons (lambda (x y) (string-append (string x) y)))
|
|
|
|
(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))
|
2021-05-09 03:43:59 +03:00
|
|
|
'() ; EOF
|
|
|
|
(cons (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))
|
|
|
|
|
2021-07-15 17:02:43 +03:00
|
|
|
(define (blodwen-set-thread-data ty a)
|
2020-05-18 16:55:43 +03:00
|
|
|
(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
|
2021-07-02 15:13:50 +03:00
|
|
|
; With thanks to Alain Zscheile (@zseri) for help with understanding condition
|
|
|
|
; variables, and figuring out where the problems were and how to solve them.
|
2021-02-05 19:16:20 +03:00
|
|
|
|
2021-07-02 15:13:50 +03:00
|
|
|
(define-record channel (read-mut read-cv read-box val-cv val-box))
|
2021-02-05 19:16:20 +03:00
|
|
|
|
|
|
|
(define (blodwen-make-channel ty)
|
|
|
|
(make-channel
|
2021-07-02 15:13:50 +03:00
|
|
|
(make-mutex)
|
|
|
|
(make-condition)
|
|
|
|
(box #t)
|
|
|
|
(make-condition)
|
|
|
|
(box '())
|
|
|
|
))
|
2021-02-05 19:16:20 +03:00
|
|
|
|
2021-07-02 15:13:50 +03:00
|
|
|
; block on the read status using read-cv until the value has been read
|
|
|
|
(define (channel-put-while-helper chan)
|
|
|
|
(let ([read-mut (channel-read-mut chan)]
|
|
|
|
[read-box (channel-read-box chan)]
|
|
|
|
[read-cv (channel-read-cv chan)]
|
|
|
|
)
|
|
|
|
(if (unbox read-box)
|
|
|
|
(void) ; val has been read, so everything is fine
|
|
|
|
(begin ; otherwise, block/spin with cv
|
|
|
|
(condition-wait read-cv read-mut)
|
|
|
|
(channel-put-while-helper chan)
|
|
|
|
)
|
2021-02-05 19:16:20 +03:00
|
|
|
)))
|
|
|
|
|
|
|
|
(define (blodwen-channel-put ty chan val)
|
2021-07-02 15:13:50 +03:00
|
|
|
(with-mutex (channel-read-mut chan)
|
|
|
|
(channel-put-while-helper chan)
|
|
|
|
(let ([read-box (channel-read-box chan)]
|
|
|
|
[val-box (channel-val-box chan)]
|
|
|
|
)
|
|
|
|
(set-box! val-box val)
|
|
|
|
(set-box! read-box #f)
|
|
|
|
))
|
|
|
|
(condition-signal (channel-val-cv chan))
|
|
|
|
)
|
|
|
|
|
|
|
|
; block on the value until it has been set
|
|
|
|
(define (channel-get-while-helper chan)
|
|
|
|
(let ([read-mut (channel-read-mut chan)]
|
|
|
|
[read-box (channel-read-box chan)]
|
|
|
|
[val-cv (channel-val-cv chan)]
|
|
|
|
)
|
|
|
|
(if (unbox read-box)
|
|
|
|
(begin
|
|
|
|
(condition-wait val-cv read-mut)
|
|
|
|
(channel-get-while-helper chan)
|
|
|
|
)
|
|
|
|
(void)
|
|
|
|
)))
|
|
|
|
|
|
|
|
(define (blodwen-channel-get ty chan)
|
|
|
|
(mutex-acquire (channel-read-mut chan))
|
|
|
|
(channel-get-while-helper chan)
|
|
|
|
(let* ([val-box (channel-val-box chan)]
|
|
|
|
[read-box (channel-read-box chan)]
|
|
|
|
[read-cv (channel-read-cv chan)]
|
|
|
|
[the-val (unbox val-box)]
|
|
|
|
)
|
|
|
|
(set-box! val-box '())
|
|
|
|
(set-box! read-box #t)
|
|
|
|
(mutex-release (channel-read-mut chan))
|
|
|
|
(condition-signal read-cv)
|
|
|
|
the-val))
|
2021-02-05 19:16:20 +03:00
|
|
|
|
|
|
|
;; 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-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))
|
|
|
|
|
|
|
|
;; 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))))))
|
2021-07-25 15:49:21 +03:00
|
|
|
|
|
|
|
;; For creating and reading back scheme objects
|
|
|
|
|
|
|
|
; read a scheme string and evaluate it, returning 'Just result' on success
|
|
|
|
; TODO: catch exception!
|
|
|
|
(define (blodwen-eval-scheme str)
|
Experimental Scheme based evaluator (#1956)
This is for compiled evaluation at compile-time, for full normalisation. You can try it by setting the evaluation mode to scheme (that is, :set eval scheme at the REPL). It's certainly an order of magnitude faster than the standard evaluator, based on my playing around with it, although still quite a bit slower than compilation for various reasons, including:
* It has to evaluate under binders, and therefore deal with blocked symbols
* It has to maintain enough information to be able to read back a Term from the evaluated scheme object, which means retaining things like types and other metadata
* We can't do a lot of the optimisations we'd do for runtime evaluation particularly setting things up so we don't need to do arity checking
Also added a new option evaltiming (set with :set evaltiming) to display how long evaluation itself takes, which is handy for checking performance.
I also don't think we should aim to replace the standard evaluator, in general, at least not for a while, because that will involve rewriting a lot of things and working out how to make it work as Call By Name (which is clearly possible, but fiddly).
Still, it's going to be interesting to experiment with it! I think it will be a good idea to use it for elaborator reflection and type providers when we eventually get around to implementing them.
Original commit details:
* Add ability to evaluate open terms via Scheme
Still lots of polish and more formal testing to do here before we can
use it in practice, but you can still use ':scheme <term>' at the REPL
to evaluate an expression by compiling to scheme then reading back the
result.
Also added 'evaltiming' option at the REPL, which, when set, displays
how long normalisaton takes (doesn't count resugaring, just the
normalisation step).
* Add scheme evaluation mode
Different when evaluating everything, vs only evaluating visible things.
We want the latter when type checking, the former at the REPL.
* Bring support.rkt up to date
A couple of missing things required for interfacing with scheme objects
* More Scheme readback machinery
We need these things in the next version so that the next-but-one
version can have a scheme evaluator!
* Add top level interface to scheme based normaliser
Also check it's available - currently chez only - and revert to the
default slow normaliser if it's not.
* Bring Context up to date with changes in main
* Now need Idris 0.5.0 to build
* Add SNF type for scheme values
This will allow us to incrementally evaluate under lambdas, which will
be useful for elaborator reflection and type providers.
* Add Quote for scheme evaluator
So, we can now get a weak head normal form, and evaluate the scope of
a binder when we have an argument to plug in, or just quote back the
whole thing.
* Add new 'scheme' evaluator mode at the REPL
Replacing the temporary 'TmpScheme', this is a better way to try out the
scheme based evaluator
* Fix name generation for new UN format
* Add scheme evaluator support to Racket
* Add another scheme eval test
With metavariables this time
* evaltiming now times execution too
This was handy for finding out the difference between the scheme based
evaluator and compilation. Compilation was something like 20 times
faster in my little test, so that'd be about 4-500 times faster than the
standard evaluator. Ouch!
* Fix whitespace errors
* Error handling when trying to evaluate Scheme
2021-09-24 22:38:55 +03:00
|
|
|
(guard
|
|
|
|
(x [#t '()]) ; Nothing on failure
|
|
|
|
(box (eval (read (open-input-string str)))))
|
|
|
|
); box == Just
|
2021-07-25 15:49:21 +03:00
|
|
|
|
|
|
|
(define (blodwen-eval-okay obj)
|
|
|
|
(if (null? obj)
|
|
|
|
0
|
|
|
|
1))
|
|
|
|
|
|
|
|
(define (blodwen-get-eval-result obj)
|
|
|
|
(unbox obj))
|
|
|
|
|
|
|
|
(define (blodwen-debug-scheme obj)
|
|
|
|
(display obj) (newline))
|
|
|
|
|
|
|
|
(define (blodwen-is-number obj)
|
|
|
|
(if (number? obj) 1 0))
|
|
|
|
|
|
|
|
(define (blodwen-is-integer obj)
|
2021-08-05 17:06:55 +03:00
|
|
|
(if (and (number? obj) (exact? obj)) 1 0))
|
2021-07-25 15:49:21 +03:00
|
|
|
|
|
|
|
(define (blodwen-is-float obj)
|
|
|
|
(if (flonum? obj) 1 0))
|
|
|
|
|
|
|
|
(define (blodwen-is-char obj)
|
|
|
|
(if (char? obj) 1 0))
|
|
|
|
|
|
|
|
(define (blodwen-is-string obj)
|
|
|
|
(if (string? obj) 1 0))
|
|
|
|
|
|
|
|
(define (blodwen-is-procedure obj)
|
|
|
|
(if (procedure? obj) 1 0))
|
|
|
|
|
|
|
|
(define (blodwen-is-symbol obj)
|
|
|
|
(if (symbol? obj) 1 0))
|
|
|
|
|
|
|
|
(define (blodwen-is-vector obj)
|
|
|
|
(if (vector? obj) 1 0))
|
|
|
|
|
|
|
|
(define (blodwen-is-nil obj)
|
|
|
|
(if (null? obj) 1 0))
|
|
|
|
|
|
|
|
(define (blodwen-is-pair obj)
|
|
|
|
(if (pair? obj) 1 0))
|
|
|
|
|
2021-08-05 17:06:55 +03:00
|
|
|
(define (blodwen-is-box obj)
|
|
|
|
(if (box? obj) 1 0))
|
|
|
|
|
2021-07-25 15:49:21 +03:00
|
|
|
(define (blodwen-make-symbol str)
|
|
|
|
(string->symbol str))
|
|
|
|
|
|
|
|
; The below rely on checking that the objects are the right type first.
|
|
|
|
|
|
|
|
(define (blodwen-vector-ref obj i)
|
|
|
|
(vector-ref obj i))
|
|
|
|
|
|
|
|
(define (blodwen-vector-length obj)
|
|
|
|
(vector-length obj))
|
|
|
|
|
2021-08-05 17:06:55 +03:00
|
|
|
(define (blodwen-vector-list obj)
|
|
|
|
(vector->list obj))
|
|
|
|
|
|
|
|
(define (blodwen-unbox obj)
|
|
|
|
(unbox obj))
|
|
|
|
|
2021-07-25 15:49:21 +03:00
|
|
|
(define (blodwen-apply obj arg)
|
|
|
|
(obj arg))
|
|
|
|
|
2021-08-05 17:06:55 +03:00
|
|
|
(define (blodwen-force obj)
|
|
|
|
(obj))
|
|
|
|
|
2021-07-25 15:49:21 +03:00
|
|
|
(define (blodwen-read-symbol sym)
|
|
|
|
(symbol->string sym))
|
|
|
|
|
|
|
|
(define (blodwen-id x) x)
|