mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-23 22:22:07 +03:00
[ new ] Missing integer type interfaces (#1629)
Co-authored-by: Guillaume ALLAIS <guillaume.allais@ens-lyon.org>
This commit is contained in:
parent
76715acc1f
commit
6ed266d306
@ -84,8 +84,8 @@ Bits Bits8 where
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits8 x . fromInteger . cast . fst
|
||||
shiftL x = prim__shl_Bits8 x . fromInteger . cast . fst
|
||||
shiftR x = prim__shr_Bits8 x . cast . fst
|
||||
shiftL x = prim__shl_Bits8 x . cast . fst
|
||||
complement = xor 0xff
|
||||
oneBits = 0xff
|
||||
|
||||
@ -98,8 +98,8 @@ Bits Bits16 where
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits16 x . fromInteger . cast . fst
|
||||
shiftL x = prim__shl_Bits16 x . fromInteger . cast . fst
|
||||
shiftR x = prim__shr_Bits16 x . cast . fst
|
||||
shiftL x = prim__shl_Bits16 x . cast . fst
|
||||
complement = xor 0xffff
|
||||
oneBits = 0xffff
|
||||
|
||||
@ -112,11 +112,25 @@ Bits Bits32 where
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits32 x . fromInteger . cast . fst
|
||||
shiftL x = prim__shl_Bits32 x . fromInteger . cast . fst
|
||||
shiftR x = prim__shr_Bits32 x . cast . fst
|
||||
shiftL x = prim__shl_Bits32 x . cast . fst
|
||||
complement = xor 0xffffffff
|
||||
oneBits = 0xffffffff
|
||||
|
||||
public export %inline
|
||||
Bits Bits64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
(.&.) = prim__and_Bits64
|
||||
(.|.) = prim__or_Bits64
|
||||
xor = prim__xor_Bits64
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits64 x . cast . fst
|
||||
shiftL x = prim__shl_Bits64 x . cast . fst
|
||||
complement = xor 0xffffffffffffffff
|
||||
oneBits = 0xffffffffffffffff
|
||||
|
||||
public export %inline
|
||||
Bits Int where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
@ -131,6 +145,76 @@ Bits Int where
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Int8 where
|
||||
Index = Subset Nat (`LT` 8)
|
||||
(.&.) = prim__and_Int8
|
||||
(.|.) = prim__or_Int8
|
||||
xor = prim__xor_Int8
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int8 x . cast . fst
|
||||
shiftL x = prim__shl_Int8 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Int16 where
|
||||
Index = Subset Nat (`LT` 16)
|
||||
(.&.) = prim__and_Int16
|
||||
(.|.) = prim__or_Int16
|
||||
xor = prim__xor_Int16
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int16 x . cast . fst
|
||||
shiftL x = prim__shl_Int16 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Int32 where
|
||||
Index = Subset Nat (`LT` 32)
|
||||
(.&.) = prim__and_Int32
|
||||
(.|.) = prim__or_Int32
|
||||
xor = prim__xor_Int32
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int32 x . cast . fst
|
||||
shiftL x = prim__shl_Int32 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Int64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
(.&.) = prim__and_Int64
|
||||
(.|.) = prim__or_Int64
|
||||
xor = prim__xor_Int64
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int64 x . cast . fst
|
||||
shiftL x = prim__shl_Int64 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Integer where
|
||||
Index = Nat
|
||||
(.&.) = prim__and_Integer
|
||||
(.|.) = prim__or_Integer
|
||||
xor = prim__xor_Integer
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Integer x . natToInteger
|
||||
shiftL x = prim__shl_Integer x . natToInteger
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- FiniteBits
|
||||
--------------------------------------------------------------------------------
|
||||
@ -157,7 +241,7 @@ FiniteBits Bits8 where
|
||||
let x1 = (x0 .&. 0x55) + ((x0 `shiftR` fromNat 1) .&. 0x55)
|
||||
x2 = (x1 .&. 0x33) + ((x1 `shiftR` fromNat 2) .&. 0x33)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F)
|
||||
in fromInteger $ cast x3
|
||||
in cast x3
|
||||
|
||||
public export %inline
|
||||
FiniteBits Bits16 where
|
||||
@ -170,7 +254,7 @@ FiniteBits Bits16 where
|
||||
x2 = (x1 .&. 0x3333) + ((x1 `shiftR` fromNat 2) .&. 0x3333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F)
|
||||
x4 = (x3 * 0x0101) `shiftR` fromNat 8
|
||||
in fromInteger $ cast x4
|
||||
in cast x4
|
||||
|
||||
public export %inline
|
||||
FiniteBits Bits32 where
|
||||
@ -183,7 +267,22 @@ FiniteBits Bits32 where
|
||||
x2 = (x1 .&. 0x33333333) + ((x1 `shiftR` fromNat 2) .&. 0x33333333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F0F0F)
|
||||
x4 = (x3 * 0x01010101) `shiftR` fromNat 24
|
||||
in fromInteger $ cast x4
|
||||
in cast x4
|
||||
|
||||
public export %inline
|
||||
FiniteBits Bits64 where
|
||||
bitSize = 64
|
||||
bitsToIndex = id
|
||||
|
||||
popCount x0 =
|
||||
-- see https://stackoverflow.com/questions/109023/how-to-count-the-number-of-set-bits-in-a-64-bit-integer
|
||||
let x1 = (x0 .&. 0x5555555555555555) +
|
||||
((x0 `shiftR` fromNat 1) .&. 0x5555555555555555)
|
||||
x2 = (x1 .&. 0x3333333333333333)
|
||||
+ ((x1 `shiftR` fromNat 2) .&. 0x3333333333333333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` fromNat 56
|
||||
in cast x4
|
||||
|
||||
public export %inline
|
||||
FiniteBits Int where
|
||||
@ -204,4 +303,81 @@ FiniteBits Int where
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F0F0F0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` fromNat 56
|
||||
x5 = if x < 0 then x4 + 1 else x4
|
||||
in fromInteger $ cast x5
|
||||
in cast x5
|
||||
|
||||
public export %inline
|
||||
FiniteBits Int8 where
|
||||
bitSize = 8
|
||||
bitsToIndex = id
|
||||
|
||||
popCount x =
|
||||
-- see https://stackoverflow.com/questions/109023/how-to-count-the-number-of-set-bits-in-a-32-bit-integer
|
||||
-- We have to treat negative numbers separately in order to
|
||||
-- prevent overflows in the first addition.
|
||||
-- The top bit is therefore cleared and 1 is added in the end
|
||||
-- in case of a negative number
|
||||
let x0 = x `clearBit` fromNat 7
|
||||
x1 = (x0 .&. 0x55) + ((x0 `shiftR` fromNat 1) .&. 0x55)
|
||||
x2 = (x1 .&. 0x33) + ((x1 `shiftR` fromNat 2) .&. 0x33)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F)
|
||||
x4 = if x < 0 then x3 + 1 else x3
|
||||
in cast x4
|
||||
|
||||
public export %inline
|
||||
FiniteBits Int16 where
|
||||
bitSize = 16
|
||||
bitsToIndex = id
|
||||
|
||||
popCount x =
|
||||
-- see https://stackoverflow.com/questions/109023/how-to-count-the-number-of-set-bits-in-a-32-bit-integer
|
||||
-- We have to treat negative numbers separately in order to
|
||||
-- prevent overflows in the first addition.
|
||||
-- The top bit is therefore cleared and 1 is added in the end
|
||||
-- in case of a negative number
|
||||
let x0 = x `clearBit` fromNat 15
|
||||
x1 = (x0 .&. 0x5555) + ((x0 `shiftR` fromNat 1) .&. 0x5555)
|
||||
x2 = (x1 .&. 0x3333) + ((x1 `shiftR` fromNat 2) .&. 0x3333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F)
|
||||
x4 = (x3 * 0x0101) `shiftR` fromNat 8
|
||||
x5 = if x < 0 then x4 + 1 else x4
|
||||
in cast x5
|
||||
|
||||
public export %inline
|
||||
FiniteBits Int32 where
|
||||
bitSize = 32
|
||||
bitsToIndex = id
|
||||
|
||||
popCount x =
|
||||
-- see https://stackoverflow.com/questions/109023/how-to-count-the-number-of-set-bits-in-a-32-bit-integer
|
||||
-- We have to treat negative numbers separately in order to
|
||||
-- prevent overflows in the first addition.
|
||||
-- The top bit is therefore cleared and 1 is added in the end
|
||||
-- in case of a negative number
|
||||
let x0 = x `clearBit` fromNat 31
|
||||
x1 = (x0 .&. 0x55555555) + ((x0 `shiftR` fromNat 1) .&. 0x55555555)
|
||||
x2 = (x1 .&. 0x33333333) + ((x1 `shiftR` fromNat 2) .&. 0x33333333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F0F0F)
|
||||
x4 = (x3 * 0x01010101) `shiftR` fromNat 24
|
||||
x5 = if x < 0 then x4 + 1 else x4
|
||||
in cast x5
|
||||
|
||||
public export %inline
|
||||
FiniteBits Int64 where
|
||||
bitSize = 64
|
||||
bitsToIndex = id
|
||||
|
||||
popCount x =
|
||||
-- see https://stackoverflow.com/questions/109023/how-to-count-the-number-of-set-bits-in-a-32-bit-integer
|
||||
-- We have to treat negative numbers separately in order to
|
||||
-- prevent overflows in the first addition.
|
||||
-- The top bit is therefore cleared and 1 is added in the end
|
||||
-- in case of a negative number
|
||||
let x0 = x `clearBit` fromNat 63
|
||||
x1 = (x0 .&. 0x5555555555555555)
|
||||
+ ((x0 `shiftR` fromNat 1) .&. 0x5555555555555555)
|
||||
x2 = (x1 .&. 0x3333333333333333)
|
||||
+ ((x1 `shiftR` fromNat 2) .&. 0x3333333333333333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F0F0F0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` fromNat 56
|
||||
x5 = if x < 0 then x4 + 1 else x4
|
||||
in cast x5
|
||||
|
@ -43,6 +43,42 @@ export
|
||||
Cast Double String where
|
||||
cast = prim__cast_DoubleString
|
||||
|
||||
export
|
||||
Cast Nat String where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Int8 String where
|
||||
cast = prim__cast_Int8String
|
||||
|
||||
export
|
||||
Cast Int16 String where
|
||||
cast = prim__cast_Int16String
|
||||
|
||||
export
|
||||
Cast Int32 String where
|
||||
cast = prim__cast_Int32String
|
||||
|
||||
export
|
||||
Cast Int64 String where
|
||||
cast = prim__cast_Int64String
|
||||
|
||||
export
|
||||
Cast Bits8 String where
|
||||
cast = prim__cast_Bits8String
|
||||
|
||||
export
|
||||
Cast Bits16 String where
|
||||
cast = prim__cast_Bits16String
|
||||
|
||||
export
|
||||
Cast Bits32 String where
|
||||
cast = prim__cast_Bits32String
|
||||
|
||||
export
|
||||
Cast Bits64 String where
|
||||
cast = prim__cast_Bits64String
|
||||
|
||||
-- To Integer
|
||||
|
||||
export
|
||||
@ -81,6 +117,22 @@ export
|
||||
Cast Bits64 Integer where
|
||||
cast = prim__cast_Bits64Integer
|
||||
|
||||
export
|
||||
Cast Int8 Integer where
|
||||
cast = prim__cast_Int8Integer
|
||||
|
||||
export
|
||||
Cast Int16 Integer where
|
||||
cast = prim__cast_Int16Integer
|
||||
|
||||
export
|
||||
Cast Int32 Integer where
|
||||
cast = prim__cast_Int32Integer
|
||||
|
||||
export
|
||||
Cast Int64 Integer where
|
||||
cast = prim__cast_Int64Integer
|
||||
|
||||
-- To Int
|
||||
|
||||
export
|
||||
@ -119,12 +171,68 @@ export
|
||||
Cast Bits64 Int where
|
||||
cast = prim__cast_Bits64Int
|
||||
|
||||
export
|
||||
Cast Int8 Int where
|
||||
cast = prim__cast_Int8Int
|
||||
|
||||
export
|
||||
Cast Int16 Int where
|
||||
cast = prim__cast_Int16Int
|
||||
|
||||
export
|
||||
Cast Int32 Int where
|
||||
cast = prim__cast_Int32Int
|
||||
|
||||
export
|
||||
Cast Int64 Int where
|
||||
cast = prim__cast_Int64Int
|
||||
|
||||
-- To Char
|
||||
|
||||
export
|
||||
Cast Int Char where
|
||||
cast = prim__cast_IntChar
|
||||
|
||||
export
|
||||
Cast Integer Char where
|
||||
cast = prim__cast_IntegerChar
|
||||
|
||||
export
|
||||
Cast Nat Char where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Char where
|
||||
cast = prim__cast_Bits8Char
|
||||
|
||||
export
|
||||
Cast Bits16 Char where
|
||||
cast = prim__cast_Bits16Char
|
||||
|
||||
export
|
||||
Cast Bits32 Char where
|
||||
cast = prim__cast_Bits32Char
|
||||
|
||||
export
|
||||
Cast Bits64 Char where
|
||||
cast = prim__cast_Bits64Char
|
||||
|
||||
export
|
||||
Cast Int8 Char where
|
||||
cast = prim__cast_Int8Char
|
||||
|
||||
export
|
||||
Cast Int16 Char where
|
||||
cast = prim__cast_Int16Char
|
||||
|
||||
export
|
||||
Cast Int32 Char where
|
||||
cast = prim__cast_Int32Char
|
||||
|
||||
export
|
||||
Cast Int64 Char where
|
||||
cast = prim__cast_Int64Char
|
||||
|
||||
-- To Double
|
||||
|
||||
export
|
||||
@ -143,6 +251,38 @@ export
|
||||
Cast Nat Double where
|
||||
cast = prim__cast_IntegerDouble . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Double where
|
||||
cast = prim__cast_Bits8Double
|
||||
|
||||
export
|
||||
Cast Bits16 Double where
|
||||
cast = prim__cast_Bits16Double
|
||||
|
||||
export
|
||||
Cast Bits32 Double where
|
||||
cast = prim__cast_Bits32Double
|
||||
|
||||
export
|
||||
Cast Bits64 Double where
|
||||
cast = prim__cast_Bits64Double
|
||||
|
||||
export
|
||||
Cast Int8 Double where
|
||||
cast = prim__cast_Int8Double
|
||||
|
||||
export
|
||||
Cast Int16 Double where
|
||||
cast = prim__cast_Int16Double
|
||||
|
||||
export
|
||||
Cast Int32 Double where
|
||||
cast = prim__cast_Int32Double
|
||||
|
||||
export
|
||||
Cast Int64 Double where
|
||||
cast = prim__cast_Int64Double
|
||||
|
||||
|
||||
-- To Bits8
|
||||
|
||||
@ -166,6 +306,38 @@ export
|
||||
Cast Bits64 Bits8 where
|
||||
cast = prim__cast_Bits64Bits8
|
||||
|
||||
export
|
||||
Cast String Bits8 where
|
||||
cast = prim__cast_StringBits8
|
||||
|
||||
export
|
||||
Cast Double Bits8 where
|
||||
cast = prim__cast_DoubleBits8
|
||||
|
||||
export
|
||||
Cast Char Bits8 where
|
||||
cast = prim__cast_CharBits8
|
||||
|
||||
export
|
||||
Cast Nat Bits8 where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Int8 Bits8 where
|
||||
cast = prim__cast_Int8Bits8
|
||||
|
||||
export
|
||||
Cast Int16 Bits8 where
|
||||
cast = prim__cast_Int16Bits8
|
||||
|
||||
export
|
||||
Cast Int32 Bits8 where
|
||||
cast = prim__cast_Int32Bits8
|
||||
|
||||
export
|
||||
Cast Int64 Bits8 where
|
||||
cast = prim__cast_Int64Bits8
|
||||
|
||||
|
||||
-- To Bits16
|
||||
|
||||
@ -189,6 +361,38 @@ export
|
||||
Cast Bits64 Bits16 where
|
||||
cast = prim__cast_Bits64Bits16
|
||||
|
||||
export
|
||||
Cast String Bits16 where
|
||||
cast = prim__cast_StringBits16
|
||||
|
||||
export
|
||||
Cast Double Bits16 where
|
||||
cast = prim__cast_DoubleBits16
|
||||
|
||||
export
|
||||
Cast Char Bits16 where
|
||||
cast = prim__cast_CharBits16
|
||||
|
||||
export
|
||||
Cast Nat Bits16 where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Int8 Bits16 where
|
||||
cast = prim__cast_Int8Bits16
|
||||
|
||||
export
|
||||
Cast Int16 Bits16 where
|
||||
cast = prim__cast_Int16Bits16
|
||||
|
||||
export
|
||||
Cast Int32 Bits16 where
|
||||
cast = prim__cast_Int32Bits16
|
||||
|
||||
export
|
||||
Cast Int64 Bits16 where
|
||||
cast = prim__cast_Int64Bits16
|
||||
|
||||
|
||||
-- To Bits32
|
||||
|
||||
@ -212,6 +416,38 @@ export
|
||||
Cast Bits64 Bits32 where
|
||||
cast = prim__cast_Bits64Bits32
|
||||
|
||||
export
|
||||
Cast String Bits32 where
|
||||
cast = prim__cast_StringBits32
|
||||
|
||||
export
|
||||
Cast Double Bits32 where
|
||||
cast = prim__cast_DoubleBits32
|
||||
|
||||
export
|
||||
Cast Char Bits32 where
|
||||
cast = prim__cast_CharBits32
|
||||
|
||||
export
|
||||
Cast Nat Bits32 where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Int8 Bits32 where
|
||||
cast = prim__cast_Int8Bits32
|
||||
|
||||
export
|
||||
Cast Int16 Bits32 where
|
||||
cast = prim__cast_Int16Bits32
|
||||
|
||||
export
|
||||
Cast Int32 Bits32 where
|
||||
cast = prim__cast_Int32Bits32
|
||||
|
||||
export
|
||||
Cast Int64 Bits32 where
|
||||
cast = prim__cast_Int64Bits32
|
||||
|
||||
-- To Bits64
|
||||
|
||||
export
|
||||
@ -233,3 +469,305 @@ Cast Bits16 Bits64 where
|
||||
export
|
||||
Cast Bits32 Bits64 where
|
||||
cast = prim__cast_Bits32Bits64
|
||||
|
||||
export
|
||||
Cast String Bits64 where
|
||||
cast = prim__cast_StringBits64
|
||||
|
||||
export
|
||||
Cast Double Bits64 where
|
||||
cast = prim__cast_DoubleBits64
|
||||
|
||||
export
|
||||
Cast Char Bits64 where
|
||||
cast = prim__cast_CharBits64
|
||||
|
||||
export
|
||||
Cast Nat Bits64 where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Int8 Bits64 where
|
||||
cast = prim__cast_Int8Bits64
|
||||
|
||||
export
|
||||
Cast Int16 Bits64 where
|
||||
cast = prim__cast_Int16Bits64
|
||||
|
||||
export
|
||||
Cast Int32 Bits64 where
|
||||
cast = prim__cast_Int32Bits64
|
||||
|
||||
export
|
||||
Cast Int64 Bits64 where
|
||||
cast = prim__cast_Int64Bits64
|
||||
|
||||
-- To Int8
|
||||
|
||||
export
|
||||
Cast String Int8 where
|
||||
cast = prim__cast_StringInt8
|
||||
|
||||
export
|
||||
Cast Double Int8 where
|
||||
cast = prim__cast_DoubleInt8
|
||||
|
||||
export
|
||||
Cast Char Int8 where
|
||||
cast = prim__cast_CharInt8
|
||||
|
||||
export
|
||||
Cast Int Int8 where
|
||||
cast = prim__cast_IntInt8
|
||||
|
||||
export
|
||||
Cast Integer Int8 where
|
||||
cast = prim__cast_IntegerInt8
|
||||
|
||||
export
|
||||
Cast Nat Int8 where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int8 where
|
||||
cast = prim__cast_Bits8Int8
|
||||
|
||||
export
|
||||
Cast Bits16 Int8 where
|
||||
cast = prim__cast_Bits16Int8
|
||||
|
||||
export
|
||||
Cast Bits32 Int8 where
|
||||
cast = prim__cast_Bits32Int8
|
||||
|
||||
export
|
||||
Cast Bits64 Int8 where
|
||||
cast = prim__cast_Bits64Int8
|
||||
|
||||
export
|
||||
Cast Int16 Int8 where
|
||||
cast = prim__cast_Int16Int8
|
||||
|
||||
export
|
||||
Cast Int32 Int8 where
|
||||
cast = prim__cast_Int32Int8
|
||||
|
||||
export
|
||||
Cast Int64 Int8 where
|
||||
cast = prim__cast_Int64Int8
|
||||
|
||||
-- To Int16
|
||||
|
||||
export
|
||||
Cast String Int16 where
|
||||
cast = prim__cast_StringInt16
|
||||
|
||||
export
|
||||
Cast Double Int16 where
|
||||
cast = prim__cast_DoubleInt16
|
||||
|
||||
export
|
||||
Cast Char Int16 where
|
||||
cast = prim__cast_CharInt16
|
||||
|
||||
export
|
||||
Cast Int Int16 where
|
||||
cast = prim__cast_IntInt16
|
||||
|
||||
export
|
||||
Cast Integer Int16 where
|
||||
cast = prim__cast_IntegerInt16
|
||||
|
||||
export
|
||||
Cast Nat Int16 where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int16 where
|
||||
cast = prim__cast_Bits8Int16
|
||||
|
||||
export
|
||||
Cast Bits16 Int16 where
|
||||
cast = prim__cast_Bits16Int16
|
||||
|
||||
export
|
||||
Cast Bits32 Int16 where
|
||||
cast = prim__cast_Bits32Int16
|
||||
|
||||
export
|
||||
Cast Bits64 Int16 where
|
||||
cast = prim__cast_Bits64Int16
|
||||
|
||||
export
|
||||
Cast Int8 Int16 where
|
||||
cast = prim__cast_Int8Int16
|
||||
|
||||
export
|
||||
Cast Int32 Int16 where
|
||||
cast = prim__cast_Int32Int16
|
||||
|
||||
export
|
||||
Cast Int64 Int16 where
|
||||
cast = prim__cast_Int64Int16
|
||||
|
||||
-- To Int32
|
||||
|
||||
export
|
||||
Cast String Int32 where
|
||||
cast = prim__cast_StringInt32
|
||||
|
||||
export
|
||||
Cast Double Int32 where
|
||||
cast = prim__cast_DoubleInt32
|
||||
|
||||
export
|
||||
Cast Char Int32 where
|
||||
cast = prim__cast_CharInt32
|
||||
|
||||
export
|
||||
Cast Int Int32 where
|
||||
cast = prim__cast_IntInt32
|
||||
|
||||
export
|
||||
Cast Integer Int32 where
|
||||
cast = prim__cast_IntegerInt32
|
||||
|
||||
export
|
||||
Cast Nat Int32 where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int32 where
|
||||
cast = prim__cast_Bits8Int32
|
||||
|
||||
export
|
||||
Cast Bits16 Int32 where
|
||||
cast = prim__cast_Bits16Int32
|
||||
|
||||
export
|
||||
Cast Bits32 Int32 where
|
||||
cast = prim__cast_Bits32Int32
|
||||
|
||||
export
|
||||
Cast Bits64 Int32 where
|
||||
cast = prim__cast_Bits64Int32
|
||||
|
||||
export
|
||||
Cast Int8 Int32 where
|
||||
cast = prim__cast_Int8Int32
|
||||
|
||||
export
|
||||
Cast Int16 Int32 where
|
||||
cast = prim__cast_Int16Int32
|
||||
|
||||
export
|
||||
Cast Int64 Int32 where
|
||||
cast = prim__cast_Int64Int32
|
||||
|
||||
-- To Int64
|
||||
|
||||
export
|
||||
Cast String Int64 where
|
||||
cast = prim__cast_StringInt64
|
||||
|
||||
export
|
||||
Cast Double Int64 where
|
||||
cast = prim__cast_DoubleInt64
|
||||
|
||||
export
|
||||
Cast Char Int64 where
|
||||
cast = prim__cast_CharInt64
|
||||
|
||||
export
|
||||
Cast Int Int64 where
|
||||
cast = prim__cast_IntInt64
|
||||
|
||||
export
|
||||
Cast Integer Int64 where
|
||||
cast = prim__cast_IntegerInt64
|
||||
|
||||
export
|
||||
Cast Nat Int64 where
|
||||
cast = cast . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int64 where
|
||||
cast = prim__cast_Bits8Int64
|
||||
|
||||
export
|
||||
Cast Bits16 Int64 where
|
||||
cast = prim__cast_Bits16Int64
|
||||
|
||||
export
|
||||
Cast Bits32 Int64 where
|
||||
cast = prim__cast_Bits32Int64
|
||||
|
||||
export
|
||||
Cast Bits64 Int64 where
|
||||
cast = prim__cast_Bits64Int64
|
||||
|
||||
export
|
||||
Cast Int8 Int64 where
|
||||
cast = prim__cast_Int8Int64
|
||||
|
||||
export
|
||||
Cast Int16 Int64 where
|
||||
cast = prim__cast_Int16Int64
|
||||
|
||||
export
|
||||
Cast Int32 Int64 where
|
||||
cast = prim__cast_Int32Int64
|
||||
|
||||
-- To Nat
|
||||
|
||||
export
|
||||
Cast String Nat where
|
||||
cast = integerToNat . cast
|
||||
|
||||
export
|
||||
Cast Double Nat where
|
||||
cast = integerToNat . cast
|
||||
|
||||
export
|
||||
Cast Char Nat where
|
||||
cast = believe_me . cast {to = Integer}
|
||||
|
||||
export
|
||||
Cast Int Nat where
|
||||
cast = integerToNat . cast
|
||||
|
||||
export
|
||||
Cast Integer Nat where
|
||||
cast = integerToNat
|
||||
|
||||
export
|
||||
Cast Bits8 Nat where
|
||||
cast = believe_me . cast {to = Integer}
|
||||
|
||||
export
|
||||
Cast Bits16 Nat where
|
||||
cast = believe_me . cast {to = Integer}
|
||||
|
||||
export
|
||||
Cast Bits32 Nat where
|
||||
cast = believe_me . cast {to = Integer}
|
||||
|
||||
export
|
||||
Cast Bits64 Nat where
|
||||
cast = believe_me . cast {to = Integer}
|
||||
|
||||
export
|
||||
Cast Int8 Nat where
|
||||
cast = integerToNat . cast
|
||||
|
||||
export
|
||||
Cast Int16 Nat where
|
||||
cast = integerToNat . cast
|
||||
|
||||
export
|
||||
Cast Int32 Nat where
|
||||
cast = integerToNat . cast
|
||||
|
||||
export
|
||||
Cast Int64 Nat where
|
||||
cast = integerToNat . cast
|
||||
|
@ -58,6 +58,22 @@ public export
|
||||
Eq Bits64 where
|
||||
x == y = intToBool (prim__eq_Bits64 x y)
|
||||
|
||||
public export
|
||||
Eq Int8 where
|
||||
x == y = intToBool (prim__eq_Int8 x y)
|
||||
|
||||
public export
|
||||
Eq Int16 where
|
||||
x == y = intToBool (prim__eq_Int16 x y)
|
||||
|
||||
public export
|
||||
Eq Int32 where
|
||||
x == y = intToBool (prim__eq_Int32 x y)
|
||||
|
||||
public export
|
||||
Eq Int64 where
|
||||
x == y = intToBool (prim__eq_Int64 x y)
|
||||
|
||||
public export
|
||||
Eq Double where
|
||||
x == y = intToBool (prim__eq_Double x y)
|
||||
@ -175,6 +191,34 @@ Ord Bits64 where
|
||||
(>) x y = intToBool (prim__gt_Bits64 x y)
|
||||
(>=) x y = intToBool (prim__gte_Bits64 x y)
|
||||
|
||||
public export
|
||||
Ord Int8 where
|
||||
(<) x y = intToBool (prim__lt_Int8 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int8 x y)
|
||||
(>) x y = intToBool (prim__gt_Int8 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int8 x y)
|
||||
|
||||
public export
|
||||
Ord Int16 where
|
||||
(<) x y = intToBool (prim__lt_Int16 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int16 x y)
|
||||
(>) x y = intToBool (prim__gt_Int16 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int16 x y)
|
||||
|
||||
public export
|
||||
Ord Int32 where
|
||||
(<) x y = intToBool (prim__lt_Int32 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int32 x y)
|
||||
(>) x y = intToBool (prim__gt_Int32 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int32 x y)
|
||||
|
||||
public export
|
||||
Ord Int64 where
|
||||
(<) x y = intToBool (prim__lt_Int64 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int64 x y)
|
||||
(>) x y = intToBool (prim__gt_Int64 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int64 x y)
|
||||
|
||||
public export
|
||||
Ord Double where
|
||||
(<) x y = intToBool (prim__lt_Double x y)
|
||||
@ -184,8 +228,6 @@ Ord Double where
|
||||
|
||||
public export
|
||||
Ord String where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_String x y)
|
||||
(<=) x y = intToBool (prim__lte_String x y)
|
||||
(>) x y = intToBool (prim__gt_String x y)
|
||||
@ -193,8 +235,6 @@ Ord String where
|
||||
|
||||
public export
|
||||
Ord Char where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Char x y)
|
||||
(<=) x y = intToBool (prim__lte_Char x y)
|
||||
(>) x y = intToBool (prim__gt_Char x y)
|
||||
|
@ -122,6 +122,114 @@ Integral Int where
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int x y
|
||||
|
||||
-- Int8
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
public export
|
||||
Neg Int8 where
|
||||
negate x = prim__sub_Int8 0 x
|
||||
(-) = prim__sub_Int8
|
||||
|
||||
public export
|
||||
Abs Int8 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int8 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int8 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int8 x y
|
||||
|
||||
-- Int16
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
public export
|
||||
Neg Int16 where
|
||||
negate x = prim__sub_Int16 0 x
|
||||
(-) = prim__sub_Int16
|
||||
|
||||
public export
|
||||
Abs Int16 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int16 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int16 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int16 x y
|
||||
|
||||
-- Int32
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
public export
|
||||
Neg Int32 where
|
||||
negate x = prim__sub_Int32 0 x
|
||||
(-) = prim__sub_Int32
|
||||
|
||||
public export
|
||||
Abs Int32 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int32 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int32 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int32 x y
|
||||
|
||||
-- Int64
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
public export
|
||||
Neg Int64 where
|
||||
negate x = prim__sub_Int64 0 x
|
||||
(-) = prim__sub_Int64
|
||||
|
||||
public export
|
||||
Abs Int64 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int64 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int64 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int64 x y
|
||||
|
||||
-- Bits8
|
||||
|
||||
%inline
|
||||
@ -131,6 +239,24 @@ Num Bits8 where
|
||||
(*) = prim__mul_Bits8
|
||||
fromInteger = prim__cast_IntegerBits8
|
||||
|
||||
public export
|
||||
Neg Bits8 where
|
||||
negate x = prim__sub_Bits8 0 x
|
||||
(-) = prim__sub_Bits8
|
||||
|
||||
public export
|
||||
Abs Bits8 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Bits8 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits8 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits8 x y
|
||||
|
||||
-- Bits16
|
||||
|
||||
%inline
|
||||
@ -140,6 +266,24 @@ Num Bits16 where
|
||||
(*) = prim__mul_Bits16
|
||||
fromInteger = prim__cast_IntegerBits16
|
||||
|
||||
public export
|
||||
Neg Bits16 where
|
||||
negate x = prim__sub_Bits16 0 x
|
||||
(-) = prim__sub_Bits16
|
||||
|
||||
public export
|
||||
Abs Bits16 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Bits16 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits16 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits16 x y
|
||||
|
||||
-- Bits32
|
||||
|
||||
%inline
|
||||
@ -149,6 +293,24 @@ Num Bits32 where
|
||||
(*) = prim__mul_Bits32
|
||||
fromInteger = prim__cast_IntegerBits32
|
||||
|
||||
public export
|
||||
Neg Bits32 where
|
||||
negate x = prim__sub_Bits32 0 x
|
||||
(-) = prim__sub_Bits32
|
||||
|
||||
public export
|
||||
Abs Bits32 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Bits32 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits32 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits32 x y
|
||||
|
||||
-- Bits64
|
||||
|
||||
%inline
|
||||
@ -158,6 +320,24 @@ Num Bits64 where
|
||||
(*) = prim__mul_Bits64
|
||||
fromInteger = prim__cast_IntegerBits64
|
||||
|
||||
public export
|
||||
Neg Bits64 where
|
||||
negate x = prim__sub_Bits64 0 x
|
||||
(-) = prim__sub_Bits64
|
||||
|
||||
public export
|
||||
Abs Bits64 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Bits64 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits64 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits64 x y
|
||||
|
||||
-- Double
|
||||
|
||||
public export
|
||||
|
@ -124,6 +124,22 @@ export
|
||||
Show Bits64 where
|
||||
showPrec = primNumShow prim__cast_Bits64String
|
||||
|
||||
export
|
||||
Show Int8 where
|
||||
showPrec = primNumShow prim__cast_Int8String
|
||||
|
||||
export
|
||||
Show Int16 where
|
||||
showPrec = primNumShow prim__cast_Int16String
|
||||
|
||||
export
|
||||
Show Int32 where
|
||||
showPrec = primNumShow prim__cast_Int32String
|
||||
|
||||
export
|
||||
Show Int64 where
|
||||
showPrec = primNumShow prim__cast_Int64String
|
||||
|
||||
export
|
||||
Show Double where
|
||||
showPrec = primNumShow prim__cast_DoubleString
|
||||
|
@ -1,85 +1,5 @@
|
||||
import Data.List
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int8 where
|
||||
show = prim__cast_Int8String
|
||||
|
||||
public export
|
||||
Eq Int8 where
|
||||
x == y = intToBool (prim__eq_Int8 x y)
|
||||
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
Neg Int8 where
|
||||
(-) = prim__sub_Int8
|
||||
negate = prim__sub_Int8 0
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int16 where
|
||||
show = prim__cast_Int16String
|
||||
|
||||
public export
|
||||
Eq Int16 where
|
||||
x == y = intToBool (prim__eq_Int16 x y)
|
||||
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
Neg Int16 where
|
||||
(-) = prim__sub_Int16
|
||||
negate = prim__sub_Int16 0
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int32 where
|
||||
show = prim__cast_Int32String
|
||||
|
||||
public export
|
||||
Eq Int32 where
|
||||
x == y = intToBool (prim__eq_Int32 x y)
|
||||
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
Neg Int32 where
|
||||
(-) = prim__sub_Int32
|
||||
negate = prim__sub_Int32 0
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int64 where
|
||||
show = prim__cast_Int64String
|
||||
|
||||
public export
|
||||
Eq Int64 where
|
||||
x == y = intToBool (prim__eq_Int64 x y)
|
||||
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
Neg Int64 where
|
||||
(-) = prim__sub_Int64
|
||||
negate = prim__sub_Int64 0
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Tests
|
||||
--------------------------------------------------------------------------------
|
||||
|
@ -74,421 +74,6 @@
|
||||
|
||||
import Data.List
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int8 where
|
||||
show = prim__cast_Int8String
|
||||
|
||||
public export
|
||||
Eq Int8 where
|
||||
x == y = intToBool (prim__eq_Int8 x y)
|
||||
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
Neg Int8 where
|
||||
(-) = prim__sub_Int8
|
||||
negate = prim__sub_Int8 0
|
||||
|
||||
Cast Int8 Bits8 where
|
||||
cast = prim__cast_Int8Bits8
|
||||
|
||||
Cast Int8 Bits16 where
|
||||
cast = prim__cast_Int8Bits16
|
||||
|
||||
Cast Int8 Bits32 where
|
||||
cast = prim__cast_Int8Bits32
|
||||
|
||||
Cast Int8 Bits64 where
|
||||
cast = prim__cast_Int8Bits64
|
||||
|
||||
Cast Int8 Int16 where
|
||||
cast = prim__cast_Int8Int16
|
||||
|
||||
Cast Int8 Int32 where
|
||||
cast = prim__cast_Int8Int32
|
||||
|
||||
Cast Int8 Int64 where
|
||||
cast = prim__cast_Int8Int64
|
||||
|
||||
Cast Int8 Int where
|
||||
cast = prim__cast_Int8Int
|
||||
|
||||
Cast Int8 Integer where
|
||||
cast = prim__cast_Int8Integer
|
||||
|
||||
Cast Int8 String where
|
||||
cast = prim__cast_Int8String
|
||||
|
||||
Cast Int8 Char where
|
||||
cast = prim__cast_Int8Char
|
||||
|
||||
Cast Int8 Double where
|
||||
cast = prim__cast_Int8Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
public export
|
||||
Eq Int16 where
|
||||
x == y = intToBool (prim__eq_Int16 x y)
|
||||
|
||||
Show Int16 where
|
||||
show = prim__cast_Int16String
|
||||
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
Neg Int16 where
|
||||
(-) = prim__sub_Int16
|
||||
negate = prim__sub_Int16 0
|
||||
|
||||
Cast Int16 Bits8 where
|
||||
cast = prim__cast_Int16Bits8
|
||||
|
||||
Cast Int16 Bits16 where
|
||||
cast = prim__cast_Int16Bits16
|
||||
|
||||
Cast Int16 Bits32 where
|
||||
cast = prim__cast_Int16Bits32
|
||||
|
||||
Cast Int16 Bits64 where
|
||||
cast = prim__cast_Int16Bits64
|
||||
|
||||
Cast Int16 Int8 where
|
||||
cast = prim__cast_Int16Int8
|
||||
|
||||
Cast Int16 Int32 where
|
||||
cast = prim__cast_Int16Int32
|
||||
|
||||
Cast Int16 Int64 where
|
||||
cast = prim__cast_Int16Int64
|
||||
|
||||
Cast Int16 Int where
|
||||
cast = prim__cast_Int16Int
|
||||
|
||||
Cast Int16 Integer where
|
||||
cast = prim__cast_Int16Integer
|
||||
|
||||
Cast Int16 String where
|
||||
cast = prim__cast_Int16String
|
||||
|
||||
Cast Int16 Char where
|
||||
cast = prim__cast_Int16Char
|
||||
|
||||
Cast Int16 Double where
|
||||
cast = prim__cast_Int16Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int32 where
|
||||
show = prim__cast_Int32String
|
||||
|
||||
public export
|
||||
Eq Int32 where
|
||||
x == y = intToBool (prim__eq_Int32 x y)
|
||||
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
Neg Int32 where
|
||||
(-) = prim__sub_Int32
|
||||
negate = prim__sub_Int32 0
|
||||
|
||||
Cast Int32 Bits8 where
|
||||
cast = prim__cast_Int32Bits8
|
||||
|
||||
Cast Int32 Bits16 where
|
||||
cast = prim__cast_Int32Bits16
|
||||
|
||||
Cast Int32 Bits32 where
|
||||
cast = prim__cast_Int32Bits32
|
||||
|
||||
Cast Int32 Bits64 where
|
||||
cast = prim__cast_Int32Bits64
|
||||
|
||||
Cast Int32 Int8 where
|
||||
cast = prim__cast_Int32Int8
|
||||
|
||||
Cast Int32 Int16 where
|
||||
cast = prim__cast_Int32Int16
|
||||
|
||||
Cast Int32 Int64 where
|
||||
cast = prim__cast_Int32Int64
|
||||
|
||||
Cast Int32 Int where
|
||||
cast = prim__cast_Int32Int
|
||||
|
||||
Cast Int32 Integer where
|
||||
cast = prim__cast_Int32Integer
|
||||
|
||||
Cast Int32 String where
|
||||
cast = prim__cast_Int32String
|
||||
|
||||
Cast Int32 Char where
|
||||
cast = prim__cast_Int32Char
|
||||
|
||||
Cast Int32 Double where
|
||||
cast = prim__cast_Int32Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int64 where
|
||||
show = prim__cast_Int64String
|
||||
|
||||
public export
|
||||
Eq Int64 where
|
||||
x == y = intToBool (prim__eq_Int64 x y)
|
||||
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
Neg Int64 where
|
||||
(-) = prim__sub_Int64
|
||||
negate = prim__sub_Int64 0
|
||||
|
||||
Cast Int64 Bits8 where
|
||||
cast = prim__cast_Int64Bits8
|
||||
|
||||
Cast Int64 Bits16 where
|
||||
cast = prim__cast_Int64Bits16
|
||||
|
||||
Cast Int64 Bits32 where
|
||||
cast = prim__cast_Int64Bits32
|
||||
|
||||
Cast Int64 Bits64 where
|
||||
cast = prim__cast_Int64Bits64
|
||||
|
||||
Cast Int64 Int8 where
|
||||
cast = prim__cast_Int64Int8
|
||||
|
||||
Cast Int64 Int16 where
|
||||
cast = prim__cast_Int64Int16
|
||||
|
||||
Cast Int64 Int32 where
|
||||
cast = prim__cast_Int64Int32
|
||||
|
||||
Cast Int64 Int where
|
||||
cast = prim__cast_Int64Int
|
||||
|
||||
Cast Int64 Integer where
|
||||
cast = prim__cast_Int64Integer
|
||||
|
||||
Cast Int64 String where
|
||||
cast = prim__cast_Int64String
|
||||
|
||||
Cast Int64 Char where
|
||||
cast = prim__cast_Int64Char
|
||||
|
||||
Cast Int64 Double where
|
||||
cast = prim__cast_Int64Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Int Int8 where
|
||||
cast = prim__cast_IntInt8
|
||||
|
||||
Cast Int Int16 where
|
||||
cast = prim__cast_IntInt16
|
||||
|
||||
Cast Int Int32 where
|
||||
cast = prim__cast_IntInt32
|
||||
|
||||
Cast Int Int64 where
|
||||
cast = prim__cast_IntInt64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Integer
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Integer Int8 where
|
||||
cast = prim__cast_IntegerInt8
|
||||
|
||||
Cast Integer Int16 where
|
||||
cast = prim__cast_IntegerInt16
|
||||
|
||||
Cast Integer Int32 where
|
||||
cast = prim__cast_IntegerInt32
|
||||
|
||||
Cast Integer Int64 where
|
||||
cast = prim__cast_IntegerInt64
|
||||
|
||||
Cast Integer Char where
|
||||
cast = prim__cast_IntegerChar
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Bits8 Int8 where
|
||||
cast = prim__cast_Bits8Int8
|
||||
|
||||
Cast Bits8 Int16 where
|
||||
cast = prim__cast_Bits8Int16
|
||||
|
||||
Cast Bits8 Int32 where
|
||||
cast = prim__cast_Bits8Int32
|
||||
|
||||
Cast Bits8 Int64 where
|
||||
cast = prim__cast_Bits8Int64
|
||||
|
||||
Cast Bits8 String where
|
||||
cast = prim__cast_Bits8String
|
||||
|
||||
Cast Bits8 Char where
|
||||
cast = prim__cast_Bits8Char
|
||||
|
||||
Cast Bits8 Double where
|
||||
cast = prim__cast_Bits8Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Bits16 Int8 where
|
||||
cast = prim__cast_Bits16Int8
|
||||
|
||||
Cast Bits16 Int16 where
|
||||
cast = prim__cast_Bits16Int16
|
||||
|
||||
Cast Bits16 Int32 where
|
||||
cast = prim__cast_Bits16Int32
|
||||
|
||||
Cast Bits16 Int64 where
|
||||
cast = prim__cast_Bits16Int64
|
||||
|
||||
Cast Bits16 String where
|
||||
cast = prim__cast_Bits16String
|
||||
|
||||
Cast Bits16 Char where
|
||||
cast = prim__cast_Bits16Char
|
||||
|
||||
Cast Bits16 Double where
|
||||
cast = prim__cast_Bits16Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Bits32 Int8 where
|
||||
cast = prim__cast_Bits32Int8
|
||||
|
||||
Cast Bits32 Int16 where
|
||||
cast = prim__cast_Bits32Int16
|
||||
|
||||
Cast Bits32 Int32 where
|
||||
cast = prim__cast_Bits32Int32
|
||||
|
||||
Cast Bits32 Int64 where
|
||||
cast = prim__cast_Bits32Int64
|
||||
|
||||
Cast Bits32 String where
|
||||
cast = prim__cast_Bits32String
|
||||
|
||||
Cast Bits32 Char where
|
||||
cast = prim__cast_Bits32Char
|
||||
|
||||
Cast Bits32 Double where
|
||||
cast = prim__cast_Bits32Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Bits64 Int8 where
|
||||
cast = prim__cast_Bits64Int8
|
||||
|
||||
Cast Bits64 Int16 where
|
||||
cast = prim__cast_Bits64Int16
|
||||
|
||||
Cast Bits64 Int32 where
|
||||
cast = prim__cast_Bits64Int32
|
||||
|
||||
Cast Bits64 Int64 where
|
||||
cast = prim__cast_Bits64Int64
|
||||
|
||||
Cast Bits64 String where
|
||||
cast = prim__cast_Bits64String
|
||||
|
||||
Cast Bits64 Char where
|
||||
cast = prim__cast_Bits64Char
|
||||
|
||||
Cast Bits64 Double where
|
||||
cast = prim__cast_Bits64Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- String
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast String Bits8 where
|
||||
cast = prim__cast_StringBits8
|
||||
|
||||
Cast String Bits16 where
|
||||
cast = prim__cast_StringBits16
|
||||
|
||||
Cast String Bits32 where
|
||||
cast = prim__cast_StringBits32
|
||||
|
||||
Cast String Bits64 where
|
||||
cast = prim__cast_StringBits64
|
||||
|
||||
Cast String Int8 where
|
||||
cast = prim__cast_StringInt8
|
||||
|
||||
Cast String Int16 where
|
||||
cast = prim__cast_StringInt16
|
||||
|
||||
Cast String Int32 where
|
||||
cast = prim__cast_StringInt32
|
||||
|
||||
Cast String Int64 where
|
||||
cast = prim__cast_StringInt64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Double
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Double Bits8 where
|
||||
cast = prim__cast_DoubleBits8
|
||||
|
||||
Cast Double Bits16 where
|
||||
cast = prim__cast_DoubleBits16
|
||||
|
||||
Cast Double Bits32 where
|
||||
cast = prim__cast_DoubleBits32
|
||||
|
||||
Cast Double Bits64 where
|
||||
cast = prim__cast_DoubleBits64
|
||||
|
||||
Cast Double Int8 where
|
||||
cast = prim__cast_DoubleInt8
|
||||
|
||||
Cast Double Int16 where
|
||||
cast = prim__cast_DoubleInt16
|
||||
|
||||
Cast Double Int32 where
|
||||
cast = prim__cast_DoubleInt32
|
||||
|
||||
Cast Double Int64 where
|
||||
cast = prim__cast_DoubleInt64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Tests
|
||||
--------------------------------------------------------------------------------
|
||||
|
@ -1,94 +0,0 @@
|
||||
module IntBits
|
||||
|
||||
import IntCasts
|
||||
import IntEqOrd
|
||||
import IntNum
|
||||
|
||||
import Data.Bits
|
||||
import Data.DPair
|
||||
|
||||
-- This file to be deleted once these interfaces are added to base
|
||||
|
||||
public export %inline
|
||||
Bits Int8 where
|
||||
Index = Subset Nat (`LT` 8)
|
||||
(.&.) = prim__and_Int8
|
||||
(.|.) = prim__or_Int8
|
||||
xor = prim__xor_Int8
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int8 x . cast . fst
|
||||
shiftL x = prim__shl_Int8 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int16 where
|
||||
Index = Subset Nat (`LT` 16)
|
||||
(.&.) = prim__and_Int16
|
||||
(.|.) = prim__or_Int16
|
||||
xor = prim__xor_Int16
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int16 x . cast . fst
|
||||
shiftL x = prim__shl_Int16 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int32 where
|
||||
Index = Subset Nat (`LT` 32)
|
||||
(.&.) = prim__and_Int32
|
||||
(.|.) = prim__or_Int32
|
||||
xor = prim__xor_Int32
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int32 x . cast . fst
|
||||
shiftL x = prim__shl_Int32 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
(.&.) = prim__and_Int64
|
||||
(.|.) = prim__or_Int64
|
||||
xor = prim__xor_Int64
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int64 x . cast . fst
|
||||
shiftL x = prim__shl_Int64 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Bits64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
(.&.) = prim__and_Bits64
|
||||
(.|.) = prim__or_Bits64
|
||||
xor = prim__xor_Bits64
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits64 x . fromInteger . cast . fst
|
||||
shiftL x = prim__shl_Bits64 x . fromInteger . cast . fst
|
||||
complement = xor 0xffffffffffffffff
|
||||
oneBits = 0xffffffffffffffff
|
||||
|
||||
public export %inline
|
||||
Bits Integer where
|
||||
Index = Nat
|
||||
(.&.) = prim__and_Integer
|
||||
(.|.) = prim__or_Integer
|
||||
xor = prim__xor_Integer
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Integer x . cast
|
||||
shiftL x = prim__shl_Integer x . cast
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
@ -1,383 +0,0 @@
|
||||
module IntCasts
|
||||
|
||||
import IntNum
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
-- To String
|
||||
|
||||
export
|
||||
Cast Int8 String where
|
||||
cast = prim__cast_Int8String
|
||||
|
||||
export
|
||||
Cast Int16 String where
|
||||
cast = prim__cast_Int16String
|
||||
|
||||
export
|
||||
Cast Int32 String where
|
||||
cast = prim__cast_Int32String
|
||||
|
||||
export
|
||||
Cast Int64 String where
|
||||
cast = prim__cast_Int64String
|
||||
|
||||
-- To Integer
|
||||
|
||||
export
|
||||
Cast Int8 Integer where
|
||||
cast = prim__cast_Int8Integer
|
||||
|
||||
export
|
||||
Cast Int16 Integer where
|
||||
cast = prim__cast_Int16Integer
|
||||
|
||||
export
|
||||
Cast Int32 Integer where
|
||||
cast = prim__cast_Int32Integer
|
||||
|
||||
export
|
||||
Cast Int64 Integer where
|
||||
cast = prim__cast_Int64Integer
|
||||
|
||||
-- To Int
|
||||
|
||||
export
|
||||
Cast Int8 Int where
|
||||
cast = prim__cast_Int8Int
|
||||
|
||||
export
|
||||
Cast Int16 Int where
|
||||
cast = prim__cast_Int16Int
|
||||
|
||||
export
|
||||
Cast Int32 Int where
|
||||
cast = prim__cast_Int32Int
|
||||
|
||||
export
|
||||
Cast Int64 Int where
|
||||
cast = prim__cast_Int64Int
|
||||
|
||||
-- To Int8
|
||||
|
||||
export
|
||||
Cast Int Int8 where
|
||||
cast = prim__cast_IntInt8
|
||||
|
||||
export
|
||||
Cast Int16 Int8 where
|
||||
cast = prim__cast_Int16Int8
|
||||
|
||||
export
|
||||
Cast Int32 Int8 where
|
||||
cast = prim__cast_Int32Int8
|
||||
|
||||
export
|
||||
Cast Int64 Int8 where
|
||||
cast = prim__cast_Int64Int8
|
||||
|
||||
export
|
||||
Cast Integer Int8 where
|
||||
cast = prim__cast_IntegerInt8
|
||||
|
||||
export
|
||||
Cast Char Int8 where
|
||||
cast = prim__cast_CharInt8
|
||||
|
||||
export
|
||||
Cast Double Int8 where
|
||||
cast = prim__cast_DoubleInt8
|
||||
|
||||
export
|
||||
Cast String Int8 where
|
||||
cast = prim__cast_StringInt8
|
||||
|
||||
export
|
||||
Cast Nat Int8 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int8 where
|
||||
cast = prim__cast_Bits8Int8
|
||||
|
||||
export
|
||||
Cast Bits16 Int8 where
|
||||
cast = prim__cast_Bits16Int8
|
||||
|
||||
export
|
||||
Cast Bits32 Int8 where
|
||||
cast = prim__cast_Bits32Int8
|
||||
|
||||
export
|
||||
Cast Bits64 Int8 where
|
||||
cast = prim__cast_Bits64Int8
|
||||
|
||||
-- To Int16
|
||||
|
||||
export
|
||||
Cast Int Int16 where
|
||||
cast = prim__cast_IntInt16
|
||||
|
||||
export
|
||||
Cast Int8 Int16 where
|
||||
cast = prim__cast_Int8Int16
|
||||
|
||||
export
|
||||
Cast Int32 Int16 where
|
||||
cast = prim__cast_Int32Int16
|
||||
|
||||
export
|
||||
Cast Int64 Int16 where
|
||||
cast = prim__cast_Int64Int16
|
||||
|
||||
export
|
||||
Cast Integer Int16 where
|
||||
cast = prim__cast_IntegerInt16
|
||||
|
||||
export
|
||||
Cast Char Int16 where
|
||||
cast = prim__cast_CharInt16
|
||||
|
||||
export
|
||||
Cast Double Int16 where
|
||||
cast = prim__cast_DoubleInt16
|
||||
|
||||
export
|
||||
Cast String Int16 where
|
||||
cast = prim__cast_StringInt16
|
||||
|
||||
export
|
||||
Cast Nat Int16 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int16 where
|
||||
cast = prim__cast_Bits8Int16
|
||||
|
||||
export
|
||||
Cast Bits16 Int16 where
|
||||
cast = prim__cast_Bits16Int16
|
||||
|
||||
export
|
||||
Cast Bits32 Int16 where
|
||||
cast = prim__cast_Bits32Int16
|
||||
|
||||
export
|
||||
Cast Bits64 Int16 where
|
||||
cast = prim__cast_Bits64Int16
|
||||
|
||||
-- To Int32
|
||||
|
||||
export
|
||||
Cast Int Int32 where
|
||||
cast = prim__cast_IntInt32
|
||||
|
||||
export
|
||||
Cast Int8 Int32 where
|
||||
cast = prim__cast_Int8Int32
|
||||
|
||||
export
|
||||
Cast Int16 Int32 where
|
||||
cast = prim__cast_Int16Int32
|
||||
|
||||
export
|
||||
Cast Int64 Int32 where
|
||||
cast = prim__cast_Int64Int32
|
||||
|
||||
export
|
||||
Cast Integer Int32 where
|
||||
cast = prim__cast_IntegerInt32
|
||||
|
||||
export
|
||||
Cast Char Int32 where
|
||||
cast = prim__cast_CharInt32
|
||||
|
||||
export
|
||||
Cast Double Int32 where
|
||||
cast = prim__cast_DoubleInt32
|
||||
|
||||
export
|
||||
Cast String Int32 where
|
||||
cast = prim__cast_StringInt32
|
||||
|
||||
export
|
||||
Cast Nat Int32 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int32 where
|
||||
cast = prim__cast_Bits8Int32
|
||||
|
||||
export
|
||||
Cast Bits16 Int32 where
|
||||
cast = prim__cast_Bits16Int32
|
||||
|
||||
export
|
||||
Cast Bits32 Int32 where
|
||||
cast = prim__cast_Bits32Int32
|
||||
|
||||
export
|
||||
Cast Bits64 Int32 where
|
||||
cast = prim__cast_Bits64Int32
|
||||
|
||||
-- To Int64
|
||||
|
||||
export
|
||||
Cast Int Int64 where
|
||||
cast = prim__cast_IntInt64
|
||||
|
||||
export
|
||||
Cast Int8 Int64 where
|
||||
cast = prim__cast_Int8Int64
|
||||
|
||||
export
|
||||
Cast Int16 Int64 where
|
||||
cast = prim__cast_Int16Int64
|
||||
|
||||
export
|
||||
Cast Int32 Int64 where
|
||||
cast = prim__cast_Int32Int64
|
||||
|
||||
export
|
||||
Cast Integer Int64 where
|
||||
cast = prim__cast_IntegerInt64
|
||||
|
||||
export
|
||||
Cast Char Int64 where
|
||||
cast = prim__cast_CharInt64
|
||||
|
||||
export
|
||||
Cast Double Int64 where
|
||||
cast = prim__cast_DoubleInt64
|
||||
|
||||
export
|
||||
Cast String Int64 where
|
||||
cast = prim__cast_StringInt64
|
||||
|
||||
export
|
||||
Cast Nat Int64 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int64 where
|
||||
cast = prim__cast_Bits8Int64
|
||||
|
||||
export
|
||||
Cast Bits16 Int64 where
|
||||
cast = prim__cast_Bits16Int64
|
||||
|
||||
export
|
||||
Cast Bits32 Int64 where
|
||||
cast = prim__cast_Bits32Int64
|
||||
|
||||
export
|
||||
Cast Bits64 Int64 where
|
||||
cast = prim__cast_Bits64Int64
|
||||
|
||||
-- To Char
|
||||
|
||||
export
|
||||
Cast Int8 Char where
|
||||
cast = prim__cast_Int8Char
|
||||
|
||||
export
|
||||
Cast Int16 Char where
|
||||
cast = prim__cast_Int16Char
|
||||
|
||||
export
|
||||
Cast Int32 Char where
|
||||
cast = prim__cast_Int32Char
|
||||
|
||||
export
|
||||
Cast Int64 Char where
|
||||
cast = prim__cast_Int64Char
|
||||
|
||||
-- To Double
|
||||
|
||||
export
|
||||
Cast Int8 Double where
|
||||
cast = prim__cast_Int8Double
|
||||
|
||||
export
|
||||
Cast Int16 Double where
|
||||
cast = prim__cast_Int16Double
|
||||
|
||||
export
|
||||
Cast Int32 Double where
|
||||
cast = prim__cast_Int32Double
|
||||
|
||||
export
|
||||
Cast Int64 Double where
|
||||
cast = prim__cast_Int64Double
|
||||
|
||||
-- To Bits8
|
||||
|
||||
export
|
||||
Cast Int8 Bits8 where
|
||||
cast = prim__cast_Int8Bits8
|
||||
|
||||
export
|
||||
Cast Int16 Bits8 where
|
||||
cast = prim__cast_Int16Bits8
|
||||
|
||||
export
|
||||
Cast Int32 Bits8 where
|
||||
cast = prim__cast_Int32Bits8
|
||||
|
||||
export
|
||||
Cast Int64 Bits8 where
|
||||
cast = prim__cast_Int64Bits8
|
||||
|
||||
-- To Bits16
|
||||
|
||||
export
|
||||
Cast Int8 Bits16 where
|
||||
cast = prim__cast_Int8Bits16
|
||||
|
||||
export
|
||||
Cast Int16 Bits16 where
|
||||
cast = prim__cast_Int16Bits16
|
||||
|
||||
export
|
||||
Cast Int32 Bits16 where
|
||||
cast = prim__cast_Int32Bits16
|
||||
|
||||
export
|
||||
Cast Int64 Bits16 where
|
||||
cast = prim__cast_Int64Bits16
|
||||
|
||||
-- To Bits32
|
||||
|
||||
export
|
||||
Cast Int8 Bits32 where
|
||||
cast = prim__cast_Int8Bits32
|
||||
|
||||
export
|
||||
Cast Int16 Bits32 where
|
||||
cast = prim__cast_Int16Bits32
|
||||
|
||||
export
|
||||
Cast Int32 Bits32 where
|
||||
cast = prim__cast_Int32Bits32
|
||||
|
||||
export
|
||||
Cast Int64 Bits32 where
|
||||
cast = prim__cast_Int64Bits32
|
||||
|
||||
-- To Bits64
|
||||
|
||||
export
|
||||
Cast Int8 Bits64 where
|
||||
cast = prim__cast_Int8Bits64
|
||||
|
||||
export
|
||||
Cast Int16 Bits64 where
|
||||
cast = prim__cast_Int16Bits64
|
||||
|
||||
export
|
||||
Cast Int32 Bits64 where
|
||||
cast = prim__cast_Int32Bits64
|
||||
|
||||
export
|
||||
Cast Int64 Bits64 where
|
||||
cast = prim__cast_Int64Bits64
|
@ -1,53 +0,0 @@
|
||||
module IntEqOrd
|
||||
|
||||
public export
|
||||
Eq Int8 where
|
||||
x == y = intToBool (prim__eq_Int8 x y)
|
||||
|
||||
public export
|
||||
Eq Int16 where
|
||||
x == y = intToBool (prim__eq_Int16 x y)
|
||||
|
||||
public export
|
||||
Eq Int32 where
|
||||
x == y = intToBool (prim__eq_Int32 x y)
|
||||
|
||||
public export
|
||||
Eq Int64 where
|
||||
x == y = intToBool (prim__eq_Int64 x y)
|
||||
|
||||
public export
|
||||
Ord Int8 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int8 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int8 x y)
|
||||
(>) x y = intToBool (prim__gt_Int8 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int8 x y)
|
||||
|
||||
public export
|
||||
Ord Int16 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int16 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int16 x y)
|
||||
(>) x y = intToBool (prim__gt_Int16 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int16 x y)
|
||||
|
||||
public export
|
||||
Ord Int32 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int32 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int32 x y)
|
||||
(>) x y = intToBool (prim__gt_Int32 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int32 x y)
|
||||
|
||||
public export
|
||||
Ord Int64 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int64 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int64 x y)
|
||||
(>) x y = intToBool (prim__gt_Int64 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int64 x y)
|
@ -1,193 +0,0 @@
|
||||
module IntNum
|
||||
|
||||
import IntEqOrd
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
-- Int8
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
public export
|
||||
Neg Int8 where
|
||||
negate x = prim__sub_Int8 0 x
|
||||
(-) = prim__sub_Int8
|
||||
|
||||
public export
|
||||
Abs Int8 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int8 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int8 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int8 x y
|
||||
|
||||
-- Int16
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
public export
|
||||
Neg Int16 where
|
||||
negate x = prim__sub_Int16 0 x
|
||||
(-) = prim__sub_Int16
|
||||
|
||||
public export
|
||||
Abs Int16 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int16 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int16 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int16 x y
|
||||
|
||||
-- Int32
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
public export
|
||||
Neg Int32 where
|
||||
negate x = prim__sub_Int32 0 x
|
||||
(-) = prim__sub_Int32
|
||||
|
||||
public export
|
||||
Abs Int32 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int32 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int32 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int32 x y
|
||||
|
||||
-- Int64
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
public export
|
||||
Neg Int64 where
|
||||
negate x = prim__sub_Int64 0 x
|
||||
(-) = prim__sub_Int64
|
||||
|
||||
public export
|
||||
Abs Int64 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int64 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int64 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int64 x y
|
||||
|
||||
-- Bits8
|
||||
|
||||
public export
|
||||
Neg Bits8 where
|
||||
negate x = prim__sub_Bits8 0 x
|
||||
(-) = prim__sub_Bits8
|
||||
|
||||
public export
|
||||
Abs Bits8 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits8 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits8 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits8 x y
|
||||
|
||||
-- Bits16
|
||||
|
||||
public export
|
||||
Neg Bits16 where
|
||||
negate x = prim__sub_Bits16 0 x
|
||||
(-) = prim__sub_Bits16
|
||||
|
||||
public export
|
||||
Abs Bits16 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits16 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits16 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits16 x y
|
||||
|
||||
-- Bits32
|
||||
|
||||
public export
|
||||
Neg Bits32 where
|
||||
negate x = prim__sub_Bits32 0 x
|
||||
(-) = prim__sub_Bits32
|
||||
|
||||
public export
|
||||
Abs Bits32 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits32 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits32 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits32 x y
|
||||
|
||||
-- Bits64
|
||||
|
||||
public export
|
||||
Neg Bits64 where
|
||||
negate x = prim__sub_Bits64 0 x
|
||||
(-) = prim__sub_Bits64
|
||||
|
||||
public export
|
||||
Abs Bits64 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits64 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits64 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits64 x y
|
@ -1,26 +0,0 @@
|
||||
module IntShow
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
firstCharIs : (Char -> Bool) -> String -> Bool
|
||||
firstCharIs p "" = False
|
||||
firstCharIs p str = p (assert_total (prim__strHead str))
|
||||
|
||||
primNumShow : (a -> String) -> Prec -> a -> String
|
||||
primNumShow f d x = let str = f x in showParens (d >= PrefixMinus && firstCharIs (== '-') str) str
|
||||
|
||||
export
|
||||
Show Int8 where
|
||||
showPrec = primNumShow prim__cast_Int8String
|
||||
|
||||
export
|
||||
Show Int16 where
|
||||
showPrec = primNumShow prim__cast_Int16String
|
||||
|
||||
export
|
||||
Show Int32 where
|
||||
showPrec = primNumShow prim__cast_Int32String
|
||||
|
||||
export
|
||||
Show Int64 where
|
||||
showPrec = primNumShow prim__cast_Int64String
|
@ -1,11 +1,5 @@
|
||||
module TestIntegers
|
||||
|
||||
import IntBits
|
||||
import IntCasts
|
||||
import IntEqOrd
|
||||
import IntNum
|
||||
import IntShow
|
||||
|
||||
import Data.Bits
|
||||
import Data.List.Quantifiers
|
||||
|
||||
|
@ -158,146 +158,6 @@ check (MkOp name op opInt allowZero $ MkIntType type signed bits mi ma) =
|
||||
then Nothing
|
||||
else Just #"Error when calculating \#{show x} \#{name} \#{show y} for \#{type}: Expected \#{show resMod} but got \#{show resA} (unrounded: \#{show resInteger})"#
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int8 where
|
||||
show = prim__cast_Int8String
|
||||
|
||||
Cast Int8 Integer where
|
||||
cast = prim__cast_Int8Integer
|
||||
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
Neg Int8 where
|
||||
(-) = prim__sub_Int8
|
||||
negate = prim__sub_Int8 0
|
||||
|
||||
Integral Int8 where
|
||||
div = prim__div_Int8
|
||||
mod = prim__mod_Int8
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int16 where
|
||||
show = prim__cast_Int16String
|
||||
|
||||
Cast Int16 Integer where
|
||||
cast = prim__cast_Int16Integer
|
||||
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
Neg Int16 where
|
||||
(-) = prim__sub_Int16
|
||||
negate = prim__sub_Int16 0
|
||||
|
||||
Integral Int16 where
|
||||
div = prim__div_Int16
|
||||
mod = prim__mod_Int16
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int32 where
|
||||
show = prim__cast_Int32String
|
||||
|
||||
Cast Int32 Integer where
|
||||
cast = prim__cast_Int32Integer
|
||||
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
Neg Int32 where
|
||||
(-) = prim__sub_Int32
|
||||
negate = prim__sub_Int32 0
|
||||
|
||||
Integral Int32 where
|
||||
div = prim__div_Int32
|
||||
mod = prim__mod_Int32
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int64 where
|
||||
show = prim__cast_Int64String
|
||||
|
||||
Cast Int64 Integer where
|
||||
cast = prim__cast_Int64Integer
|
||||
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
Neg Int64 where
|
||||
(-) = prim__sub_Int64
|
||||
negate = prim__sub_Int64 0
|
||||
|
||||
Integral Int64 where
|
||||
div = prim__div_Int64
|
||||
mod = prim__mod_Int64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Neg Bits8 where
|
||||
(-) = prim__sub_Bits8
|
||||
negate = prim__sub_Bits8 0
|
||||
|
||||
Integral Bits8 where
|
||||
div = prim__div_Bits8
|
||||
mod = prim__mod_Bits8
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Neg Bits16 where
|
||||
(-) = prim__sub_Bits16
|
||||
negate = prim__sub_Bits16 0
|
||||
|
||||
Integral Bits16 where
|
||||
div = prim__div_Bits16
|
||||
mod = prim__mod_Bits16
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Neg Bits32 where
|
||||
(-) = prim__sub_Bits32
|
||||
negate = prim__sub_Bits32 0
|
||||
|
||||
Integral Bits32 where
|
||||
div = prim__div_Bits32
|
||||
mod = prim__mod_Bits32
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Neg Bits64 where
|
||||
(-) = prim__sub_Bits64
|
||||
negate = prim__sub_Bits64 0
|
||||
|
||||
Integral Bits64 where
|
||||
div = prim__div_Bits64
|
||||
mod = prim__mod_Bits64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Main
|
||||
--------------------------------------------------------------------------------
|
||||
|
@ -46,6 +46,10 @@ Main> Prelude.Show : Type -> Type
|
||||
Show Bits16
|
||||
Show Bits32
|
||||
Show Bits64
|
||||
Show Int8
|
||||
Show Int16
|
||||
Show Int32
|
||||
Show Int64
|
||||
Show Double
|
||||
Show Char
|
||||
Show String
|
||||
|
@ -1,85 +1,5 @@
|
||||
import Data.List
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int8 where
|
||||
show = prim__cast_Int8String
|
||||
|
||||
public export
|
||||
Eq Int8 where
|
||||
x == y = intToBool (prim__eq_Int8 x y)
|
||||
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
Neg Int8 where
|
||||
(-) = prim__sub_Int8
|
||||
negate = prim__sub_Int8 0
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int16 where
|
||||
show = prim__cast_Int16String
|
||||
|
||||
public export
|
||||
Eq Int16 where
|
||||
x == y = intToBool (prim__eq_Int16 x y)
|
||||
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
Neg Int16 where
|
||||
(-) = prim__sub_Int16
|
||||
negate = prim__sub_Int16 0
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int32 where
|
||||
show = prim__cast_Int32String
|
||||
|
||||
public export
|
||||
Eq Int32 where
|
||||
x == y = intToBool (prim__eq_Int32 x y)
|
||||
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
Neg Int32 where
|
||||
(-) = prim__sub_Int32
|
||||
negate = prim__sub_Int32 0
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int64 where
|
||||
show = prim__cast_Int64String
|
||||
|
||||
public export
|
||||
Eq Int64 where
|
||||
x == y = intToBool (prim__eq_Int64 x y)
|
||||
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
Neg Int64 where
|
||||
(-) = prim__sub_Int64
|
||||
negate = prim__sub_Int64 0
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Tests
|
||||
--------------------------------------------------------------------------------
|
||||
|
@ -74,421 +74,6 @@
|
||||
|
||||
import Data.List
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int8 where
|
||||
show = prim__cast_Int8String
|
||||
|
||||
public export
|
||||
Eq Int8 where
|
||||
x == y = intToBool (prim__eq_Int8 x y)
|
||||
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
Neg Int8 where
|
||||
(-) = prim__sub_Int8
|
||||
negate = prim__sub_Int8 0
|
||||
|
||||
Cast Int8 Bits8 where
|
||||
cast = prim__cast_Int8Bits8
|
||||
|
||||
Cast Int8 Bits16 where
|
||||
cast = prim__cast_Int8Bits16
|
||||
|
||||
Cast Int8 Bits32 where
|
||||
cast = prim__cast_Int8Bits32
|
||||
|
||||
Cast Int8 Bits64 where
|
||||
cast = prim__cast_Int8Bits64
|
||||
|
||||
Cast Int8 Int16 where
|
||||
cast = prim__cast_Int8Int16
|
||||
|
||||
Cast Int8 Int32 where
|
||||
cast = prim__cast_Int8Int32
|
||||
|
||||
Cast Int8 Int64 where
|
||||
cast = prim__cast_Int8Int64
|
||||
|
||||
Cast Int8 Int where
|
||||
cast = prim__cast_Int8Int
|
||||
|
||||
Cast Int8 Integer where
|
||||
cast = prim__cast_Int8Integer
|
||||
|
||||
Cast Int8 String where
|
||||
cast = prim__cast_Int8String
|
||||
|
||||
Cast Int8 Char where
|
||||
cast = prim__cast_Int8Char
|
||||
|
||||
Cast Int8 Double where
|
||||
cast = prim__cast_Int8Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
public export
|
||||
Eq Int16 where
|
||||
x == y = intToBool (prim__eq_Int16 x y)
|
||||
|
||||
Show Int16 where
|
||||
show = prim__cast_Int16String
|
||||
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
Neg Int16 where
|
||||
(-) = prim__sub_Int16
|
||||
negate = prim__sub_Int16 0
|
||||
|
||||
Cast Int16 Bits8 where
|
||||
cast = prim__cast_Int16Bits8
|
||||
|
||||
Cast Int16 Bits16 where
|
||||
cast = prim__cast_Int16Bits16
|
||||
|
||||
Cast Int16 Bits32 where
|
||||
cast = prim__cast_Int16Bits32
|
||||
|
||||
Cast Int16 Bits64 where
|
||||
cast = prim__cast_Int16Bits64
|
||||
|
||||
Cast Int16 Int8 where
|
||||
cast = prim__cast_Int16Int8
|
||||
|
||||
Cast Int16 Int32 where
|
||||
cast = prim__cast_Int16Int32
|
||||
|
||||
Cast Int16 Int64 where
|
||||
cast = prim__cast_Int16Int64
|
||||
|
||||
Cast Int16 Int where
|
||||
cast = prim__cast_Int16Int
|
||||
|
||||
Cast Int16 Integer where
|
||||
cast = prim__cast_Int16Integer
|
||||
|
||||
Cast Int16 String where
|
||||
cast = prim__cast_Int16String
|
||||
|
||||
Cast Int16 Char where
|
||||
cast = prim__cast_Int16Char
|
||||
|
||||
Cast Int16 Double where
|
||||
cast = prim__cast_Int16Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int32 where
|
||||
show = prim__cast_Int32String
|
||||
|
||||
public export
|
||||
Eq Int32 where
|
||||
x == y = intToBool (prim__eq_Int32 x y)
|
||||
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
Neg Int32 where
|
||||
(-) = prim__sub_Int32
|
||||
negate = prim__sub_Int32 0
|
||||
|
||||
Cast Int32 Bits8 where
|
||||
cast = prim__cast_Int32Bits8
|
||||
|
||||
Cast Int32 Bits16 where
|
||||
cast = prim__cast_Int32Bits16
|
||||
|
||||
Cast Int32 Bits32 where
|
||||
cast = prim__cast_Int32Bits32
|
||||
|
||||
Cast Int32 Bits64 where
|
||||
cast = prim__cast_Int32Bits64
|
||||
|
||||
Cast Int32 Int8 where
|
||||
cast = prim__cast_Int32Int8
|
||||
|
||||
Cast Int32 Int16 where
|
||||
cast = prim__cast_Int32Int16
|
||||
|
||||
Cast Int32 Int64 where
|
||||
cast = prim__cast_Int32Int64
|
||||
|
||||
Cast Int32 Int where
|
||||
cast = prim__cast_Int32Int
|
||||
|
||||
Cast Int32 Integer where
|
||||
cast = prim__cast_Int32Integer
|
||||
|
||||
Cast Int32 String where
|
||||
cast = prim__cast_Int32String
|
||||
|
||||
Cast Int32 Char where
|
||||
cast = prim__cast_Int32Char
|
||||
|
||||
Cast Int32 Double where
|
||||
cast = prim__cast_Int32Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int64 where
|
||||
show = prim__cast_Int64String
|
||||
|
||||
public export
|
||||
Eq Int64 where
|
||||
x == y = intToBool (prim__eq_Int64 x y)
|
||||
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
Neg Int64 where
|
||||
(-) = prim__sub_Int64
|
||||
negate = prim__sub_Int64 0
|
||||
|
||||
Cast Int64 Bits8 where
|
||||
cast = prim__cast_Int64Bits8
|
||||
|
||||
Cast Int64 Bits16 where
|
||||
cast = prim__cast_Int64Bits16
|
||||
|
||||
Cast Int64 Bits32 where
|
||||
cast = prim__cast_Int64Bits32
|
||||
|
||||
Cast Int64 Bits64 where
|
||||
cast = prim__cast_Int64Bits64
|
||||
|
||||
Cast Int64 Int8 where
|
||||
cast = prim__cast_Int64Int8
|
||||
|
||||
Cast Int64 Int16 where
|
||||
cast = prim__cast_Int64Int16
|
||||
|
||||
Cast Int64 Int32 where
|
||||
cast = prim__cast_Int64Int32
|
||||
|
||||
Cast Int64 Int where
|
||||
cast = prim__cast_Int64Int
|
||||
|
||||
Cast Int64 Integer where
|
||||
cast = prim__cast_Int64Integer
|
||||
|
||||
Cast Int64 String where
|
||||
cast = prim__cast_Int64String
|
||||
|
||||
Cast Int64 Char where
|
||||
cast = prim__cast_Int64Char
|
||||
|
||||
Cast Int64 Double where
|
||||
cast = prim__cast_Int64Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Int Int8 where
|
||||
cast = prim__cast_IntInt8
|
||||
|
||||
Cast Int Int16 where
|
||||
cast = prim__cast_IntInt16
|
||||
|
||||
Cast Int Int32 where
|
||||
cast = prim__cast_IntInt32
|
||||
|
||||
Cast Int Int64 where
|
||||
cast = prim__cast_IntInt64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Integer
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Integer Int8 where
|
||||
cast = prim__cast_IntegerInt8
|
||||
|
||||
Cast Integer Int16 where
|
||||
cast = prim__cast_IntegerInt16
|
||||
|
||||
Cast Integer Int32 where
|
||||
cast = prim__cast_IntegerInt32
|
||||
|
||||
Cast Integer Int64 where
|
||||
cast = prim__cast_IntegerInt64
|
||||
|
||||
Cast Integer Char where
|
||||
cast = prim__cast_IntegerChar
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Bits8 Int8 where
|
||||
cast = prim__cast_Bits8Int8
|
||||
|
||||
Cast Bits8 Int16 where
|
||||
cast = prim__cast_Bits8Int16
|
||||
|
||||
Cast Bits8 Int32 where
|
||||
cast = prim__cast_Bits8Int32
|
||||
|
||||
Cast Bits8 Int64 where
|
||||
cast = prim__cast_Bits8Int64
|
||||
|
||||
Cast Bits8 String where
|
||||
cast = prim__cast_Bits8String
|
||||
|
||||
Cast Bits8 Char where
|
||||
cast = prim__cast_Bits8Char
|
||||
|
||||
Cast Bits8 Double where
|
||||
cast = prim__cast_Bits8Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Bits16 Int8 where
|
||||
cast = prim__cast_Bits16Int8
|
||||
|
||||
Cast Bits16 Int16 where
|
||||
cast = prim__cast_Bits16Int16
|
||||
|
||||
Cast Bits16 Int32 where
|
||||
cast = prim__cast_Bits16Int32
|
||||
|
||||
Cast Bits16 Int64 where
|
||||
cast = prim__cast_Bits16Int64
|
||||
|
||||
Cast Bits16 String where
|
||||
cast = prim__cast_Bits16String
|
||||
|
||||
Cast Bits16 Char where
|
||||
cast = prim__cast_Bits16Char
|
||||
|
||||
Cast Bits16 Double where
|
||||
cast = prim__cast_Bits16Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Bits32 Int8 where
|
||||
cast = prim__cast_Bits32Int8
|
||||
|
||||
Cast Bits32 Int16 where
|
||||
cast = prim__cast_Bits32Int16
|
||||
|
||||
Cast Bits32 Int32 where
|
||||
cast = prim__cast_Bits32Int32
|
||||
|
||||
Cast Bits32 Int64 where
|
||||
cast = prim__cast_Bits32Int64
|
||||
|
||||
Cast Bits32 String where
|
||||
cast = prim__cast_Bits32String
|
||||
|
||||
Cast Bits32 Char where
|
||||
cast = prim__cast_Bits32Char
|
||||
|
||||
Cast Bits32 Double where
|
||||
cast = prim__cast_Bits32Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Bits64 Int8 where
|
||||
cast = prim__cast_Bits64Int8
|
||||
|
||||
Cast Bits64 Int16 where
|
||||
cast = prim__cast_Bits64Int16
|
||||
|
||||
Cast Bits64 Int32 where
|
||||
cast = prim__cast_Bits64Int32
|
||||
|
||||
Cast Bits64 Int64 where
|
||||
cast = prim__cast_Bits64Int64
|
||||
|
||||
Cast Bits64 String where
|
||||
cast = prim__cast_Bits64String
|
||||
|
||||
Cast Bits64 Char where
|
||||
cast = prim__cast_Bits64Char
|
||||
|
||||
Cast Bits64 Double where
|
||||
cast = prim__cast_Bits64Double
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- String
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast String Bits8 where
|
||||
cast = prim__cast_StringBits8
|
||||
|
||||
Cast String Bits16 where
|
||||
cast = prim__cast_StringBits16
|
||||
|
||||
Cast String Bits32 where
|
||||
cast = prim__cast_StringBits32
|
||||
|
||||
Cast String Bits64 where
|
||||
cast = prim__cast_StringBits64
|
||||
|
||||
Cast String Int8 where
|
||||
cast = prim__cast_StringInt8
|
||||
|
||||
Cast String Int16 where
|
||||
cast = prim__cast_StringInt16
|
||||
|
||||
Cast String Int32 where
|
||||
cast = prim__cast_StringInt32
|
||||
|
||||
Cast String Int64 where
|
||||
cast = prim__cast_StringInt64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Double
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Cast Double Bits8 where
|
||||
cast = prim__cast_DoubleBits8
|
||||
|
||||
Cast Double Bits16 where
|
||||
cast = prim__cast_DoubleBits16
|
||||
|
||||
Cast Double Bits32 where
|
||||
cast = prim__cast_DoubleBits32
|
||||
|
||||
Cast Double Bits64 where
|
||||
cast = prim__cast_DoubleBits64
|
||||
|
||||
Cast Double Int8 where
|
||||
cast = prim__cast_DoubleInt8
|
||||
|
||||
Cast Double Int16 where
|
||||
cast = prim__cast_DoubleInt16
|
||||
|
||||
Cast Double Int32 where
|
||||
cast = prim__cast_DoubleInt32
|
||||
|
||||
Cast Double Int64 where
|
||||
cast = prim__cast_DoubleInt64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Tests
|
||||
--------------------------------------------------------------------------------
|
||||
|
@ -1,94 +0,0 @@
|
||||
module IntBits
|
||||
|
||||
import IntCasts
|
||||
import IntEqOrd
|
||||
import IntNum
|
||||
|
||||
import Data.Bits
|
||||
import Data.DPair
|
||||
|
||||
-- This file to be deleted once these interfaces are added to base
|
||||
|
||||
public export %inline
|
||||
Bits Int8 where
|
||||
Index = Subset Nat (`LT` 8)
|
||||
(.&.) = prim__and_Int8
|
||||
(.|.) = prim__or_Int8
|
||||
xor = prim__xor_Int8
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int8 x . cast . fst
|
||||
shiftL x = prim__shl_Int8 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int16 where
|
||||
Index = Subset Nat (`LT` 16)
|
||||
(.&.) = prim__and_Int16
|
||||
(.|.) = prim__or_Int16
|
||||
xor = prim__xor_Int16
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int16 x . cast . fst
|
||||
shiftL x = prim__shl_Int16 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int32 where
|
||||
Index = Subset Nat (`LT` 32)
|
||||
(.&.) = prim__and_Int32
|
||||
(.|.) = prim__or_Int32
|
||||
xor = prim__xor_Int32
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int32 x . cast . fst
|
||||
shiftL x = prim__shl_Int32 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
(.&.) = prim__and_Int64
|
||||
(.|.) = prim__or_Int64
|
||||
xor = prim__xor_Int64
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int64 x . cast . fst
|
||||
shiftL x = prim__shl_Int64 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Bits64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
(.&.) = prim__and_Bits64
|
||||
(.|.) = prim__or_Bits64
|
||||
xor = prim__xor_Bits64
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits64 x . fromInteger . cast . fst
|
||||
shiftL x = prim__shl_Bits64 x . fromInteger . cast . fst
|
||||
complement = xor 0xffffffffffffffff
|
||||
oneBits = 0xffffffffffffffff
|
||||
|
||||
public export %inline
|
||||
Bits Integer where
|
||||
Index = Nat
|
||||
(.&.) = prim__and_Integer
|
||||
(.|.) = prim__or_Integer
|
||||
xor = prim__xor_Integer
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Integer x . cast
|
||||
shiftL x = prim__shl_Integer x . cast
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
@ -1,383 +0,0 @@
|
||||
module IntCasts
|
||||
|
||||
import IntNum
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
-- To String
|
||||
|
||||
export
|
||||
Cast Int8 String where
|
||||
cast = prim__cast_Int8String
|
||||
|
||||
export
|
||||
Cast Int16 String where
|
||||
cast = prim__cast_Int16String
|
||||
|
||||
export
|
||||
Cast Int32 String where
|
||||
cast = prim__cast_Int32String
|
||||
|
||||
export
|
||||
Cast Int64 String where
|
||||
cast = prim__cast_Int64String
|
||||
|
||||
-- To Integer
|
||||
|
||||
export
|
||||
Cast Int8 Integer where
|
||||
cast = prim__cast_Int8Integer
|
||||
|
||||
export
|
||||
Cast Int16 Integer where
|
||||
cast = prim__cast_Int16Integer
|
||||
|
||||
export
|
||||
Cast Int32 Integer where
|
||||
cast = prim__cast_Int32Integer
|
||||
|
||||
export
|
||||
Cast Int64 Integer where
|
||||
cast = prim__cast_Int64Integer
|
||||
|
||||
-- To Int
|
||||
|
||||
export
|
||||
Cast Int8 Int where
|
||||
cast = prim__cast_Int8Int
|
||||
|
||||
export
|
||||
Cast Int16 Int where
|
||||
cast = prim__cast_Int16Int
|
||||
|
||||
export
|
||||
Cast Int32 Int where
|
||||
cast = prim__cast_Int32Int
|
||||
|
||||
export
|
||||
Cast Int64 Int where
|
||||
cast = prim__cast_Int64Int
|
||||
|
||||
-- To Int8
|
||||
|
||||
export
|
||||
Cast Int Int8 where
|
||||
cast = prim__cast_IntInt8
|
||||
|
||||
export
|
||||
Cast Int16 Int8 where
|
||||
cast = prim__cast_Int16Int8
|
||||
|
||||
export
|
||||
Cast Int32 Int8 where
|
||||
cast = prim__cast_Int32Int8
|
||||
|
||||
export
|
||||
Cast Int64 Int8 where
|
||||
cast = prim__cast_Int64Int8
|
||||
|
||||
export
|
||||
Cast Integer Int8 where
|
||||
cast = prim__cast_IntegerInt8
|
||||
|
||||
export
|
||||
Cast Char Int8 where
|
||||
cast = prim__cast_CharInt8
|
||||
|
||||
export
|
||||
Cast Double Int8 where
|
||||
cast = prim__cast_DoubleInt8
|
||||
|
||||
export
|
||||
Cast String Int8 where
|
||||
cast = prim__cast_StringInt8
|
||||
|
||||
export
|
||||
Cast Nat Int8 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int8 where
|
||||
cast = prim__cast_Bits8Int8
|
||||
|
||||
export
|
||||
Cast Bits16 Int8 where
|
||||
cast = prim__cast_Bits16Int8
|
||||
|
||||
export
|
||||
Cast Bits32 Int8 where
|
||||
cast = prim__cast_Bits32Int8
|
||||
|
||||
export
|
||||
Cast Bits64 Int8 where
|
||||
cast = prim__cast_Bits64Int8
|
||||
|
||||
-- To Int16
|
||||
|
||||
export
|
||||
Cast Int Int16 where
|
||||
cast = prim__cast_IntInt16
|
||||
|
||||
export
|
||||
Cast Int8 Int16 where
|
||||
cast = prim__cast_Int8Int16
|
||||
|
||||
export
|
||||
Cast Int32 Int16 where
|
||||
cast = prim__cast_Int32Int16
|
||||
|
||||
export
|
||||
Cast Int64 Int16 where
|
||||
cast = prim__cast_Int64Int16
|
||||
|
||||
export
|
||||
Cast Integer Int16 where
|
||||
cast = prim__cast_IntegerInt16
|
||||
|
||||
export
|
||||
Cast Char Int16 where
|
||||
cast = prim__cast_CharInt16
|
||||
|
||||
export
|
||||
Cast Double Int16 where
|
||||
cast = prim__cast_DoubleInt16
|
||||
|
||||
export
|
||||
Cast String Int16 where
|
||||
cast = prim__cast_StringInt16
|
||||
|
||||
export
|
||||
Cast Nat Int16 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int16 where
|
||||
cast = prim__cast_Bits8Int16
|
||||
|
||||
export
|
||||
Cast Bits16 Int16 where
|
||||
cast = prim__cast_Bits16Int16
|
||||
|
||||
export
|
||||
Cast Bits32 Int16 where
|
||||
cast = prim__cast_Bits32Int16
|
||||
|
||||
export
|
||||
Cast Bits64 Int16 where
|
||||
cast = prim__cast_Bits64Int16
|
||||
|
||||
-- To Int32
|
||||
|
||||
export
|
||||
Cast Int Int32 where
|
||||
cast = prim__cast_IntInt32
|
||||
|
||||
export
|
||||
Cast Int8 Int32 where
|
||||
cast = prim__cast_Int8Int32
|
||||
|
||||
export
|
||||
Cast Int16 Int32 where
|
||||
cast = prim__cast_Int16Int32
|
||||
|
||||
export
|
||||
Cast Int64 Int32 where
|
||||
cast = prim__cast_Int64Int32
|
||||
|
||||
export
|
||||
Cast Integer Int32 where
|
||||
cast = prim__cast_IntegerInt32
|
||||
|
||||
export
|
||||
Cast Char Int32 where
|
||||
cast = prim__cast_CharInt32
|
||||
|
||||
export
|
||||
Cast Double Int32 where
|
||||
cast = prim__cast_DoubleInt32
|
||||
|
||||
export
|
||||
Cast String Int32 where
|
||||
cast = prim__cast_StringInt32
|
||||
|
||||
export
|
||||
Cast Nat Int32 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int32 where
|
||||
cast = prim__cast_Bits8Int32
|
||||
|
||||
export
|
||||
Cast Bits16 Int32 where
|
||||
cast = prim__cast_Bits16Int32
|
||||
|
||||
export
|
||||
Cast Bits32 Int32 where
|
||||
cast = prim__cast_Bits32Int32
|
||||
|
||||
export
|
||||
Cast Bits64 Int32 where
|
||||
cast = prim__cast_Bits64Int32
|
||||
|
||||
-- To Int64
|
||||
|
||||
export
|
||||
Cast Int Int64 where
|
||||
cast = prim__cast_IntInt64
|
||||
|
||||
export
|
||||
Cast Int8 Int64 where
|
||||
cast = prim__cast_Int8Int64
|
||||
|
||||
export
|
||||
Cast Int16 Int64 where
|
||||
cast = prim__cast_Int16Int64
|
||||
|
||||
export
|
||||
Cast Int32 Int64 where
|
||||
cast = prim__cast_Int32Int64
|
||||
|
||||
export
|
||||
Cast Integer Int64 where
|
||||
cast = prim__cast_IntegerInt64
|
||||
|
||||
export
|
||||
Cast Char Int64 where
|
||||
cast = prim__cast_CharInt64
|
||||
|
||||
export
|
||||
Cast Double Int64 where
|
||||
cast = prim__cast_DoubleInt64
|
||||
|
||||
export
|
||||
Cast String Int64 where
|
||||
cast = prim__cast_StringInt64
|
||||
|
||||
export
|
||||
Cast Nat Int64 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int64 where
|
||||
cast = prim__cast_Bits8Int64
|
||||
|
||||
export
|
||||
Cast Bits16 Int64 where
|
||||
cast = prim__cast_Bits16Int64
|
||||
|
||||
export
|
||||
Cast Bits32 Int64 where
|
||||
cast = prim__cast_Bits32Int64
|
||||
|
||||
export
|
||||
Cast Bits64 Int64 where
|
||||
cast = prim__cast_Bits64Int64
|
||||
|
||||
-- To Char
|
||||
|
||||
export
|
||||
Cast Int8 Char where
|
||||
cast = prim__cast_Int8Char
|
||||
|
||||
export
|
||||
Cast Int16 Char where
|
||||
cast = prim__cast_Int16Char
|
||||
|
||||
export
|
||||
Cast Int32 Char where
|
||||
cast = prim__cast_Int32Char
|
||||
|
||||
export
|
||||
Cast Int64 Char where
|
||||
cast = prim__cast_Int64Char
|
||||
|
||||
-- To Double
|
||||
|
||||
export
|
||||
Cast Int8 Double where
|
||||
cast = prim__cast_Int8Double
|
||||
|
||||
export
|
||||
Cast Int16 Double where
|
||||
cast = prim__cast_Int16Double
|
||||
|
||||
export
|
||||
Cast Int32 Double where
|
||||
cast = prim__cast_Int32Double
|
||||
|
||||
export
|
||||
Cast Int64 Double where
|
||||
cast = prim__cast_Int64Double
|
||||
|
||||
-- To Bits8
|
||||
|
||||
export
|
||||
Cast Int8 Bits8 where
|
||||
cast = prim__cast_Int8Bits8
|
||||
|
||||
export
|
||||
Cast Int16 Bits8 where
|
||||
cast = prim__cast_Int16Bits8
|
||||
|
||||
export
|
||||
Cast Int32 Bits8 where
|
||||
cast = prim__cast_Int32Bits8
|
||||
|
||||
export
|
||||
Cast Int64 Bits8 where
|
||||
cast = prim__cast_Int64Bits8
|
||||
|
||||
-- To Bits16
|
||||
|
||||
export
|
||||
Cast Int8 Bits16 where
|
||||
cast = prim__cast_Int8Bits16
|
||||
|
||||
export
|
||||
Cast Int16 Bits16 where
|
||||
cast = prim__cast_Int16Bits16
|
||||
|
||||
export
|
||||
Cast Int32 Bits16 where
|
||||
cast = prim__cast_Int32Bits16
|
||||
|
||||
export
|
||||
Cast Int64 Bits16 where
|
||||
cast = prim__cast_Int64Bits16
|
||||
|
||||
-- To Bits32
|
||||
|
||||
export
|
||||
Cast Int8 Bits32 where
|
||||
cast = prim__cast_Int8Bits32
|
||||
|
||||
export
|
||||
Cast Int16 Bits32 where
|
||||
cast = prim__cast_Int16Bits32
|
||||
|
||||
export
|
||||
Cast Int32 Bits32 where
|
||||
cast = prim__cast_Int32Bits32
|
||||
|
||||
export
|
||||
Cast Int64 Bits32 where
|
||||
cast = prim__cast_Int64Bits32
|
||||
|
||||
-- To Bits64
|
||||
|
||||
export
|
||||
Cast Int8 Bits64 where
|
||||
cast = prim__cast_Int8Bits64
|
||||
|
||||
export
|
||||
Cast Int16 Bits64 where
|
||||
cast = prim__cast_Int16Bits64
|
||||
|
||||
export
|
||||
Cast Int32 Bits64 where
|
||||
cast = prim__cast_Int32Bits64
|
||||
|
||||
export
|
||||
Cast Int64 Bits64 where
|
||||
cast = prim__cast_Int64Bits64
|
@ -1,53 +0,0 @@
|
||||
module IntEqOrd
|
||||
|
||||
public export
|
||||
Eq Int8 where
|
||||
x == y = intToBool (prim__eq_Int8 x y)
|
||||
|
||||
public export
|
||||
Eq Int16 where
|
||||
x == y = intToBool (prim__eq_Int16 x y)
|
||||
|
||||
public export
|
||||
Eq Int32 where
|
||||
x == y = intToBool (prim__eq_Int32 x y)
|
||||
|
||||
public export
|
||||
Eq Int64 where
|
||||
x == y = intToBool (prim__eq_Int64 x y)
|
||||
|
||||
public export
|
||||
Ord Int8 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int8 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int8 x y)
|
||||
(>) x y = intToBool (prim__gt_Int8 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int8 x y)
|
||||
|
||||
public export
|
||||
Ord Int16 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int16 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int16 x y)
|
||||
(>) x y = intToBool (prim__gt_Int16 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int16 x y)
|
||||
|
||||
public export
|
||||
Ord Int32 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int32 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int32 x y)
|
||||
(>) x y = intToBool (prim__gt_Int32 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int32 x y)
|
||||
|
||||
public export
|
||||
Ord Int64 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int64 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int64 x y)
|
||||
(>) x y = intToBool (prim__gt_Int64 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int64 x y)
|
@ -1,193 +0,0 @@
|
||||
module IntNum
|
||||
|
||||
import IntEqOrd
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
-- Int8
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
public export
|
||||
Neg Int8 where
|
||||
negate x = prim__sub_Int8 0 x
|
||||
(-) = prim__sub_Int8
|
||||
|
||||
public export
|
||||
Abs Int8 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int8 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int8 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int8 x y
|
||||
|
||||
-- Int16
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
public export
|
||||
Neg Int16 where
|
||||
negate x = prim__sub_Int16 0 x
|
||||
(-) = prim__sub_Int16
|
||||
|
||||
public export
|
||||
Abs Int16 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int16 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int16 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int16 x y
|
||||
|
||||
-- Int32
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
public export
|
||||
Neg Int32 where
|
||||
negate x = prim__sub_Int32 0 x
|
||||
(-) = prim__sub_Int32
|
||||
|
||||
public export
|
||||
Abs Int32 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int32 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int32 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int32 x y
|
||||
|
||||
-- Int64
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
public export
|
||||
Neg Int64 where
|
||||
negate x = prim__sub_Int64 0 x
|
||||
(-) = prim__sub_Int64
|
||||
|
||||
public export
|
||||
Abs Int64 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int64 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int64 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int64 x y
|
||||
|
||||
-- Bits8
|
||||
|
||||
public export
|
||||
Neg Bits8 where
|
||||
negate x = prim__sub_Bits8 0 x
|
||||
(-) = prim__sub_Bits8
|
||||
|
||||
public export
|
||||
Abs Bits8 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits8 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits8 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits8 x y
|
||||
|
||||
-- Bits16
|
||||
|
||||
public export
|
||||
Neg Bits16 where
|
||||
negate x = prim__sub_Bits16 0 x
|
||||
(-) = prim__sub_Bits16
|
||||
|
||||
public export
|
||||
Abs Bits16 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits16 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits16 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits16 x y
|
||||
|
||||
-- Bits32
|
||||
|
||||
public export
|
||||
Neg Bits32 where
|
||||
negate x = prim__sub_Bits32 0 x
|
||||
(-) = prim__sub_Bits32
|
||||
|
||||
public export
|
||||
Abs Bits32 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits32 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits32 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits32 x y
|
||||
|
||||
-- Bits64
|
||||
|
||||
public export
|
||||
Neg Bits64 where
|
||||
negate x = prim__sub_Bits64 0 x
|
||||
(-) = prim__sub_Bits64
|
||||
|
||||
public export
|
||||
Abs Bits64 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits64 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits64 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits64 x y
|
@ -1,26 +0,0 @@
|
||||
module IntShow
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
firstCharIs : (Char -> Bool) -> String -> Bool
|
||||
firstCharIs p "" = False
|
||||
firstCharIs p str = p (assert_total (prim__strHead str))
|
||||
|
||||
primNumShow : (a -> String) -> Prec -> a -> String
|
||||
primNumShow f d x = let str = f x in showParens (d >= PrefixMinus && firstCharIs (== '-') str) str
|
||||
|
||||
export
|
||||
Show Int8 where
|
||||
showPrec = primNumShow prim__cast_Int8String
|
||||
|
||||
export
|
||||
Show Int16 where
|
||||
showPrec = primNumShow prim__cast_Int16String
|
||||
|
||||
export
|
||||
Show Int32 where
|
||||
showPrec = primNumShow prim__cast_Int32String
|
||||
|
||||
export
|
||||
Show Int64 where
|
||||
showPrec = primNumShow prim__cast_Int64String
|
@ -1,11 +1,5 @@
|
||||
module TestIntegers
|
||||
|
||||
import IntBits
|
||||
import IntCasts
|
||||
import IntEqOrd
|
||||
import IntNum
|
||||
import IntShow
|
||||
|
||||
import Data.Bits
|
||||
import Data.List.Quantifiers
|
||||
|
||||
|
@ -156,146 +156,6 @@ check (MkOp name op opInt allowZero $ MkIntType type signed bits mi ma) =
|
||||
then Nothing
|
||||
else Just #"Error when calculating \#{show x} \#{name} \#{show y} for \#{type}: Expected \#{show resMod} but got \#{show resA} (unrounded: \#{show resInteger})"#
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int8 where
|
||||
show = prim__cast_Int8String
|
||||
|
||||
Cast Int8 Integer where
|
||||
cast = prim__cast_Int8Integer
|
||||
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
Neg Int8 where
|
||||
(-) = prim__sub_Int8
|
||||
negate = prim__sub_Int8 0
|
||||
|
||||
Integral Int8 where
|
||||
div = prim__div_Int8
|
||||
mod = prim__mod_Int8
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int16 where
|
||||
show = prim__cast_Int16String
|
||||
|
||||
Cast Int16 Integer where
|
||||
cast = prim__cast_Int16Integer
|
||||
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
Neg Int16 where
|
||||
(-) = prim__sub_Int16
|
||||
negate = prim__sub_Int16 0
|
||||
|
||||
Integral Int16 where
|
||||
div = prim__div_Int16
|
||||
mod = prim__mod_Int16
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int32 where
|
||||
show = prim__cast_Int32String
|
||||
|
||||
Cast Int32 Integer where
|
||||
cast = prim__cast_Int32Integer
|
||||
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
Neg Int32 where
|
||||
(-) = prim__sub_Int32
|
||||
negate = prim__sub_Int32 0
|
||||
|
||||
Integral Int32 where
|
||||
div = prim__div_Int32
|
||||
mod = prim__mod_Int32
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Int64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Show Int64 where
|
||||
show = prim__cast_Int64String
|
||||
|
||||
Cast Int64 Integer where
|
||||
cast = prim__cast_Int64Integer
|
||||
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
Neg Int64 where
|
||||
(-) = prim__sub_Int64
|
||||
negate = prim__sub_Int64 0
|
||||
|
||||
Integral Int64 where
|
||||
div = prim__div_Int64
|
||||
mod = prim__mod_Int64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits8
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Neg Bits8 where
|
||||
(-) = prim__sub_Bits8
|
||||
negate = prim__sub_Bits8 0
|
||||
|
||||
Integral Bits8 where
|
||||
div = prim__div_Bits8
|
||||
mod = prim__mod_Bits8
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits16
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Neg Bits16 where
|
||||
(-) = prim__sub_Bits16
|
||||
negate = prim__sub_Bits16 0
|
||||
|
||||
Integral Bits16 where
|
||||
div = prim__div_Bits16
|
||||
mod = prim__mod_Bits16
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits32
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Neg Bits32 where
|
||||
(-) = prim__sub_Bits32
|
||||
negate = prim__sub_Bits32 0
|
||||
|
||||
Integral Bits32 where
|
||||
div = prim__div_Bits32
|
||||
mod = prim__mod_Bits32
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Bits64
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
Neg Bits64 where
|
||||
(-) = prim__sub_Bits64
|
||||
negate = prim__sub_Bits64 0
|
||||
|
||||
Integral Bits64 where
|
||||
div = prim__div_Bits64
|
||||
mod = prim__mod_Bits64
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Main
|
||||
--------------------------------------------------------------------------------
|
||||
|
@ -1,94 +0,0 @@
|
||||
module IntBits
|
||||
|
||||
import IntCasts
|
||||
import IntEqOrd
|
||||
import IntNum
|
||||
|
||||
import Data.Bits
|
||||
import Data.DPair
|
||||
|
||||
-- This file to be deleted once these interfaces are added to base
|
||||
|
||||
public export %inline
|
||||
Bits Int8 where
|
||||
Index = Subset Nat (`LT` 8)
|
||||
(.&.) = prim__and_Int8
|
||||
(.|.) = prim__or_Int8
|
||||
xor = prim__xor_Int8
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int8 x . cast . fst
|
||||
shiftL x = prim__shl_Int8 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int16 where
|
||||
Index = Subset Nat (`LT` 16)
|
||||
(.&.) = prim__and_Int16
|
||||
(.|.) = prim__or_Int16
|
||||
xor = prim__xor_Int16
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int16 x . cast . fst
|
||||
shiftL x = prim__shl_Int16 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int32 where
|
||||
Index = Subset Nat (`LT` 32)
|
||||
(.&.) = prim__and_Int32
|
||||
(.|.) = prim__or_Int32
|
||||
xor = prim__xor_Int32
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int32 x . cast . fst
|
||||
shiftL x = prim__shl_Int32 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Int64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
(.&.) = prim__and_Int64
|
||||
(.|.) = prim__or_Int64
|
||||
xor = prim__xor_Int64
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int64 x . cast . fst
|
||||
shiftL x = prim__shl_Int64 x . cast . fst
|
||||
complement = xor (-1)
|
||||
oneBits = -1
|
||||
|
||||
public export %inline
|
||||
Bits Bits64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
(.&.) = prim__and_Bits64
|
||||
(.|.) = prim__or_Bits64
|
||||
xor = prim__xor_Bits64
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits64 x . fromInteger . cast . fst
|
||||
shiftL x = prim__shl_Bits64 x . fromInteger . cast . fst
|
||||
complement = xor 0xffffffffffffffff
|
||||
oneBits = 0xffffffffffffffff
|
||||
|
||||
public export %inline
|
||||
Bits Integer where
|
||||
Index = Nat
|
||||
(.&.) = prim__and_Integer
|
||||
(.|.) = prim__or_Integer
|
||||
xor = prim__xor_Integer
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Integer x . cast
|
||||
shiftL x = prim__shl_Integer x . cast
|
||||
complement = xor (-1)
|
||||
oneBits = (-1)
|
@ -1,383 +0,0 @@
|
||||
module IntCasts
|
||||
|
||||
import IntNum
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
-- To String
|
||||
|
||||
export
|
||||
Cast Int8 String where
|
||||
cast = prim__cast_Int8String
|
||||
|
||||
export
|
||||
Cast Int16 String where
|
||||
cast = prim__cast_Int16String
|
||||
|
||||
export
|
||||
Cast Int32 String where
|
||||
cast = prim__cast_Int32String
|
||||
|
||||
export
|
||||
Cast Int64 String where
|
||||
cast = prim__cast_Int64String
|
||||
|
||||
-- To Integer
|
||||
|
||||
export
|
||||
Cast Int8 Integer where
|
||||
cast = prim__cast_Int8Integer
|
||||
|
||||
export
|
||||
Cast Int16 Integer where
|
||||
cast = prim__cast_Int16Integer
|
||||
|
||||
export
|
||||
Cast Int32 Integer where
|
||||
cast = prim__cast_Int32Integer
|
||||
|
||||
export
|
||||
Cast Int64 Integer where
|
||||
cast = prim__cast_Int64Integer
|
||||
|
||||
-- To Int
|
||||
|
||||
export
|
||||
Cast Int8 Int where
|
||||
cast = prim__cast_Int8Int
|
||||
|
||||
export
|
||||
Cast Int16 Int where
|
||||
cast = prim__cast_Int16Int
|
||||
|
||||
export
|
||||
Cast Int32 Int where
|
||||
cast = prim__cast_Int32Int
|
||||
|
||||
export
|
||||
Cast Int64 Int where
|
||||
cast = prim__cast_Int64Int
|
||||
|
||||
-- To Int8
|
||||
|
||||
export
|
||||
Cast Int Int8 where
|
||||
cast = prim__cast_IntInt8
|
||||
|
||||
export
|
||||
Cast Int16 Int8 where
|
||||
cast = prim__cast_Int16Int8
|
||||
|
||||
export
|
||||
Cast Int32 Int8 where
|
||||
cast = prim__cast_Int32Int8
|
||||
|
||||
export
|
||||
Cast Int64 Int8 where
|
||||
cast = prim__cast_Int64Int8
|
||||
|
||||
export
|
||||
Cast Integer Int8 where
|
||||
cast = prim__cast_IntegerInt8
|
||||
|
||||
export
|
||||
Cast Char Int8 where
|
||||
cast = prim__cast_CharInt8
|
||||
|
||||
export
|
||||
Cast Double Int8 where
|
||||
cast = prim__cast_DoubleInt8
|
||||
|
||||
export
|
||||
Cast String Int8 where
|
||||
cast = prim__cast_StringInt8
|
||||
|
||||
export
|
||||
Cast Nat Int8 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int8 where
|
||||
cast = prim__cast_Bits8Int8
|
||||
|
||||
export
|
||||
Cast Bits16 Int8 where
|
||||
cast = prim__cast_Bits16Int8
|
||||
|
||||
export
|
||||
Cast Bits32 Int8 where
|
||||
cast = prim__cast_Bits32Int8
|
||||
|
||||
export
|
||||
Cast Bits64 Int8 where
|
||||
cast = prim__cast_Bits64Int8
|
||||
|
||||
-- To Int16
|
||||
|
||||
export
|
||||
Cast Int Int16 where
|
||||
cast = prim__cast_IntInt16
|
||||
|
||||
export
|
||||
Cast Int8 Int16 where
|
||||
cast = prim__cast_Int8Int16
|
||||
|
||||
export
|
||||
Cast Int32 Int16 where
|
||||
cast = prim__cast_Int32Int16
|
||||
|
||||
export
|
||||
Cast Int64 Int16 where
|
||||
cast = prim__cast_Int64Int16
|
||||
|
||||
export
|
||||
Cast Integer Int16 where
|
||||
cast = prim__cast_IntegerInt16
|
||||
|
||||
export
|
||||
Cast Char Int16 where
|
||||
cast = prim__cast_CharInt16
|
||||
|
||||
export
|
||||
Cast Double Int16 where
|
||||
cast = prim__cast_DoubleInt16
|
||||
|
||||
export
|
||||
Cast String Int16 where
|
||||
cast = prim__cast_StringInt16
|
||||
|
||||
export
|
||||
Cast Nat Int16 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int16 where
|
||||
cast = prim__cast_Bits8Int16
|
||||
|
||||
export
|
||||
Cast Bits16 Int16 where
|
||||
cast = prim__cast_Bits16Int16
|
||||
|
||||
export
|
||||
Cast Bits32 Int16 where
|
||||
cast = prim__cast_Bits32Int16
|
||||
|
||||
export
|
||||
Cast Bits64 Int16 where
|
||||
cast = prim__cast_Bits64Int16
|
||||
|
||||
-- To Int32
|
||||
|
||||
export
|
||||
Cast Int Int32 where
|
||||
cast = prim__cast_IntInt32
|
||||
|
||||
export
|
||||
Cast Int8 Int32 where
|
||||
cast = prim__cast_Int8Int32
|
||||
|
||||
export
|
||||
Cast Int16 Int32 where
|
||||
cast = prim__cast_Int16Int32
|
||||
|
||||
export
|
||||
Cast Int64 Int32 where
|
||||
cast = prim__cast_Int64Int32
|
||||
|
||||
export
|
||||
Cast Integer Int32 where
|
||||
cast = prim__cast_IntegerInt32
|
||||
|
||||
export
|
||||
Cast Char Int32 where
|
||||
cast = prim__cast_CharInt32
|
||||
|
||||
export
|
||||
Cast Double Int32 where
|
||||
cast = prim__cast_DoubleInt32
|
||||
|
||||
export
|
||||
Cast String Int32 where
|
||||
cast = prim__cast_StringInt32
|
||||
|
||||
export
|
||||
Cast Nat Int32 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int32 where
|
||||
cast = prim__cast_Bits8Int32
|
||||
|
||||
export
|
||||
Cast Bits16 Int32 where
|
||||
cast = prim__cast_Bits16Int32
|
||||
|
||||
export
|
||||
Cast Bits32 Int32 where
|
||||
cast = prim__cast_Bits32Int32
|
||||
|
||||
export
|
||||
Cast Bits64 Int32 where
|
||||
cast = prim__cast_Bits64Int32
|
||||
|
||||
-- To Int64
|
||||
|
||||
export
|
||||
Cast Int Int64 where
|
||||
cast = prim__cast_IntInt64
|
||||
|
||||
export
|
||||
Cast Int8 Int64 where
|
||||
cast = prim__cast_Int8Int64
|
||||
|
||||
export
|
||||
Cast Int16 Int64 where
|
||||
cast = prim__cast_Int16Int64
|
||||
|
||||
export
|
||||
Cast Int32 Int64 where
|
||||
cast = prim__cast_Int32Int64
|
||||
|
||||
export
|
||||
Cast Integer Int64 where
|
||||
cast = prim__cast_IntegerInt64
|
||||
|
||||
export
|
||||
Cast Char Int64 where
|
||||
cast = prim__cast_CharInt64
|
||||
|
||||
export
|
||||
Cast Double Int64 where
|
||||
cast = prim__cast_DoubleInt64
|
||||
|
||||
export
|
||||
Cast String Int64 where
|
||||
cast = prim__cast_StringInt64
|
||||
|
||||
export
|
||||
Cast Nat Int64 where
|
||||
cast = fromInteger . natToInteger
|
||||
|
||||
export
|
||||
Cast Bits8 Int64 where
|
||||
cast = prim__cast_Bits8Int64
|
||||
|
||||
export
|
||||
Cast Bits16 Int64 where
|
||||
cast = prim__cast_Bits16Int64
|
||||
|
||||
export
|
||||
Cast Bits32 Int64 where
|
||||
cast = prim__cast_Bits32Int64
|
||||
|
||||
export
|
||||
Cast Bits64 Int64 where
|
||||
cast = prim__cast_Bits64Int64
|
||||
|
||||
-- To Char
|
||||
|
||||
export
|
||||
Cast Int8 Char where
|
||||
cast = prim__cast_Int8Char
|
||||
|
||||
export
|
||||
Cast Int16 Char where
|
||||
cast = prim__cast_Int16Char
|
||||
|
||||
export
|
||||
Cast Int32 Char where
|
||||
cast = prim__cast_Int32Char
|
||||
|
||||
export
|
||||
Cast Int64 Char where
|
||||
cast = prim__cast_Int64Char
|
||||
|
||||
-- To Double
|
||||
|
||||
export
|
||||
Cast Int8 Double where
|
||||
cast = prim__cast_Int8Double
|
||||
|
||||
export
|
||||
Cast Int16 Double where
|
||||
cast = prim__cast_Int16Double
|
||||
|
||||
export
|
||||
Cast Int32 Double where
|
||||
cast = prim__cast_Int32Double
|
||||
|
||||
export
|
||||
Cast Int64 Double where
|
||||
cast = prim__cast_Int64Double
|
||||
|
||||
-- To Bits8
|
||||
|
||||
export
|
||||
Cast Int8 Bits8 where
|
||||
cast = prim__cast_Int8Bits8
|
||||
|
||||
export
|
||||
Cast Int16 Bits8 where
|
||||
cast = prim__cast_Int16Bits8
|
||||
|
||||
export
|
||||
Cast Int32 Bits8 where
|
||||
cast = prim__cast_Int32Bits8
|
||||
|
||||
export
|
||||
Cast Int64 Bits8 where
|
||||
cast = prim__cast_Int64Bits8
|
||||
|
||||
-- To Bits16
|
||||
|
||||
export
|
||||
Cast Int8 Bits16 where
|
||||
cast = prim__cast_Int8Bits16
|
||||
|
||||
export
|
||||
Cast Int16 Bits16 where
|
||||
cast = prim__cast_Int16Bits16
|
||||
|
||||
export
|
||||
Cast Int32 Bits16 where
|
||||
cast = prim__cast_Int32Bits16
|
||||
|
||||
export
|
||||
Cast Int64 Bits16 where
|
||||
cast = prim__cast_Int64Bits16
|
||||
|
||||
-- To Bits32
|
||||
|
||||
export
|
||||
Cast Int8 Bits32 where
|
||||
cast = prim__cast_Int8Bits32
|
||||
|
||||
export
|
||||
Cast Int16 Bits32 where
|
||||
cast = prim__cast_Int16Bits32
|
||||
|
||||
export
|
||||
Cast Int32 Bits32 where
|
||||
cast = prim__cast_Int32Bits32
|
||||
|
||||
export
|
||||
Cast Int64 Bits32 where
|
||||
cast = prim__cast_Int64Bits32
|
||||
|
||||
-- To Bits64
|
||||
|
||||
export
|
||||
Cast Int8 Bits64 where
|
||||
cast = prim__cast_Int8Bits64
|
||||
|
||||
export
|
||||
Cast Int16 Bits64 where
|
||||
cast = prim__cast_Int16Bits64
|
||||
|
||||
export
|
||||
Cast Int32 Bits64 where
|
||||
cast = prim__cast_Int32Bits64
|
||||
|
||||
export
|
||||
Cast Int64 Bits64 where
|
||||
cast = prim__cast_Int64Bits64
|
@ -1,53 +0,0 @@
|
||||
module IntEqOrd
|
||||
|
||||
public export
|
||||
Eq Int8 where
|
||||
x == y = intToBool (prim__eq_Int8 x y)
|
||||
|
||||
public export
|
||||
Eq Int16 where
|
||||
x == y = intToBool (prim__eq_Int16 x y)
|
||||
|
||||
public export
|
||||
Eq Int32 where
|
||||
x == y = intToBool (prim__eq_Int32 x y)
|
||||
|
||||
public export
|
||||
Eq Int64 where
|
||||
x == y = intToBool (prim__eq_Int64 x y)
|
||||
|
||||
public export
|
||||
Ord Int8 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int8 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int8 x y)
|
||||
(>) x y = intToBool (prim__gt_Int8 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int8 x y)
|
||||
|
||||
public export
|
||||
Ord Int16 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int16 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int16 x y)
|
||||
(>) x y = intToBool (prim__gt_Int16 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int16 x y)
|
||||
|
||||
public export
|
||||
Ord Int32 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int32 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int32 x y)
|
||||
(>) x y = intToBool (prim__gt_Int32 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int32 x y)
|
||||
|
||||
public export
|
||||
Ord Int64 where
|
||||
compare x y = if x < y then LT else if x == y then EQ else GT
|
||||
|
||||
(<) x y = intToBool (prim__lt_Int64 x y)
|
||||
(<=) x y = intToBool (prim__lte_Int64 x y)
|
||||
(>) x y = intToBool (prim__gt_Int64 x y)
|
||||
(>=) x y = intToBool (prim__gte_Int64 x y)
|
@ -1,193 +0,0 @@
|
||||
module IntNum
|
||||
|
||||
import IntEqOrd
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
-- Int8
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int8 where
|
||||
(+) = prim__add_Int8
|
||||
(*) = prim__mul_Int8
|
||||
fromInteger = prim__cast_IntegerInt8
|
||||
|
||||
public export
|
||||
Neg Int8 where
|
||||
negate x = prim__sub_Int8 0 x
|
||||
(-) = prim__sub_Int8
|
||||
|
||||
public export
|
||||
Abs Int8 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int8 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int8 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int8 x y
|
||||
|
||||
-- Int16
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int16 where
|
||||
(+) = prim__add_Int16
|
||||
(*) = prim__mul_Int16
|
||||
fromInteger = prim__cast_IntegerInt16
|
||||
|
||||
public export
|
||||
Neg Int16 where
|
||||
negate x = prim__sub_Int16 0 x
|
||||
(-) = prim__sub_Int16
|
||||
|
||||
public export
|
||||
Abs Int16 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int16 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int16 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int16 x y
|
||||
|
||||
-- Int32
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int32 where
|
||||
(+) = prim__add_Int32
|
||||
(*) = prim__mul_Int32
|
||||
fromInteger = prim__cast_IntegerInt32
|
||||
|
||||
public export
|
||||
Neg Int32 where
|
||||
negate x = prim__sub_Int32 0 x
|
||||
(-) = prim__sub_Int32
|
||||
|
||||
public export
|
||||
Abs Int32 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int32 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int32 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int32 x y
|
||||
|
||||
-- Int64
|
||||
|
||||
%inline
|
||||
public export
|
||||
Num Int64 where
|
||||
(+) = prim__add_Int64
|
||||
(*) = prim__mul_Int64
|
||||
fromInteger = prim__cast_IntegerInt64
|
||||
|
||||
public export
|
||||
Neg Int64 where
|
||||
negate x = prim__sub_Int64 0 x
|
||||
(-) = prim__sub_Int64
|
||||
|
||||
public export
|
||||
Abs Int64 where
|
||||
abs x = if x < 0 then -x else x
|
||||
|
||||
public export
|
||||
Integral Int64 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Int64 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Int64 x y
|
||||
|
||||
-- Bits8
|
||||
|
||||
public export
|
||||
Neg Bits8 where
|
||||
negate x = prim__sub_Bits8 0 x
|
||||
(-) = prim__sub_Bits8
|
||||
|
||||
public export
|
||||
Abs Bits8 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits8 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits8 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits8 x y
|
||||
|
||||
-- Bits16
|
||||
|
||||
public export
|
||||
Neg Bits16 where
|
||||
negate x = prim__sub_Bits16 0 x
|
||||
(-) = prim__sub_Bits16
|
||||
|
||||
public export
|
||||
Abs Bits16 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits16 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits16 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits16 x y
|
||||
|
||||
-- Bits32
|
||||
|
||||
public export
|
||||
Neg Bits32 where
|
||||
negate x = prim__sub_Bits32 0 x
|
||||
(-) = prim__sub_Bits32
|
||||
|
||||
public export
|
||||
Abs Bits32 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits32 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits32 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits32 x y
|
||||
|
||||
-- Bits64
|
||||
|
||||
public export
|
||||
Neg Bits64 where
|
||||
negate x = prim__sub_Bits64 0 x
|
||||
(-) = prim__sub_Bits64
|
||||
|
||||
public export
|
||||
Abs Bits64 where
|
||||
abs = id
|
||||
|
||||
public export
|
||||
Integral Bits64 where
|
||||
div x y
|
||||
= case y == 0 of
|
||||
False => prim__div_Bits64 x y
|
||||
mod x y
|
||||
= case y == 0 of
|
||||
False => prim__mod_Bits64 x y
|
@ -1,26 +0,0 @@
|
||||
module IntShow
|
||||
|
||||
-- This file to be deleted once these interfaces are added to prelude
|
||||
|
||||
firstCharIs : (Char -> Bool) -> String -> Bool
|
||||
firstCharIs p "" = False
|
||||
firstCharIs p str = p (assert_total (prim__strHead str))
|
||||
|
||||
primNumShow : (a -> String) -> Prec -> a -> String
|
||||
primNumShow f d x = let str = f x in showParens (d >= PrefixMinus && firstCharIs (== '-') str) str
|
||||
|
||||
export
|
||||
Show Int8 where
|
||||
showPrec = primNumShow prim__cast_Int8String
|
||||
|
||||
export
|
||||
Show Int16 where
|
||||
showPrec = primNumShow prim__cast_Int16String
|
||||
|
||||
export
|
||||
Show Int32 where
|
||||
showPrec = primNumShow prim__cast_Int32String
|
||||
|
||||
export
|
||||
Show Int64 where
|
||||
showPrec = primNumShow prim__cast_Int64String
|
@ -1,11 +1,5 @@
|
||||
module TestIntegers
|
||||
|
||||
import IntBits
|
||||
import IntCasts
|
||||
import IntEqOrd
|
||||
import IntNum
|
||||
import IntShow
|
||||
|
||||
import Data.Bits
|
||||
import Data.List.Quantifiers
|
||||
|
||||
|
Loading…
Reference in New Issue
Block a user