[ fix ] euclidian mod and div for scheme backends

This commit is contained in:
stefan-hoeck 2022-03-03 05:55:22 +01:00 committed by G. Allais
parent 405f80946b
commit 51eb854a60
6 changed files with 63 additions and 16 deletions

View File

@ -78,7 +78,7 @@ mul (Just $ Unsigned n) x y = op "bu*" [x, y, show n]
mul _ x y = op "*" [x, y]
div : Maybe IntKind -> String -> String -> String
div (Just $ Signed Unlimited) x y = op "quotient" [x, y]
div (Just $ Signed Unlimited) x y = op "blodwen-euclidDiv" [x, y]
div (Just $ Signed $ P n) x y = op "bs/" [x, y, show (n-1)]
div (Just $ Unsigned n) x y = op "bu/" [x, y, show n]
div _ x y = op "/" [x, y]
@ -136,7 +136,7 @@ schOp (Add ty) [x, y] = pure $ add (intKind ty) x y
schOp (Sub ty) [x, y] = pure $ sub (intKind ty) x y
schOp (Mul ty) [x, y] = pure $ mul (intKind ty) x y
schOp (Div ty) [x, y] = pure $ div (intKind ty) x y
schOp (Mod ty) [x, y] = pure $ op "remainder" [x, y]
schOp (Mod ty) [x, y] = pure $ op "blodwen-euclidMod" [x, y]
schOp (Neg ty) [x] = pure $ op "-" [x]
schOp (ShiftL ty) [x, y] = pure $ shl (intKind ty) x y
schOp (ShiftR ty) [x, y] = pure $ op "blodwen-shr" [x, y]

View File

@ -27,6 +27,18 @@
(define (blodwen-toUnsignedInt x bits)
(logand x (sub1 (ash 1 bits))))
(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)))
(define bu+ (lambda (x y bits) (blodwen-toUnsignedInt (+ x y) bits)))
(define bu- (lambda (x y bits) (blodwen-toUnsignedInt (- x y) bits)))
@ -36,7 +48,7 @@
(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)))
(define bs/ (lambda (x y bits) (blodwen-toSignedInt (quotient x y) bits)))
(define bs/ (lambda (x y bits) (blodwen-toSignedInt (blodwen-euclidDiv x y) bits)))
(define (integer->bits8 x) (logand x (sub1 (ash 1 8))))
(define (integer->bits16 x) (logand x (sub1 (ash 1 16))))

View File

@ -33,6 +33,19 @@
(define (blodwen-toUnsignedInt x bits)
(bitwise-and x (sub1 (arithmetic-shift 1 bits))))
(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)))
(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)))
@ -41,7 +54,7 @@
(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)))
(define bs/ (lambda (x y bits) (blodwen-toSignedInt (quotient x y) bits)))
(define bs/ (lambda (x y bits) (blodwen-toSignedInt (blodwen-euclidDiv x y) bits)))
; To match Chez
(define (add1 x) (+ x 1))

View File

@ -24,6 +24,19 @@
(define (blodwen-toUnsignedInt x bits)
(bitwise-and x (sub1 (arithmetic-shift 1 bits))))
(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)))
(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)))
@ -32,7 +45,7 @@
(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)))
(define bs/ (lambda (x y bits) (blodwen-toSignedInt (quotient x y) bits)))
(define bs/ (lambda (x y bits) (blodwen-toSignedInt (blodwen-euclidDiv x y) bits)))
; To match Chez
(define (fxadd1 x) (unsafe-fx+ x 1))

View File

@ -116,17 +116,17 @@ main = do
putStrLn "Multiply:"
put $ imapProperty Num (* 2) ints
-- putStrLn "Divide:"
-- put $ imapProperty Integral (`div` 3) ints
-- put $ imapProperty Integral (`div` -3) ints
--
-- putStrLn "Mod:"
-- put $ imapProperty Integral (`mod` 17) ints
-- put $ imapProperty Integral (`mod` -17) ints
--
-- putStrLn "Division Euclidean:"
-- put $ imapProperty Integral (\x => (x `div` 19) * 19 + (x `mod` 19)) ints
-- put $ imapProperty Integral (\x => (x `div` -19) * -19 + (x `mod` -19)) ints
putStrLn "Divide:"
put $ imapProperty Integral (`div` 3) ints
put $ imapProperty Integral (`div` -3) ints
putStrLn "Mod:"
put $ imapProperty Integral (`mod` 17) ints
put $ imapProperty Integral (`mod` -17) ints
putStrLn "Division Euclidean:"
put $ imapProperty Integral (\x => (x `div` 19) * 19 + (x `mod` 19)) ints
put $ imapProperty Integral (\x => (x `div` -19) * -19 + (x `mod` -19)) ints
putStrLn "Shift:"
put $ imapProperty SomeBits (`shiftL` one) ints

View File

@ -20,6 +20,15 @@ Negate:
[0, 206, 1, 0, 55536, 1, 0, 3294967296, 1, 0, 13446744073709551616, 1, -128, -50, -127, -32768, -10000, -32767, -2147483648, -500000000, -2147483647, -9223372036854775808, -1000000000000000000, -9223372036854775807, -9223372036854775808, -1000000000000000000, -9223372036854775807, 9223372036854775809, 0, -9223372036854775808]
Multiply:
[0, 100, 254, 0, 20000, 65534, 0, 2000000000, 4294967294, 0, 10000000000000000000, 18446744073709551614, 0, 100, -2, 0, 20000, -2, 0, 1000000000, -2, 0, 2000000000000000000, -2, 0, 2000000000000000000, -2, -18446744073709551618, 0, 18446744073709551616]
Divide:
[0, 16, 85, 0, 3333, 21845, 0, 333333333, 1431655765, 0, 1666666666666666666, 6148914691236517205, -43, 16, 42, -10923, 3333, 10922, -715827883, 166666666, 715827882, -3074457345618258603, 333333333333333333, 3074457345618258602, -3074457345618258603, 333333333333333333, 3074457345618258602, -3074457345618258603, 0, 3074457345618258602]
[0, 0, 1, 0, 0, 1, 0, 0, 1, 0, 0, 1, 43, -16, -42, 10923, -3333, -10922, 715827883, -166666666, -715827882, 3074457345618258603, -333333333333333333, -3074457345618258602, 3074457345618258603, -333333333333333333, -3074457345618258602, 3074457345618258603, 0, -3074457345618258602]
Mod:
[0, 16, 0, 0, 4, 0, 0, 7, 0, 0, 7, 0, 8, 16, 8, 8, 4, 8, 8, 12, 8, 8, 15, 8, 8, 15, 8, 7, 0, 9]
[0, 50, 16, 0, 10000, 16, 0, 1000000000, 16, 0, 5000000000000000000, 16, 8, 16, 8, 8, 4, 8, 8, 12, 8, 8, 15, 8, 8, 15, 8, 7, 0, 9]
Division Euclidean:
[0, 50, 255, 0, 10000, 65535, 0, 1000000000, 4294967295, 0, 5000000000000000000, 18446744073709551615, -128, 50, 127, -32768, 10000, 32767, -2147483648, 500000000, 2147483647, -9223372036854775808, 1000000000000000000, 9223372036854775807, -9223372036854775808, 1000000000000000000, 9223372036854775807, -9223372036854775809, 0, 9223372036854775808]
[0, 50, 255, 0, 10000, 65535, 0, 1000000000, 4294967295, 0, 5000000000000000000, 18446744073709551615, -128, 50, 127, -32768, 10000, 32767, -2147483648, 500000000, 2147483647, -9223372036854775808, 1000000000000000000, 9223372036854775807, -9223372036854775808, 1000000000000000000, 9223372036854775807, -9223372036854775809, 0, 9223372036854775808]
Shift:
[0, 100, 254, 0, 20000, 65534, 0, 2000000000, 4294967294, 0, 10000000000000000000, 18446744073709551614, 0, 100, -2, 0, 20000, -2, 0, 1000000000, -2, 0, 2000000000000000000, -2, 0, 2000000000000000000, -2, -18446744073709551618, 0, 18446744073709551616]
[0, 25, 127, 0, 5000, 32767, 0, 500000000, 2147483647, 0, 2500000000000000000, 9223372036854775807, -64, 25, 63, -16384, 5000, 16383, -1073741824, 250000000, 1073741823, -4611686018427387904, 500000000000000000, 4611686018427387903, -4611686018427387904, 500000000000000000, 4611686018427387903, -4611686018427387905, 0, 4611686018427387904]