mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-27 13:40:15 +03:00
[ base ] Use Fin n
as index in Bits
(#2192)
This commit is contained in:
parent
d3aed0404c
commit
a09c5082c5
@ -333,7 +333,7 @@ now uses a safer type for the number of shifts:
|
||||
|
||||
randoms : Int -> Stream Int
|
||||
randoms seed = let seed' = 1664525 * seed + 1013904223 in
|
||||
(seed' `shiftR` fromNat 2) :: randoms seed'
|
||||
(seed' `shiftR` 2) :: randoms seed'
|
||||
|
||||
|
||||
In ``ArithCmd.idr``, update ``DivBy``, ``randoms``, and ``import Data.Bits``
|
||||
|
@ -1,7 +1,6 @@
|
||||
module Data.Bits
|
||||
|
||||
import Data.DPair
|
||||
import public Data.Nat
|
||||
import public Data.Fin
|
||||
|
||||
%default total
|
||||
|
||||
@ -10,19 +9,6 @@ infixl 7 .&.
|
||||
infixl 6 `xor`
|
||||
infixl 5 .|.
|
||||
|
||||
||| Utility for using bitwise operations at compile
|
||||
||| time.
|
||||
|||
|
||||
||| ```idris example
|
||||
||| the Bits8 13 `shiftL` fromNat 3
|
||||
||| ```
|
||||
export
|
||||
fromNat : (k : Nat)
|
||||
-> {n : Nat}
|
||||
-> {auto 0 prf : lt k n === True}
|
||||
-> Subset Nat (`LT` n)
|
||||
fromNat k = Element k (ltReflectsLT k n prf)
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Interface Bits
|
||||
--------------------------------------------------------------------------------
|
||||
@ -55,6 +41,7 @@ interface Bits a where
|
||||
|
||||
||| Returns the bitwise complement of a value.
|
||||
complement : a -> a
|
||||
complement = xor oneBits
|
||||
|
||||
||| The value with all bits set..
|
||||
oneBits : a
|
||||
@ -77,128 +64,119 @@ interface Bits a where
|
||||
|
||||
public export %inline
|
||||
Bits Bits8 where
|
||||
Index = Subset Nat (`LT` 8)
|
||||
Index = Fin 8
|
||||
(.&.) = prim__and_Bits8
|
||||
(.|.) = prim__or_Bits8
|
||||
xor = prim__xor_Bits8
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits8 x . cast . fst
|
||||
shiftL x = prim__shl_Bits8 x . cast . fst
|
||||
complement = xor 0xff
|
||||
shiftR x = prim__shr_Bits8 x . cast . finToNat
|
||||
shiftL x = prim__shl_Bits8 x . cast . finToNat
|
||||
oneBits = 0xff
|
||||
|
||||
public export %inline
|
||||
Bits Bits16 where
|
||||
Index = Subset Nat (`LT` 16)
|
||||
Index = Fin 16
|
||||
(.&.) = prim__and_Bits16
|
||||
(.|.) = prim__or_Bits16
|
||||
xor = prim__xor_Bits16
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits16 x . cast . fst
|
||||
shiftL x = prim__shl_Bits16 x . cast . fst
|
||||
complement = xor 0xffff
|
||||
shiftR x = prim__shr_Bits16 x . cast . finToNat
|
||||
shiftL x = prim__shl_Bits16 x . cast . finToNat
|
||||
oneBits = 0xffff
|
||||
|
||||
public export %inline
|
||||
Bits Bits32 where
|
||||
Index = Subset Nat (`LT` 32)
|
||||
Index = Fin 32
|
||||
(.&.) = prim__and_Bits32
|
||||
(.|.) = prim__or_Bits32
|
||||
xor = prim__xor_Bits32
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Bits32 x . cast . fst
|
||||
shiftL x = prim__shl_Bits32 x . cast . fst
|
||||
complement = xor 0xffffffff
|
||||
shiftR x = prim__shr_Bits32 x . cast . finToNat
|
||||
shiftL x = prim__shl_Bits32 x . cast . finToNat
|
||||
oneBits = 0xffffffff
|
||||
|
||||
public export %inline
|
||||
Bits Bits64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
Index = Fin 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
|
||||
shiftR x = prim__shr_Bits64 x . cast . finToNat
|
||||
shiftL x = prim__shl_Bits64 x . cast . finToNat
|
||||
oneBits = 0xffffffffffffffff
|
||||
|
||||
public export %inline
|
||||
Bits Int where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
Index = Fin 64
|
||||
(.&.) = prim__and_Int
|
||||
(.|.) = prim__or_Int
|
||||
xor = prim__xor_Int
|
||||
bit = (1 `shiftL`)
|
||||
zeroBits = 0
|
||||
testBit x i = (x .&. bit i) /= 0
|
||||
shiftR x = prim__shr_Int x . cast . fst
|
||||
shiftL x = prim__shl_Int x . cast . fst
|
||||
complement = xor (-1)
|
||||
shiftR x = prim__shr_Int x . cast . finToNat
|
||||
shiftL x = prim__shl_Int x . cast . finToNat
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Int8 where
|
||||
Index = Subset Nat (`LT` 8)
|
||||
Index = Fin 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)
|
||||
shiftR x = prim__shr_Int8 x . cast . finToNat
|
||||
shiftL x = prim__shl_Int8 x . cast . finToNat
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Int16 where
|
||||
Index = Subset Nat (`LT` 16)
|
||||
Index = Fin 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)
|
||||
shiftR x = prim__shr_Int16 x . cast . finToNat
|
||||
shiftL x = prim__shl_Int16 x . cast . finToNat
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Int32 where
|
||||
Index = Subset Nat (`LT` 32)
|
||||
Index = Fin 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)
|
||||
shiftR x = prim__shr_Int32 x . cast . finToNat
|
||||
shiftL x = prim__shl_Int32 x . cast . finToNat
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
Bits Int64 where
|
||||
Index = Subset Nat (`LT` 64)
|
||||
Index = Fin 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)
|
||||
shiftR x = prim__shr_Int64 x . cast . finToNat
|
||||
shiftL x = prim__shl_Int64 x . cast . finToNat
|
||||
oneBits = (-1)
|
||||
|
||||
public export %inline
|
||||
@ -212,7 +190,6 @@ Bits Integer where
|
||||
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)
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
@ -225,7 +202,7 @@ interface Bits a => FiniteBits a where
|
||||
bitSize : Nat
|
||||
|
||||
||| Properly correlates `bitSize` and `Index`.
|
||||
bitsToIndex : Subset Nat (`LT` bitSize) -> Index {a}
|
||||
bitsToIndex : Fin bitSize -> Index {a}
|
||||
|
||||
||| Return the number of set bits in the argument. This number is
|
||||
||| known as the population count or the Hamming weight.
|
||||
@ -238,9 +215,9 @@ FiniteBits Bits8 where
|
||||
|
||||
popCount x0 =
|
||||
-- see https://stackoverflow.com/questions/109023/how-to-count-the-number-of-set-bits-in-a-32-bit-integer
|
||||
let x1 = (x0 .&. 0x55) + ((x0 `shiftR` fromNat 1) .&. 0x55)
|
||||
x2 = (x1 .&. 0x33) + ((x1 `shiftR` fromNat 2) .&. 0x33)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F)
|
||||
let x1 = (x0 .&. 0x55) + ((x0 `shiftR` 1) .&. 0x55)
|
||||
x2 = (x1 .&. 0x33) + ((x1 `shiftR` 2) .&. 0x33)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F)
|
||||
in cast x3
|
||||
|
||||
public export %inline
|
||||
@ -250,10 +227,10 @@ FiniteBits Bits16 where
|
||||
|
||||
popCount x0 =
|
||||
-- see https://stackoverflow.com/questions/109023/how-to-count-the-number-of-set-bits-in-a-32-bit-integer
|
||||
let 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
|
||||
let x1 = (x0 .&. 0x5555) + ((x0 `shiftR` 1) .&. 0x5555)
|
||||
x2 = (x1 .&. 0x3333) + ((x1 `shiftR` 2) .&. 0x3333)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F0F)
|
||||
x4 = (x3 * 0x0101) `shiftR` 8
|
||||
in cast x4
|
||||
|
||||
public export %inline
|
||||
@ -263,10 +240,10 @@ FiniteBits Bits32 where
|
||||
|
||||
popCount x0 =
|
||||
-- see https://stackoverflow.com/questions/109023/how-to-count-the-number-of-set-bits-in-a-32-bit-integer
|
||||
let 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
|
||||
let x1 = (x0 .&. 0x55555555) + ((x0 `shiftR` 1) .&. 0x55555555)
|
||||
x2 = (x1 .&. 0x33333333) + ((x1 `shiftR` 2) .&. 0x33333333)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F0F0F0F)
|
||||
x4 = (x3 * 0x01010101) `shiftR` 24
|
||||
in cast x4
|
||||
|
||||
public export %inline
|
||||
@ -277,11 +254,11 @@ FiniteBits Bits64 where
|
||||
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)
|
||||
((x0 `shiftR` 1) .&. 0x5555555555555555)
|
||||
x2 = (x1 .&. 0x3333333333333333)
|
||||
+ ((x1 `shiftR` fromNat 2) .&. 0x3333333333333333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` fromNat 56
|
||||
+ ((x1 `shiftR` 2) .&. 0x3333333333333333)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` 56
|
||||
in cast x4
|
||||
|
||||
public export %inline
|
||||
@ -295,13 +272,13 @@ FiniteBits Int where
|
||||
-- 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
|
||||
let x0 = x `clearBit` 63
|
||||
x1 = (x0 .&. 0x5555555555555555)
|
||||
+ ((x0 `shiftR` fromNat 1) .&. 0x5555555555555555)
|
||||
+ ((x0 `shiftR` 1) .&. 0x5555555555555555)
|
||||
x2 = (x1 .&. 0x3333333333333333)
|
||||
+ ((x1 `shiftR` fromNat 2) .&. 0x3333333333333333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F0F0F0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` fromNat 56
|
||||
+ ((x1 `shiftR` 2) .&. 0x3333333333333333)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F0F0F0F0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` 56
|
||||
x5 = if x < 0 then x4 + 1 else x4
|
||||
in cast x5
|
||||
|
||||
@ -316,10 +293,10 @@ FiniteBits Int8 where
|
||||
-- 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)
|
||||
let x0 = x `clearBit` 7
|
||||
x1 = (x0 .&. 0x55) + ((x0 `shiftR` 1) .&. 0x55)
|
||||
x2 = (x1 .&. 0x33) + ((x1 `shiftR` 2) .&. 0x33)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F)
|
||||
x4 = if x < 0 then x3 + 1 else x3
|
||||
in cast x4
|
||||
|
||||
@ -334,11 +311,11 @@ FiniteBits Int16 where
|
||||
-- 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
|
||||
let x0 = x `clearBit` 15
|
||||
x1 = (x0 .&. 0x5555) + ((x0 `shiftR` 1) .&. 0x5555)
|
||||
x2 = (x1 .&. 0x3333) + ((x1 `shiftR` 2) .&. 0x3333)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F0F)
|
||||
x4 = (x3 * 0x0101) `shiftR` 8
|
||||
x5 = if x < 0 then x4 + 1 else x4
|
||||
in cast x5
|
||||
|
||||
@ -353,11 +330,11 @@ FiniteBits Int32 where
|
||||
-- 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
|
||||
let x0 = x `clearBit` 31
|
||||
x1 = (x0 .&. 0x55555555) + ((x0 `shiftR` 1) .&. 0x55555555)
|
||||
x2 = (x1 .&. 0x33333333) + ((x1 `shiftR` 2) .&. 0x33333333)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F0F0F0F)
|
||||
x4 = (x3 * 0x01010101) `shiftR` 24
|
||||
x5 = if x < 0 then x4 + 1 else x4
|
||||
in cast x5
|
||||
|
||||
@ -372,12 +349,12 @@ FiniteBits Int64 where
|
||||
-- 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
|
||||
let x0 = x `clearBit` 63
|
||||
x1 = (x0 .&. 0x5555555555555555)
|
||||
+ ((x0 `shiftR` fromNat 1) .&. 0x5555555555555555)
|
||||
+ ((x0 `shiftR` 1) .&. 0x5555555555555555)
|
||||
x2 = (x1 .&. 0x3333333333333333)
|
||||
+ ((x1 `shiftR` fromNat 2) .&. 0x3333333333333333)
|
||||
x3 = ((x2 + (x2 `shiftR` fromNat 4)) .&. 0x0F0F0F0F0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` fromNat 56
|
||||
+ ((x1 `shiftR` 2) .&. 0x3333333333333333)
|
||||
x3 = ((x2 + (x2 `shiftR` 4)) .&. 0x0F0F0F0F0F0F0F0F)
|
||||
x4 = (x3 * 0x0101010101010101) `shiftR` 56
|
||||
x5 = if x < 0 then x4 + 1 else x4
|
||||
in cast x5
|
||||
|
@ -40,7 +40,7 @@ b16ToHexString n =
|
||||
14 => "E"
|
||||
15 => "F"
|
||||
other => assert_total $
|
||||
b16ToHexString (n `shiftR` fromNat 4) ++
|
||||
b16ToHexString (n `shiftR` 4) ++
|
||||
b16ToHexString (n .&. 15)
|
||||
|
||||
private
|
||||
|
@ -3,6 +3,11 @@ module Libraries.Utils.Hex
|
||||
import Data.Bits
|
||||
import Data.List
|
||||
|
||||
-- Those three imports are for compatibility and should be removed after release of 0.6.0
|
||||
import Data.DPair
|
||||
import Data.Nat
|
||||
import Data.Fin
|
||||
|
||||
%default total
|
||||
|
||||
hexDigit : Bits64 -> Char
|
||||
@ -24,6 +29,16 @@ hexDigit 14 = 'e'
|
||||
hexDigit 15 = 'f'
|
||||
hexDigit _ = 'X' -- TMP HACK: Ideally we'd have a bounds proof, generated below
|
||||
|
||||
-- `i4` is to be replaced with a `4` literal after release of 0.6.0
|
||||
namespace Old
|
||||
export
|
||||
i4 : Subset Nat (`LT` 64)
|
||||
i4 = Element (the Nat 4) %search
|
||||
namespace New
|
||||
export
|
||||
i4 : Fin 64
|
||||
i4 = 4
|
||||
|
||||
||| Convert a Bits64 value into a list of (lower case) hexadecimal characters
|
||||
export
|
||||
asHex : Bits64 -> String
|
||||
@ -32,7 +47,7 @@ asHex n = pack $ asHex' n []
|
||||
where
|
||||
asHex' : Bits64 -> List Char -> List Char
|
||||
asHex' 0 hex = hex
|
||||
asHex' n hex = asHex' (assert_smaller n (n `shiftR` fromNat 4)) (hexDigit (n .&. 0xf) :: hex)
|
||||
asHex' n hex = asHex' (assert_smaller n (n `shiftR` i4)) (hexDigit (n .&. 0xf) :: hex)
|
||||
|
||||
export
|
||||
leftPad : Char -> Nat -> String -> String
|
||||
|
@ -4,6 +4,8 @@ import Data.List
|
||||
import Data.Stream
|
||||
import Decidable.Equality
|
||||
|
||||
import Data.Fin
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- Utilities
|
||||
--------------------------------------------------------------------------------
|
||||
@ -38,38 +40,38 @@ powsOf2 n = take n (iterate (*2) 1)
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
shiftRBits8 : List Bits8
|
||||
shiftRBits8 = map (`shiftR` fromNat 1) (powsOf2 8 ++ [b8max])
|
||||
shiftRBits8 = map (`shiftR` 1) (powsOf2 8 ++ [b8max])
|
||||
|
||||
shiftRBits16 : List Bits16
|
||||
shiftRBits16 = map (`shiftR` fromNat 1) (powsOf2 16 ++ [b16max])
|
||||
shiftRBits16 = map (`shiftR` 1) (powsOf2 16 ++ [b16max])
|
||||
|
||||
shiftRBits32 : List Bits32
|
||||
shiftRBits32 = map (`shiftR` fromNat 1) (powsOf2 32 ++ [b32max])
|
||||
shiftRBits32 = map (`shiftR` 1) (powsOf2 32 ++ [b32max])
|
||||
|
||||
shiftRInt : List Int
|
||||
shiftRInt = map (`shiftR` fromNat 1) (powsOf2 63 ++ [intmax])
|
||||
shiftRInt = map (`shiftR` 1) (powsOf2 63 ++ [intmax])
|
||||
|
||||
shiftRNegativeInt : List Int
|
||||
shiftRNegativeInt = map (`shiftR` fromNat 1) (map negate (powsOf2 63) ++ [intmin])
|
||||
shiftRNegativeInt = map (`shiftR` 1) (map negate (powsOf2 63) ++ [intmin])
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- shiftL
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
shiftLBits8 : List Bits8
|
||||
shiftLBits8 = map (`shiftL` fromNat 1) (0 :: powsOf2 7)
|
||||
shiftLBits8 = map (`shiftL` 1) (0 :: powsOf2 7)
|
||||
|
||||
shiftLBits16 : List Bits16
|
||||
shiftLBits16 = map (`shiftL` fromNat 1) (0 :: powsOf2 15)
|
||||
shiftLBits16 = map (`shiftL` 1) (0 :: powsOf2 15)
|
||||
|
||||
shiftLBits32 : List Bits32
|
||||
shiftLBits32 = map (`shiftL` fromNat 1) (0 :: powsOf2 31)
|
||||
shiftLBits32 = map (`shiftL` 1) (0 :: powsOf2 31)
|
||||
|
||||
shiftLInt : List Int
|
||||
shiftLInt = map (`shiftL` fromNat 1) (0 :: powsOf2 62)
|
||||
shiftLInt = map (`shiftL` 1) (0 :: powsOf2 62)
|
||||
|
||||
shiftLNegativeInt : List Int
|
||||
shiftLNegativeInt = map (`shiftL` fromNat 1) (map negate (powsOf2 62))
|
||||
shiftLNegativeInt = map (`shiftL` 1) (map negate (powsOf2 62))
|
||||
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
@ -238,70 +240,65 @@ xorNegativeInt = [ (-11) `xor` intmax
|
||||
-- bit
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
fromNatMay : (n : Nat) -> (k : Nat) -> Maybe (Subset Nat (`Nat.LT` n))
|
||||
fromNatMay n k with (decEq (lt k n) True)
|
||||
fromNatMay n k | (Yes refl) = Just $ fromNat k
|
||||
fromNatMay n k | (No _) = Nothing
|
||||
|
||||
bitBits8 : List Bits8
|
||||
bitBits8 = map bit $ mapMaybe (fromNatMay 8) [0..7]
|
||||
bitBits8 = map bit $ mapMaybe (`natToFin` 8) [0..7]
|
||||
|
||||
bitBits16 : List Bits16
|
||||
bitBits16 = map bit $ mapMaybe (fromNatMay 16) [0..15]
|
||||
bitBits16 = map bit $ mapMaybe (`natToFin` 16) [0..15]
|
||||
|
||||
bitBits32 : List Bits32
|
||||
bitBits32 = map bit $ mapMaybe (fromNatMay 32) [0..31]
|
||||
bitBits32 = map bit $ mapMaybe (`natToFin` 32) [0..31]
|
||||
|
||||
bitInt : List Int
|
||||
bitInt = map bit $ mapMaybe (fromNatMay 64) [0..63]
|
||||
bitInt = map bit $ mapMaybe (`natToFin` 64) [0..63]
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- complementBit
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
complementBitBits8 : List Bits8
|
||||
complementBitBits8 = map (`complementBit` fromNat 1) bitBits8
|
||||
complementBitBits8 = map (`complementBit` 1) bitBits8
|
||||
|
||||
complementBitBits16 : List Bits16
|
||||
complementBitBits16 = map (`complementBit` fromNat 1) bitBits16
|
||||
complementBitBits16 = map (`complementBit` 1) bitBits16
|
||||
|
||||
complementBitBits32 : List Bits32
|
||||
complementBitBits32 = map (`complementBit` fromNat 1) bitBits32
|
||||
complementBitBits32 = map (`complementBit` 1) bitBits32
|
||||
|
||||
complementBitInt : List Int
|
||||
complementBitInt = map (`complementBit` fromNat 1) bitInt
|
||||
complementBitInt = map (`complementBit` 1) bitInt
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- clearBit
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
clearBitBits8 : List Bits8
|
||||
clearBitBits8 = map (`clearBit` fromNat 5) bitBits8
|
||||
clearBitBits8 = map (`clearBit` 5) bitBits8
|
||||
|
||||
clearBitBits16 : List Bits16
|
||||
clearBitBits16 = map (`clearBit` fromNat 5) bitBits16
|
||||
clearBitBits16 = map (`clearBit` 5) bitBits16
|
||||
|
||||
clearBitBits32 : List Bits32
|
||||
clearBitBits32 = map (`clearBit` fromNat 5) bitBits32
|
||||
clearBitBits32 = map (`clearBit` 5) bitBits32
|
||||
|
||||
clearBitInt : List Int
|
||||
clearBitInt = map (`clearBit` fromNat 5) bitInt
|
||||
clearBitInt = map (`clearBit` 5) bitInt
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- setBit
|
||||
--------------------------------------------------------------------------------
|
||||
|
||||
setBitBits8 : List Bits8
|
||||
setBitBits8 = map (`setBit` fromNat 1) bitBits8
|
||||
setBitBits8 = map (`setBit` 1) bitBits8
|
||||
|
||||
setBitBits16 : List Bits16
|
||||
setBitBits16 = map (`setBit` fromNat 1) bitBits16
|
||||
setBitBits16 = map (`setBit` 1) bitBits16
|
||||
|
||||
setBitBits32 : List Bits32
|
||||
setBitBits32 = map (`setBit` fromNat 1) bitBits32
|
||||
setBitBits32 = map (`setBit` 1) bitBits32
|
||||
|
||||
setBitInt : List Int
|
||||
setBitInt = map (`setBit` fromNat 1) bitInt
|
||||
setBitInt = map (`setBit` 1) bitInt
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- testBit
|
||||
@ -309,23 +306,23 @@ setBitInt = map (`setBit` fromNat 1) bitInt
|
||||
|
||||
testBitBits8 : List Bool
|
||||
testBitBits8 = map (testBit (the Bits8 0xaa))
|
||||
$ mapMaybe (fromNatMay 8) [0..7]
|
||||
$ mapMaybe (`natToFin` 8) [0..7]
|
||||
|
||||
testBitBits16 : List Bool
|
||||
testBitBits16 = map (testBit (the Bits16 0xaaaa))
|
||||
$ mapMaybe (fromNatMay 16) [0..15]
|
||||
$ mapMaybe (`natToFin` 16) [0..15]
|
||||
|
||||
testBitBits32 : List Bool
|
||||
testBitBits32 = map (testBit (the Bits32 0xaaaaaaaa))
|
||||
$ mapMaybe (fromNatMay 32) [0..31]
|
||||
$ mapMaybe (`natToFin` 32) [0..31]
|
||||
|
||||
testBitInt : List Bool
|
||||
testBitInt = map (testBit (the Int 0xaaaaaaaaaaaaaaa))
|
||||
$ mapMaybe (fromNatMay 64) [0..63]
|
||||
$ mapMaybe (`natToFin` 64) [0..63]
|
||||
|
||||
testBitNegInt : List Bool
|
||||
testBitNegInt = map (testBit (negate $ the Int 0xaaaaaaaaaaaaaaa))
|
||||
$ mapMaybe (fromNatMay 64) [0..63]
|
||||
$ mapMaybe (`natToFin` 64) [0..63]
|
||||
|
||||
--------------------------------------------------------------------------------
|
||||
-- popCount
|
||||
|
@ -15,31 +15,31 @@ interface Num a => Bits a => SomeBits a where
|
||||
one : Index {a}
|
||||
|
||||
SomeBits Bits8 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits16 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits32 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits64 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int8 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int16 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int32 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int64 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Integer where
|
||||
one = 1
|
||||
|
@ -15,31 +15,31 @@ interface Num a => Bits a => SomeBits a where
|
||||
one : Index {a}
|
||||
|
||||
SomeBits Bits8 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits16 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits32 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits64 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int8 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int16 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int32 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int64 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Integer where
|
||||
one = 1
|
||||
|
@ -15,31 +15,31 @@ interface Num a => Bits a => SomeBits a where
|
||||
one : Index {a}
|
||||
|
||||
SomeBits Bits8 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits16 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits32 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Bits64 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int8 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int16 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int32 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int64 where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Int where
|
||||
one = fromNat 1
|
||||
one = 1
|
||||
|
||||
SomeBits Integer where
|
||||
one = 1
|
||||
|
@ -15,7 +15,7 @@ quiz (num1 :: num2 :: nums) score
|
||||
|
||||
randoms : Int -> Stream Int
|
||||
randoms seed = let seed' = 1664525 * seed + 1013904223 in
|
||||
(seed' `shiftR` fromNat 2) :: randoms seed'
|
||||
(seed' `shiftR` 2) :: randoms seed'
|
||||
|
||||
arithInputs : Int -> Stream Int
|
||||
arithInputs seed = map bound (randoms seed)
|
||||
|
@ -34,7 +34,7 @@ run Dry p = pure Nothing
|
||||
|
||||
randoms : Int -> Stream Int
|
||||
randoms seed = let seed' = 1664525 * seed + 1013904223 in
|
||||
(seed' `shiftR` fromNat 2) :: randoms seed'
|
||||
(seed' `shiftR` 2) :: randoms seed'
|
||||
|
||||
arithInputs : Int -> Stream Int
|
||||
arithInputs seed = map bound (randoms seed)
|
||||
|
@ -55,7 +55,7 @@ run Dry p = pure Nothing
|
||||
|
||||
randoms : Int -> Stream Int
|
||||
randoms seed = let seed' = 1664525 * seed + 1013904223 in
|
||||
(seed' `shiftR` fromNat 2) :: randoms seed'
|
||||
(seed' `shiftR` 2) :: randoms seed'
|
||||
|
||||
arithInputs : Int -> Stream Int
|
||||
arithInputs seed = map bound (randoms seed)
|
||||
|
@ -24,7 +24,7 @@ run Dry p = putStrLn "Out of fuel"
|
||||
|
||||
randoms : Int -> Stream Int
|
||||
randoms seed = let seed' = 1664525 * seed + 1013904223 in
|
||||
(seed' `shiftR` fromNat 2) :: randoms seed'
|
||||
(seed' `shiftR` 2) :: randoms seed'
|
||||
|
||||
arithInputs : Int -> Stream Int
|
||||
arithInputs seed = map bound (randoms seed)
|
||||
|
@ -71,7 +71,7 @@ namespace CommandDo
|
||||
|
||||
randoms : Int -> Stream Int
|
||||
randoms seed = let seed' = 1664525 * seed + 1013904223 in
|
||||
(seed' `shiftR` fromNat 2) :: randoms seed'
|
||||
(seed' `shiftR` 2) :: randoms seed'
|
||||
|
||||
runCommand : Stream Int -> GameState -> Command a ->
|
||||
IO (a, Stream Int, GameState)
|
||||
|
Loading…
Reference in New Issue
Block a user