diff --git a/effects/Effect/Memory.idr b/effects/Effect/Memory.idr index 62097bc0fd..3af13de168 100644 --- a/effects/Effect/Memory.idr +++ b/effects/Effect/Memory.idr @@ -48,7 +48,7 @@ do_malloc size with (fromInteger (cast size) == size) private do_memset : Ptr -> Nat -> Bits8 -> Nat -> IO () do_memset ptr offset c size - = mkForeign (FFun "idris_memset" [FPtr, FInt, FChar, FInt] FUnit) + = mkForeign (FFun "idris_memset" [FPtr, FInt, FAny Bits8, FInt] FUnit) ptr (cast offset) c (cast size) private @@ -73,7 +73,7 @@ private do_poke : Ptr -> Nat -> Vect Bits8 size -> IO () do_poke _ _ [] = return () do_poke ptr offset (b::bs) - = do mkForeign (FFun "idris_poke" [FPtr, FInt, FChar] FUnit) ptr (cast offset) b + = do mkForeign (FFun "idris_poke" [FPtr, FInt, FAny Bits8] FUnit) ptr (cast offset) b do_poke ptr (S offset) bs instance Handler RawMemory (IOExcept String) where diff --git a/lib/IO.idr b/lib/IO.idr index 318289c30d..e188e49fa4 100644 --- a/lib/IO.idr +++ b/lib/IO.idr @@ -21,28 +21,12 @@ io_return x = prim__IO x run__IO : IO () -> IO () run__IO v = io_bind v (\v' => io_return v') -data IntTy = ITNative | IT8 | IT16 | IT32 | IT64 -data FTy = FIntT IntTy | FFloat | FString | FPtr | FAny Type | FUnit - -FInt : FTy -FInt = FIntT ITNative - -FChar : FTy -FChar = FIntT IT8 - -FShort : FTy -FShort = FIntT IT16 - -FLong : FTy -FLong = FIntT IT64 +data FTy = FInt | FFloat | FChar | FString | FPtr | FAny Type | FUnit interpFTy : FTy -> Type -interpFTy (FIntT ITNative) = Int -interpFTy (FIntT IT8) = Bits8 -interpFTy (FIntT IT16) = Bits16 -interpFTy (FIntT IT32) = Bits32 -interpFTy (FIntT IT64) = Bits64 +interpFTy FInt = Int interpFTy FFloat = Float +interpFTy FChar = Char interpFTy FString = String interpFTy FPtr = Ptr interpFTy (FAny t) = t diff --git a/lib/Prelude.idr b/lib/Prelude.idr index b886476851..514d2a0753 100644 --- a/lib/Prelude.idr +++ b/lib/Prelude.idr @@ -251,11 +251,11 @@ getLine = return (prim__readString prim__stdin) partial putChar : Char -> IO () -putChar c = mkForeign (FFun "putchar" [FInt] FUnit) (cast c) +putChar c = mkForeign (FFun "putchar" [FChar] FUnit) c partial getChar : IO Char -getChar = fmap cast $ mkForeign (FFun "getchar" [] FInt) +getChar = mkForeign (FFun "getchar" [] FChar) ---- some basic file handling @@ -323,7 +323,7 @@ ferror (FHandle h) = do err <- do_ferror h partial nullPtr : Ptr -> IO Bool -nullPtr p = do ok <- mkForeign (FFun "isNull" [FPtr] FInt) p +nullPtr p = do ok <- mkForeign (FFun "isNull" [FPtr] FInt) p return (ok /= 0); partial diff --git a/rts/idris_bitstring.c b/rts/idris_bitstring.c index 6f22288898..316996854e 100644 --- a/rts/idris_bitstring.c +++ b/rts/idris_bitstring.c @@ -60,34 +60,6 @@ VAL idris_castB32Int(VM *vm, VAL a) { return MKINT((i_int)a->info.bits32); } -VAL idris_b8const(VM *vm, uint8_t a) { - VAL cl = allocate(vm, sizeof(Closure), 0); - SETTY(cl, BITS8); - cl->info.bits8 = a; - return cl; -} - -VAL idris_b16const(VM *vm, uint16_t a) { - VAL cl = allocate(vm, sizeof(Closure), 0); - SETTY(cl, BITS16); - cl->info.bits16 = a; - return cl; -} - -VAL idris_b32const(VM *vm, uint32_t a) { - VAL cl = allocate(vm, sizeof(Closure), 0); - SETTY(cl, BITS32); - cl->info.bits32 = a; - return cl; -} - -VAL idris_b64const(VM *vm, uint64_t a) { - VAL cl = allocate(vm, sizeof(Closure), 0); - SETTY(cl, BITS64); - cl->info.bits64 = a; - return cl; -} - VAL idris_b8Plus(VM *vm, VAL a, VAL b) { VAL cl = allocate(vm, sizeof(Closure), 0); SETTY(cl, BITS8); diff --git a/rts/idris_bitstring.h b/rts/idris_bitstring.h index 8bf1db8d0e..c46edf846a 100644 --- a/rts/idris_bitstring.h +++ b/rts/idris_bitstring.h @@ -11,10 +11,6 @@ VAL idris_b16(VM *vm, VAL a); VAL idris_b32(VM *vm, VAL a); VAL idris_b64(VM *vm, VAL a); VAL idris_castB32Int(VM *vm, VAL a); -VAL idris_b8const(VM *vm, uint8_t a); -VAL idris_b16const(VM *vm, uint16_t a); -VAL idris_b32const(VM *vm, uint32_t a); -VAL idris_b64const(VM *vm, uint64_t a); VAL idris_b8Plus(VM *vm, VAL a, VAL b); VAL idris_b8Minus(VM *vm, VAL a, VAL b); diff --git a/rts/idris_rts.h b/rts/idris_rts.h index b3c4116887..e868630dcd 100644 --- a/rts/idris_rts.h +++ b/rts/idris_rts.h @@ -125,7 +125,6 @@ typedef intptr_t i_int; #define ISINT(x) ((((i_int)x)&1) == 1) #define INTOP(op,x,y) MKINT((i_int)((((i_int)x)>>1) op (((i_int)y)>>1))) -#define UINTOP(op,x,y) MKINT((i_int)((((uintptr_t)x)>>1) op (((uintptr_t)y)>>1))) #define FLOATOP(op,x,y) MKFLOAT(vm, ((GETFLOAT(x)) op (GETFLOAT(y)))) #define FLOATBOP(op,x,y) MKINT((i_int)(((GETFLOAT(x)) op (GETFLOAT(y))))) #define ADD(x,y) (void*)(((i_int)x)+(((i_int)y)-1)) diff --git a/src/Core/Execute.hs b/src/Core/Execute.hs index 73eee9131d..9d4edc5269 100644 --- a/src/Core/Execute.hs +++ b/src/Core/Execute.hs @@ -3,9 +3,6 @@ module Core.Execute (execute) where import Idris.AbsSyntax import Idris.AbsSyntaxTree -import IRTS.Lang( IntTy(..) - , intTyToConst - , FType(..)) import Core.TT import Core.Evaluate @@ -428,16 +425,18 @@ chooseAlt _ [] = Nothing -idrisType :: FType -> ExecVal +data FTy = FInt | FFloat | FChar | FString | FPtr | FUnit deriving (Show, Read) + +idrisType :: FTy -> ExecVal idrisType FUnit = EP Ref unitTy EErased idrisType ft = EConstant (idr ft) - where idr (FInt ty) = intTyToConst ty - idr FDouble = FlType + where idr FInt = IType + idr FFloat = FlType idr FChar = ChType idr FString = StrType idr FPtr = PtrType -data Foreign = FFun String [FType] FType deriving Show +data Foreign = FFun String [FTy] FTy deriving Show call :: Foreign -> [ExecVal] -> Exec (Maybe ExecVal) @@ -447,19 +446,11 @@ call (FFun name argTypes retType) args = Nothing -> return Nothing Just f -> do res <- call' f args retType return . Just . ioWrap $ res - where call' :: ForeignFun -> [ExecVal] -> FType -> Exec ExecVal - call' (Fun _ h) args (FInt ITNative) = do res <- execIO $ callFFI h retCInt (prepArgs args) - return (EConstant (I (fromIntegral res))) - call' (Fun _ h) args (FInt IT8) = do res <- execIO $ callFFI h retCChar (prepArgs args) - return (EConstant (B8 (fromIntegral res))) - call' (Fun _ h) args (FInt IT16) = do res <- execIO $ callFFI h retCWchar (prepArgs args) - return (EConstant (B16 (fromIntegral res))) - call' (Fun _ h) args (FInt IT32) = do res <- execIO $ callFFI h retCInt (prepArgs args) - return (EConstant (B32 (fromIntegral res))) - call' (Fun _ h) args (FInt IT64) = do res <- execIO $ callFFI h retCLong (prepArgs args) - return (EConstant (B64 (fromIntegral res))) - call' (Fun _ h) args FDouble = do res <- execIO $ callFFI h retCDouble (prepArgs args) - return (EConstant (Fl (realToFrac res))) + where call' :: ForeignFun -> [ExecVal] -> FTy -> Exec ExecVal + call' (Fun _ h) args FInt = do res <- execIO $ callFFI h retCInt (prepArgs args) + return (EConstant (I (fromIntegral res))) + call' (Fun _ h) args FFloat = do res <- execIO $ callFFI h retCDouble (prepArgs args) + return (EConstant (Fl (realToFrac res))) call' (Fun _ h) args FChar = do res <- execIO $ callFFI h retCChar (prepArgs args) return (EConstant (Ch (castCCharToChar res))) call' (Fun _ h) args FString = do res <- execIO $ callFFI h retCString (prepArgs args) @@ -476,10 +467,6 @@ call (FFun name argTypes retType) args = prepArgs = map prepArg prepArg (EConstant (I i)) = argCInt (fromIntegral i) - prepArg (EConstant (B8 i)) = argCChar (fromIntegral i) - prepArg (EConstant (B16 i)) = argCWchar (fromIntegral i) - prepArg (EConstant (B32 i)) = argCInt (fromIntegral i) - prepArg (EConstant (B64 i)) = argCLong (fromIntegral i) prepArg (EConstant (Fl f)) = argCDouble (realToFrac f) prepArg (EConstant (Ch c)) = argCChar (castCharToCChar c) -- FIXME - castCharToCChar only safe for first 256 chars prepArg (EConstant (Str s)) = argString s @@ -497,18 +484,11 @@ foreignFromTT t = case (unApplyV t) of return $ FFun name argFTy retFTy _ -> trace "failed to construct ffun" Nothing -getFTy :: ExecVal -> Maybe FType -getFTy (EApp (EP _ (UN "FInt") _) (EP _ (UN intTy) _)) = - case intTy of - "ITNative" -> Just $ FInt ITNative - "IT8" -> Just $ FInt IT8 - "IT16" -> Just $ FInt IT16 - "IT32" -> Just $ FInt IT32 - "IT64" -> Just $ FInt IT64 - _ -> Nothing +getFTy :: ExecVal -> Maybe FTy getFTy (EP _ (UN t) _) = case t of - "FFloat" -> Just FDouble + "FInt" -> Just FInt + "FFloat" -> Just FFloat "FChar" -> Just FChar "FString" -> Just FString "FPtr" -> Just FPtr diff --git a/src/IRTS/CodegenC.hs b/src/IRTS/CodegenC.hs index 2a7ccb8ba3..f841325465 100644 --- a/src/IRTS/CodegenC.hs +++ b/src/IRTS/CodegenC.hs @@ -10,7 +10,6 @@ import Paths_idris import Util.System import Data.Char -import Data.List (intercalate) import System.Process import System.Exit import System.IO @@ -106,10 +105,10 @@ bcc i (ASSIGNCONST l c) mkConst (Fl f) = "MKFLOAT(vm, " ++ show f ++ ")" mkConst (Ch c) = "MKINT(" ++ show (fromEnum c) ++ ")" mkConst (Str s) = "MKSTR(vm, " ++ show s ++ ")" - mkConst (B8 x) = "idris_b8const(vm, " ++ show x ++ ")" - mkConst (B16 x) = "idris_b16const(vm, " ++ show x ++ ")" - mkConst (B32 x) = "idris_b32const(vm, " ++ show x ++ ")" - mkConst (B64 x) = "idris_b64const(vm, " ++ show x ++ ")" + mkConst (B8 b) = "MKB8(vm, " ++ show b ++")" + mkConst (B16 b) = "MKB16(vm, " ++ show b ++ ")" + mkConst (B32 b) = "MKB32(vm, " ++ show b ++ ")" + mkConst (B64 b) = "MKB64(vm, " ++ show b ++ ")" mkConst _ = "MKINT(42424242)" bcc i (UPDATE l r) = indent i ++ creg l ++ " = " ++ creg r ++ ";\n" bcc i (MKCON l tag args) @@ -227,10 +226,7 @@ bcc i (NULL r) = indent i ++ creg r ++ " = NULL;\n" -- clear, so it'll be GCed bcc i (ERROR str) = indent i ++ "fprintf(stderr, " ++ show str ++ "); assert(0); exit(-1);" -- bcc i _ = indent i ++ "// not done yet\n" - - -c_irts (FInt ITNative) l x = l ++ "MKINT((i_int)(" ++ x ++ "))" -c_irts (FInt ty) l x = l ++ "idris_b" ++ show (intTyWidth ty) ++ "const(vm, " ++ x ++ ")" +c_irts FInt l x = l ++ "MKINT((i_int)(" ++ x ++ "))" c_irts FChar l x = l ++ "MKINT((i_int)(" ++ x ++ "))" c_irts FString l x = l ++ "MKSTR(vm, " ++ x ++ ")" c_irts FUnit l x = x @@ -238,8 +234,7 @@ c_irts FPtr l x = l ++ "MKPTR(vm, " ++ x ++ ")" c_irts FDouble l x = l ++ "MKFLOAT(vm, " ++ x ++ ")" c_irts FAny l x = l ++ x -irts_c (FInt ITNative) x = "GETINT(" ++ x ++ ")" -irts_c (FInt ty) x = "(" ++ x ++ "->info.bits" ++ show (intTyWidth ty) ++ ")" +irts_c FInt x = "GETINT(" ++ x ++ ")" irts_c FChar x = "GETINT(" ++ x ++ ")" irts_c FString x = "GETSTR(" ++ x ++ ")" irts_c FUnit x = x @@ -247,30 +242,22 @@ irts_c FPtr x = "GETPTR(" ++ x ++ ")" irts_c FDouble x = "GETFLOAT(" ++ x ++ ")" irts_c FAny x = x -bitOp v op ty args = v ++ "idris_b" ++ show (intTyWidth ty) ++ op ++ "(vm, " ++ intercalate ", " (map creg args) ++ ")" - -bitCoerce v op input output arg - = v ++ "idris_b" ++ show (intTyWidth input) ++ op ++ show (intTyWidth output) ++ "(vm, " ++ creg arg ++ ")" - -doOp v (LPlus ITNative) [l, r] = v ++ "ADD(" ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LMinus ITNative) [l, r] = v ++ "INTOP(-," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LTimes ITNative) [l, r] = v ++ "MULT(" ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LUDiv ITNative) [l, r] = v ++ "UINTOP(/," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LSDiv ITNative) [l, r] = v ++ "INTOP(/," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LURem ITNative) [l, r] = v ++ "UINTOP(%," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LSRem ITNative) [l, r] = v ++ "INTOP(%," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LAnd ITNative) [l, r] = v ++ "INTOP(&," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LOr ITNative) [l, r] = v ++ "INTOP(|," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LXOr ITNative) [l, r] = v ++ "INTOP(^," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LSHL ITNative) [l, r] = v ++ "INTOP(<<," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LLSHR ITNative) [l, r] = v ++ "UINTOP(>>," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LASHR ITNative) [l, r] = v ++ "INTOP(>>," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LCompl ITNative) [x] = v ++ "INTOP(~," ++ creg x ++ ")" -doOp v (LEq ITNative) [l, r] = v ++ "INTOP(==," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LLt ITNative) [l, r] = v ++ "INTOP(<," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LLe ITNative) [l, r] = v ++ "INTOP(<=," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LGt ITNative) [l, r] = v ++ "INTOP(>," ++ creg l ++ ", " ++ creg r ++ ")" -doOp v (LGe ITNative) [l, r] = v ++ "INTOP(>=," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LPlus [l, r] = v ++ "ADD(" ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LMinus [l, r] = v ++ "INTOP(-," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LTimes [l, r] = v ++ "MULT(" ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LDiv [l, r] = v ++ "INTOP(/," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LMod [l, r] = v ++ "INTOP(%," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LAnd [l, r] = v ++ "INTOP(&," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LOr [l, r] = v ++ "INTOP(|," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LXOr [l, r] = v ++ "INTOP(^," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LSHL [l, r] = v ++ "INTOP(<<," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LSHR [l, r] = v ++ "INTOP(>>," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LCompl [x] = v ++ "INTOP(~," ++ creg x ++ ")" +doOp v LEq [l, r] = v ++ "INTOP(==," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LLt [l, r] = v ++ "INTOP(<," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LLe [l, r] = v ++ "INTOP(<=," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LGt [l, r] = v ++ "INTOP(>," ++ creg l ++ ", " ++ creg r ++ ")" +doOp v LGe [l, r] = v ++ "INTOP(>=," ++ creg l ++ ", " ++ creg r ++ ")" doOp v LFPlus [l, r] = v ++ "FLOATOP(+," ++ creg l ++ ", " ++ creg r ++ ")" doOp v LFMinus [l, r] = v ++ "FLOATOP(-," ++ creg l ++ ", " ++ creg r ++ ")" @@ -321,34 +308,136 @@ doOp v LIntB64 [x] = v ++ "idris_b64(vm, " ++ creg x ++ ")" doOp v LB32Int [x] = v ++ "idris_castB32Int(vm, " ++ creg x ++ ")" -doOp v (LLt ty) [x, y] = bitOp v "Lt" ty [x, y] -doOp v (LLe ty) [x, y] = bitOp v "Lte" ty [x, y] -doOp v (LEq ty) [x, y] = bitOp v "Eq" ty [x, y] -doOp v (LGe ty) [x, y] = bitOp v "Gte" ty [x, y] -doOp v (LGt ty) [x, y] = bitOp v "Gt" ty [x, y] - -doOp v (LSHL ty) [x, y] = bitOp v "Shl" ty [x, y] -doOp v (LLSHR ty) [x, y] = bitOp v "Shr" ty [x, y] -doOp v (LASHR ty) [x, y] = bitOp v "AShr" ty [x, y] -doOp v (LAnd ty) [x, y] = bitOp v "And" ty [x, y] -doOp v (LOr ty) [x, y] = bitOp v "Or" ty [x, y] -doOp v (LXOr ty) [x, y] = bitOp v "Xor" ty [x, y] -doOp v (LCompl ty) [x] = bitOp v "Compl" ty [x] - -doOp v (LPlus ty) [x, y] = bitOp v "Plus" ty [x, y] -doOp v (LMinus ty) [x, y] = bitOp v "Minus" ty [x, y] -doOp v (LTimes ty) [x, y] = bitOp v "Times" ty [x, y] -doOp v (LUDiv ty) [x, y] = bitOp v "UDiv" ty [x, y] -doOp v (LSDiv ty) [x, y] = bitOp v "SDiv" ty [x, y] -doOp v (LURem ty) [x, y] = bitOp v "URem" ty [x, y] -doOp v (LSRem ty) [x, y] = bitOp v "SRem" ty [x, y] - -doOp v (LSExt from to) [x] - | intTyWidth from < intTyWidth to = bitCoerce v "S" from to x -doOp v (LZExt from to) [x] - | intTyWidth from < intTyWidth to = bitCoerce v "Z" from to x -doOp v (LTrunc from to) [x] - | intTyWidth from > intTyWidth to = bitCoerce v "T" from to x +doOp v LB8Lt [x, y] = v ++ "idris_b8Lt(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Lte [x, y] = v ++ "idris_b8Lte(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Eq [x, y] = v ++ "idris_b8Eq(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Gte [x, y] = v ++ "idris_b8Gte(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Gt [x, y] = v ++ "idris_b8Gt(vm, " ++ creg x ++ "," ++ creg y ++ ")" + +doOp v LB8Shl [x, y] = v ++ "idris_b8Shl(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8LShr [x, y] = v ++ "idris_b8Shr(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8AShr [x, y] = v ++ "idris_b8AShr(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8And [x, y] = v ++ "idris_b8And(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Or [x, y] = v ++ "idris_b8Or(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Xor [x, y] = v ++ "idris_b8Xor(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Compl [x] = v ++ "idris_b8Compl(vm, " ++ creg x ++ ")" + +doOp v LB8Plus [x, y] = v ++ "idris_b8Plus(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Minus [x, y] = v ++ "idris_b8Minus(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8Times [x, y] = v ++ "idris_b8Times(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8UDiv [x, y] = v ++ "idris_b8UDiv(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8SDiv [x, y] = v ++ "idris_b8SDiv(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8URem [x, y] = v ++ "idris_b8URem(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB8SRem [x, y] = v ++ "idris_b8SRem(vm, " ++ creg x ++ "," ++ creg y ++ ")" + +doOp v LB8Z16 [x] = v ++ "idris_b8Z16(vm, " ++ creg x ++ ")" +doOp v LB8Z32 [x] = v ++ "idris_b8Z32(vm, " ++ creg x ++ ")" +doOp v LB8Z64 [x] = v ++ "idris_b8Z64(vm, " ++ creg x ++ ")" +doOp v LB8S16 [x] = v ++ "idris_b8S16(vm, " ++ creg x ++ ")" +doOp v LB8S32 [x] = v ++ "idris_b8S32(vm, " ++ creg x ++ ")" +doOp v LB8S64 [x] = v ++ "idris_b8S64(vm, " ++ creg x ++ ")" + +doOp v LB16Lt [x, y] = v ++ "idris_b16Lt(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Lte [x, y] = v ++ "idris_b16Lte(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Eq [x, y] = v ++ "idris_b16Eq(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Gte [x, y] = v ++ "idris_b16Gte(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Gt [x, y] = v ++ "idris_b16Gt(vm, " ++ creg x ++ "," ++ creg y ++ ")" + +doOp v LB16Shl [x, y] = v ++ "idris_b16Shl(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16LShr [x, y] = v ++ "idris_b16Shr(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16AShr [x, y] = v ++ "idris_b16AShr(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16And [x, y] = v ++ "idris_b16And(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Or [x, y] = v ++ "idris_b16Or(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Xor [x, y] = v ++ "idris_b16Xor(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Compl [x] = v ++ "idris_b16Compl(vm, " ++ creg x ++ ")" + +doOp v LB16Plus [x, y] = + v ++ "idris_b16Plus(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Minus [x, y] = + v ++ "idris_b16Minus(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16Times [x, y] = + v ++ "idris_b16Times(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16UDiv [x, y] = + v ++ "idris_b16UDiv(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16SDiv [x, y] = + v ++ "idris_b16SDiv(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16URem [x, y] = + v ++ "idris_b16URem(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB16SRem [x, y] = + v ++ "idris_b16SRem(vm, " ++ creg x ++ "," ++ creg y ++ ")" + +doOp v LB16Z32 [x] = v ++ "idris_b16Z32(vm, " ++ creg x ++ ")" +doOp v LB16Z64 [x] = v ++ "idris_b16Z64(vm, " ++ creg x ++ ")" +doOp v LB16S32 [x] = v ++ "idris_b16S32(vm, " ++ creg x ++ ")" +doOp v LB16S64 [x] = v ++ "idris_b16S64(vm, " ++ creg x ++ ")" +doOp v LB16T8 [x] = v ++ "idris_b16T8(vm, " ++ creg x ++ ")" + +doOp v LB32Lt [x, y] = v ++ "idris_b32Lt(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Lte [x, y] = v ++ "idris_b32Lte(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Eq [x, y] = v ++ "idris_b32Eq(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Gte [x, y] = v ++ "idris_b32Gte(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Gt [x, y] = v ++ "idris_b32Gt(vm, " ++ creg x ++ "," ++ creg y ++ ")" + +doOp v LB32Shl [x, y] = v ++ "idris_b32Shl(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32LShr [x, y] = v ++ "idris_b32Shr(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32AShr [x, y] = v ++ "idris_b32AShr(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32And [x, y] = v ++ "idris_b32And(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Or [x, y] = v ++ "idris_b32Or(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Xor [x, y] = v ++ "idris_b32Xor(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Compl [x] = v ++ "idris_b32Compl(vm, " ++ creg x ++ ")" + +doOp v LB32Plus [x, y] = + v ++ "idris_b32Plus(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Minus [x, y] = + v ++ "idris_b32Minus(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32Times [x, y] = + v ++ "idris_b32Times(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32UDiv [x, y] = + v ++ "idris_b32UDiv(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32SDiv [x, y] = + v ++ "idris_b32SDiv(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32URem [x, y] = + v ++ "idris_b32URem(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB32SRem [x, y] = + v ++ "idris_b32SRem(vm, " ++ creg x ++ "," ++ creg y ++ ")" + +doOp v LB32Z64 [x] = v ++ "idris_b32Z64(vm, " ++ creg x ++ ")" +doOp v LB32S64 [x] = v ++ "idris_b32S64(vm, " ++ creg x ++ ")" +doOp v LB32T8 [x] = v ++ "idris_b32T8(vm, " ++ creg x ++ ")" +doOp v LB32T16 [x] = v ++ "idris_b32T16(vm, " ++ creg x ++ ")" + +doOp v LB64Lt [x, y] = v ++ "idris_b64Lt(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Lte [x, y] = v ++ "idris_b64Lte(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Eq [x, y] = v ++ "idris_b64Eq(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Gte [x, y] = v ++ "idris_b64Gte(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Gt [x, y] = v ++ "idris_b64Gt(vm, " ++ creg x ++ "," ++ creg y ++ ")" + +doOp v LB64Shl [x, y] = v ++ "idris_b64Shl(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64LShr [x, y] = v ++ "idris_b64Shr(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64AShr [x, y] = v ++ "idris_b64AShr(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64And [x, y] = v ++ "idris_b64And(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Or [x, y] = v ++ "idris_b64Or(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Xor [x, y] = v ++ "idris_b64Xor(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Compl [x] = v ++ "idris_b64Compl(vm, " ++ creg x ++ ")" + +doOp v LB64Plus [x, y] = + v ++ "idris_b64Plus(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Minus [x, y] = + v ++ "idris_b64Minus(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64Times [x, y] = + v ++ "idris_b64Times(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64UDiv [x, y] = + v ++ "idris_b64UDiv(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64SDiv [x, y] = + v ++ "idris_b64SDiv(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64URem [x, y] = + v ++ "idris_b64URem(vm, " ++ creg x ++ "," ++ creg y ++ ")" +doOp v LB64SRem [x, y] = + v ++ "idris_b64SRem(vm, " ++ creg x ++ "," ++ creg y ++ ")" + +doOp v LB64T8 [x] = v ++ "idris_b64T8(vm, " ++ creg x ++ ")" +doOp v LB64T16 [x] = v ++ "idris_b64T16(vm, " ++ creg x ++ ")" +doOp v LB64T32 [x] = v ++ "idris_b64T32(vm, " ++ creg x ++ ")" doOp v LFExp [x] = v ++ "MKFLOAT(exp(GETFLOAT(" ++ creg x ++ ")))" doOp v LFLog [x] = v ++ "MKFLOAT(log(GETFLOAT(" ++ creg x ++ ")))" diff --git a/src/IRTS/CodegenJava.hs b/src/IRTS/CodegenJava.hs index 4f478d91a8..07bce086a3 100644 --- a/src/IRTS/CodegenJava.hs +++ b/src/IRTS/CodegenJava.hs @@ -468,27 +468,7 @@ integerType :: ClassType integerType = ClassType [(Ident "Integer", [])] longType :: ClassType -longType = ClassType [(Ident "Long", [])] - -intTyToIdent IT8 = Ident "Byte" -intTyToIdent IT16 = Ident "Short" -intTyToIdent IT32 = Ident "Int" -intTyToIdent IT64 = Ident "Long" - -intTyToClass :: IntTy -> ClassType -intTyToClass ty = ClassType [(intTyToIdent ty, [])] - -intTyToMethod :: IntTy -> String -intTyToMethod IT8 = "byteValue" -intTyToMethod IT16 = "shortValue" -intTyToMethod IT32 = "intValue" -intTyToMethod IT64 = "longValue" - -intTyToPrimTy :: IntTy -> PrimType -intTyToPrimTy IT8 = ByteT -intTyToPrimTy IT16 = ShortT -intTyToPrimTy IT32 = IntT -intTyToPrimTy IT64 = LongT +longType = ClassType [(Ident "Integer", [])] bigIntegerType :: ClassType bigIntegerType = ClassType [(Ident "BigInteger", [])] @@ -982,7 +962,7 @@ mkStringAtIndex var indexExp = <$> mkVarAccess (Just stringType) var mkForeignType :: FType -> Maybe ClassType -mkForeignType (FInt ty) = return (intTyToClass ty) +mkForeignType FInt = return integerType mkForeignType FChar = return integerType mkForeignType FString = return stringType mkForeignType FPtr = return objectType @@ -991,14 +971,14 @@ mkForeignType FAny = return objectType mkForeignType FUnit = Nothing mkForeignVarAccess :: FType -> LVar -> CodeGeneration Exp -mkForeignVarAccess (FInt ty) var = - (\ var -> MethodInv $ PrimaryMethodCall var +mkForeignVarAccess FInt var = + (\ var -> MethodInv $ PrimaryMethodCall (var) [] - (Ident (intTyToMethod ty)) + (Ident "intValue") [] ) <$> mkVarAccess (Just integerType) var -mkForeignVarAccess FChar var = Cast (PrimType CharT) <$> mkForeignVarAccess (FInt IT32) var +mkForeignVarAccess FChar var = Cast (PrimType CharT) <$> mkForeignVarAccess FInt var mkForeignVarAccess FDouble var = (\ var -> MethodInv $ PrimaryMethodCall (var) [] @@ -1009,12 +989,12 @@ mkForeignVarAccess FDouble var = mkForeignVarAccess otherType var = mkVarAccess (mkForeignType otherType) var mkFromForeignType :: FType -> Exp -> Exp -mkFromForeignType (FInt ty) from = - MethodInv $ TypeMethodCall (J.Name [intTyToIdent ty]) +mkFromForeignType FInt from = + MethodInv $ TypeMethodCall (J.Name [Ident "Integer"]) [] (Ident "valueOf") [from] -mkFromForeignType FChar from = mkFromForeignType (FInt IT32) from +mkFromForeignType FChar from = mkFromForeignType FInt from mkFromForeignType FDouble from = MethodInv $ TypeMethodCall (J.Name [Ident "Double"]) [] @@ -1271,22 +1251,24 @@ mkExp (SConst (PtrType)) = return $ mkClass objectType mkExp (SConst (VoidType)) = return $ mkClass voidType mkExp (SConst (Forgot)) = return $ mkClass objectType mkExp (SForeign _ fType meth args) = mkForeignInvoke fType meth args -mkExp (SOp (LPlus ITNative) args) = mkExp (SOp (LPlus IT32) args) -mkExp (SOp (LMinus ITNative) args) = mkExp (SOp (LMinus IT32) args) -mkExp (SOp (LTimes ITNative) args) = mkExp (SOp (LTimes IT32) args) -mkExp (SOp (LSDiv ITNative) args) = mkExp (SOp (LSDiv IT32) args) -mkExp (SOp (LSRem ITNative) args) = mkExp (SOp (LSRem IT32) args) -mkExp (SOp (LAnd ITNative) args) = mkExp (SOp (LAnd IT32) args) -mkExp (SOp (LOr ITNative) args) = mkExp (SOp (LOr IT32) args) -mkExp (SOp (LXOr ITNative) args) = mkExp (SOp (LXOr IT32) args) -mkExp (SOp (LCompl ITNative) args) = mkExp (SOp (LCompl IT32) args) -mkExp (SOp (LSHL ITNative) args) = mkExp (SOp (LSHL IT32) args) -mkExp (SOp (LASHR ITNative) args) = mkExp (SOp (LASHR IT32) args) -mkExp (SOp (LEq ITNative) args) = mkExp (SOp (LEq IT32) args) -mkExp (SOp (LLt ITNative) args) = mkExp (SOp (LLt IT32) args) -mkExp (SOp (LLe ITNative) args) = mkExp (SOp (LLe IT32) args) -mkExp (SOp (LGt ITNative) args) = mkExp (SOp (LGt IT32) args) -mkExp (SOp (LGe ITNative) args) = mkExp (SOp (LGe IT32) args) +mkExp (SOp LPlus args) = mkBinOpExp integerType Add args +mkExp (SOp LMinus args) = mkBinOpExp integerType Sub args +mkExp (SOp LTimes args) = mkBinOpExp integerType Mult args +mkExp (SOp LDiv args) = mkBinOpExp integerType Div args +mkExp (SOp LMod args) = mkBinOpExp integerType Rem args +mkExp (SOp LAnd args) = mkBinOpExp integerType And args +mkExp (SOp LOr args) = mkBinOpExp integerType Or args +mkExp (SOp LXOr args) = mkBinOpExp integerType Xor args +mkExp (SOp LCompl [var]) = + PreBitCompl <$> mkVarAccess (Just $ integerType) var +mkExp (SOp LSHL args) = mkBinOpExp integerType LShift args +mkExp (SOp LSHR args) = mkBinOpExp integerType RShift args +mkExp (SOp LEq args) = + mkMethodOpChain1 (mkBoolToNumber integerType) objectType "equals" args +mkExp (SOp LLt args) = mkLogicalBinOpExp integerType LThan args +mkExp (SOp LLe args) = mkLogicalBinOpExp integerType LThanE args +mkExp (SOp LGt args) = mkLogicalBinOpExp integerType GThan args +mkExp (SOp LGe args) = mkLogicalBinOpExp integerType GThanE args mkExp (SOp LFPlus args) = mkBinOpExp doubleType Add args mkExp (SOp LFMinus args) = mkBinOpExp doubleType Sub args mkExp (SOp LFTimes args) = mkBinOpExp doubleType Mult args @@ -1387,46 +1369,203 @@ mkExp (SOp LPrintNum [arg]) = mkExp (SOp LPrintStr [arg]) = mkSystemOutPrint <$> (mkVarAccess (Just stringType) arg) mkExp (SOp LReadStr [arg]) = mkExp (SForeign LANG_C FString "idris_readStr" [(FPtr, arg)]) -mkExp (SOp (LLt ty) args) = mkLogicalBinOpExp (intTyToClass ty) LThan args -mkExp (SOp (LLe ty) args) = mkLogicalBinOpExp (intTyToClass ty) LThanE args -mkExp (SOp (LEq ty) args) = - mkMethodOpChain1 (mkBoolToNumber (intTyToClass ty)) (intTyToClass ty) "equals" args -mkExp (SOp (LGt ty) args) = mkLogicalBinOpExp (intTyToClass ty) GThan args -mkExp (SOp (LGe ty) args) = mkLogicalBinOpExp (intTyToClass ty) GThanE args -mkExp (SOp (LPlus ty) args) = mkBinOpExp (intTyToClass ty) Add args -mkExp (SOp (LMinus ty) args) = mkBinOpExp (intTyToClass ty) Sub args -mkExp (SOp (LTimes ty) args) = mkBinOpExp (intTyToClass ty) Mult args -mkExp (SOp (LUDiv ty) args) = mkBinOpExpConv (intTyToMethod ty) (intTyToPrimTy ty) (intTyToClass ty) Div args -mkExp (SOp (LSDiv ty) args) = mkBinOpExp (intTyToClass ty) Div args -mkExp (SOp (LURem ty) args) = mkBinOpExpConv (intTyToMethod ty) (intTyToPrimTy ty) (intTyToClass ty) Rem args -mkExp (SOp (LSRem ty) args) = mkBinOpExp (intTyToClass ty) Rem args -mkExp (SOp (LSHL ty) args) = mkBinOpExp (intTyToClass ty) LShift args -mkExp (SOp (LLSHR ty) args) = mkBinOpExp (intTyToClass ty) RRShift args -mkExp (SOp (LASHR ty) args) = mkBinOpExp (intTyToClass ty) RShift args -mkExp (SOp (LAnd ty) args) = mkBinOpExp (intTyToClass ty) And args -mkExp (SOp (LOr ty) args) = mkBinOpExp (intTyToClass ty) Or args -mkExp (SOp (LXOr ty) args) = mkBinOpExp (intTyToClass ty) Xor args -mkExp (SOp (LCompl ty) [var]) = PreBitCompl <$> mkVarAccess (Just $ intTyToClass ty) var -mkExp (SOp (LZExt from to) [var]) - | intTyWidth from < intTyWidth to - = mkZeroExt (intTyToMethod to) (intTyWidth from) (intTyToClass from) (intTyToClass to) var -mkExp (SOp (LSExt from to) [var]) - | intTyWidth from < intTyWidth to - = mkSignedExt (intTyToMethod to) (intTyToClass from) (intTyToClass to) var -mkExp (SOp (LTrunc from to) [var]) - | intTyWidth from > intTyWidth to - = (\ var -> MethodInv $ - TypeMethodCall (J.Name [intTyToIdent to]) +mkExp (SOp LB8Lt args) = mkLogicalBinOpExp byteType LThan args +mkExp (SOp LB8Lte args) = mkLogicalBinOpExp byteType LThanE args +mkExp (SOp LB8Eq args) = + mkMethodOpChain1 (mkBoolToNumber byteType) byteType "equals" args +mkExp (SOp LB8Gt args) = mkLogicalBinOpExp byteType GThan args +mkExp (SOp LB8Gte args) = mkLogicalBinOpExp byteType GThanE args +mkExp (SOp LB8Plus args) = mkBinOpExp byteType Add args +mkExp (SOp LB8Minus args) = mkBinOpExp byteType Sub args +mkExp (SOp LB8Times args) = mkBinOpExp byteType Mult args +mkExp (SOp LB8UDiv args) = mkBinOpExpConv "shortValue" ShortT byteType Div args +mkExp (SOp LB8SDiv args) = mkBinOpExp byteType Div args +mkExp (SOp LB8URem args) = mkBinOpExpConv "shortValue" ShortT byteType Rem args +mkExp (SOp LB8SRem args) = mkBinOpExp byteType Rem args +mkExp (SOp LB8Shl args) = mkBinOpExp byteType LShift args +mkExp (SOp LB8LShr args) = mkBinOpExp byteType RRShift args +mkExp (SOp LB8AShr args) = mkBinOpExp byteType RShift args +mkExp (SOp LB8And args) = mkBinOpExp byteType And args +mkExp (SOp LB8Or args) = mkBinOpExp byteType Or args +mkExp (SOp LB8Xor args) = mkBinOpExp byteType Xor args +mkExp (SOp LB8Compl [var]) = + PreBitCompl <$> mkVarAccess (Just $ byteType) var +mkExp (SOp LB8Z16 [var]) = mkZeroExt "shortValue" 8 byteType shortType var +mkExp (SOp LB8Z32 [var]) = mkZeroExt "intValue" 8 byteType integerType var +mkExp (SOp LB8Z64 [var]) = mkZeroExt "longValue" 8 byteType longType var +mkExp (SOp LB8S16 [var]) = mkSignedExt "shortValue" byteType shortType var +mkExp (SOp LB8S32 [var]) = mkSignedExt "intValue" byteType integerType var +mkExp (SOp LB8S64 [var]) = mkSignedExt "longValue" byteType longType var +mkExp (SOp LB16Lt args) = mkLogicalBinOpExp shortType LThan args +mkExp (SOp LB16Lte args) = mkLogicalBinOpExp shortType LThanE args +mkExp (SOp LB16Eq args) = + mkMethodOpChain1 (mkBoolToNumber shortType) shortType "equals" args +mkExp (SOp LB16Gt args) = mkLogicalBinOpExp shortType GThan args +mkExp (SOp LB16Gte args) = mkLogicalBinOpExp shortType GThanE args +mkExp (SOp LB16Plus args) = mkBinOpExp shortType Add args +mkExp (SOp LB16Minus args) = mkBinOpExp shortType Sub args +mkExp (SOp LB16Times args) = mkBinOpExp shortType Mult args +mkExp (SOp LB16UDiv args) = mkBinOpExpConv "intValue" IntT shortType Div args +mkExp (SOp LB16SDiv args) = mkBinOpExp shortType Div args +mkExp (SOp LB16URem args) = mkBinOpExpConv "intValue" IntT shortType Rem args +mkExp (SOp LB16SRem args) = mkBinOpExp shortType Rem args +mkExp (SOp LB16Shl args) = mkBinOpExp shortType LShift args +mkExp (SOp LB16LShr args) = mkBinOpExp shortType RRShift args +mkExp (SOp LB16AShr args) = mkBinOpExp shortType RShift args +mkExp (SOp LB16And args) = mkBinOpExp shortType And args +mkExp (SOp LB16Or args) = mkBinOpExp shortType Or args +mkExp (SOp LB16Xor args) = mkBinOpExp shortType Xor args +mkExp (SOp LB16Compl [var]) = + PreBitCompl <$> mkVarAccess (Just $ shortType) var +mkExp (SOp LB16Z32 [var]) = mkZeroExt "intValue" 16 shortType integerType var +mkExp (SOp LB16Z64 [var]) = mkZeroExt "longValue" 16 shortType longType var +mkExp (SOp LB16S32 [var]) = mkSignedExt "intValue" shortType integerType var +mkExp (SOp LB16S64 [var]) = mkSignedExt "longValue" shortType longType var +mkExp (SOp LB16T8 [var]) = + (\ var -> MethodInv $ + TypeMethodCall (J.Name [Ident "Byte"]) [] (Ident "valueOf") [ MethodInv - $ PrimaryMethodCall var [] (Ident (intTyToMethod to)) [] ] + $ PrimaryMethodCall var [] (Ident "byteValue") [] ] + ) + <$> mkVarAccess (Just $ shortType) var +mkExp (SOp LB32Lt args) = mkLogicalBinOpExp integerType LThan args +mkExp (SOp LB32Lte args) = mkLogicalBinOpExp integerType LThanE args +mkExp (SOp LB32Eq args) = + mkMethodOpChain1 (mkBoolToNumber integerType) integerType "equals" args +mkExp (SOp LB32Gt args) = mkLogicalBinOpExp integerType GThan args +mkExp (SOp LB32Gte args) = mkLogicalBinOpExp integerType GThanE args +mkExp (SOp LB32Plus args) = mkBinOpExp integerType Add args +mkExp (SOp LB32Minus args) = mkBinOpExp integerType Sub args +mkExp (SOp LB32Times args) = mkBinOpExp integerType Mult args +mkExp (SOp LB32UDiv args) = mkBinOpExpConv "longValue" LongT integerType Div args +mkExp (SOp LB32SDiv args) = mkBinOpExp integerType Div args +mkExp (SOp LB32URem args) = mkBinOpExpConv "longValue" LongT integerType Rem args +mkExp (SOp LB32SRem args) = mkBinOpExp integerType Rem args +mkExp (SOp LB32Shl args) = mkBinOpExp integerType LShift args +mkExp (SOp LB32LShr args) = mkBinOpExp integerType RRShift args +mkExp (SOp LB32AShr args) = mkBinOpExp integerType RShift args +mkExp (SOp LB32And args) = mkBinOpExp integerType And args +mkExp (SOp LB32Or args) = mkBinOpExp integerType Or args +mkExp (SOp LB32Xor args) = mkBinOpExp integerType Xor args +mkExp (SOp LB32Compl [var]) = + PreBitCompl <$> mkVarAccess (Just $ integerType) var +mkExp (SOp LB32Z64 [var]) = mkZeroExt "longValue" 32 integerType longType var +mkExp (SOp LB32S64 [var]) = mkSignedExt "longValue" integerType longType var +mkExp (SOp LB32T8 [var]) = + (\ var -> MethodInv + $ TypeMethodCall (J.Name [Ident "Byte"]) + [] + (Ident "valueOf") + [ MethodInv $ PrimaryMethodCall var [] (Ident "byteValue") [] ] + ) + <$> mkVarAccess (Just $ integerType) var +mkExp (SOp LB32T16 [var]) = + (\ var -> MethodInv + $ TypeMethodCall (J.Name [Ident "Short"]) + [] + (Ident "valueOf") + [ MethodInv $ PrimaryMethodCall var [] (Ident "shortValue") [] ] + ) + <$> mkVarAccess (Just $ integerType) var +mkExp (SOp LB64Lt args) = mkLogicalBinOpExp longType LThan args +mkExp (SOp LB64Lte args) = mkLogicalBinOpExp longType LThanE args +mkExp (SOp LB64Eq args) = + mkMethodOpChain1 (mkBoolToNumber longType) longType "equals" args +mkExp (SOp LB64Gt args) = mkLogicalBinOpExp longType GThan args +mkExp (SOp LB64Gte args) = mkLogicalBinOpExp longType GThanE args +mkExp (SOp LB64Plus args) = mkBinOpExp longType Add args +mkExp (SOp LB64Minus args) = mkBinOpExp longType Sub args +mkExp (SOp LB64Times args) = mkBinOpExp longType Mult args +mkExp (SOp LB64UDiv (arg:args)) = do + (arg:args) <- mapM (mkVarAccess (Just longType)) (arg:args) + return $ foldl' (\ exp arg -> + MethodInv $ PrimaryMethodCall + ( MethodInv $ PrimaryMethodCall + ( MethodInv $ TypeMethodCall (J.Name [Ident "BigInteger"]) + [] + (Ident "valueOf") + [ exp ] + ) + [] + (Ident "divide") + [ MethodInv $ TypeMethodCall (J.Name [Ident "BigInteger"]) + [] + (Ident "valueOf") + [ arg ] + ] + ) + [] + (Ident "longValue") + [] + ) + arg + args +mkExp (SOp LB64SDiv args) = mkBinOpExp longType Div args +mkExp (SOp LB64URem (arg:args)) = do + (arg:args) <- mapM (mkVarAccess (Just longType)) (arg:args) + return $ foldl' (\ exp arg -> + MethodInv $ PrimaryMethodCall + ( MethodInv $ PrimaryMethodCall + ( MethodInv $ TypeMethodCall (J.Name [Ident "BigInteger"]) + [] + (Ident "valueOf") + [ exp ] + ) + [] + (Ident "remainder") + [ MethodInv $ TypeMethodCall (J.Name [Ident "BigInteger"]) + [] + (Ident "valueOf") + [ arg ] + ] + ) + [] + (Ident "longValue") + [] + ) + arg + args +mkExp (SOp LB64SRem args) = mkBinOpExp longType Rem args +mkExp (SOp LB64Shl args) = mkBinOpExp longType LShift args +mkExp (SOp LB64LShr args) = mkBinOpExp longType RRShift args +mkExp (SOp LB64AShr args) = mkBinOpExp longType RShift args +mkExp (SOp LB64And args) = mkBinOpExp longType And args +mkExp (SOp LB64Or args) = mkBinOpExp longType Or args +mkExp (SOp LB64Xor args) = mkBinOpExp longType Xor args +mkExp (SOp LB64Compl [var]) = + PreBitCompl <$> mkVarAccess (Just $ longType) var +mkExp (SOp LB64T8 [var]) = + (\ var -> MethodInv + $ TypeMethodCall (J.Name [Ident "Byte"]) + [] + (Ident "valueOf") + [ MethodInv + $ PrimaryMethodCall var [] (Ident "byteValue") [] + ] + ) + <$> mkVarAccess (Just $ longType) var +mkExp (SOp LB64T16 [var]) = + (\ var -> MethodInv + $ TypeMethodCall (J.Name [Ident "Short"]) + [] + (Ident "valueOf") + [ MethodInv $ PrimaryMethodCall var [] (Ident "shortValue") [] ] + ) + <$> mkVarAccess (Just $ longType) var +mkExp (SOp LB64T32 [var]) = + (\ var -> MethodInv + $ TypeMethodCall (J.Name [Ident "Integer"]) + [] + (Ident "valueOf") + [ MethodInv $ PrimaryMethodCall var [] (Ident "intValue") [] ] ) - <$> mkVarAccess (Just $ intTyToClass from) var -mkExp (SOp LIntB8 [arg]) = mkExp (SOp (LTrunc IT32 IT8) [arg]) -mkExp (SOp LIntB16 [arg]) = mkExp (SOp (LTrunc IT32 IT16) [arg]) + <$> mkVarAccess (Just $ longType) var +mkExp (SOp LIntB8 [arg]) = mkExp (SOp LB32T8 [arg]) +mkExp (SOp LIntB16 [arg]) = mkExp (SOp LB32T16 [arg]) mkExp (SOp LIntB32 [arg]) = mkExp (SV arg) -mkExp (SOp LIntB64 [arg]) = mkExp (SOp (LSExt IT32 IT64) [arg]) +mkExp (SOp LIntB64 [arg]) = mkExp (SOp LB32S64 [arg]) mkExp (SOp LB32Int [arg]) = mkExp (SV arg) mkExp (SOp LFExp [arg]) = mkMathFun "exp" arg mkExp (SOp LFLog [arg]) = mkMathFun "log" arg diff --git a/src/IRTS/CodegenJavaScript.hs b/src/IRTS/CodegenJavaScript.hs index ffceb8b24d..bb4e3a4bc3 100644 --- a/src/IRTS/CodegenJavaScript.hs +++ b/src/IRTS/CodegenJavaScript.hs @@ -203,37 +203,37 @@ translateExpression (SApp True name vars) = ++ ";\n});" translateExpression (SOp op vars) - | (LPlus _) <- op + | LPlus <- op , (lhs:rhs:_) <- vars = translateBinaryOp "+" lhs rhs - | (LMinus _) <- op + | LMinus <- op , (lhs:rhs:_) <- vars = translateBinaryOp "-" lhs rhs - | (LTimes _) <- op + | LTimes <- op , (lhs:rhs:_) <- vars = translateBinaryOp "*" lhs rhs - | (LSDiv _) <- op + | LDiv <- op , (lhs:rhs:_) <- vars = translateBinaryOp "/" lhs rhs - | (LSRem _) <- op + | LMod <- op , (lhs:rhs:_) <- vars = translateBinaryOp "%" lhs rhs - | (LEq _) <- op + | LEq <- op , (lhs:rhs:_) <- vars = translateBinaryOp "==" lhs rhs - | (LLt _) <- op + | LLt <- op , (lhs:rhs:_) <- vars = translateBinaryOp "<" lhs rhs - | (LLe _) <- op + | LLe <- op , (lhs:rhs:_) <- vars = translateBinaryOp "<=" lhs rhs - | (LGt _) <- op + | LGt <- op , (lhs:rhs:_) <- vars = translateBinaryOp ">" lhs rhs - | (LGe _) <- op + | LGe <- op , (lhs:rhs:_) <- vars = translateBinaryOp ">=" lhs rhs - | (LAnd _) <- op + | LAnd <- op , (lhs:rhs:_) <- vars = translateBinaryOp "&" lhs rhs - | (LOr _) <- op + | LOr <- op , (lhs:rhs:_) <- vars = translateBinaryOp "|" lhs rhs - | (LXOr _) <- op + | LXOr <- op , (lhs:rhs:_) <- vars = translateBinaryOp "^" lhs rhs - | (LSHL _) <- op + | LSHL <- op , (lhs:rhs:_) <- vars = translateBinaryOp "<<" rhs lhs - | (LASHR _) <- op + | LSHR <- op , (lhs:rhs:_) <- vars = translateBinaryOp ">>" rhs lhs - | (LCompl _) <- op + | LCompl <- op , (arg:_) <- vars = '~' : translateVariableName arg | LBPlus <- op diff --git a/src/IRTS/Compiler.hs b/src/IRTS/Compiler.hs index 2dd27f905f..083a4a8061 100644 --- a/src/IRTS/Compiler.hs +++ b/src/IRTS/Compiler.hs @@ -284,21 +284,15 @@ getFTypes tm = case unApply tm of _ -> Nothing mkIty' (P _ (UN ty) _) = mkIty ty -mkIty' (App (P _ (UN "FIntT") _) (P _ (UN intTy) _)) = mkIntIty intTy mkIty' _ = FAny +mkIty "FInt" = FInt mkIty "FFloat" = FDouble mkIty "FChar" = FChar mkIty "FString" = FString mkIty "FPtr" = FPtr mkIty "FUnit" = FUnit -mkIntIty "ITNative" = FInt ITNative -mkIntIty "IT8" = FInt IT8 -mkIntIty "IT16" = FInt IT16 -mkIntIty "IT32" = FInt IT32 -mkIntIty "IT64" = FInt IT64 - zname = NS (UN "O") ["Nat","Prelude"] sname = NS (UN "S") ["Nat","Prelude"] diff --git a/src/IRTS/LParser.hs b/src/IRTS/LParser.hs index bd6eabdb75..d2e834878b 100644 --- a/src/IRTS/LParser.hs +++ b/src/IRTS/LParser.hs @@ -92,16 +92,16 @@ pLDecl = do reserved "data" pLExp = buildExpressionParser optable pLExp' -optable = [[binary "*" (\x y -> LOp (LTimes ITNative) [x,y]) AssocLeft, - binary "/" (\x y -> LOp (LSDiv ITNative) [x,y]) AssocLeft, +optable = [[binary "*" (\x y -> LOp LTimes [x,y]) AssocLeft, + binary "/" (\x y -> LOp LDiv [x,y]) AssocLeft, binary "*." (\x y -> LOp LFTimes [x,y]) AssocLeft, binary "/." (\x y -> LOp LFDiv [x,y]) AssocLeft, binary "*:" (\x y -> LOp LBTimes [x,y]) AssocLeft, binary "/:" (\x y -> LOp LBDiv [x,y]) AssocLeft ], [ - binary "+" (\x y -> LOp (LPlus ITNative) [x,y]) AssocLeft, - binary "-" (\x y -> LOp (LMinus ITNative) [x,y]) AssocLeft, + binary "+" (\x y -> LOp LPlus [x,y]) AssocLeft, + binary "-" (\x y -> LOp LMinus [x,y]) AssocLeft, binary "++" (\x y -> LOp LStrConcat [x,y]) AssocLeft, binary "+." (\x y -> LOp LFPlus [x,y]) AssocLeft, binary "-." (\x y -> LOp LFMinus [x,y]) AssocLeft, @@ -109,15 +109,15 @@ optable = [[binary "*" (\x y -> LOp (LTimes ITNative) [x,y]) AssocLeft, binary "-:" (\x y -> LOp LBMinus [x,y]) AssocLeft ], [ - binary "==" (\x y -> LOp (LEq ITNative) [x, y]) AssocNone, + binary "==" (\x y -> LOp LEq [x, y]) AssocNone, binary "==." (\x y -> LOp LFEq [x, y]) AssocNone, - binary "<" (\x y -> LOp (LLt ITNative) [x, y]) AssocNone, + binary "<" (\x y -> LOp LLt [x, y]) AssocNone, binary "<." (\x y -> LOp LFLt [x, y]) AssocNone, - binary ">" (\x y -> LOp (LGt ITNative) [x, y]) AssocNone, + binary ">" (\x y -> LOp LGt [x, y]) AssocNone, binary ">." (\x y -> LOp LFGt [x, y]) AssocNone, - binary "<=" (\x y -> LOp (LLe ITNative) [x, y]) AssocNone, + binary "<=" (\x y -> LOp LLe [x, y]) AssocNone, binary "<=." (\x y -> LOp LFLe [x, y]) AssocNone, - binary ">=" (\x y -> LOp (LGe ITNative) [x, y]) AssocNone, + binary ">=" (\x y -> LOp LGe [x, y]) AssocNone, binary ">=." (\x y -> LOp LFGe [x, y]) AssocNone, binary "==:" (\x y -> LOp LBEq [x, y]) AssocNone, @@ -164,7 +164,7 @@ pLExp' = try (do lchar '%'; pCast) pLang = do reserved "C"; return LANG_C -pType = do reserved "Int"; return (FInt ITNative) +pType = do reserved "Int"; return FInt <|> do reserved "Float"; return FDouble <|> do reserved "String"; return FString <|> do reserved "Unit"; return FUnit diff --git a/src/IRTS/Lang.hs b/src/IRTS/Lang.hs index 8ff3948451..cc3563c02f 100644 --- a/src/IRTS/Lang.hs +++ b/src/IRTS/Lang.hs @@ -28,29 +28,9 @@ data LExp = LV LVar -- Primitive operators. Backends are not *required* to implement all -- of these, but should report an error if they are unable -data IntTy = ITNative | IT8 | IT16 | IT32 | IT64 - deriving (Show, Eq) - -intTyWidth :: IntTy -> Int -intTyWidth IT8 = 8 -intTyWidth IT16 = 16 -intTyWidth IT32 = 32 -intTyWidth IT64 = 64 -intTyWidth ITNative = error "IRTS.Lang.intTyWidth: Native integer width is unspecified" - -intTyToConst :: IntTy -> Const -intTyToConst IT8 = B8Type -intTyToConst IT16 = B16Type -intTyToConst IT32 = B32Type -intTyToConst IT64 = B64Type -intTyToConst ITNative = IType - -data PrimFn = LPlus IntTy | LMinus IntTy | LTimes IntTy - | LUDiv IntTy | LSDiv IntTy | LURem IntTy | LSRem IntTy - | LAnd IntTy | LOr IntTy | LXOr IntTy | LCompl IntTy - | LSHL IntTy | LLSHR IntTy | LASHR IntTy - | LEq IntTy | LLt IntTy | LLe IntTy | LGt IntTy | LGe IntTy - | LSExt IntTy IntTy | LZExt IntTy IntTy | LTrunc IntTy IntTy +data PrimFn = LPlus | LMinus | LTimes | LDiv | LMod + | LAnd | LOr | LXOr | LCompl | LSHL| LSHR + | LEq | LLt | LLe | LGt | LGe | LFPlus | LFMinus | LFTimes | LFDiv | LFEq | LFLt | LFLe | LFGt | LFGe | LBPlus | LBMinus | LBDec | LBTimes | LBDiv | LBMod @@ -60,6 +40,22 @@ data PrimFn = LPlus IntTy | LMinus IntTy | LTimes IntTy | LIntBig | LBigInt | LStrBig | LBigStr | LChInt | LIntCh | LPrintNum | LPrintStr | LReadStr + | LB8Lt | LB8Lte | LB8Eq | LB8Gt | LB8Gte + | LB8Plus | LB8Minus | LB8Times | LB8UDiv | LB8SDiv | LB8URem | LB8SRem + | LB8Shl | LB8LShr | LB8AShr | LB8And | LB8Or | LB8Xor | LB8Compl + | LB8Z16 | LB8Z32 | LB8Z64 | LB8S16 | LB8S32 | LB8S64 -- Zero/Sign extension + | LB16Lt | LB16Lte | LB16Eq | LB16Gt | LB16Gte + | LB16Plus | LB16Minus | LB16Times | LB16UDiv | LB16SDiv | LB16URem | LB16SRem + | LB16Shl | LB16LShr | LB16AShr | LB16And | LB16Or | LB16Xor | LB16Compl + | LB16Z32 | LB16Z64 | LB16S32 | LB16S64 | LB16T8 -- and Truncation + | LB32Lt | LB32Lte | LB32Eq | LB32Gt | LB32Gte + | LB32Plus | LB32Minus | LB32Times | LB32UDiv | LB32SDiv | LB32URem | LB32SRem + | LB32Shl | LB32LShr | LB32AShr | LB32And | LB32Or | LB32Xor | LB32Compl + | LB32Z64 | LB32S64 | LB32T8 | LB32T16 + | LB64Lt | LB64Lte | LB64Eq | LB64Gt | LB64Gte + | LB64Plus | LB64Minus | LB64Times | LB64UDiv | LB64SDiv | LB64URem | LB64SRem + | LB64Shl | LB64LShr | LB64AShr | LB64And | LB64Or | LB64Xor | LB64Compl + | LB64T8 | LB64T16 | LB64T32 | LIntB8 | LIntB16 | LIntB32 | LIntB64 | LB32Int | LFExp | LFLog | LFSin | LFCos | LFTan | LFASin | LFACos | LFATan @@ -80,7 +76,7 @@ data PrimFn = LPlus IntTy | LMinus IntTy | LTimes IntTy data FLang = LANG_C deriving (Show, Eq) -data FType = FInt IntTy | FChar | FString | FUnit | FPtr | FDouble | FAny +data FType = FInt | FChar | FString | FUnit | FPtr | FDouble | FAny deriving (Show, Eq) data LAlt = LConCase Int Name [Name] LExp diff --git a/src/Idris/Primitives.hs b/src/Idris/Primitives.hs index 206971ad98..b18e9f4989 100644 --- a/src/Idris/Primitives.hs +++ b/src/Idris/Primitives.hs @@ -35,68 +35,247 @@ partial = Partial NotCovering primitives = -- operators [Prim (UN "prim__addInt") (ty [IType, IType] IType) 2 (iBin (+)) - (2, LPlus ITNative) total, + (2, LPlus) total, Prim (UN "prim__subInt") (ty [IType, IType] IType) 2 (iBin (-)) - (2, LMinus ITNative) total, + (2, LMinus) total, Prim (UN "prim__mulInt") (ty [IType, IType] IType) 2 (iBin (*)) - (2, LTimes ITNative) total, + (2, LTimes) total, Prim (UN "prim__divInt") (ty [IType, IType] IType) 2 (iBin div) - (2, LSDiv ITNative) partial, + (2, LDiv) partial, Prim (UN "prim__modInt") (ty [IType, IType] IType) 2 (iBin mod) - (2, LSRem ITNative) partial, + (2, LMod) partial, Prim (UN "prim__andInt") (ty [IType, IType] IType) 2 (iBin (.&.)) - (2, LAnd ITNative) partial, + (2, LAnd) partial, Prim (UN "prim__orInt") (ty [IType, IType] IType) 2 (iBin (.|.)) - (2, LOr ITNative) partial, + (2, LOr) partial, Prim (UN "prim__xorInt") (ty [IType, IType] IType) 2 (iBin xor) - (2, LXOr ITNative) partial, + (2, LXOr) partial, Prim (UN "prim__shLInt") (ty [IType, IType] IType) 2 (iBin shiftL) - (2, LSHL ITNative) partial, + (2, LSHL) partial, Prim (UN "prim__shRInt") (ty [IType, IType] IType) 2 (iBin shiftR) - (2, LASHR ITNative) partial, + (2, LSHR) partial, Prim (UN "prim__complInt") (ty [IType] IType) 1 (iUn complement) - (1, LCompl ITNative) partial, + (1, LCompl) partial, Prim (UN "prim__eqInt") (ty [IType, IType] IType) 2 (biBin (==)) - (2, LEq ITNative) total, + (2, LEq) total, Prim (UN "prim__ltInt") (ty [IType, IType] IType) 2 (biBin (<)) - (2, LLt ITNative) total, + (2, LLt) total, Prim (UN "prim__lteInt") (ty [IType, IType] IType) 2 (biBin (<=)) - (2, LLe ITNative) total, + (2, LLe) total, Prim (UN "prim__gtInt") (ty [IType, IType] IType) 2 (biBin (>)) - (2, LGt ITNative) total, + (2, LGt) total, Prim (UN "prim__gteInt") (ty [IType, IType] IType) 2 (biBin (>=)) - (2, LGe ITNative) total, + (2, LGe) total, Prim (UN "prim__eqChar") (ty [ChType, ChType] IType) 2 (bcBin (==)) - (2, LEq ITNative) total, + (2, LEq) total, Prim (UN "prim__ltChar") (ty [ChType, ChType] IType) 2 (bcBin (<)) - (2, LLt ITNative) total, + (2, LLt) total, Prim (UN "prim__lteChar") (ty [ChType, ChType] IType) 2 (bcBin (<=)) - (2, LLe ITNative) total, + (2, LLe) total, Prim (UN "prim__gtChar") (ty [ChType, ChType] IType) 2 (bcBin (>)) - (2, LGt ITNative) total, + (2, LGt) total, Prim (UN "prim__gteChar") (ty [ChType, ChType] IType) 2 (bcBin (>=)) - (2, LGe ITNative) total, - - iCoerce IT8 IT16 "zext" zext LZExt, - iCoerce IT8 IT32 "zext" zext LZExt, - iCoerce IT8 IT64 "zext" zext LZExt, - iCoerce IT16 IT32 "zext" zext LZExt, - iCoerce IT16 IT64 "zext" zext LZExt, - iCoerce IT32 IT64 "zext" zext LZExt, - - iCoerce IT8 IT16 "sext" sext LSExt, - iCoerce IT8 IT32 "sext" sext LSExt, - iCoerce IT8 IT64 "sext" sext LSExt, - iCoerce IT16 IT32 "sext" sext LSExt, - iCoerce IT16 IT64 "sext" sext LSExt, - iCoerce IT32 IT64 "sext" sext LSExt, - - iCoerce IT16 IT8 "trunc" trunc LTrunc, - iCoerce IT32 IT8 "trunc" trunc LTrunc, - iCoerce IT64 IT8 "trunc" trunc LTrunc, - iCoerce IT32 IT16 "trunc" trunc LTrunc, - iCoerce IT64 IT16 "trunc" trunc LTrunc, - iCoerce IT64 IT32 "trunc" trunc LTrunc, + (2, LGe) total, + + Prim (UN "prim__ltB8") (ty [B8Type, B8Type] IType) 2 (b8Cmp (<)) + (2, LB8Lt) total, + Prim (UN "prim__lteB8") (ty [B8Type, B8Type] IType) 2 (b8Cmp (<=)) + (2, LB8Lte) total, + Prim (UN "prim__eqB8") (ty [B8Type, B8Type] IType) 2 (b8Cmp (==)) + (2, LB8Eq) total, + Prim (UN "prim__gteB8") (ty [B8Type, B8Type] IType) 2 (b8Cmp (>)) + (2, LB8Gte) total, + Prim (UN "prim__gtB8") (ty [B8Type, B8Type] IType) 2 (b8Cmp (>=)) + (2, LB8Gt) total, + Prim (UN "prim__addB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin (+)) + (2, LB8Plus) total, + Prim (UN "prim__subB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin (-)) + (2, LB8Minus) total, + Prim (UN "prim__mulB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin (*)) + (2, LB8Times) total, + Prim (UN "prim__udivB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin div) + (2, LB8UDiv) partial, + Prim (UN "prim__sdivB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin s8div) + (2, LB8SDiv) partial, + Prim (UN "prim__uremB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin rem) + (2, LB8URem) partial, + Prim (UN "prim__sremB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin s8rem) + (2, LB8SRem) partial, + Prim (UN "prim__shlB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin (\x y -> shiftL x (fromIntegral y))) + (2, LB8Shl) total, + Prim (UN "prim__lshrB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin (\x y -> shiftR x (fromIntegral y))) + (2, LB8AShr) total, + Prim (UN "prim__ashrB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin (\x y -> fromIntegral $ + shiftR (fromIntegral x :: Int8) + (fromIntegral y))) + (2, LB8LShr) total, + Prim (UN "prim__andB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin (.&.)) + (2, LB8And) total, + Prim (UN "prim__orB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin (.|.)) + (2, LB8Or) total, + Prim (UN "prim__xorB8") (ty [B8Type, B8Type] B8Type) 2 (b8Bin xor) + (2, LB8Xor) total, + Prim (UN "prim__complB8") (ty [B8Type] B8Type) 1 (b8Un complement) + (1, LB8Compl) total, + Prim (UN "prim__zextB8_16") (ty [B8Type] B16Type) 1 zext8_16 + (1, LB8Z16) total, + Prim (UN "prim__zextB8_32") (ty [B8Type] B32Type) 1 zext8_32 + (1, LB8Z32) total, + Prim (UN "prim__zextB8_64") (ty [B8Type] B64Type) 1 zext8_64 + (1, LB8Z64) total, + Prim (UN "prim__sextB8_16") (ty [B8Type] B16Type) 1 sext8_16 + (1, LB8Z16) total, + Prim (UN "prim__sextB8_32") (ty [B8Type] B32Type) 1 sext8_32 + (1, LB8Z32) total, + Prim (UN "prim__sextB8_64") (ty [B8Type] B64Type) 1 sext8_64 + (1, LB8Z64) total, + + Prim (UN "prim__ltB16") (ty [B16Type, B16Type] IType) 2 (b16Cmp (<)) + (2, LB16Lt) total, + Prim (UN "prim__lteB16") (ty [B16Type, B16Type] IType) 2 (b16Cmp (<=)) + (2, LB16Lte) total, + Prim (UN "prim__eqB16") (ty [B16Type, B16Type] IType) 2 (b16Cmp (==)) + (2, LB16Eq) total, + Prim (UN "prim__gteB16") (ty [B16Type, B16Type] IType) 2 (b16Cmp (>)) + (2, LB16Gte) total, + Prim (UN "prim__gtB16") (ty [B16Type, B16Type] IType) 2 (b16Cmp (>=)) + (2, LB16Gt) total, + Prim (UN "prim__addB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin (+)) + (2, LB16Plus) total, + Prim (UN "prim__subB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin (-)) + (2, LB16Minus) total, + Prim (UN "prim__mulB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin (*)) + (2, LB16Times) total, + Prim (UN "prim__udivB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin div) + (2, LB16UDiv) partial, + Prim (UN "prim__sdivB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin s16div) + (2, LB16SDiv) partial, + Prim (UN "prim__uremB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin rem) + (2, LB16URem) partial, + Prim (UN "prim__sremB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin s16rem) + (2, LB16SRem) partial, + Prim (UN "prim__shlB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin (\x y -> shiftL x (fromIntegral y))) + (2, LB16Shl) total, + Prim (UN "prim__lshrB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin (\x y -> shiftR x (fromIntegral y))) + (2, LB16AShr) total, + Prim (UN "prim__ashrB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin (\x y -> fromIntegral $ + shiftR (fromIntegral x :: Int16) + (fromIntegral y))) + (2, LB16LShr) total, + Prim (UN "prim__andB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin (.&.)) + (2, LB16And) total, + Prim (UN "prim__orB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin (.|.)) + (2, LB16Or) total, + Prim (UN "prim__xorB16") (ty [B16Type, B16Type] B16Type) 2 (b16Bin xor) + (2, LB16Xor) total, + Prim (UN "prim__complB16") (ty [B16Type] B16Type) 1 (b16Un complement) + (1, LB16Compl) total, + Prim (UN "prim__zextB16_32") (ty [B16Type] B32Type) 1 zext16_32 + (1, LB16Z32) total, + Prim (UN "prim__zextB16_64") (ty [B16Type] B64Type) 1 zext16_64 + (1, LB16Z64) total, + Prim (UN "prim__sextB16_32") (ty [B16Type] B32Type) 1 sext16_32 + (1, LB16Z32) total, + Prim (UN "prim__sextB16_64") (ty [B16Type] B64Type) 1 sext16_64 + (1, LB16Z64) total, + Prim (UN "prim__truncB16_8") (ty [B16Type] B8Type) 1 trunc16_8 + (1, LB16T8) total, + + Prim (UN "prim__ltB32") (ty [B32Type, B32Type] IType) 2 (b32Cmp (<)) + (2, LB32Lt) total, + Prim (UN "prim__lteB32") (ty [B32Type, B32Type] IType) 2 (b32Cmp (<=)) + (2, LB32Lte) total, + Prim (UN "prim__eqB32") (ty [B32Type, B32Type] IType) 2 (b32Cmp (==)) + (2, LB32Eq) total, + Prim (UN "prim__gteB32") (ty [B32Type, B32Type] IType) 2 (b32Cmp (>)) + (2, LB32Gte) total, + Prim (UN "prim__gtB32") (ty [B32Type, B32Type] IType) 2 (b32Cmp (>=)) + (2, LB32Gt) total, + Prim (UN "prim__addB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin (+)) + (2, LB32Plus) total, + Prim (UN "prim__subB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin (-)) + (2, LB32Minus) total, + Prim (UN "prim__mulB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin (*)) + (2, LB32Times) total, + Prim (UN "prim__udivB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin div) + (2, LB32UDiv) partial, + Prim (UN "prim__sdivB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin s32div) + (2, LB32SDiv) partial, + Prim (UN "prim__uremB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin rem) + (2, LB32URem) partial, + Prim (UN "prim__sremB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin s32rem) + (2, LB32SRem) partial, + Prim (UN "prim__shlB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin (\x y -> shiftL x (fromIntegral y))) + (2, LB32Shl) total, + Prim (UN "prim__lshrB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin (\x y -> shiftR x (fromIntegral y))) + (2, LB32AShr) total, + Prim (UN "prim__ashrB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin (\x y -> fromIntegral $ + shiftR (fromIntegral x :: Int32) + (fromIntegral y))) + (2, LB32LShr) total, + Prim (UN "prim__andB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin (.&.)) + (2, LB32And) total, + Prim (UN "prim__orB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin (.|.)) + (2, LB32Or) total, + Prim (UN "prim__xorB32") (ty [B32Type, B32Type] B32Type) 2 (b32Bin xor) + (2, LB32Xor) total, + Prim (UN "prim__complB32") (ty [B32Type] B32Type) 1 (b32Un complement) + (1, LB32Compl) total, + Prim (UN "prim__zextB32_64") (ty [B32Type] B64Type) 1 zext32_64 + (1, LB32Z64) total, + Prim (UN "prim__sextB32_64") (ty [B32Type] B64Type) 1 sext32_64 + (1, LB32Z64) total, + Prim (UN "prim__truncB32_8") (ty [B32Type] B8Type) 1 trunc32_8 + (1, LB32T8) total, + Prim (UN "prim__truncB32_16") (ty [B32Type] B16Type) 1 trunc32_16 + (1, LB32T16) total, + + Prim (UN "prim__ltB64") (ty [B64Type, B64Type] IType) 2 (b64Cmp (<)) + (2, LB64Lt) total, + Prim (UN "prim__lteB64") (ty [B64Type, B64Type] IType) 2 (b64Cmp (<=)) + (2, LB64Lte) total, + Prim (UN "prim__eqB64") (ty [B64Type, B64Type] IType) 2 (b64Cmp (==)) + (2, LB64Eq) total, + Prim (UN "prim__gteB64") (ty [B64Type, B64Type] IType) 2 (b64Cmp (>)) + (2, LB64Gte) total, + Prim (UN "prim__gtB64") (ty [B64Type, B64Type] IType) 2 (b64Cmp (>=)) + (2, LB64Gt) total, + Prim (UN "prim__addB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin (+)) + (2, LB64Plus) total, + Prim (UN "prim__subB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin (-)) + (2, LB64Minus) total, + Prim (UN "prim__mulB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin (*)) + (2, LB64Times) total, + Prim (UN "prim__udivB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin div) + (2, LB64UDiv) partial, + Prim (UN "prim__sdivB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin s64div) + (2, LB64SDiv) partial, + Prim (UN "prim__uremB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin rem) + (2, LB64URem) partial, + Prim (UN "prim__sremB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin s64rem) + (2, LB64SRem) partial, + Prim (UN "prim__shlB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin (\x y -> shiftL x (fromIntegral y))) + (2, LB64Shl) total, + Prim (UN "prim__lshrB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin (\x y -> shiftR x (fromIntegral y))) + (2, LB64AShr) total, + Prim (UN "prim__ashrB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin (\x y -> fromIntegral $ + shiftR (fromIntegral x :: Int64) + (fromIntegral y))) + (2, LB64LShr) total, + Prim (UN "prim__andB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin (.&.)) + (2, LB64And) total, + Prim (UN "prim__orB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin (.|.)) + (2, LB64Or) total, + Prim (UN "prim__xorB64") (ty [B64Type, B64Type] B64Type) 2 (b64Bin xor) + (2, LB64Xor) total, + Prim (UN "prim__complB64") (ty [B64Type] B64Type) 1 (b64Un complement) + (1, LB64Compl) total, + Prim (UN "prim__truncB64_8") (ty [B64Type] B8Type) 1 trunc64_8 + (1, LB64T8) total, + Prim (UN "prim__truncB64_16") (ty [B64Type] B16Type) 1 trunc64_16 + (1, LB64T16) total, + Prim (UN "prim__truncB64_32") (ty [B64Type] B32Type) 1 trunc64_32 + (1, LB64T32) total, Prim (UN "prim__intToB8") (ty [IType] B8Type) 1 intToBits8 (1, LIntB8) total, @@ -223,36 +402,7 @@ primitives = (0, LStdIn) partial, Prim (UN "prim__believe_me") believeTy 3 (p_believeMe) (3, LNoOp) partial -- ahem - ] ++ concatMap intOps [IT8, IT16, IT32, IT64] - -intOps ity = - [ iCmp ity "lt" (bCmp ity (<)) LLt total - , iCmp ity "lte" (bCmp ity (<=)) LLe total - , iCmp ity "eq" (bCmp ity (==)) LEq total - , iCmp ity "gte" (bCmp ity (>=)) LGe total - , iCmp ity "gt" (bCmp ity (>)) LGt total - , iBinOp ity "add" (bitBin ity (+)) LPlus total - , iBinOp ity "sub" (bitBin ity (-)) LMinus total - , iBinOp ity "mul" (bitBin ity (*)) LTimes total - , iBinOp ity "udiv" (bitBin ity div) LUDiv partial - , iBinOp ity "sdiv" (bsdiv ity) LSDiv partial - , iBinOp ity "urem" (bitBin ity rem) LURem partial - , iBinOp ity "srem" (bsrem ity) LSRem partial - , iBinOp ity "shl" (bitBin ity (\x y -> shiftL x (fromIntegral y))) LSHL total - , iBinOp ity "lshr" (bitBin ity (\x y -> shiftR x (fromIntegral y))) LLSHR total - , iBinOp ity "ashr" (bashr ity) LASHR total - , iBinOp ity "and" (bitBin ity (.&.)) LAnd total - , iBinOp ity "or" (bitBin ity (.|.)) LOr total - , iBinOp ity "xor" (bitBin ity (xor)) LXOr total - , iUnOp ity "compl" (bUn ity complement) LCompl total - ] - -iCmp ity op impl irop totality = Prim (UN $ "prim__" ++ op ++ "B" ++ show (intTyWidth ity)) (ty (replicate 2 $ intTyToConst ity) IType) 2 impl (2, irop ity) totality -iBinOp ity op impl irop totality = Prim (UN $ "prim__" ++ op ++ "B" ++ show (intTyWidth ity)) (ty (replicate 2 $ intTyToConst ity) (intTyToConst ity)) 2 impl (2, irop ity) totality -iUnOp ity op impl irop totality = Prim (UN $ "prim__" ++ op ++ "B" ++ show (intTyWidth ity)) (ty [intTyToConst ity] (intTyToConst ity)) 1 impl (1, irop ity) totality -iCoerce from to op impl irop = - Prim (UN $ "prim__" ++ op ++ "B" ++ show (intTyWidth from) ++ "_" ++ show (intTyWidth to)) - (ty [intTyToConst from] (intTyToConst to)) 1 (impl from to) (1, irop from to) total + ] p_believeMe [_,_,x] = Just x p_believeMe _ = Nothing @@ -291,96 +441,132 @@ bsBin _ _ = Nothing sBin op [VConstant (Str x), VConstant (Str y)] = Just $ VConstant (Str (op x y)) sBin _ _ = Nothing -bsrem IT8 [VConstant (B8 x), VConstant (B8 y)] - = Just $ VConstant (B8 (fromIntegral (fromIntegral x `rem` fromIntegral y :: Int8))) -bsrem IT16 [VConstant (B16 x), VConstant (B16 y)] - = Just $ VConstant (B16 (fromIntegral (fromIntegral x `rem` fromIntegral y :: Int16))) -bsrem IT32 [VConstant (B32 x), VConstant (B32 y)] - = Just $ VConstant (B32 (fromIntegral (fromIntegral x `rem` fromIntegral y :: Int32))) -bsrem IT64 [VConstant (B64 x), VConstant (B64 y)] - = Just $ VConstant (B64 (fromIntegral (fromIntegral x `rem` fromIntegral y :: Int64))) -bsrem ITNative [VConstant (I x), VConstant (I y)] = Just $ VConstant (I (x `rem` y)) -bsrem _ _ = Nothing - -bsdiv IT8 [VConstant (B8 x), VConstant (B8 y)] - = Just $ VConstant (B8 (fromIntegral (fromIntegral x `div` fromIntegral y :: Int8))) -bsdiv IT16 [VConstant (B16 x), VConstant (B16 y)] - = Just $ VConstant (B16 (fromIntegral (fromIntegral x `div` fromIntegral y :: Int16))) -bsdiv IT32 [VConstant (B32 x), VConstant (B32 y)] - = Just $ VConstant (B32 (fromIntegral (fromIntegral x `div` fromIntegral y :: Int32))) -bsdiv IT64 [VConstant (B64 x), VConstant (B64 y)] - = Just $ VConstant (B64 (fromIntegral (fromIntegral x `div` fromIntegral y :: Int64))) -bsdiv ITNative [VConstant (I x), VConstant (I y)] = Just $ VConstant (I (x `div` y)) -bsdiv _ _ = Nothing - -bashr IT8 [VConstant (B8 x), VConstant (B8 y)] - = Just $ VConstant (B8 (fromIntegral (fromIntegral x `shiftR` fromIntegral y :: Int8))) -bashr IT16 [VConstant (B16 x), VConstant (B16 y)] - = Just $ VConstant (B16 (fromIntegral (fromIntegral x `shiftR` fromIntegral y :: Int16))) -bashr IT32 [VConstant (B32 x), VConstant (B32 y)] - = Just $ VConstant (B32 (fromIntegral (fromIntegral x `shiftR` fromIntegral y :: Int32))) -bashr IT64 [VConstant (B64 x), VConstant (B64 y)] - = Just $ VConstant (B64 (fromIntegral (fromIntegral x `shiftR` fromIntegral y :: Int64))) -bashr ITNative [VConstant (I x), VConstant (I y)] = Just $ VConstant (I (x `shiftR` y)) -bashr _ _ = Nothing - -bUn :: IntTy -> (forall a. Bits a => a -> a) -> [Value] -> Maybe Value -bUn IT8 op [VConstant (B8 x)] = Just $ VConstant (B8 (op x)) -bUn IT16 op [VConstant (B16 x)] = Just $ VConstant (B16 (op x)) -bUn IT32 op [VConstant (B32 x)] = Just $ VConstant (B32 (op x)) -bUn IT64 op [VConstant (B64 x)] = Just $ VConstant (B64 (op x)) -bUn _ _ _ = Nothing - -bitBin :: IntTy -> (forall a. (Bits a, Integral a) => a -> a -> a) -> [Value] -> Maybe Value -bitBin IT8 op [VConstant (B8 x), VConstant (B8 y)] = Just $ VConstant (B8 (op x y)) -bitBin IT16 op [VConstant (B16 x), VConstant (B16 y)] = Just $ VConstant (B16 (op x y)) -bitBin IT32 op [VConstant (B32 x), VConstant (B32 y)] = Just $ VConstant (B32 (op x y)) -bitBin IT64 op [VConstant (B64 x), VConstant (B64 y)] = Just $ VConstant (B64 (op x y)) -bitBin _ _ _ = Nothing - -bCmp :: IntTy -> (forall a. Ord a => a -> a -> Bool) -> [Value] -> Maybe Value -bCmp IT8 op [VConstant (B8 x), VConstant (B8 y)] = Just $ VConstant (I (if (op x y) then 1 else 0)) -bCmp IT16 op [VConstant (B16 x), VConstant (B16 y)] = Just $ VConstant (I (if (op x y) then 1 else 0)) -bCmp IT32 op [VConstant (B32 x), VConstant (B32 y)] = Just $ VConstant (I (if (op x y) then 1 else 0)) -bCmp IT64 op [VConstant (B64 x), VConstant (B64 y)] = Just $ VConstant (I (if (op x y) then 1 else 0)) -bCmp _ _ _ = Nothing - -toB IT8 x = VConstant (B8 (fromIntegral x)) -toB IT16 x = VConstant (B16 (fromIntegral x)) -toB IT32 x = VConstant (B32 (fromIntegral x)) -toB IT64 x = VConstant (B64 (fromIntegral x)) - -zext IT8 IT16 [VConstant (B8 x)] = Just $ VConstant (B16 (fromIntegral x)) -zext IT8 IT32 [VConstant (B8 x)] = Just $ VConstant (B32 (fromIntegral x)) -zext IT8 IT64 [VConstant (B8 x)] = Just $ VConstant (B64 (fromIntegral x)) -zext IT16 IT32 [VConstant (B16 x)] = Just $ VConstant (B32 (fromIntegral x)) -zext IT16 IT64 [VConstant (B16 x)] = Just $ VConstant (B64 (fromIntegral x)) -zext IT32 IT64 [VConstant (B32 x)] = Just $ VConstant (B64 (fromIntegral x)) -zext _ _ _ = Nothing - -sext IT8 IT16 [VConstant (B8 x)] = Just $ VConstant (B16 (fromIntegral (fromIntegral x :: Int8))) -sext IT8 IT32 [VConstant (B8 x)] = Just $ VConstant (B32 (fromIntegral (fromIntegral x :: Int8))) -sext IT8 IT64 [VConstant (B8 x)] = Just $ VConstant (B64 (fromIntegral (fromIntegral x :: Int8))) -sext IT16 IT32 [VConstant (B16 x)] = Just $ VConstant (B32 (fromIntegral (fromIntegral x :: Int16))) -sext IT16 IT64 [VConstant (B16 x)] = Just $ VConstant (B64 (fromIntegral (fromIntegral x :: Int16))) -sext IT32 IT64 [VConstant (B32 x)] = Just $ VConstant (B64 (fromIntegral (fromIntegral x :: Int32))) -sext _ _ _ = Nothing - -trunc IT16 IT8 [VConstant (B16 x)] = Just $ VConstant (B8 (fromIntegral x)) -trunc IT32 IT8 [VConstant (B32 x)] = Just $ VConstant (B8 (fromIntegral x)) -trunc IT64 IT8 [VConstant (B64 x)] = Just $ VConstant (B8 (fromIntegral x)) -trunc IT32 IT16 [VConstant (B32 x)] = Just $ VConstant (B16 (fromIntegral x)) -trunc IT64 IT16 [VConstant (B64 x)] = Just $ VConstant (B16 (fromIntegral x)) -trunc IT64 IT32 [VConstant (B64 x)] = Just $ VConstant (B32 (fromIntegral x)) -trunc _ _ _ = Nothing - -intToBits8 [VConstant (I x)] = Just $ toB IT8 x +s8div :: Word8 -> Word8 -> Word8 +s8div x y = fromIntegral (fromIntegral x `div` fromIntegral y :: Int8) + +s16div :: Word16 -> Word16 -> Word16 +s16div x y = fromIntegral (fromIntegral x `div` fromIntegral y :: Int16) + +s32div :: Word32 -> Word32 -> Word32 +s32div x y = fromIntegral (fromIntegral x `div` fromIntegral y :: Int32) + +s64div :: Word64 -> Word64 -> Word64 +s64div x y = fromIntegral (fromIntegral x `div` fromIntegral y :: Int64) + +s8rem :: Word8 -> Word8 -> Word8 +s8rem x y = fromIntegral (fromIntegral x `rem` fromIntegral y :: Int8) + +s16rem :: Word16 -> Word16 -> Word16 +s16rem x y = fromIntegral (fromIntegral x `rem` fromIntegral y :: Int16) + +s32rem :: Word32 -> Word32 -> Word32 +s32rem x y = fromIntegral (fromIntegral x `rem` fromIntegral y :: Int32) + +s64rem :: Word64 -> Word64 -> Word64 +s64rem x y = fromIntegral (fromIntegral x `rem` fromIntegral y :: Int64) + +b8Bin op [VConstant (B8 x), VConstant (B8 y)] = Just $ VConstant (B8 (op x y)) +b8Bin _ _ = Nothing + +b8Un op [VConstant (B8 x)] = Just $ VConstant (B8 (op x)) +b8Un _ _ = Nothing + +b16Bin op [VConstant (B16 x), VConstant (B16 y)] = Just $ VConstant (B16 (op x y)) +b16Bin _ _ = Nothing + +b16Un op [VConstant (B16 x)] = Just $ VConstant (B16 (op x)) +b16Un _ _ = Nothing + +b32Bin op [VConstant (B32 x), VConstant (B32 y)] = Just $ VConstant (B32 (op x y)) +b32Bin _ _ = Nothing + +b32Un op [VConstant (B32 x)] = Just $ VConstant (B32 (op x)) +b32Un _ _ = Nothing + +b64Bin op [VConstant (B64 x), VConstant (B64 y)] = Just $ VConstant (B64 (op x y)) +b64Bin _ _ = Nothing + +b64Un op [VConstant (B64 x)] = Just $ VConstant (B64 (op x)) +b64Un _ _ = Nothing + +b8Cmp op [VConstant (B8 x), VConstant (B8 y)] = Just $ VConstant (I (if (op x y) then 1 else 0)) +b8Cmp _ _ = Nothing + +b16Cmp op [VConstant (B16 x), VConstant (B16 y)] = Just $ VConstant (I (if (op x y) then 1 else 0)) +b16Cmp _ _ = Nothing + +b32Cmp op [VConstant (B32 x), VConstant (B32 y)] = Just $ VConstant (I (if (op x y) then 1 else 0)) +b32Cmp _ _ = Nothing + +b64Cmp op [VConstant (B64 x), VConstant (B64 y)] = Just $ VConstant (I (if (op x y) then 1 else 0)) +b64Cmp _ _ = Nothing + +toB8 x = VConstant (B8 (fromIntegral x)) +toB16 x = VConstant (B16 (fromIntegral x)) +toB32 x = VConstant (B32 (fromIntegral x)) +toB64 x = VConstant (B64 (fromIntegral x)) + +zext8_16 [VConstant (B8 x)] = Just $ toB16 x +zext8_16 _ = Nothing + +zext8_32 [VConstant (B8 x)] = Just $ toB32 x +zext8_32 _ = Nothing + +zext8_64 [VConstant (B8 x)] = Just $ toB64 x +zext8_64 _ = Nothing + +sext8_16 [VConstant (B8 x)] = Just $ toB16 (fromIntegral x :: Int8) +sext8_16 _ = Nothing + +sext8_32 [VConstant (B8 x)] = Just $ toB32 (fromIntegral x :: Int8) +sext8_32 _ = Nothing + +sext8_64 [VConstant (B8 x)] = Just $ toB64 (fromIntegral x :: Int8) +sext8_64 _ = Nothing + +zext16_32 [VConstant (B16 x)] = Just $ toB32 x +zext16_32 _ = Nothing + +zext16_64 [VConstant (B16 x)] = Just $ toB64 x +zext16_64 _ = Nothing + +sext16_32 [VConstant (B16 x)] = Just $ toB32 (fromIntegral x :: Int16) +sext16_32 _ = Nothing + +sext16_64 [VConstant (B16 x)] = Just $ toB64 (fromIntegral x :: Int16) +sext16_64 _ = Nothing + +trunc16_8 [VConstant (B16 x)] = Just $ toB8 x +trunc16_8 _ = Nothing + +zext32_64 [VConstant (B32 x)] = Just $ toB64 x +zext32_64 _ = Nothing + +sext32_64 [VConstant (B32 x)] = Just $ toB64 (fromIntegral x :: Int32) +sext32_64 _ = Nothing + +trunc32_8 [VConstant (B32 x)] = Just $ toB8 x +trunc32_8 _ = Nothing + +trunc32_16 [VConstant (B32 x)] = Just $ toB16 x +trunc32_16 _ = Nothing + +trunc64_8 [VConstant (B64 x)] = Just $ toB8 x +trunc64_8 _ = Nothing + +trunc64_16 [VConstant (B64 x)] = Just $ toB16 x +trunc64_16 _ = Nothing + +trunc64_32 [VConstant (B64 x)] = Just $ toB32 x +trunc64_32 _ = Nothing + +intToBits8 [VConstant (I x)] = Just $ toB8 x intToBits8 _ = Nothing -intToBits16 [VConstant (I x)] = Just $ toB IT16 x +intToBits16 [VConstant (I x)] = Just $ toB16 x intToBits16 _ = Nothing -intToBits32 [VConstant (I x)] = Just $ toB IT32 x +intToBits32 [VConstant (I x)] = Just $ toB32 x intToBits32 _ = Nothing -intToBits64 [VConstant (I x)] = Just $ toB IT64 x +intToBits64 [VConstant (I x)] = Just $ toB64 x intToBits64 _ = Nothing bits32ToInt [VConstant (B32 x)] = Just $ VConstant (I (fromIntegral x))