mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 12:42:02 +03:00
Adding missing scheme support
Sorry!
This commit is contained in:
parent
80450bbc8e
commit
052713b645
224
support/chez/support.ss
Normal file
224
support/chez/support.ss
Normal file
@ -0,0 +1,224 @@
|
||||
(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)))))))
|
||||
(define b+ (lambda (x y bits) (remainder (+ x y) (expt 2 bits))))
|
||||
(define b- (lambda (x y bits) (remainder (- x y) (expt 2 bits))))
|
||||
(define b* (lambda (x y bits) (remainder (* x y) (expt 2 bits))))
|
||||
(define b/ (lambda (x y bits) (remainder (exact-floor (/ x y)) (expt 2 bits))))
|
||||
|
||||
(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))))
|
||||
(define cast-string-int
|
||||
(lambda (x)
|
||||
(floor (cast-num (string->number (destroy-prefix x))))))
|
||||
(define exact-floor
|
||||
(lambda (x)
|
||||
(inexact->exact (floor x))))
|
||||
(define cast-string-double
|
||||
(lambda (x)
|
||||
(cast-num (string->number (destroy-prefix x)))))
|
||||
(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))))
|
||||
|
||||
(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))
|
||||
|
||||
(define (blodwen-buffer-setint buf loc val)
|
||||
(bytevector-s32-set! buf loc val (native-endianness)))
|
||||
|
||||
(define (blodwen-buffer-getint buf loc)
|
||||
(bytevector-s32-ref buf loc (native-endianness)))
|
||||
|
||||
(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))
|
||||
|
||||
(define (blodwen-read-bytevec fname)
|
||||
(guard
|
||||
(x (#t #f))
|
||||
(let* [(h (open-file-input-port fname
|
||||
(file-options)
|
||||
(buffer-mode line) #f))
|
||||
(vec (get-bytevector-all h))]
|
||||
(begin (close-port h)
|
||||
vec))))
|
||||
|
||||
(define (blodwen-isbytevec v)
|
||||
(if (bytevector? v)
|
||||
0
|
||||
-1))
|
||||
|
||||
(define (blodwen-write-bytevec fname vec max)
|
||||
(guard
|
||||
(x (#t -1))
|
||||
(let* [(h (open-file-output-port fname (file-options no-fail)
|
||||
(buffer-mode line) #f))]
|
||||
(begin (put-bytevector h vec 0 max)
|
||||
(close-port h)
|
||||
0))))
|
||||
|
||||
|
||||
|
||||
;; Threads
|
||||
|
||||
(define blodwen-thread-data (make-thread-parameter #f))
|
||||
|
||||
(define (blodwen-thread p)
|
||||
(fork-thread (lambda () (p (vector 0)))))
|
||||
|
||||
(define (blodwen-get-thread-data ty)
|
||||
(blodwen-thread-data))
|
||||
|
||||
(define (blodwen-set-thread-data a)
|
||||
(blodwen-thread-data a))
|
||||
|
||||
(define (blodwen-mutex) (make-mutex))
|
||||
(define (blodwen-lock m) (mutex-acquire m))
|
||||
(define (blodwen-unlock m) (mutex-release m))
|
||||
(define (blodwen-thisthread) (get-thread-id))
|
||||
|
||||
(define (blodwen-condition) (make-condition))
|
||||
(define (blodwen-condition-wait c m) (condition-wait c m))
|
||||
(define (blodwen-condition-wait-timeout c m t) (condition-wait c m t))
|
||||
(define (blodwen-condition-signal c) (condition-signal c))
|
||||
(define (blodwen-condition-broadcast c) (condition-broadcast c))
|
||||
|
||||
(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))
|
||||
|
||||
(define (blodwen-args)
|
||||
(define (blodwen-build-args args)
|
||||
(if (null? args)
|
||||
(vector 0) ; Prelude.List
|
||||
(vector 1 (car args) (blodwen-build-args (cdr args)))))
|
||||
(blodwen-build-args (command-line)))
|
||||
|
||||
(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)))]))
|
69
support/gambit/foreign.scm
Normal file
69
support/gambit/foreign.scm
Normal file
@ -0,0 +1,69 @@
|
||||
;; Copyright 2009 Marc Feeley <feeley@iro.umontreal.ca>
|
||||
;; Licensed under Apache 2.0 or LGPL 2.1
|
||||
|
||||
(define-macro (define-c-struct type . fields)
|
||||
(let* ((type-str
|
||||
(symbol->string type))
|
||||
(struct-type-str
|
||||
(string-append "struct " type-str))
|
||||
(release-type-str
|
||||
(string-append "release_" type-str))
|
||||
(sym
|
||||
(lambda strs (string->symbol (apply string-append strs))))
|
||||
(type*
|
||||
(sym type-str "*"))
|
||||
(type*/nonnull
|
||||
(sym type-str "*/nonnull"))
|
||||
(type*/release-rc
|
||||
(sym type-str "*/release-rc"))
|
||||
(expansion
|
||||
`(begin
|
||||
|
||||
;; Release function which is called when the object
|
||||
;; is no longer accessible from the Scheme world
|
||||
|
||||
(c-declare
|
||||
,(string-append
|
||||
"static ___SCMOBJ " release-type-str "(void* ptr)\n"
|
||||
"{\n"
|
||||
" ___EXT(___release_rc)(ptr);\n"
|
||||
" return ___FIX(___NO_ERR);\n"
|
||||
"}\n"))
|
||||
|
||||
;; C types
|
||||
|
||||
(c-define-type ,type (struct ,type-str))
|
||||
(c-define-type ,type* (pointer ,type (,type*)))
|
||||
(c-define-type ,type*/nonnull (nonnull-pointer ,type (,type*)))
|
||||
(c-define-type ,type*/release-rc (nonnull-pointer ,type (,type*) ,release-type-str))
|
||||
|
||||
;; Type allocator procedure
|
||||
|
||||
(define ,(sym "alloc-" type-str)
|
||||
(c-lambda ()
|
||||
,type*/release-rc
|
||||
,(string-append "___result_voidstar = ___EXT(___alloc_rc)(sizeof(" struct-type-str "));")))
|
||||
|
||||
;; Field getters
|
||||
|
||||
,@(map (lambda (field)
|
||||
(let ((field-str (symbol->string (car field)))
|
||||
(field-type (cadr field)))
|
||||
`(define ,(sym type-str "-" field-str)
|
||||
(c-lambda (,type*/nonnull)
|
||||
,field-type
|
||||
,(string-append "___result = ___arg1->" field-str ";")))))
|
||||
fields)
|
||||
|
||||
;; Field setters
|
||||
|
||||
,@(map (lambda (field)
|
||||
(let ((field-str (symbol->string (car field)))
|
||||
(field-type (cadr field)))
|
||||
`(define ,(sym type-str "-" field-str "-set!")
|
||||
(c-lambda (,type*/nonnull ,field-type)
|
||||
void
|
||||
,(string-append "___arg1->" field-str " = ___arg2;")))))
|
||||
fields))))
|
||||
|
||||
expansion))
|
128
support/gambit/support.scm
Normal file
128
support/gambit/support.scm
Normal file
@ -0,0 +1,128 @@
|
||||
;; TODO Convert to macro
|
||||
(define (blodwen-read-args desc)
|
||||
(if (fx= (vector-ref desc 0) 0)
|
||||
'()
|
||||
(cons (vector-ref desc 2)
|
||||
(blodwen-read-args (vector-ref desc 3)))))
|
||||
|
||||
(define-macro (b+ x y bits)
|
||||
(if (exact-integer? bits)
|
||||
`(remainder (+ ,x ,y) ,(arithmetic-shift 1 bits))
|
||||
`(remainder (+ ,x ,y) (arithmetic-shift 1 ,bits))))
|
||||
(define-macro (b- x y bits)
|
||||
(if (exact-integer? bits)
|
||||
`(remainder (- ,x ,y) ,(arithmetic-shift 1 bits))
|
||||
`(remainder (- ,x ,y) (arithmetic-shift 1 ,bits))))
|
||||
(define-macro (b* x y bits)
|
||||
(if (exact-integer? bits)
|
||||
`(remainder (* ,x ,y) ,(arithmetic-shift 1 bits))
|
||||
`(remainder (* ,x ,y) (arithmetic-shift 1 ,bits))))
|
||||
(define-macro (b/ x y bits)
|
||||
(if (exact-integer? bits)
|
||||
`(remainder (floor (/ ,x ,y)) ,(arithmetic-shift 1 bits))
|
||||
`(remainder (floor (/ ,x ,y)) (arithmetic-shift 1 ,bits))))
|
||||
|
||||
(define-macro (blodwen-and . args) `(bitwise-and ,@args))
|
||||
(define-macro (blodwen-or . args) `(bitwise-ior ,@args))
|
||||
(define-macro (blodwen-xor . args) `(bitwise-xor ,@args))
|
||||
(define-macro (blodwen-shl x y) `(arithmetic-shift ,x ,y))
|
||||
(define-macro (blodwen-shr x y) `(arithmetic-shift ,x (- ,y)))
|
||||
|
||||
(define-macro (exact-floor x)
|
||||
(let ((s (gensym)))
|
||||
`(let ((,s ,x))
|
||||
(if (flonum? ,s) (##flonum->exact-int ,s) (##floor ,s)))))
|
||||
|
||||
;; TODO Convert to macro
|
||||
(define (cast-string-double x)
|
||||
(define (cast-num x)
|
||||
(if (number? x) x 0))
|
||||
(define (destroy-prefix x)
|
||||
(if (or (string=? x "") (char=? (string-ref x 0) #\#)) "" x))
|
||||
(cast-num (string->number (destroy-prefix x))))
|
||||
|
||||
(define-macro (cast-string-int x)
|
||||
`(floor (cast-string-double ,x)))
|
||||
|
||||
(define-macro (string-cons x y)
|
||||
`(string-append (string ,x) ,y))
|
||||
|
||||
(define-macro (string-reverse x)
|
||||
`(list->string (reverse! (string->list ,x))))
|
||||
|
||||
;; TODO Convert to macro
|
||||
(define (string-substr off len s)
|
||||
(let* ((start (fxmax 0 off))
|
||||
(end (fxmin (fx+ start (fxmax 0 len))
|
||||
(string-length s))))
|
||||
(substring s start end)))
|
||||
|
||||
(define-macro (get-tag x) `(vector-ref ,x 0))
|
||||
|
||||
;; These two are only used in this file
|
||||
(define-macro (either-left x) `(vector 0 ,x))
|
||||
(define-macro (either-right x) `(vector 1 ,x))
|
||||
|
||||
(define-macro (blodwen-error-quit msg)
|
||||
`(begin
|
||||
(display ,msg)
|
||||
(newline)
|
||||
(exit 1)))
|
||||
|
||||
(define (blodwen-get-line p)
|
||||
(if (input-port? p)
|
||||
(let ((str (read-line p)))
|
||||
(if (eof-object? str) "" str))
|
||||
""))
|
||||
|
||||
(define (blodwen-get-char p)
|
||||
(if (input-port? p)
|
||||
(let ((chr (read-char p)))
|
||||
(if (eof-object? chr) #\nul chr))
|
||||
#\nul))
|
||||
|
||||
;; Threads
|
||||
|
||||
(define (blodwen-thread p)
|
||||
(thread-start! (make-thread (lambda () (p '#(0))))))
|
||||
|
||||
(define (blodwen-get-thread-data ty)
|
||||
(let ((data (thread-specific (current-thread))))
|
||||
(if (eq? data #!void) #f data)))
|
||||
|
||||
(define (blodwen-set-thread-data a)
|
||||
(thread-specific-set! (current-thread) a))
|
||||
|
||||
(define blodwen-mutex make-mutex)
|
||||
(define blodwen-lock mutex-lock!)
|
||||
(define blodwen-unlock mutex-unlock!)
|
||||
(define blodwen-thisthread current-thread)
|
||||
|
||||
(define blodwen-condition make-condition-variable)
|
||||
(define blodwen-condition-signal condition-variable-signal!)
|
||||
(define blodwen-condition-broadcast condition-variable-broadcast!)
|
||||
(define (blodwen-condition-wait c m)
|
||||
(mutex-unlock! m c)
|
||||
(mutex-lock! m))
|
||||
(define (blodwen-condition-wait-timeout c m t) ; XXX
|
||||
(mutex-unlock! m c t)
|
||||
(mutex-lock! m))
|
||||
|
||||
(define blodwen-sleep thread-sleep!)
|
||||
(define (blodwen-usleep s) (thread-sleep! (/ s 1e6)))
|
||||
|
||||
(define (blodwen-time)
|
||||
(exact-floor (time->seconds (current-time))))
|
||||
|
||||
(define (blodwen-args)
|
||||
(define (blodwen-build-args args)
|
||||
(if (null? args)
|
||||
(vector 0) ; Prelude.List
|
||||
(vector 1 (car args) (blodwen-build-args (cdr args)))))
|
||||
(blodwen-build-args (cdr (command-line))))
|
||||
|
||||
(define (blodwen-hasenv var)
|
||||
(if (getenv var #f) 1 0))
|
||||
|
||||
(define (blodwen-system cmd)
|
||||
(fxarithmetic-shift-right (shell-command cmd) 8))
|
210
support/racket/support.rkt
Normal file
210
support/racket/support.rkt
Normal file
@ -0,0 +1,210 @@
|
||||
(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)))))))
|
||||
(define b+ (lambda (x y bits) (remainder (+ x y) (expt 2 bits))))
|
||||
(define b- (lambda (x y bits) (remainder (- x y) (expt 2 bits))))
|
||||
(define b* (lambda (x y bits) (remainder (* x y) (expt 2 bits))))
|
||||
(define b/ (lambda (x y bits) (remainder (exact-floor (/ x y)) (expt 2 bits))))
|
||||
|
||||
(define blodwen-shl (lambda (x y) (arithmetic-shift x y)))
|
||||
(define blodwen-shr (lambda (x y) (arithmetic-shift x (- y))))
|
||||
(define blodwen-and (lambda (x y) (bitwise-and x y)))
|
||||
(define blodwen-or (lambda (x y) (bitwise-ior x y)))
|
||||
(define blodwen-xor (lambda (x y) (bitwise-xor 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))))
|
||||
(define cast-string-int
|
||||
(lambda (x)
|
||||
(floor (cast-num (string->number (destroy-prefix x))))))
|
||||
(define cast-string-double
|
||||
(lambda (x)
|
||||
(cast-num (string->number (destroy-prefix x)))))
|
||||
(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))))
|
||||
(substring s b end)))
|
||||
|
||||
(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 (read-line p)))
|
||||
(if (eof-object? str)
|
||||
""
|
||||
str))
|
||||
void))
|
||||
|
||||
(define (blodwen-get-char p)
|
||||
(if (port? p)
|
||||
(let ((chr (read-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))
|
||||
|
||||
(define (blodwen-buffer-setint buf loc val)
|
||||
(bytevector-s32-set! buf loc val (native-endianness)))
|
||||
|
||||
(define (blodwen-buffer-getint buf loc)
|
||||
(bytevector-s32-ref buf loc (native-endianness)))
|
||||
|
||||
(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))
|
||||
|
||||
(define (blodwen-read-bytevec fname)
|
||||
(with-handlers
|
||||
([(lambda (x) #t) (lambda (exn) (make-bytevector 0))])
|
||||
(let* [(h (open-file-input-port fname
|
||||
(file-options)
|
||||
(buffer-mode line) #f))
|
||||
(vec (get-bytevector-all h))]
|
||||
(begin (close-port h)
|
||||
vec))))
|
||||
|
||||
(define (blodwen-isbytevec v)
|
||||
(if (bytevector? v)
|
||||
0
|
||||
-1))
|
||||
|
||||
(define (blodwen-write-bytevec fname vec max)
|
||||
(with-handlers
|
||||
([(lambda (x) #t) (lambda (exn) -1)])
|
||||
(let* [(h (open-file-output-port fname (file-options no-fail)
|
||||
(buffer-mode line) #f))]
|
||||
(begin (put-bytevector h vec 0 max)
|
||||
(close-port h)
|
||||
0))))
|
||||
|
||||
|
||||
;; Threads
|
||||
|
||||
(define blodwen-thread-data (make-thread-cell #f))
|
||||
|
||||
(define (blodwen-thread p)
|
||||
(thread (lambda () (p (vector 0)))))
|
||||
|
||||
(define (blodwen-get-thread-data ty)
|
||||
(thread-cell-ref blodwen-thread-data))
|
||||
|
||||
(define (blodwen-set-thread-data a)
|
||||
(thread-cell-set! blodwen-thread-data a))
|
||||
|
||||
(define (blodwen-mutex) (make-semaphore 1))
|
||||
(define (blodwen-lock m) (semaphore-post m))
|
||||
(define (blodwen-unlock m) (semaphore-wait m))
|
||||
(define (blodwen-thisthread) (current-thread))
|
||||
|
||||
(define (blodwen-condition) (make-channel))
|
||||
(define (blodwen-condition-wait c m)
|
||||
(blodwen-unlock m) ;; consistency with interface for posix condition variables
|
||||
(sync c)
|
||||
(blodwen-lock m))
|
||||
(define (blodwen-condition-wait-timeout c m t)
|
||||
(blodwen-unlock m) ;; consistency with interface for posix condition variables
|
||||
(sync/timeout t c)
|
||||
(blodwen-lock m))
|
||||
(define (blodwen-condition-signal c)
|
||||
(channel-put c 'ready))
|
||||
|
||||
(define (blodwen-sleep s) (sleep s))
|
||||
|
||||
(define (blodwen-time) (current-seconds))
|
||||
|
||||
(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) 0) ;; unsupported
|
||||
(define (blodwen-clock-time-gcreal) 0) ;; unsupported
|
||||
(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))
|
||||
|
||||
(define (blodwen-args)
|
||||
(define (blodwen-build-args args)
|
||||
(if (null? args)
|
||||
(vector 0) ; Prelude.List
|
||||
(vector 1 (car args) (blodwen-build-args (cdr args)))))
|
||||
(blodwen-build-args
|
||||
(cons (path->string (find-system-path 'run-file))
|
||||
(vector->list (current-command-line-arguments)))))
|
||||
|
||||
(define (blodwen-system cmd)
|
||||
(if (system cmd)
|
||||
0
|
||||
1))
|
||||
|
||||
;; Randoms
|
||||
(random-seed (date*-nanosecond (current-date))) ; initialize random seed
|
||||
|
||||
(define (blodwen-random-seed s) (random-seed s))
|
||||
(define blodwen-random
|
||||
(case-lambda
|
||||
;; no argument, pick a real value from [0, 1.0)
|
||||
[() (random)]
|
||||
;; single argument k, pick an integral value from [0, k)
|
||||
[(k) (if (> k 0)
|
||||
(random k)
|
||||
(raise 'blodwen-random-invalid-range-argument))]))
|
Loading…
Reference in New Issue
Block a user