[RefC] Unbox small integers. (#3181)

This commit is contained in:
Hattori, Hiroki 2024-03-21 21:32:37 +09:00 committed by GitHub
parent 7ce4c45e82
commit ddc634b1b2
No known key found for this signature in database
GPG Key ID: B5690EEEBB952194
25 changed files with 982 additions and 1828 deletions

View File

@ -33,14 +33,6 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
* Elaborator scripts were made to be able to access the visibility modifier of a
definition, via `getVis`.
### Backend changes
#### RefC
* Compiler can emit precise reference counting instructions where a reference
is dropped as soon as possible. This allows you to reuse unique variables and
optimize memory consumption.
### Compiler changes
* The compiler now differentiates between "package search path" and "package
@ -51,8 +43,14 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
environment variable adds to the "Package Search Paths." Functionally this is
not a breaking change.
### Backend changes
#### RefC Backend
* Compiler can emit precise reference counting instructions where a reference
is dropped as soon as possible. This allows you to reuse unique variables and
optimize memory consumption.
* Fix invalid memory read onf strSubStr.
* Fix memory leaks of IORef. Now that IORef holds values by itself,
@ -67,6 +65,13 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
* Special constructors such as Nil and Nothing were eliminated and assigned to
NULL.
* Unbox Bits32,Bits16,Bits8,Int32,Int16,Int8. These types are now packed into
Value*. Now, RefC backend requires at least 32 bits for pointers.
16-bit CPUs are no longer supported. And we expect the address returned by
malloc to be aligned with at least 32 bits. Otherwise it cause a runtime error.
* Rename C function to avoid confliction. But only a part.
#### NodeJS Backend
* The NodeJS executable output to `build/exec/` now has its executable bit set.

View File

@ -11,7 +11,7 @@ OLD_WIN ?= 0
RANLIB ?= ranlib
AR ?= ar
CFLAGS := -Wall $(CFLAGS)
CFLAGS += -Wall
CPPFLAGS := $(CPPFLAGS)
LDFLAGS := $(LDFLAGS)

View File

@ -291,7 +291,7 @@ freeVariables (AApp _ _ closure arg) = fromList [closure, arg]
freeVariables (ALet _ var value body) =
union (freeVariables value) (delete (ALocal var) $ freeVariables body)
freeVariables (ACon _ _ _ _ args) = fromList args
freeVariables (AOp _ _ _ args) = fromList $ foldl (\acc, elem => elem :: acc) [] args
freeVariables (AOp _ _ _ args) = fromList $ toList args
freeVariables (AExtPrim _ _ _ args) = fromList args
freeVariables (AConCase _ sc alts mDef) =
let altsAnf =

View File

@ -28,6 +28,7 @@ import Libraries.Utils.Path
%default covering
showcCleanStringChar : Char -> String -> String
showcCleanStringChar ' ' = ("_" ++)
showcCleanStringChar '!' = ("_bang" ++)
@ -137,93 +138,69 @@ cPrimType Bits16Type = "Bits16"
cPrimType Bits32Type = "Bits32"
cPrimType Bits64Type = "Bits64"
cPrimType StringType = "string"
cPrimType CharType = "char"
cPrimType DoubleType = "double"
cPrimType CharType = "Char"
cPrimType DoubleType = "Double"
cPrimType WorldType = "void"
cConstant : Constant -> String
cConstant (I x) = "(Value*)makeInt64("++ showIntMin x ++")"
cConstant (I8 x) = "(Value*)makeInt8(INT8_C("++ show x ++"))"
cConstant (I16 x) = "(Value*)makeInt16(INT16_C("++ show x ++"))"
cConstant (I32 x) = "(Value*)makeInt32(INT32_C("++ show x ++"))"
cConstant (I64 x) = "(Value*)makeInt64("++ showInt64Min x ++")"
cConstant (BI x) = "(Value*)makeIntegerLiteral(\""++ show x ++"\")"
cConstant (B8 x) = "(Value*)makeBits8(UINT8_C("++ show x ++"))"
cConstant (B16 x) = "(Value*)makeBits16(UINT16_C("++ show x ++"))"
cConstant (B32 x) = "(Value*)makeBits32(UINT32_C("++ show x ++"))"
cConstant (B64 x) = "(Value*)makeBits64(UINT64_C("++ show x ++"))"
cConstant (Db x) = "(Value*)makeDouble("++ show x ++")"
cConstant (Ch x) = "(Value*)makeChar("++ escapeChar x ++")"
cConstant (Str x) = "(Value*)makeString("++ cStringQuoted x ++")"
cConstant (I x) = "idris2_mkInt64("++ showIntMin x ++")"
cConstant (I8 x) = "idris2_mkInt8(INT8_C("++ show x ++"))"
cConstant (I16 x) = "idris2_mkInt16(INT16_C("++ show x ++"))"
cConstant (I32 x) = "idris2_mkInt32(INT32_C("++ show x ++"))"
cConstant (I64 x) = "idris2_mkInt64("++ showInt64Min x ++")"
cConstant (BI x) = "(Value*)idris2_mkIntegerLiteral(\""++ show x ++"\")"
cConstant (B8 x) = "idris2_mkBits8(UINT8_C("++ show x ++"))"
cConstant (B16 x) = "idris2_mkBits16(UINT16_C("++ show x ++"))"
cConstant (B32 x) = "idris2_mkBits32(UINT32_C("++ show x ++"))"
cConstant (B64 x) = "idris2_mkBits64(UINT64_C("++ show x ++"))"
cConstant (Db x) = "idris2_mkDouble("++ show x ++")"
cConstant (Ch x) = "idris2_mkChar("++ escapeChar x ++")"
cConstant (Str x) = "(Value*)idris2_mkString("++ cStringQuoted x ++")"
cConstant (PrT t) = cPrimType t
cConstant WorldVal = "(Value*)NULL"
extractConstant : Constant -> String
extractConstant (I x) = show x
extractConstant (I8 x) = show x
extractConstant (I16 x) = show x
extractConstant (I32 x) = show x
extractConstant (I64 x) = show x
extractConstant (BI x) = show x
extractConstant (Db x) = show x
extractConstant (Ch x) = show x
extractConstant (Str x) = cStringQuoted x
extractConstant (B8 x) = show x
extractConstant (B16 x) = show x
extractConstant (B32 x) = show x
extractConstant (B64 x) = show x
extractConstant c = assert_total $ idris_crash ("INTERNAL ERROR: Unable to extract constant: " ++ cConstant c)
-- not really total but this way this internal error does not contaminate everything else
||| Generate scheme for a plain function.
plainOp : String -> List String -> String
plainOp op args = op ++ "(" ++ (showSep ", " args) ++ ")"
||| Generate scheme for a primitive function.
cOp : {0 arity : Nat} -> PrimFn arity -> Vect arity String -> String
cOp (Neg ty) [x] = "negate_" ++ cPrimType ty ++ "(" ++ x ++ ")"
cOp (Neg ty) [x] = "idris2_negate_" ++ cPrimType ty ++ "(" ++ x ++ ")"
cOp StrLength [x] = "stringLength(" ++ x ++ ")"
cOp StrHead [x] = "head(" ++ x ++ ")"
cOp StrTail [x] = "tail(" ++ x ++ ")"
cOp StrReverse [x] = "reverse(" ++ x ++ ")"
cOp (Cast i o) [x] = "cast_" ++ (cPrimType i) ++ "_to_" ++ (cPrimType o) ++ "(" ++ x ++ ")"
cOp DoubleExp [x] = "(Value*)makeDouble(exp(unpackDouble(" ++ x ++ ")))"
cOp DoubleLog [x] = "(Value*)makeDouble(log(unpackDouble(" ++ x ++ ")))"
cOp DoublePow [x, y] = "(Value*)makeDouble(pow(unpackDouble(" ++ x ++ "), unpackDouble(" ++ y ++ ")))"
cOp DoubleSin [x] = "(Value*)makeDouble(sin(unpackDouble(" ++ x ++ ")))"
cOp DoubleCos [x] = "(Value*)makeDouble(cos(unpackDouble(" ++ x ++ ")))"
cOp DoubleTan [x] = "(Value*)makeDouble(tan(unpackDouble(" ++ x ++ ")))"
cOp DoubleASin [x] = "(Value*)makeDouble(asin(unpackDouble(" ++ x ++ ")))"
cOp DoubleACos [x] = "(Value*)makeDouble(acos(unpackDouble(" ++ x ++ ")))"
cOp DoubleATan [x] = "(Value*)makeDouble(atan(unpackDouble(" ++ x ++ ")))"
cOp DoubleSqrt [x] = "(Value*)makeDouble(sqrt(unpackDouble(" ++ x ++ ")))"
cOp DoubleFloor [x] = "(Value*)makeDouble(floor(unpackDouble(" ++ x ++ ")))"
cOp DoubleCeiling [x] = "(Value*)makeDouble(ceil(unpackDouble(" ++ x ++ ")))"
cOp (Add ty) [x, y] = "add_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Sub ty) [x, y] = "sub_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Mul ty) [x, y] = "mul_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Div ty) [x, y] = "div_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Mod ty) [x, y] = "mod_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (ShiftL ty) [x, y] = "shiftl_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (ShiftR ty) [x, y] = "shiftr_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (BAnd ty) [x, y] = "and_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (BOr ty) [x, y] = "or_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (BXOr ty) [x, y] = "xor_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (LT ty) [x, y] = "lt_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (GT ty) [x, y] = "gt_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (EQ ty) [x, y] = "eq_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (LTE ty) [x, y] = "lte_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (GTE ty) [x, y] = "gte_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Cast i o) [x] = "idris2_cast_" ++ (cPrimType i) ++ "_to_" ++ (cPrimType o) ++ "(" ++ x ++ ")"
cOp DoubleExp [x] = "idris2_mkDouble(exp(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleLog [x] = "idris2_mkDouble(log(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoublePow [x, y] = "idris2_mkDouble(pow(idris2_vp_to_Double(" ++ x ++ "), idris2_vp_to_Double(" ++ y ++ ")))"
cOp DoubleSin [x] = "idris2_mkDouble(sin(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleCos [x] = "idris2_mkDouble(cos(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleTan [x] = "idris2_mkDouble(tan(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleASin [x] = "idris2_mkDouble(asin(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleACos [x] = "idris2_mkDouble(acos(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleATan [x] = "idris2_mkDouble(atan(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleSqrt [x] = "idris2_mkDouble(sqrt(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleFloor [x] = "idris2_mkDouble(floor(idris2_vp_to_Double(" ++ x ++ ")))"
cOp DoubleCeiling [x] = "idris2_mkDouble(ceil(idris2_vp_to_Double(" ++ x ++ ")))"
cOp (Add ty) [x, y] = "idris2_add_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Sub ty) [x, y] = "idris2_sub_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Mul ty) [x, y] = "idris2_mul_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Div ty) [x, y] = "idris2_div_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (Mod ty) [x, y] = "idris2_mod_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (ShiftL ty) [x, y] = "idris2_shiftl_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (ShiftR ty) [x, y] = "idris2_shiftr_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (BAnd ty) [x, y] = "idris2_and_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (BOr ty) [x, y] = "idris2_or_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (BXOr ty) [x, y] = "idris2_xor_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (LT ty) [x, y] = "idris2_lt_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (GT ty) [x, y] = "idris2_gt_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (EQ ty) [x, y] = "idris2_eq_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (LTE ty) [x, y] = "idris2_lte_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp (GTE ty) [x, y] = "idris2_gte_" ++ cPrimType ty ++ "(" ++ x ++ ", " ++ y ++ ")"
cOp StrIndex [x, i] = "strIndex(" ++ x ++ ", " ++ i ++ ")"
cOp StrCons [x, y] = "strCons(" ++ x ++ ", " ++ y ++ ")"
cOp StrAppend [x, y] = "strAppend(" ++ x ++ ", " ++ y ++ ")"
cOp StrSubstr [x, y, z] = "strSubstr(" ++ x ++ ", " ++ y ++ ", " ++ z ++ ")"
cOp StrSubstr [x, y, z] = "strSubstr(" ++ x ++ ", " ++ y ++ ", " ++ z ++ ")"
cOp BelieveMe [_, _, x] = "newReference(" ++ x ++ ")"
cOp Crash [_, msg] = "idris2_crash(" ++ msg ++ ");"
cOp fn args = plainOp (show fn) (toList args)
cOp fn args = show fn ++ "(" ++ (showSep ", " $ toList args) ++ ")"
varName : AVar -> String
varName (ALocal i) = "var_" ++ (show i)
@ -257,16 +234,15 @@ Output = DList String
------------------------------------------------------------------------
getNextCounter : {auto a : Ref ArgCounter Nat} -> Core Nat
getNextCounter : {auto a : Ref ArgCounter Nat} -> Core String
getNextCounter = do
c <- get ArgCounter
put ArgCounter (S c)
pure c
pure $ show c
getNewVarThatWillNotBeFreedAtEndOfBlock : {auto a : Ref ArgCounter Nat} -> Core String
getNewVarThatWillNotBeFreedAtEndOfBlock = do
c <- getNextCounter
pure $ "tmp_" ++ show c
pure $ "tmp_" ++ !(getNextCounter)
maxLineLengthForComment : Nat
@ -309,11 +285,6 @@ emit fc line = do
(Yes _) => flip snoc (lJust indentedLine maxLineLengthForComment ' ' ++ " " ++ comment)
(No _) => flip appendR [indentedLine, (lJust "" maxLineLengthForComment ' ' ++ " " ++ comment)]
addHeader : {auto h : Ref HeaderFiles (SortedSet String)}
-> String
-> Core ()
addHeader = update HeaderFiles . insert
applyFunctionToVars : {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> String
@ -356,8 +327,7 @@ makeArglist : {auto a : Ref ArgCounter Nat}
-> List AVar
-> Core String
makeArglist missing xs = do
c <- getNextCounter
let arglist = "arglist_" ++ (show c)
let arglist = "arglist_" ++ !(getNextCounter)
emit EmptyFC $ "Value_Arglist *"
++ arglist
++ " = newArglist(" ++ show missing
@ -409,22 +379,21 @@ integer_switch (MkAConstAlt c _ :: _) =
(Ch x) => True
_ => False
const2Integer : Constant -> Integer -> Integer
const2Integer : Constant -> Integer -> String
const2Integer c i =
case c of
(I x) => cast x
(I8 x) => cast x
(I16 x) => cast x
(I32 x) => cast x
(I64 x) => cast x
(BI x) => cast x
(Ch x) => cast x
(B8 x) => cast x
(B16 x) => cast x
(B32 x) => cast x
(B64 x) => cast x
_ => i
(I x) => showIntMin x
(I8 x) => "INT8_C(\{show x})"
(I16 x) => "INT16_C(\{show x})"
(I32 x) => "INT32_C(\{show x})"
(I64 x) => showInt64Min x
(BI x) => show x
(Ch x) => escapeChar x
(B8 x) => "UINT8_C(\{show x})"
(B16 x) => "UINT16_C(\{show x})"
(B32 x) => "UINT32_C(\{show x})"
(B64 x) => "UINT64_C(\{show x})"
_ => show i
data TailPositionStatus = InTailPosition | NotInTailPosition
@ -446,13 +415,6 @@ dropUnusedOwnedVars owned usedVars =
let shouldDrop = difference owned actualOwned in
(varName <$> SortedSet.toList shouldDrop, actualOwned)
locally : {auto t : Ref EnvTracker Env} -> Env -> Core () -> Core ()
locally newEnv act = do
oldEnv <- get EnvTracker
put EnvTracker newEnv
act
put EnvTracker oldEnv
-- if the constructor is unique use it, otherwise add it to should drop vars and create null constructor
addReuseConstructor : {auto a : Ref ArgCounter Nat}
-> {auto oft : Ref OutfileText Output}
@ -471,8 +433,7 @@ addReuseConstructor reuseMap sc conName conArgs consts shouldDrop actualReuseCon
if (isNothing $ SortedMap.lookup conName reuseMap)
&& contains conName consts
&& (isJust $ find (== sc) shouldDrop) then do
c <- getNextCounter
let constr = "constructor_" ++ (show c)
let constr = "constructor_" ++ !(getNextCounter)
emit EmptyFC $ "Value_Constructor* " ++ constr ++ " = NULL;"
-- If the constructor variable is unique (has 1 reference), then assign it for reuse
emit EmptyFC $ "if (isUnique(" ++ sc ++ ")) {"
@ -498,17 +459,22 @@ mutual
-> {auto e : Ref EnvTracker Env}
-> {auto oft : Ref OutfileText Output}
-> {auto il : Ref IndentLevel Nat}
-> List String -> List String
-> Env
-> String -> String -> List Int -> ANF -> TailPositionStatus
-> Core ()
concaseBody dropVars dropReuseCons returnvar expr args bdy tailPosition = do
concaseBody env returnvar expr args body tailPosition = do
increaseIndentation
_ <- foldlC (\k, arg => do
emit emptyFC "Value *var_\{show arg} = ((Value_Constructor*)\{expr})->args[\{show k}];"
pure (S k) ) 0 args
removeVars dropVars
let (shouldDrop, actualOwned) = dropUnusedOwnedVars env.owned (freeVariables body)
let usedCons = usedConstructors body
let (dropReuseCons, actualReuseMap) = dropUnusedReuseCons env.reuseMap usedCons
put EnvTracker ({owned := actualOwned, reuseMap := actualReuseMap} env)
removeVars shouldDrop
removeReuseConstructors dropReuseCons
emit emptyFC "\{returnvar} = \{!(cStatementsFromANF bdy tailPosition)};"
emit emptyFC "\{returnvar} = \{!(cStatementsFromANF body tailPosition)};"
decreaseIndentation
cStatementsFromANF : {auto a : Ref ArgCounter Nat}
@ -524,9 +490,9 @@ mutual
emit fc $ ("// start " ++ cName n ++ "(" ++ showSep ", " (map (\v => varName v) args) ++ ")")
arglist <- makeArglist 0 args
c <- getNextCounter
let f_ptr_name = "fPtr_" ++ show c
let f_ptr_name = "fPtr_" ++ c
emit fc $ "Value *(*"++ f_ptr_name ++ ")(Value_Arglist*) = " ++ cName n ++ "_arglist;"
let closure_name = "closure_" ++ show c
let closure_name = "closure_" ++ c
emit fc $ "Value *"
++ closure_name
++ " = (Value*)makeClosureFromArglist("
@ -540,8 +506,7 @@ mutual
InTailPosition => closure_name
cStatementsFromANF (AUnderApp fc n missing args) _ = do
arglist <- makeArglist missing args
c <- getNextCounter
let f_ptr_name = "closure_" ++ show c
let f_ptr_name = "closure_" ++ !(getNextCounter)
let f_ptr = "Value *(*"++ f_ptr_name ++ ")(Value_Arglist*) = "++ cName n ++ "_arglist;"
emit fc f_ptr
pure "(Value*)makeClosureFromArglist(\{f_ptr_name}, \{arglist})"
@ -585,7 +550,7 @@ mutual
emit fc "}"
pure constr
Nothing => do
let constr = "constructor_\{show !(getNextCounter)}"
let constr = "constructor_\{!(getNextCounter)}"
emit fc $ "Value_Constructor* " ++ constr ++ createNewConstructor
when (Nothing == tag) $ emit fc "\{constr}->name = idris2_constr_\{cName n};"
pure constr
@ -593,12 +558,11 @@ mutual
pure $ "(Value*)\{constr}"
cStatementsFromANF (AOp fc _ op args) _ = do
c <- getNextCounter
let resultVar = "primVar_" ++ (show c)
let resultVar = "primVar_" ++ !(getNextCounter)
let argsVect = map (avarToC !(get EnvTracker)) args
emit fc $ "Value *" ++ resultVar ++ " = " ++ cOp op argsVect ++ ";"
-- Removing arguments that apply to primitive functions
removeVars (foldl (\acc, elem => elem :: acc) [] (map varName args))
removeVars $ toList $ map varName args
pure resultVar
cStatementsFromANF (AExtPrim fc _ p args) _ = do
@ -629,8 +593,7 @@ mutual
Just tag' => emit emptyFC "\{els}if (((Value_Constructor *)\{sc'})->tag == \{show tag'} /* \{show name} */) {"
let conArgs = ALocal <$> args
let owned = if erased then delete sc env.owned else env.owned
let ownedWithArgs = union (fromList conArgs) owned
let ownedWithArgs = union (fromList conArgs) $ if erased then delete sc env.owned else env.owned
let (shouldDrop, actualOwned) = dropUnusedOwnedVars ownedWithArgs (freeVariables body)
let usedCons = usedConstructors body
let (dropReuseCons, actualReuseMap) = dropUnusedReuseCons env.reuseMap usedCons
@ -649,12 +612,8 @@ mutual
case mDef of
Nothing => pure ()
Just body => do
let (shouldDrop, actualOwned) = dropUnusedOwnedVars env.owned (freeVariables body)
let usedCons = usedConstructors body
let (dropReuseCons, actualReuseMap) = dropUnusedReuseCons env.reuseMap usedCons
emit emptyFC "} else {"
put EnvTracker ({owned := actualOwned, reuseMap := actualReuseMap} env)
concaseBody shouldDrop dropReuseCons switchReturnVar "" [] body tailPosition
concaseBody env switchReturnVar "" [] body tailPosition
emit emptyFC "}"
pure switchReturnVar
@ -666,40 +625,28 @@ mutual
case integer_switch alts of
True => do
tmpint <- getNewVarThatWillNotBeFreedAtEndOfBlock
emit emptyFC "int \{tmpint} = extractInt(\{sc'});"
emit emptyFC "int64_t \{tmpint} = idris2_extractInt(\{sc'});"
_ <- foldlC (\els, (MkAConstAlt c body) => do
let (shouldDrop, actualOwned) = dropUnusedOwnedVars env.owned (freeVariables body)
let usedCons = usedConstructors body
let (dropReuseCons, actualReuseMap) = dropUnusedReuseCons env.reuseMap usedCons
emit emptyFC "\{els}if (\{tmpint} == \{show $ const2Integer c 0}) {"
put EnvTracker ({owned := actualOwned, reuseMap := actualReuseMap} env)
concaseBody shouldDrop dropReuseCons switchReturnVar "" [] body tailPosition
emit emptyFC "\{els}if (\{tmpint} == \{const2Integer c 0}) {"
concaseBody env switchReturnVar "" [] body tailPosition
pure "} else ") "" alts
pure ()
False => do
_ <- foldlC (\els, (MkAConstAlt c body) => do
let (shouldDrop, actualOwned) = dropUnusedOwnedVars env.owned (freeVariables body)
let usedCons = usedConstructors body
let (dropReuseCons, actualReuseMap) = dropUnusedReuseCons env.reuseMap usedCons
case c of
Str x => emit emptyFC "\{els}if (! strcmp(\{cStringQuoted x}, ((Value_String *)\{sc'})->str)) {"
Db x => emit emptyFC "\{els}if (((Value_Double *)\{sc'})->d == \{show x}) {"
x => throw $ InternalError "[refc] AConstCase : unsupported type. \{show fc} \{show x}"
put EnvTracker ({owned := actualOwned, reuseMap := actualReuseMap} env)
concaseBody shouldDrop dropReuseCons switchReturnVar "" [] body tailPosition
concaseBody env switchReturnVar "" [] body tailPosition
pure "} else ") "" alts
pure ()
case def of
Nothing => pure ()
Just body => do
let (shouldDrop, actualOwned) = dropUnusedOwnedVars env.owned (freeVariables body)
let usedCons = usedConstructors body
let (dropReuseCons, actualReuseMap) = dropUnusedReuseCons env.reuseMap usedCons
emit emptyFC "} else {"
put EnvTracker ({owned := actualOwned, reuseMap := actualReuseMap} env)
concaseBody shouldDrop dropReuseCons switchReturnVar "" [] body tailPosition
concaseBody env switchReturnVar "" [] body tailPosition
emit emptyFC "}"
pure switchReturnVar
@ -779,18 +726,18 @@ data CLang = CLangC | CLangRefC
extractValue : (cLang : CLang) -> (cfType:CFType) -> (varName:String) -> String
extractValue _ CFUnit varName = "NULL"
extractValue _ CFInt varName = "((Value_Int64*)" ++ varName ++ ")->i64"
extractValue _ CFInt8 varName = "((Value_Int8*)" ++ varName ++ ")->i8"
extractValue _ CFInt16 varName = "((Value_Int16*)" ++ varName ++ ")->i16"
extractValue _ CFInt32 varName = "((Value_Int32*)" ++ varName ++ ")->i32"
extractValue _ CFInt64 varName = "((Value_Int64*)" ++ varName ++ ")->i64"
extractValue _ CFUnsigned8 varName = "((Value_Bits8*)" ++ varName ++ ")->ui8"
extractValue _ CFUnsigned16 varName = "((Value_Bits16*)" ++ varName ++ ")->ui16"
extractValue _ CFUnsigned32 varName = "((Value_Bits32*)" ++ varName ++ ")->ui32"
extractValue _ CFUnsigned64 varName = "((Value_Bits64*)" ++ varName ++ ")->ui64"
extractValue _ CFInt varName = "(idris2_vp_to_Int64(" ++ varName ++ "))"
extractValue _ CFInt8 varName = "(idris2_vp_to_Int8(" ++ varName ++ "))"
extractValue _ CFInt16 varName = "(idris2_vp_to_Int16(" ++ varName ++ "))"
extractValue _ CFInt32 varName = "(idris2_vp_to_Int32(" ++ varName ++ "))"
extractValue _ CFInt64 varName = "(idris2_vp_to_Int64(" ++ varName ++ "))"
extractValue _ CFUnsigned8 varName = "(idris2_vp_to_Bits8(" ++ varName ++ "))"
extractValue _ CFUnsigned16 varName = "(idris2_vp_to_Bits16(" ++ varName ++ "))"
extractValue _ CFUnsigned32 varName = "(idris2_vp_to_Bits32(" ++ varName ++ "))"
extractValue _ CFUnsigned64 varName = "(idris2_vp_to_Bits64(" ++ varName ++ "))"
extractValue _ CFString varName = "((Value_String*)" ++ varName ++ ")->str"
extractValue _ CFDouble varName = "((Value_Double*)" ++ varName ++ ")->d"
extractValue _ CFChar varName = "((Value_Char*)" ++ varName ++ ")->c"
extractValue _ CFDouble varName = "(idris2_vp_to_Double(" ++ varName ++ "))"
extractValue _ CFChar varName = "(idris2_vp_to_Char(" ++ varName ++ "))"
extractValue _ CFPtr varName = "((Value_Pointer*)" ++ varName ++ ")->p"
extractValue _ CFGCPtr varName = "((Value_GCPointer*)" ++ varName ++ ")->p->p"
extractValue CLangC CFBuffer varName = "((Value_Buffer*)" ++ varName ++ ")->buffer->data"
@ -804,19 +751,19 @@ extractValue _ (CFUser x xs) varName = "(Value*)" ++ varName
extractValue _ n _ = assert_total $ idris_crash ("INTERNAL ERROR: Unknonw FFI type in C backend: " ++ show n)
packCFType : (cfType:CFType) -> (varName:String) -> String
packCFType CFUnit varName = "NULL"
packCFType CFInt varName = "makeInt64(" ++ varName ++ ")"
packCFType CFInt8 varName = "makeInt8(" ++ varName ++ ")"
packCFType CFInt16 varName = "makeInt16(" ++ varName ++ ")"
packCFType CFInt32 varName = "makeInt32(" ++ varName ++ ")"
packCFType CFInt64 varName = "makeInt64(" ++ varName ++ ")"
packCFType CFUnsigned64 varName = "makeBits64(" ++ varName ++ ")"
packCFType CFUnsigned32 varName = "makeBits32(" ++ varName ++ ")"
packCFType CFUnsigned16 varName = "makeBits16(" ++ varName ++ ")"
packCFType CFUnsigned8 varName = "makeBits8(" ++ varName ++ ")"
packCFType CFString varName = "makeString(" ++ varName ++ ")"
packCFType CFDouble varName = "makeDouble(" ++ varName ++ ")"
packCFType CFChar varName = "makeChar(" ++ varName ++ ")"
packCFType CFUnit varName = "((Value *)NULL)"
packCFType CFInt varName = "idris2_mkInt64(" ++ varName ++ ")"
packCFType CFInt8 varName = "idris2_mkInt8(" ++ varName ++ ")"
packCFType CFInt16 varName = "idris2_mkInt16(" ++ varName ++ ")"
packCFType CFInt32 varName = "idris2_mkInt32(" ++ varName ++ ")"
packCFType CFInt64 varName = "idris2_mkInt64(" ++ varName ++ ")"
packCFType CFUnsigned64 varName = "idris2_mkBits64(" ++ varName ++ ")"
packCFType CFUnsigned32 varName = "idris2_mkBits32(" ++ varName ++ ")"
packCFType CFUnsigned16 varName = "idris2_mkBits16(" ++ varName ++ ")"
packCFType CFUnsigned8 varName = "idris2_mkBits8(" ++ varName ++ ")"
packCFType CFString varName = "idris2_mkString(" ++ varName ++ ")"
packCFType CFDouble varName = "idris2_mkDouble(" ++ varName ++ ")"
packCFType CFChar varName = "idris2_mkChar(" ++ varName ++ ")"
packCFType CFPtr varName = "makePointer(" ++ varName ++ ")"
packCFType CFGCPtr varName = "makePointer(" ++ varName ++ ")"
packCFType CFBuffer varName = "makeBuffer(" ++ varName ++ ")"
@ -905,7 +852,7 @@ createCFunctions n (MkAForeign ccs fargs ret) = do
else NS (mkNamespace lang) n
if isStandardFFI
then case extLibOpts of
[lib, header] => addHeader header
[lib, header] => update HeaderFiles $ insert header
_ => pure ()
else emit EmptyFC $ additionalFFIStub fctName fargs ret
let fnDef = "Value *" ++ (cName n) ++ "(" ++ showSep ", " (replicate (length fargs) "Value *") ++ ");"

View File

@ -211,11 +211,13 @@ struct filetime *idris2_fileTime(FILE *f) {
ft->mtime_sec = buf.st_mtime;
ft->ctime_sec = buf.st_ctime;
#if defined(__MACH__) || defined(__APPLE__)
#if defined(_DARWIN_C_SOURCE) || \
(!defined(_POSIX_C_SOURCE) && (defined(__MACH__) || defined(__APPLE__)))
ft->atime_nsec = buf.st_atimespec.tv_nsec;
ft->mtime_nsec = buf.st_mtimespec.tv_nsec;
ft->ctime_nsec = buf.st_ctimespec.tv_nsec;
#elif (_POSIX_VERSION >= 200809L) || defined(__FreeBSD__)
#elif (defined(_POSIX_VERSION) && _POSIX_VERSION >= 200809L) || \
defined(__FreeBSD__)
ft->atime_nsec = buf.st_atim.tv_nsec;
ft->mtime_nsec = buf.st_mtim.tv_nsec;
ft->ctime_nsec = buf.st_ctim.tv_nsec;

View File

@ -11,6 +11,11 @@
#include <time.h>
#include <unistd.h>
#if !defined(_SC_NPROCESSORS_ONLN) && defined(_DARWIN_C_SOURCE)
#include <sys/sysctl.h>
#include <sys/types.h>
#endif
int _argc;
char **_argv;
@ -149,9 +154,25 @@ int idris2_getPID() {
// get the number of processors configured
long idris2_getNProcessors() {
#ifdef _WIN32
#if defined(_WIN32)
return win32_getNProcessors();
#else
#elif defined(_POSIX_VERSION) && _POSIX_VERSION >= 200112L && \
defined(_SC_NPROCESSORS_ONLN)
// Note: Under MacOS with _POSIX_C_SOURCE, _SC_NPROCESSORS_ONLN never defined.
return sysconf(_SC_NPROCESSORS_ONLN);
#elif defined(_DARWIN_C_SOURCE) || defined(_BSD_SOURCE)
// Generic way for BSD style system.
#if defined(_DARWIN_C_SOURCE)
int xs[] = {CTL_HW, HW_LOGICALCPU};
#else
int xs[] = {CTL_HW, HW_NCPU};
#endif
int numcpus = 1;
size_t n = sizeof(numcpus);
return sysctll(xs, 2, &numcpus, &len, NULL, 0) < 0 ? 1 : numcpus;
#else
return 1
#endif
}

View File

@ -10,17 +10,12 @@
#include "buffer.h"
#define NO_TAG 0
#define BITS8_TAG 1
#define BITS16_TAG 2
#define BITS32_TAG 3
#define BITS64_TAG 4
#define INT8_TAG 5
#define INT16_TAG 6
#define INT32_TAG 7
#define INT64_TAG 8
#define INTEGER_TAG 9
#define DOUBLE_TAG 10
#define CHAR_TAG 11
#define STRING_TAG 12
#define CLOSURE_TAG 15
@ -50,15 +45,58 @@ typedef struct {
// followed by type-specific payload.
} Value;
typedef struct {
Value_header header;
uint8_t ui8;
} Value_Bits8;
/*
We expect at least 4 bytes for `Value_header` alignment, to use bit0 and bit1 of
pointer as flags.
typedef struct {
Value_header header;
uint16_t ui16;
} Value_Bits16;
RefC does not have complete static tracking of type information, so types are
identified at runtime using Value_Header's tag field. However, Int that are
pretending to be pointers cannot have that tag, so use that flag to identify
them first. Of course, this flag is not used if it is clear that Value* is
actually an Int. But places like newReference/removeReference require this flag.
*/
#define idris2_vp_is_unboxed(p) ((uintptr_t)(p)&3)
#define idris2_vp_int_shift \
((sizeof(uintptr_t) >= 8 && sizeof(Value *) >= 8) ? 32 : 16)
#define idris2_vp_to_Bits64(p) (((Value_Bits64 *)(p))->ui64)
#if !defined(UINTPTR_WIDTH)
#define idris2_vp_to_Bits32(p) \
((idris2_vp_int_shift == 16) \
? (((Value_Bits32 *)(p))->ui32) \
: ((uint32_t)((uintptr_t)(p) >> idris2_vp_int_shift)))
#define idris2_vp_to_Int32(p) \
((idris2_vp_int_shift == 16) \
? (((Value_Int32 *)(p))->i32) \
: ((int32_t)((uintptr_t)(p) >> idris2_vp_int_shift)))
#elif UINTPTR_WIDTH >= 64
// NOTE: We stole two bits from pointer. So, even if we have 64-bit CPU,
// Int64/Bits654 are not unboxable.
#define idris2_vp_to_Bits32(p) \
((uint32_t)((uintptr_t)(p) >> idris2_vp_int_shift))
#define idris2_vp_to_Int32(p) ((int32_t)((uintptr_t)(p) >> idris2_vp_int_shift))
#elif UINTPTR_WIDTH >= 32
#define idris2_vp_to_Bits32(p) (((Value_Bits32 *)(p))->ui32)
#define idris2_vp_to_Int32(p) (((Value_Int32 *)(p))->i32)
#else
#error "unsupported uintptr_t width"
#endif
#define idris2_vp_to_Bits16(p) \
((uint16_t)((uintptr_t)(p) >> idris2_vp_int_shift))
#define idris2_vp_to_Bits8(p) ((uint8_t)((uintptr_t)(p) >> idris2_vp_int_shift))
#define idris2_vp_to_Int64(p) (((Value_Int64 *)(p))->i64)
#define idris2_vp_to_Int16(p) ((int16_t)((uintptr_t)(p) >> idris2_vp_int_shift))
#define idris2_vp_to_Int8(p) ((int8_t)((uintptr_t)(p) >> idris2_vp_int_shift))
#define idris2_vp_to_Char(p) \
((unsigned char)((uintptr_t)(p) >> idris2_vp_int_shift))
#define idris2_vp_to_Double(p) (((Value_Double *)(p))->d)
#define idris2_vp_to_Bool(p) (idris2_vp_to_Int8(p))
typedef struct {
Value_header header;
@ -70,16 +108,6 @@ typedef struct {
uint64_t ui64;
} Value_Bits64;
typedef struct {
Value_header header;
int8_t i8;
} Value_Int8;
typedef struct {
Value_header header;
int16_t i16;
} Value_Int16;
typedef struct {
Value_header header;
int32_t i32;
@ -100,11 +128,6 @@ typedef struct {
double d;
} Value_Double;
typedef struct {
Value_header header;
unsigned char c;
} Value_Char;
typedef struct {
Value_header header;
char *str;

View File

@ -1,721 +1,230 @@
#include "casts.h"
#include <inttypes.h>
/* conversions from Int8 */
Value *cast_Int8_to_Bits8(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeBits8((uint8_t)from->i8);
}
Value *cast_Int8_to_Bits16(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeBits16((uint16_t)from->i8);
}
Value *cast_Int8_to_Bits32(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeBits32((uint32_t)from->i8);
}
Value *cast_Int8_to_Bits64(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeBits64((uint64_t)from->i8);
}
Value *cast_Int8_to_Int16(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeInt16((int16_t)from->i8);
}
Value *cast_Int8_to_Int32(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeInt32((int32_t)from->i8);
}
Value *cast_Int8_to_Int64(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeInt64((int64_t)from->i8);
}
Value *cast_Int8_to_Integer(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
Value_Integer *retVal = makeInteger();
mpz_set_si(retVal->i, from->i8);
Value *idris2_cast_Int8_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_si(retVal->i, idris2_vp_to_Int8(input));
return (Value *)retVal;
}
Value *cast_Int8_to_double(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeDouble((double)from->i8);
}
Value *idris2_cast_Int8_to_string(Value *input) {
int8_t x = idris2_vp_to_Int8(input);
Value *cast_Int8_to_char(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
return (Value *)makeChar((unsigned char)from->i8);
}
Value *cast_Int8_to_string(Value *input) {
Value_Int8 *from = (Value_Int8 *)input;
int l = snprintf(NULL, 0, "%" PRId8 "", from->i8);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%" PRId8 "", from->i8);
int l = snprintf(NULL, 0, "%" PRId8 "", x);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%" PRId8 "", x);
return (Value *)retVal;
}
/* conversions from Int16 */
Value *cast_Int16_to_Bits8(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeBits8((uint8_t)from->i16);
}
Value *cast_Int16_to_Bits16(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeBits16((uint16_t)from->i16);
}
Value *cast_Int16_to_Bits32(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeBits32((uint32_t)from->i16);
}
Value *cast_Int16_to_Bits64(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeBits64((uint64_t)from->i16);
}
Value *cast_Int16_to_Int8(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeInt8((int8_t)from->i16);
}
Value *cast_Int16_to_Int32(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeInt32((int32_t)from->i16);
}
Value *cast_Int16_to_Int64(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeInt64((int64_t)from->i16);
}
Value *cast_Int16_to_Integer(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
Value_Integer *retVal = makeInteger();
mpz_set_si(retVal->i, from->i16);
Value *idris2_cast_Int16_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_si(retVal->i, idris2_vp_to_Int16(input));
return (Value *)retVal;
}
Value *cast_Int16_to_double(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeDouble((double)from->i16);
}
Value *idris2_cast_Int16_to_string(Value *input) {
int16_t x = idris2_vp_to_Int16(input);
Value *cast_Int16_to_char(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
return (Value *)makeChar((unsigned char)from->i16);
}
Value *cast_Int16_to_string(Value *input) {
Value_Int16 *from = (Value_Int16 *)input;
int l = snprintf(NULL, 0, "%" PRId16 "", from->i16);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%" PRId16 "", from->i16);
int l = snprintf(NULL, 0, "%" PRId16 "", x);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%" PRId16 "", x);
return (Value *)retVal;
}
/* conversions from Int32 */
Value *cast_Int32_to_Bits8(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeBits8((uint8_t)from->i32);
}
Value *cast_Int32_to_Bits16(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeBits16((uint16_t)from->i32);
}
Value *cast_Int32_to_Bits32(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeBits32((uint32_t)from->i32);
}
Value *cast_Int32_to_Bits64(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeBits64((uint64_t)from->i32);
}
Value *cast_Int32_to_Int8(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeInt8((int8_t)from->i32);
}
Value *cast_Int32_to_Int16(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeInt16((int16_t)from->i32);
}
Value *cast_Int32_to_Int64(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeInt64((int64_t)from->i32);
}
Value *cast_Int32_to_Integer(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
Value_Integer *retVal = makeInteger();
mpz_set_si(retVal->i, from->i32);
Value *idris2_cast_Int32_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_si(retVal->i, idris2_vp_to_Int32(input));
return (Value *)retVal;
}
Value *cast_Int32_to_double(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeDouble((double)from->i32);
}
Value *idris2_cast_Int32_to_string(Value *input) {
int32_t x = idris2_vp_to_Int32(input);
Value *cast_Int32_to_char(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
return (Value *)makeChar((unsigned char)from->i32);
}
Value *cast_Int32_to_string(Value *input) {
Value_Int32 *from = (Value_Int32 *)input;
int l = snprintf(NULL, 0, "%" PRId32 "", from->i32);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%" PRId32 "", from->i32);
int l = snprintf(NULL, 0, "%" PRId32 "", x);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%" PRId32 "", x);
return (Value *)retVal;
}
/* conversions from Int64 */
Value *cast_Int64_to_Bits8(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeBits8((uint8_t)from->i64);
}
Value *cast_Int64_to_Bits16(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeBits16((uint16_t)from->i64);
}
Value *cast_Int64_to_Bits32(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeBits32((uint32_t)from->i64);
}
Value *cast_Int64_to_Bits64(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeBits64((uint64_t)from->i64);
}
Value *cast_Int64_to_Int8(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeInt8((int8_t)from->i64);
}
Value *cast_Int64_to_Int16(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeInt16((int16_t)from->i64);
}
Value *cast_Int64_to_Int32(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeInt32((int32_t)from->i64);
}
Value *cast_Int64_to_Int64(Value *input) {
// Identity cast required as Value_Int64 represents both Int and Int64 types
return newReference(input);
}
Value *cast_Int64_to_Integer(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
Value_Integer *retVal = makeInteger();
mpz_set_si(retVal->i, from->i64);
Value *idris2_cast_Int64_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_si(retVal->i, idris2_vp_to_Int64(input));
return (Value *)retVal;
}
Value *cast_Int64_to_double(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeDouble((double)from->i64);
}
Value *cast_Int64_to_char(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
return (Value *)makeChar((unsigned char)from->i64);
}
Value *cast_Int64_to_string(Value *input) {
Value_Int64 *from = (Value_Int64 *)input;
int l = snprintf(NULL, 0, "%" PRId64 "", from->i64);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%" PRId64 "", from->i64);
Value *idris2_cast_Int64_to_string(Value *input) {
int64_t from = idris2_vp_to_Int64(input);
int l = snprintf(NULL, 0, "%" PRId64 "", from);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%" PRId64 "", from);
return (Value *)retVal;
}
Value *cast_double_to_Bits8(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeBits8((uint8_t)from->d);
}
Value *cast_double_to_Bits16(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeBits16((uint16_t)from->d);
}
Value *cast_double_to_Bits32(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeBits32((uint32_t)from->d);
}
Value *cast_double_to_Bits64(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeBits64((uint64_t)from->d);
}
Value *cast_double_to_Int8(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeInt8((int8_t)from->d);
}
Value *cast_double_to_Int16(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeInt16((int16_t)from->d);
}
Value *cast_double_to_Int32(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeInt32((int32_t)from->d);
}
Value *cast_double_to_Int64(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeInt64((int64_t)from->d);
}
Value *cast_double_to_Integer(Value *input) {
Value_Double *from = (Value_Double *)input;
Value_Integer *retVal = makeInteger();
mpz_set_d(retVal->i, from->d);
Value *idris2_cast_Double_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_d(retVal->i, idris2_vp_to_Double(input));
return (Value *)retVal;
}
Value *cast_double_to_char(Value *input) {
Value_Double *from = (Value_Double *)input;
return (Value *)makeChar((unsigned char)from->d);
}
Value *idris2_cast_Double_to_string(Value *input) {
double x = idris2_vp_to_Double(input);
Value *cast_double_to_string(Value *input) {
Value_Double *from = (Value_Double *)input;
int l = snprintf(NULL, 0, "%f", from->d);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%f", from->d);
int l = snprintf(NULL, 0, "%f", x);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%f", x);
return (Value *)retVal;
}
Value *cast_char_to_Bits8(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeBits8((uint8_t)from->c);
}
Value *cast_char_to_Bits16(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeBits16((uint16_t)from->c);
}
Value *cast_char_to_Bits32(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeBits32((uint32_t)from->c);
}
Value *cast_char_to_Bits64(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeBits64((uint64_t)from->c);
}
Value *cast_char_to_Int8(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeInt8((int8_t)from->c);
}
Value *cast_char_to_Int16(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeInt16((int16_t)from->c);
}
Value *cast_char_to_Int32(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeInt32((int32_t)from->c);
}
Value *cast_char_to_Int64(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeInt64((int64_t)from->c);
}
Value *cast_char_to_Integer(Value *input) {
Value_Char *from = (Value_Char *)input;
Value_Integer *retVal = makeInteger();
mpz_set_si(retVal->i, from->c);
Value *idris2_cast_Char_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_si(retVal->i, idris2_vp_to_Char(input));
return (Value *)retVal;
}
Value *cast_char_to_double(Value *input) {
Value_Char *from = (Value_Char *)input;
return (Value *)makeDouble((unsigned char)from->c);
}
Value *cast_char_to_string(Value *input) {
Value_Char *from = (Value_Char *)input;
Value_String *retVal = makeEmptyString(2);
retVal->str[0] = from->c;
Value *idris2_cast_Char_to_string(Value *input) {
Value_String *retVal = idris2_mkEmptyString(2);
retVal->str[0] = idris2_vp_to_Char(input);
return (Value *)retVal;
}
Value *cast_string_to_Bits8(Value *input) {
Value *idris2_cast_String_to_Bits8(Value *input) {
Value_String *from = (Value_String *)input;
return (Value *)makeBits8((uint8_t)atoi(from->str));
return (Value *)idris2_mkBits8((uint8_t)atoi(from->str));
}
Value *cast_string_to_Bits16(Value *input) {
Value *idris2_cast_String_to_Bits16(Value *input) {
Value_String *from = (Value_String *)input;
return (Value *)makeBits16((uint16_t)atoi(from->str));
return (Value *)idris2_mkBits16((uint16_t)atoi(from->str));
}
Value *cast_string_to_Bits32(Value *input) {
Value *idris2_cast_String_to_Bits32(Value *input) {
Value_String *from = (Value_String *)input;
return (Value *)makeBits32((uint32_t)atoi(from->str));
return (Value *)idris2_mkBits32((uint32_t)atoi(from->str));
}
Value *cast_string_to_Bits64(Value *input) {
Value *idris2_cast_String_to_Bits64(Value *input) {
Value_String *from = (Value_String *)input;
return (Value *)makeBits64((uint64_t)atoi(from->str));
return (Value *)idris2_mkBits64((uint64_t)atoi(from->str));
}
Value *cast_string_to_Int8(Value *input) {
Value *idris2_cast_String_to_Int8(Value *input) {
Value_String *from = (Value_String *)input;
return (Value *)makeInt8((int8_t)atoi(from->str));
return (Value *)idris2_mkInt8((int8_t)atoi(from->str));
}
Value *cast_string_to_Int16(Value *input) {
Value *idris2_cast_String_to_Int16(Value *input) {
Value_String *from = (Value_String *)input;
return (Value *)makeInt16((int16_t)atoi(from->str));
return (Value *)idris2_mkInt16((int16_t)atoi(from->str));
}
Value *cast_string_to_Int32(Value *input) {
Value *idris2_cast_String_to_Int32(Value *input) {
Value_String *from = (Value_String *)input;
return (Value *)makeInt32((int32_t)atoi(from->str));
return (Value *)idris2_mkInt32((int32_t)atoi(from->str));
}
Value *cast_string_to_Int64(Value *input) {
Value *idris2_cast_String_to_Int64(Value *input) {
Value_String *from = (Value_String *)input;
return (Value *)makeInt64((int64_t)atoi(from->str));
return (Value *)idris2_mkInt64((int64_t)atoi(from->str));
}
Value *cast_string_to_Integer(Value *input) {
Value *idris2_cast_String_to_Integer(Value *input) {
Value_String *from = (Value_String *)input;
Value_Integer *retVal = makeInteger();
Value_Integer *retVal = idris2_mkInteger();
mpz_set_str(retVal->i, from->str, 10);
return (Value *)retVal;
}
Value *cast_string_to_double(Value *input) {
Value_Double *retVal = IDRIS2_NEW_VALUE(Value_Double);
retVal->header.tag = DOUBLE_TAG;
Value_String *from = (Value_String *)input;
retVal->d = atof(from->str);
return (Value *)retVal;
}
Value *cast_string_to_char(Value *input) {
Value_Char *retVal = IDRIS2_NEW_VALUE(Value_Char);
retVal->header.tag = CHAR_TAG;
Value_String *from = (Value_String *)input;
retVal->c = from->str[0];
return (Value *)retVal;
Value *idris2_cast_String_to_Double(Value *input) {
return (Value *)idris2_mkDouble(atof(((Value_String *)input)->str));
}
/* conversions from Bits8 */
Value *cast_Bits8_to_Bits16(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeBits16((uint16_t)from->ui8);
}
Value *cast_Bits8_to_Bits32(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeBits32((uint32_t)from->ui8);
}
Value *cast_Bits8_to_Bits64(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeBits64((uint64_t)from->ui8);
}
Value *cast_Bits8_to_Int8(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeInt8((int8_t)from->ui8);
}
Value *cast_Bits8_to_Int16(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeInt16((int16_t)from->ui8);
}
Value *cast_Bits8_to_Int32(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeInt32((int32_t)from->ui8);
}
Value *cast_Bits8_to_Int64(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeInt64((int64_t)from->ui8);
}
Value *cast_Bits8_to_Integer(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
Value_Integer *retVal = makeInteger();
mpz_set_ui(retVal->i, from->ui8);
Value *idris2_cast_Bits8_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_ui(retVal->i, idris2_vp_to_Bits8(input));
return (Value *)retVal;
}
Value *cast_Bits8_to_double(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeDouble((double)from->ui8);
}
Value *idris2_cast_Bits8_to_string(Value *input) {
uint8_t x = idris2_vp_to_Bits8(input);
Value *cast_Bits8_to_char(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
return (Value *)makeChar((unsigned char)from->ui8);
}
Value *cast_Bits8_to_string(Value *input) {
Value_Bits8 *from = (Value_Bits8 *)input;
int l = snprintf(NULL, 0, "%" PRIu8 "", from->ui8);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%" PRIu8 "", from->ui8);
int l = snprintf(NULL, 0, "%" PRIu8 "", x);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%" PRIu8 "", x);
return (Value *)retVal;
}
/* conversions from Bits16 */
Value *cast_Bits16_to_Bits8(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeBits8((uint8_t)from->ui16);
}
Value *cast_Bits16_to_Bits32(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeBits32((uint32_t)from->ui16);
}
Value *cast_Bits16_to_Bits64(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeBits64((uint64_t)from->ui16);
}
Value *cast_Bits16_to_Int8(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeInt8((int8_t)from->ui16);
}
Value *cast_Bits16_to_Int16(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeInt16((int16_t)from->ui16);
}
Value *cast_Bits16_to_Int32(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeInt32((int32_t)from->ui16);
}
Value *cast_Bits16_to_Int64(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeInt64((int64_t)from->ui16);
}
Value *cast_Bits16_to_Integer(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
Value_Integer *retVal = makeInteger();
mpz_set_ui(retVal->i, from->ui16);
Value *idris2_cast_Bits16_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_ui(retVal->i, idris2_vp_to_Bits16(input));
return (Value *)retVal;
}
Value *cast_Bits16_to_double(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeDouble((double)from->ui16);
}
Value *idris2_cast_Bits16_to_string(Value *input) {
uint16_t x = idris2_vp_to_Bits16(input);
Value *cast_Bits16_to_char(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
return (Value *)makeChar((unsigned char)from->ui16);
}
Value *cast_Bits16_to_string(Value *input) {
Value_Bits16 *from = (Value_Bits16 *)input;
int l = snprintf(NULL, 0, "%" PRIu16 "", from->ui16);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%" PRIu16 "", from->ui16);
int l = snprintf(NULL, 0, "%" PRIu16 "", x);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%" PRIu16 "", x);
return (Value *)retVal;
}
/* conversions from Bits32 */
Value *cast_Bits32_to_Bits8(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeBits8((uint8_t)from->ui32);
}
Value *cast_Bits32_to_Bits16(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeBits16((uint16_t)from->ui32);
}
Value *cast_Bits32_to_Bits64(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeBits64((uint64_t)from->ui32);
}
Value *cast_Bits32_to_Int8(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeInt8((int8_t)from->ui32);
}
Value *cast_Bits32_to_Int16(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeInt16((int16_t)from->ui32);
}
Value *cast_Bits32_to_Int32(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeInt32((int32_t)from->ui32);
}
Value *cast_Bits32_to_Int64(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeInt64((int64_t)from->ui32);
}
Value *cast_Bits32_to_Integer(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
Value_Integer *retVal = makeInteger();
mpz_set_ui(retVal->i, from->ui32);
Value *idris2_cast_Bits32_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_ui(retVal->i, idris2_vp_to_Bits32(input));
return (Value *)retVal;
}
Value *cast_Bits32_to_double(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeDouble((double)from->ui32);
}
Value *idris2_cast_Bits32_to_string(Value *input) {
uint32_t x = idris2_vp_to_Bits32(input);
Value *cast_Bits32_to_char(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
return (Value *)makeChar((unsigned char)from->ui32);
}
Value *cast_Bits32_to_string(Value *input) {
Value_Bits32 *from = (Value_Bits32 *)input;
int l = snprintf(NULL, 0, "%" PRIu32 "", from->ui32);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%" PRIu32 "", from->ui32);
int l = snprintf(NULL, 0, "%" PRIu32 "", x);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%" PRIu32 "", x);
return (Value *)retVal;
}
/* conversions from Bits64 */
Value *cast_Bits64_to_Bits8(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeBits8((uint8_t)from->ui64);
}
Value *cast_Bits64_to_Bits16(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeBits16((uint16_t)from->ui64);
}
Value *cast_Bits64_to_Bits32(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeBits32((uint32_t)from->ui64);
}
Value *cast_Bits64_to_Int8(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeInt8((int8_t)from->ui64);
}
Value *cast_Bits64_to_Int16(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeInt16((int16_t)from->ui64);
}
Value *cast_Bits64_to_Int32(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeInt32((int32_t)from->ui64);
}
Value *cast_Bits64_to_Int64(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeInt64((int64_t)from->ui64);
}
Value *cast_Bits64_to_Integer(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
Value_Integer *retVal = makeInteger();
mpz_set_ui(retVal->i, from->ui64);
Value *idris2_cast_Bits64_to_Integer(Value *input) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_ui(retVal->i, idris2_vp_to_Bits64(input));
return (Value *)retVal;
}
Value *cast_Bits64_to_double(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeDouble((double)from->ui64);
}
Value *idris2_cast_Bits64_to_string(Value *input) {
uint64_t x = idris2_vp_to_Bits64(input);
Value *cast_Bits64_to_char(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
return (Value *)makeChar((unsigned char)from->ui64);
}
Value *cast_Bits64_to_string(Value *input) {
Value_Bits64 *from = (Value_Bits64 *)input;
int l = snprintf(NULL, 0, "%" PRIu64 "", from->ui64);
Value_String *retVal = makeEmptyString(l + 1);
sprintf(retVal->str, "%" PRIu64 "", from->ui64);
int l = snprintf(NULL, 0, "%" PRIu64 "", x);
Value_String *retVal = idris2_mkEmptyString(l + 1);
sprintf(retVal->str, "%" PRIu64 "", x);
return (Value *)retVal;
}
@ -730,57 +239,57 @@ uint64_t mpz_get_lsb(mpz_t i, mp_bitcnt_t b) {
return retVal;
}
Value *cast_Integer_to_Bits8(Value *input) {
Value *idris2_cast_Integer_to_Bits8(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeBits8((uint8_t)mpz_get_lsb(from->i, 8));
return (Value *)idris2_mkBits8((uint8_t)mpz_get_lsb(from->i, 8));
}
Value *cast_Integer_to_Bits16(Value *input) {
Value *idris2_cast_Integer_to_Bits16(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeBits16((uint16_t)mpz_get_lsb(from->i, 16));
return (Value *)idris2_mkBits16((uint16_t)mpz_get_lsb(from->i, 16));
}
Value *cast_Integer_to_Bits32(Value *input) {
Value *idris2_cast_Integer_to_Bits32(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeBits32((uint32_t)mpz_get_lsb(from->i, 32));
return (Value *)idris2_mkBits32((uint32_t)mpz_get_lsb(from->i, 32));
}
Value *cast_Integer_to_Bits64(Value *input) {
Value *idris2_cast_Integer_to_Bits64(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeBits64((uint64_t)mpz_get_lsb(from->i, 64));
return (Value *)idris2_mkBits64((uint64_t)mpz_get_lsb(from->i, 64));
}
Value *cast_Integer_to_Int8(Value *input) {
Value *idris2_cast_Integer_to_Int8(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeInt8((int8_t)mpz_get_lsb(from->i, 8));
return (Value *)idris2_mkInt8((int8_t)mpz_get_lsb(from->i, 8));
}
Value *cast_Integer_to_Int16(Value *input) {
Value *idris2_cast_Integer_to_Int16(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeInt16((int16_t)mpz_get_lsb(from->i, 16));
return (Value *)idris2_mkInt16((int16_t)mpz_get_lsb(from->i, 16));
}
Value *cast_Integer_to_Int32(Value *input) {
Value *idris2_cast_Integer_to_Int32(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeInt32((int32_t)mpz_get_lsb(from->i, 32));
return (Value *)idris2_mkInt32((int32_t)mpz_get_lsb(from->i, 32));
}
Value *cast_Integer_to_Int64(Value *input) {
Value *idris2_cast_Integer_to_Int64(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeInt64((int64_t)mpz_get_lsb(from->i, 64));
return (Value *)idris2_mkInt64((int64_t)mpz_get_lsb(from->i, 64));
}
Value *cast_Integer_to_double(Value *input) {
Value *idris2_cast_Integer_to_Double(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeDouble(mpz_get_d(from->i));
return (Value *)idris2_mkDouble(mpz_get_d(from->i));
}
Value *cast_Integer_to_char(Value *input) {
Value *idris2_cast_Integer_to_Char(Value *input) {
Value_Integer *from = (Value_Integer *)input;
return (Value *)makeChar((unsigned char)mpz_get_lsb(from->i, 8));
return (Value *)idris2_mkChar((unsigned char)mpz_get_lsb(from->i, 8));
}
Value *cast_Integer_to_string(Value *input) {
Value *idris2_cast_Integer_to_string(Value *input) {
Value_Integer *from = (Value_Integer *)input;
Value_String *retVal = IDRIS2_NEW_VALUE(Value_String);

View File

@ -4,147 +4,226 @@
#include <gmp.h>
#include <stdio.h>
Value *cast_Int8_to_Bits8(Value *);
Value *cast_Int8_to_Bits16(Value *);
Value *cast_Int8_to_Bits32(Value *);
Value *cast_Int8_to_Bits64(Value *);
Value *cast_Int8_to_Int16(Value *);
Value *cast_Int8_to_Int32(Value *);
Value *cast_Int8_to_Int64(Value *);
Value *cast_Int8_to_Integer(Value *);
Value *cast_Int8_to_double(Value *);
Value *cast_Int8_to_char(Value *);
Value *cast_Int8_to_string(Value *);
#define idris2_cast_Int8_to_Bits8(x) (x)
#define idris2_cast_Int8_to_Bits16(x) (x)
#define idris2_cast_Int8_to_Bits32(x) (x)
#define idris2_cast_Int8_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Int8(x)))
#define idris2_cast_Int8_to_Int16(x) \
(idris2_mkInt16((int16_t)idris2_vp_to_Int8(x)))
#define idris2_cast_Int8_to_Int32(x) \
(idris2_mkInt32((int32_t)idris2_vp_to_Int8(x)))
#define idris2_cast_Int8_to_Int64(x) \
(idris2_mkInt64((int64_t)idris2_vp_to_Int8(x)))
Value *idris2_cast_Int8_to_Integer(Value *);
#define idris2_cast_Int8_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Int8(x)))
#define idris2_cast_Int8_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Int8(x)))
Value *idris2_cast_Int8_to_string(Value *);
Value *cast_Int16_to_Bits8(Value *);
Value *cast_Int16_to_Bits16(Value *);
Value *cast_Int16_to_Bits32(Value *);
Value *cast_Int16_to_Bits64(Value *);
Value *cast_Int16_to_Int8(Value *);
Value *cast_Int16_to_Int32(Value *);
Value *cast_Int16_to_Int64(Value *);
Value *cast_Int16_to_Integer(Value *);
Value *cast_Int16_to_double(Value *);
Value *cast_Int16_to_char(Value *);
Value *cast_Int16_to_string(Value *);
#define idris2_cast_Int16_to_Bits8(x) \
(idris2_mkBits8((uint8_t)idris2_vp_to_Int16(x)))
#define idris2_cast_Int16_to_Bits16(x) (x)
#define idris2_cast_Int16_to_Bits32(x) (x)
#define idris2_cast_Int16_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Int16(x)))
#define idris2_cast_Int16_to_Int8(x) \
(idris2_mkInt8((int8_t)idris2_vp_to_Int16(x)))
#define idris2_cast_Int16_to_Int32(x) (x)
#define idris2_cast_Int16_to_Int64(x) \
(idris2_mkInt64((int64_t)idris2_vp_to_Int16(x)))
Value *idris2_cast_Int16_to_Integer(Value *);
#define idris2_cast_Int16_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Int16(x)))
#define idris2_cast_Int16_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Int16(x)))
Value *idris2_cast_Int16_to_string(Value *);
Value *cast_Int32_to_Bits8(Value *);
Value *cast_Int32_to_Bits16(Value *);
Value *cast_Int32_to_Bits32(Value *);
Value *cast_Int32_to_Bits64(Value *);
Value *cast_Int32_to_Int8(Value *);
Value *cast_Int32_to_Int16(Value *);
Value *cast_Int32_to_Int64(Value *);
Value *cast_Int32_to_Integer(Value *);
Value *cast_Int32_to_double(Value *);
Value *cast_Int32_to_char(Value *);
Value *cast_Int32_to_string(Value *);
#define idris2_cast_Int32_to_Bits8(x) \
(idris2_mkBits8((uint8_t)idris2_vp_to_Int32(x)))
#define idris2_cast_Int32_to_Bits16(x) \
(idris2_mkBits16((uint16_t)idris2_vp_to_Int32(x)))
#define idris2_cast_Int32_to_Bits32(x) (x)
#define idris2_cast_Int32_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Int32(x)))
#define idris2_cast_Int32_to_Int8(x) \
(idris2_mkInt8((int8_t)idris2_vp_to_Int32(x)))
#define idris2_cast_Int32_to_Int16(x) \
(idris2_mkInt16((int16_t)idris2_vp_to_Int32(x)))
#define idris2_cast_Int32_to_Int64(x) \
(idris2_mkInt64((int64_t)idris2_vp_to_Int32(x)))
Value *idris2_cast_Int32_to_Integer(Value *);
#define idris2_cast_Int32_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Int32(x)))
#define idris2_cast_Int32_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Int32(x)))
Value *idris2_cast_Int32_to_string(Value *);
Value *cast_Int64_to_Bits8(Value *);
Value *cast_Int64_to_Bits16(Value *);
Value *cast_Int64_to_Bits32(Value *);
Value *cast_Int64_to_Bits64(Value *);
Value *cast_Int64_to_Int8(Value *);
Value *cast_Int64_to_Int16(Value *);
Value *cast_Int64_to_Int32(Value *);
Value *cast_Int64_to_Int64(Value *);
Value *cast_Int64_to_Integer(Value *);
Value *cast_Int64_to_double(Value *);
Value *cast_Int64_to_char(Value *);
Value *cast_Int64_to_string(Value *);
#define idris2_cast_Int64_to_Bits8(x) \
(idris2_mkBits8((uint8_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Bits16(x) \
(idris2_mkBits16((uint16_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Bits32(x) \
(idris2_mkBits32((uint32_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Int8(x) \
(idris2_mkInt8((int8_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Int16(x) \
(idris2_mkInt16((int16_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Int32(x) \
(idris2_mkInt32((int32_t)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Int64(x) (newReference(x))
Value *idris2_cast_Int64_to_Integer(Value *);
#define idris2_cast_Int64_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Int64(x)))
#define idris2_cast_Int64_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Int64(x)))
Value *idris2_cast_Int64_to_string(Value *);
Value *cast_double_to_Bits8(Value *);
Value *cast_double_to_Bits16(Value *);
Value *cast_double_to_Bits32(Value *);
Value *cast_double_to_Bits64(Value *);
Value *cast_double_to_Int8(Value *);
Value *cast_double_to_Int16(Value *);
Value *cast_double_to_Int32(Value *);
Value *cast_double_to_Int64(Value *);
Value *cast_double_to_Integer(Value *);
Value *cast_double_to_char(Value *);
Value *cast_double_to_string(Value *);
#define idris2_cast_Double_to_Bits8(x) \
(idris2_mkBits8((uint8_t)idris2_vp_to_Double(x)))
#define idris2_cast_Double_to_Bits16(x) \
(idris2_mkBits16((uint16_t)idris2_vp_to_Double(x)))
#define idris2_cast_Double_to_Bits32(x) \
(idris2_mkBits32((uint32_t)idris2_vp_to_Double(x)))
#define idris2_cast_Double_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Double(x)))
#define idris2_cast_Double_to_Int(x) \
(idris2_mkInt8((int8_t)idris2_vp_to_Double(x)))
#define idris2_cast_Double_to_Int16(x) \
(idris2_mkInt16((int16_t)idris2_vp_to_Double(x)))
#define idris2_cast_Double_to_Int32(x) \
(idris2_mkInt32((int32_t)idris2_vp_to_Double(x)))
#define idris2_cast_Double_to_Int64(x) \
(idris2_mkInt64((int64_t)idris2_vp_to_Double(x)))
Value *idris2_cast_Double_to_Integer(Value *);
#define idris2_cast_Double_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Double))
Value *idris2_cast_Double_to_string(Value *);
Value *cast_char_to_Bits8(Value *);
Value *cast_char_to_Bits16(Value *);
Value *cast_char_to_Bits32(Value *);
Value *cast_char_to_Bits64(Value *);
Value *cast_char_to_Int8(Value *);
Value *cast_char_to_Int16(Value *);
Value *cast_char_to_Int32(Value *);
Value *cast_char_to_Int64(Value *);
Value *cast_char_to_Integer(Value *);
Value *cast_char_to_double(Value *);
Value *cast_char_to_string(Value *);
#define idris2_cast_Char_to_Bits8(x) \
(idris2_mkBits8((uint8_t)idris2_vp_to_Char(x)))
#define idris2_cast_Char_to_Bits16(x) \
(idris2_mkBits16((uint16_t)idris2_vp_to_Char(x)))
#define idris2_cast_Char_to_Bits32(x) \
(idris2_mkBits32((uint32_t)idris2_vp_to_Char(x)))
#define idris2_cast_Char_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Char(x)))
#define idris2_cast_Char_to_Int8(x) \
(idris2_mkInt8((int8_t)idris2_vp_to_Char(x)))
#define idris2_cast_Char_to_Int16(x) \
(idris2_mkInt16((int16_t)idris2_vp_to_Char(x)))
#define idris2_cast_Char_to_Int32(x) \
(idris2_mkInt32((int32_t)idris2_vp_to_Char(x)))
#define idris2_cast_Char_to_Int64(x) \
(idris2_mkInt64((int64_t)idris2_vp_to_Char(x)))
Value *idris2_cast_Char_to_Integer(Value *);
#define idris2_cast_Char_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Char(x)))
Value *idris2_cast_Char_to_string(Value *);
Value *cast_string_to_Bits8(Value *);
Value *cast_string_to_Bits16(Value *);
Value *cast_string_to_Bits32(Value *);
Value *cast_string_to_Bits64(Value *);
Value *cast_string_to_Int8(Value *);
Value *cast_string_to_Int16(Value *);
Value *cast_string_to_Int32(Value *);
Value *cast_string_to_Int64(Value *);
Value *cast_string_to_Integer(Value *);
Value *cast_string_to_double(Value *);
Value *cast_string_to_char(Value *);
Value *idris2_cast_String_to_Bits8(Value *);
Value *idris2_cast_String_to_Bits16(Value *);
Value *idris2_cast_String_to_Bits32(Value *);
Value *idris2_cast_String_to_Bits64(Value *);
Value *idris2_cast_String_to_Int8(Value *);
Value *idris2_cast_String_to_Int16(Value *);
Value *idris2_cast_String_to_Int32(Value *);
Value *idris2_cast_String_to_Int64(Value *);
Value *idris2_cast_String_to_Integer(Value *);
Value *idris2_cast_String_to_Double(Value *);
#define idris2_cast_String_to_Char(x) \
(idris2_mkChar(((Value_String *)(x))->str[0]))
Value *cast_Bits8_to_Bits16(Value *input);
Value *cast_Bits8_to_Bits32(Value *input);
Value *cast_Bits8_to_Bits64(Value *input);
Value *cast_Bits8_to_Int8(Value *input);
Value *cast_Bits8_to_Int16(Value *input);
Value *cast_Bits8_to_Int32(Value *input);
Value *cast_Bits8_to_Int64(Value *input);
Value *cast_Bits8_to_Integer(Value *input);
Value *cast_Bits8_to_double(Value *input);
Value *cast_Bits8_to_char(Value *input);
Value *cast_Bits8_to_string(Value *input);
#define idris2_cast_Bits8_to_Bits16(x) (x)
#define idris2_cast_Bits8_to_Bits32(x) (x)
#define idris2_cast_Bits8_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Bits8(x)))
#define idris2_cast_Bits8_to_Int8(x) (x)
#define idris2_cast_Bits8_to_Int16(x) (x)
#define idris2_cast_Bits8_to_Int32(x) (x)
#define idris2_cast_Bits8_to_Int64(x) \
(idris2_mkInt64((int64_t)idris2_vp_to_Bits8(x)))
Value *idris2_cast_Bits8_to_Integer(Value *input);
#define idris2_cast_Bits8_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Bits8(x)))
#define idris2_cast_Bits8_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Bits8(x)))
Value *idris2_cast_Bits8_to_string(Value *input);
Value *cast_Bits16_to_Bits8(Value *input);
Value *cast_Bits16_to_Bits32(Value *input);
Value *cast_Bits16_to_Bits64(Value *input);
Value *cast_Bits16_to_Int8(Value *input);
Value *cast_Bits16_to_Int16(Value *input);
Value *cast_Bits16_to_Int32(Value *input);
Value *cast_Bits16_to_Int64(Value *input);
Value *cast_Bits16_to_Integer(Value *input);
Value *cast_Bits16_to_double(Value *input);
Value *cast_Bits16_to_char(Value *input);
Value *cast_Bits16_to_string(Value *input);
#define idris2_cast_Bits16_to_Bits8(x) \
(idris2_mkBits8((uint8_t)idris2_vp_to_Bits16(x)))
#define idris2_cast_Bits16_to_Bits32(x) \
(idris2_mkBits32((uint32_t)idris2_vp_to_Bits16(x)))
#define idris2_cast_Bits16_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Bits16(x)))
#define idris2_cast_Bits16_to_Int8(x) \
(idris2_mkInt8((int8_t)idris2_vp_to_Bits16(x)))
#define idris2_cast_Bits16_to_Int16(x) \
(idris2_mkInt16((int16_t)idris2_vp_to_Bits16(x)))
#define idris2_cast_Bits16_to_Int32(x) \
(idris2_mkInt32((int32_t)idris2_vp_to_Bits16(x)))
#define idris2_cast_Bits16_to_Int64(x) \
(idris2_mkInt64((int64_t)idris2_vp_to_Bits16(x)))
Value *idris2_cast_Bits16_to_Integer(Value *input);
#define idris2_cast_Bits16_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Bits16(x)))
#define idris2_cast_Bits16_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Bits16(x)))
Value *idris2_cast_Bits16_to_string(Value *input);
Value *cast_Bits32_to_Bits8(Value *input);
Value *cast_Bits32_to_Bits16(Value *input);
Value *cast_Bits32_to_Bits64(Value *input);
Value *cast_Bits32_to_Int8(Value *input);
Value *cast_Bits32_to_Int16(Value *input);
Value *cast_Bits32_to_Int32(Value *input);
Value *cast_Bits32_to_Int64(Value *input);
Value *cast_Bits32_to_Integer(Value *input);
Value *cast_Bits32_to_double(Value *input);
Value *cast_Bits32_to_char(Value *input);
Value *cast_Bits32_to_string(Value *input);
#define idris2_cast_Bits32_to_Bits8(x) \
(idris2_mkBits8((uint8_t)idris2_vp_to_Bits32(x)))
#define idris2_cast_Bits32_to_Bits16(x) \
(idris2_mkBits16((uint16_t)idris2_vp_to_Bits32(x)))
#define idris2_cast_Bits32_to_Bits64(x) \
(idris2_mkBits64((uint64_t)idris2_vp_to_Bits32(x)))
#define idris2_cast_Bits32_to_Int8(x) \
(idris2_mkBits8((int8_t)idris2_vp_to_Bits32(x)))
#define idris2_cast_Bits32_to_Int16(x) \
(idris2_mkBits16((int16_t)idris2_vp_to_Bits32(x)))
#define idris2_cast_Bits32_to_Int32(x) \
(idris2_mkBits32((int32_t)idris2_vp_to_Bits32(x)))
#define idris2_cast_Bits32_to_Int64(x) \
(idris2_mkBits64((int64_t)idris2_vp_to_Bits32(x)))
Value *idris2_cast_Bits32_to_Integer(Value *input);
#define idris2_cast_Bits32_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Bits32(x)))
#define idris2_cast_Bits32_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Bits32(x)))
Value *idris2_cast_Bits32_to_string(Value *input);
Value *cast_Bits64_to_Bits8(Value *input);
Value *cast_Bits64_to_Bits16(Value *input);
Value *cast_Bits64_to_Bits32(Value *input);
Value *cast_Bits64_to_Int8(Value *input);
Value *cast_Bits64_to_Int16(Value *input);
Value *cast_Bits64_to_Int32(Value *input);
Value *cast_Bits64_to_Int64(Value *input);
Value *cast_Bits64_to_Integer(Value *input);
Value *cast_Bits64_to_double(Value *input);
Value *cast_Bits64_to_char(Value *input);
Value *cast_Bits64_to_string(Value *input);
#define idris2_cast_Bits64_to_Bits8(x) \
(idris2_mkBits8((uint8_t)idris2_vp_to_Bits64(x)))
#define idris2_cast_Bits64_to_Bits16(x) \
(idris2_mkBits16((uint16_t)idris2_vp_to_Bits64(x)))
#define idris2_cast_Bits64_to_Bits32(x) \
(idris2_mkBits32((uint32_t)idris2_vp_to_Bits64(x)))
#define idris2_cast_Bits64_to_Int8(x) \
(idris2_mkInt8((int8_t)idris2_vp_to_Bits64(x)))
#define idris2_cast_Bits64_to_Int16(x) \
(idris2_mkInt16((int16_t)idris2_vp_to_Bits64(x)))
#define idris2_cast_Bits64_to_Int32(x) \
(idris2_mkInt32((int32_t)idris2_vp_to_Bits64(x)))
#define idris2_cast_Bits64_to_Int64(x) \
(idris2_mkInt64((int64_t)idris2_vp_to_Bits64(x)))
Value *idris2_cast_Bits64_to_Integer(Value *input);
#define idris2_cast_Bits64_to_Double(x) \
(idris2_mkDouble((double)idris2_vp_to_Bits64(x)))
#define idris2_cast_Bits64_to_Char(x) \
(idris2_mkChar((unsigned char)idris2_vp_to_Bits64(x)))
Value *idris2_cast_Bits64_to_string(Value *input);
Value *cast_Integer_to_Bits8(Value *input);
Value *cast_Integer_to_Bits16(Value *input);
Value *cast_Integer_to_Bits32(Value *input);
Value *cast_Integer_to_Bits64(Value *input);
Value *cast_Integer_to_Int8(Value *input);
Value *cast_Integer_to_Int16(Value *input);
Value *cast_Integer_to_Int32(Value *input);
Value *cast_Integer_to_Int64(Value *input);
Value *cast_Integer_to_double(Value *input);
Value *cast_Integer_to_char(Value *input);
Value *cast_Integer_to_string(Value *input);
Value *idris2_cast_Integer_to_Bits8(Value *input);
Value *idris2_cast_Integer_to_Bits16(Value *input);
Value *idris2_cast_Integer_to_Bits32(Value *input);
Value *idris2_cast_Integer_to_Bits64(Value *input);
Value *idris2_cast_Integer_to_Int8(Value *input);
Value *idris2_cast_Integer_to_Int16(Value *input);
Value *idris2_cast_Integer_to_Int32(Value *input);
Value *idris2_cast_Integer_to_Int64(Value *input);
Value *idris2_cast_Integer_to_Double(Value *input);
Value *idris2_cast_Integer_to_Char(Value *input);
Value *idris2_cast_Integer_to_string(Value *input);

View File

@ -5,11 +5,13 @@
Value *clockTimeMonotonic() { return clockTimeUtc(); }
Value *clockTimeUtc() { return (Value *)makeBits64(time(NULL) * NSEC_PER_SEC); }
Value *clockTimeUtc() {
return (Value *)idris2_mkBits64(time(NULL) * NSEC_PER_SEC);
}
Value *clockTimeProcess() {
uint64_t time_ns = clock() / CLOCKS_PER_NSEC;
return (Value *)makeBits64(time_ns);
return (Value *)idris2_mkBits64(time_ns);
}
Value *clockTimeThread() { return clockTimeProcess(); }
@ -21,11 +23,9 @@ Value *clockTimeGcReal() { return NULL; }
int clockValid(Value *clock) { return clock != NULL; }
uint64_t clockSecond(Value *clock) {
uint64_t totalNano = ((Value_Bits64 *)clock)->ui64;
return totalNano / NSEC_PER_SEC;
return idris2_vp_to_Bits64(clock) / NSEC_PER_SEC;
}
uint64_t clockNanosecond(Value *clock) {
uint64_t totalNano = ((Value_Bits64 *)clock)->ui64;
return totalNano % NSEC_PER_SEC;
return idris2_vp_to_Bits64(clock) % NSEC_PER_SEC;
}

View File

@ -2,664 +2,161 @@
#include "memoryManagement.h"
#include "runtime.h"
double unpackDouble(Value *d) { return ((Value_Double *)d)->d; }
/* add */
Value *add_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 + ((Value_Bits8 *)y)->ui8);
}
Value *add_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 +
((Value_Bits16 *)y)->ui16);
}
Value *add_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 +
((Value_Bits32 *)y)->ui32);
}
Value *add_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 +
((Value_Bits64 *)y)->ui64);
}
Value *add_Int8(Value *x, Value *y) {
return (Value *)makeInt8(((Value_Int8 *)x)->i8 + ((Value_Int8 *)y)->i8);
}
Value *add_Int16(Value *x, Value *y) {
return (Value *)makeInt16(((Value_Int16 *)x)->i16 + ((Value_Int16 *)y)->i16);
}
Value *add_Int32(Value *x, Value *y) {
return (Value *)makeInt32(((Value_Int32 *)x)->i32 + ((Value_Int32 *)y)->i32);
}
Value *add_Int64(Value *x, Value *y) {
return (Value *)makeInt64(((Value_Int64 *)x)->i64 + ((Value_Int64 *)y)->i64);
}
Value *add_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_add_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mpz_add(retVal->i, ((Value_Integer *)x)->i, ((Value_Integer *)y)->i);
return (Value *)retVal;
}
Value *add_double(Value *x, Value *y) {
return (Value *)makeDouble(((Value_Double *)x)->d + ((Value_Double *)y)->d);
}
/* sub */
Value *sub_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 - ((Value_Bits8 *)y)->ui8);
}
Value *sub_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 -
((Value_Bits16 *)y)->ui16);
}
Value *sub_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 -
((Value_Bits32 *)y)->ui32);
}
Value *sub_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 -
((Value_Bits64 *)y)->ui64);
}
Value *sub_Int8(Value *x, Value *y) {
return (Value *)makeInt8(((Value_Int8 *)x)->i8 - ((Value_Int8 *)y)->i8);
}
Value *sub_Int16(Value *x, Value *y) {
return (Value *)makeInt16(((Value_Int16 *)x)->i16 - ((Value_Int16 *)y)->i16);
}
Value *sub_Int32(Value *x, Value *y) {
return (Value *)makeInt32(((Value_Int32 *)x)->i32 - ((Value_Int32 *)y)->i32);
}
Value *sub_Int64(Value *x, Value *y) {
return (Value *)makeInt64(((Value_Int64 *)x)->i64 - ((Value_Int64 *)y)->i64);
}
Value *sub_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_sub_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mpz_sub(retVal->i, ((Value_Integer *)x)->i, ((Value_Integer *)y)->i);
return (Value *)retVal;
}
Value *sub_double(Value *x, Value *y) {
return (Value *)makeDouble(((Value_Double *)x)->d - ((Value_Double *)y)->d);
}
/* negate */
Value *negate_Bits8(Value *x) {
return (Value *)makeBits8(-((Value_Bits8 *)x)->ui8);
}
Value *negate_Bits16(Value *x) {
return (Value *)makeBits16(-((Value_Bits16 *)x)->ui16);
}
Value *negate_Bits32(Value *x) {
return (Value *)makeBits32(-((Value_Bits32 *)x)->ui32);
}
Value *negate_Bits64(Value *x) {
return (Value *)makeBits64(-((Value_Bits64 *)x)->ui64);
}
Value *negate_Int8(Value *x) {
return (Value *)makeInt8(-((Value_Int8 *)x)->i8);
}
Value *negate_Int16(Value *x) {
return (Value *)makeInt16(-((Value_Int16 *)x)->i16);
}
Value *negate_Int32(Value *x) {
return (Value *)makeInt32(-((Value_Int32 *)x)->i32);
}
Value *negate_Int64(Value *x) {
return (Value *)makeInt64(-((Value_Int64 *)x)->i64);
}
Value *negate_Integer(Value *x) {
Value_Integer *retVal = makeInteger();
Value *idris2_negate_Integer(Value *x) {
Value_Integer *retVal = idris2_mkInteger();
mpz_neg(retVal->i, ((Value_Integer *)x)->i);
return (Value *)retVal;
}
Value *negate_double(Value *x) {
return (Value *)makeDouble(-((Value_Double *)x)->d);
}
/* mul */
Value *mul_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 * ((Value_Bits8 *)y)->ui8);
}
Value *mul_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 *
((Value_Bits16 *)y)->ui16);
}
Value *mul_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 *
((Value_Bits32 *)y)->ui32);
}
Value *mul_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 *
((Value_Bits64 *)y)->ui64);
}
Value *mul_Int8(Value *x, Value *y) {
return (Value *)makeInt8(((Value_Int8 *)x)->i8 * ((Value_Int8 *)y)->i8);
}
Value *mul_Int16(Value *x, Value *y) {
return (Value *)makeInt16(((Value_Int16 *)x)->i16 * ((Value_Int16 *)y)->i16);
}
Value *mul_Int32(Value *x, Value *y) {
return (Value *)makeInt32(((Value_Int32 *)x)->i32 * ((Value_Int32 *)y)->i32);
}
Value *mul_Int64(Value *x, Value *y) {
return (Value *)makeInt64(((Value_Int64 *)x)->i64 * ((Value_Int64 *)y)->i64);
}
Value *mul_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_mul_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mpz_mul(retVal->i, ((Value_Integer *)x)->i, ((Value_Integer *)y)->i);
return (Value *)retVal;
}
Value *mul_double(Value *x, Value *y) {
return (Value *)makeDouble(((Value_Double *)x)->d * ((Value_Double *)y)->d);
}
/* div */
Value *div_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 / ((Value_Bits8 *)y)->ui8);
}
Value *div_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 /
((Value_Bits16 *)y)->ui16);
}
Value *div_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 /
((Value_Bits32 *)y)->ui32);
}
Value *div_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 /
((Value_Bits64 *)y)->ui64);
}
Value *div_Int8(Value *x, Value *y) {
Value *idris2_div_Int8(Value *x, Value *y) {
// Correction term added to convert from truncated division (C default) to
// Euclidean division For proof of correctness, see Division and Modulus for
// Computer Scientists (Daan Leijen)
// https://www.microsoft.com/en-us/research/publication/division-and-modulus-for-computer-scientists/
int8_t num = ((Value_Int8 *)x)->i8;
int8_t denom = ((Value_Int8 *)y)->i8;
int8_t num = idris2_vp_to_Int8(x);
int8_t denom = idris2_vp_to_Int8(y);
int8_t rem = num % denom;
return (Value *)makeInt8(num / denom +
((rem < 0) ? (denom < 0) ? 1 : -1 : 0));
return idris2_mkInt8(num / denom + ((rem < 0) ? (denom < 0) ? 1 : -1 : 0));
}
Value *div_Int16(Value *x, Value *y) {
Value *idris2_div_Int16(Value *x, Value *y) {
// Correction term added to convert from truncated division (C default) to
// Euclidean division For proof of correctness, see Division and Modulus for
// Computer Scientists (Daan Leijen)
// https://www.microsoft.com/en-us/research/publication/division-and-modulus-for-computer-scientists/
int16_t num = ((Value_Int16 *)x)->i16;
int16_t denom = ((Value_Int16 *)y)->i16;
int16_t num = idris2_vp_to_Int16(x);
int16_t denom = idris2_vp_to_Int16(y);
int16_t rem = num % denom;
return (Value *)makeInt16(num / denom +
((rem < 0) ? (denom < 0) ? 1 : -1 : 0));
return idris2_mkInt16(num / denom + ((rem < 0) ? (denom < 0) ? 1 : -1 : 0));
}
Value *div_Int32(Value *x, Value *y) {
Value *idris2_div_Int32(Value *x, Value *y) {
// Correction term added to convert from truncated division (C default) to
// Euclidean division For proof of correctness, see Division and Modulus for
// Computer Scientists (Daan Leijen)
// https://www.microsoft.com/en-us/research/publication/division-and-modulus-for-computer-scientists/
int32_t num = ((Value_Int32 *)x)->i32;
int32_t denom = ((Value_Int32 *)y)->i32;
int32_t num = idris2_vp_to_Int32(x);
int32_t denom = idris2_vp_to_Int32(y);
int32_t rem = num % denom;
return (Value *)makeInt32(num / denom +
((rem < 0) ? (denom < 0) ? 1 : -1 : 0));
return idris2_mkInt32(num / denom + ((rem < 0) ? (denom < 0) ? 1 : -1 : 0));
}
Value *div_Int64(Value *x, Value *y) {
Value *idris2_div_Int64(Value *x, Value *y) {
// Correction term added to convert from truncated division (C default) to
// Euclidean division For proof of correctness, see Division and Modulus for
// Computer Scientists (Daan Leijen)
// https://www.microsoft.com/en-us/research/publication/division-and-modulus-for-computer-scientists/
int64_t num = ((Value_Int64 *)x)->i64;
int64_t denom = ((Value_Int64 *)y)->i64;
int64_t num = idris2_vp_to_Int64(x);
int64_t denom = idris2_vp_to_Int64(y);
int64_t rem = num % denom;
return (Value *)makeInt64(num / denom +
((rem < 0) ? (denom < 0) ? 1 : -1 : 0));
return (Value *)idris2_mkInt64(num / denom +
((rem < 0) ? (denom < 0) ? 1 : -1 : 0));
}
Value *div_Integer(Value *x, Value *y) {
Value *idris2_div_Integer(Value *x, Value *y) {
mpz_t rem, yq;
mpz_inits(rem, yq, NULL);
mpz_mod(rem, ((Value_Integer *)x)->i, ((Value_Integer *)y)->i);
mpz_sub(yq, ((Value_Integer *)x)->i, rem);
Value_Integer *retVal = makeInteger();
Value_Integer *retVal = idris2_mkInteger();
mpz_divexact(retVal->i, yq, ((Value_Integer *)y)->i);
mpz_clears(rem, yq, NULL);
return (Value *)retVal;
}
Value *div_double(Value *x, Value *y) {
return (Value *)makeDouble(((Value_Double *)x)->d / ((Value_Double *)y)->d);
}
/* mod */
Value *mod_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 % ((Value_Bits8 *)y)->ui8);
}
Value *mod_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 %
((Value_Bits16 *)y)->ui16);
}
Value *mod_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 %
((Value_Bits32 *)y)->ui32);
}
Value *mod_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 %
((Value_Bits64 *)y)->ui64);
}
Value *mod_Int8(Value *x, Value *y) {
int8_t num = ((Value_Int8 *)x)->i8;
int8_t denom = ((Value_Int8 *)y)->i8;
Value *idris2_mod_Int8(Value *x, Value *y) {
int8_t num = idris2_vp_to_Int8(x);
int8_t denom = idris2_vp_to_Int8(y);
denom = (denom < 0) ? -denom : denom;
return (Value *)makeInt8(num % denom + (num < 0 ? denom : 0));
return (Value *)idris2_mkInt8(num % denom + (num < 0 ? denom : 0));
}
Value *mod_Int16(Value *x, Value *y) {
int16_t num = ((Value_Int16 *)x)->i16;
int16_t denom = ((Value_Int16 *)y)->i16;
Value *idris2_mod_Int16(Value *x, Value *y) {
int16_t num = idris2_vp_to_Int16(x);
int16_t denom = idris2_vp_to_Int16(y);
denom = (denom < 0) ? -denom : denom;
return (Value *)makeInt16(num % denom + (num < 0 ? denom : 0));
return (Value *)idris2_mkInt16(num % denom + (num < 0 ? denom : 0));
}
Value *mod_Int32(Value *x, Value *y) {
int32_t num = ((Value_Int32 *)x)->i32;
int32_t denom = ((Value_Int32 *)y)->i32;
Value *idris2_mod_Int32(Value *x, Value *y) {
int32_t num = idris2_vp_to_Int32(x);
int32_t denom = idris2_vp_to_Int32(y);
denom = (denom < 0) ? -denom : denom;
return (Value *)makeInt32(num % denom + (num < 0 ? denom : 0));
return (Value *)idris2_mkInt32(num % denom + (num < 0 ? denom : 0));
}
Value *mod_Int64(Value *x, Value *y) {
int64_t num = ((Value_Int64 *)x)->i64;
int64_t denom = ((Value_Int64 *)y)->i64;
Value *idris2_mod_Int64(Value *x, Value *y) {
int64_t num = idris2_vp_to_Int64(x);
int64_t denom = idris2_vp_to_Int64(y);
denom = (denom < 0) ? -denom : denom;
return (Value *)makeInt64(num % denom + (num < 0 ? denom : 0));
return (Value *)idris2_mkInt64(num % denom + (num < 0 ? denom : 0));
}
Value *mod_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_mod_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mpz_mod(retVal->i, ((Value_Integer *)x)->i, ((Value_Integer *)y)->i);
return (Value *)retVal;
}
/* shiftl */
Value *shiftl_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 << ((Value_Bits8 *)y)->ui8);
}
Value *shiftl_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16
<< ((Value_Bits16 *)y)->ui16);
}
Value *shiftl_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32
<< ((Value_Bits32 *)y)->ui32);
}
Value *shiftl_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64
<< ((Value_Bits64 *)y)->ui64);
}
Value *shiftl_Int8(Value *x, Value *y) {
return (Value *)makeInt8(((Value_Int8 *)x)->i8 << ((Value_Int8 *)y)->i8);
}
Value *shiftl_Int16(Value *x, Value *y) {
return (Value *)makeInt16(((Value_Int16 *)x)->i16 << ((Value_Int16 *)y)->i16);
}
Value *shiftl_Int32(Value *x, Value *y) {
return (Value *)makeInt32(((Value_Int32 *)x)->i32 << ((Value_Int32 *)y)->i32);
}
Value *shiftl_Int64(Value *x, Value *y) {
return (Value *)makeInt64(((Value_Int64 *)x)->i64 << ((Value_Int64 *)y)->i64);
}
Value *shiftl_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_shiftl_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mp_bitcnt_t cnt = (mp_bitcnt_t)mpz_get_ui(((Value_Integer *)y)->i);
mpz_mul_2exp(retVal->i, ((Value_Integer *)x)->i, cnt);
return (Value *)retVal;
}
/* shiftr */
Value *shiftr_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 >> ((Value_Bits8 *)y)->ui8);
}
Value *shiftr_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 >>
((Value_Bits16 *)y)->ui16);
}
Value *shiftr_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 >>
((Value_Bits32 *)y)->ui32);
}
Value *shiftr_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 >>
((Value_Bits64 *)y)->ui64);
}
Value *shiftr_Int8(Value *x, Value *y) {
return (Value *)makeInt8(((Value_Int8 *)x)->i8 >> ((Value_Int8 *)y)->i8);
}
Value *shiftr_Int16(Value *x, Value *y) {
return (Value *)makeInt16(((Value_Int16 *)x)->i16 >> ((Value_Int16 *)y)->i16);
}
Value *shiftr_Int32(Value *x, Value *y) {
return (Value *)makeInt32(((Value_Int32 *)x)->i32 >> ((Value_Int32 *)y)->i32);
}
Value *shiftr_Int64(Value *x, Value *y) {
return (Value *)makeInt64(((Value_Int64 *)x)->i64 >> ((Value_Int64 *)y)->i64);
}
Value *shiftr_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_shiftr_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mp_bitcnt_t cnt = (mp_bitcnt_t)mpz_get_ui(((Value_Integer *)y)->i);
mpz_fdiv_q_2exp(retVal->i, ((Value_Integer *)x)->i, cnt);
return (Value *)retVal;
}
/* and */
Value *and_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 & ((Value_Bits8 *)y)->ui8);
}
Value *and_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 &
((Value_Bits16 *)y)->ui16);
}
Value *and_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 &
((Value_Bits32 *)y)->ui32);
}
Value *and_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 &
((Value_Bits64 *)y)->ui64);
}
Value *and_Int8(Value *x, Value *y) {
return (Value *)makeInt8(((Value_Int8 *)x)->i8 & ((Value_Int8 *)y)->i8);
}
Value *and_Int16(Value *x, Value *y) {
return (Value *)makeInt16(((Value_Int16 *)x)->i16 & ((Value_Int16 *)y)->i16);
}
Value *and_Int32(Value *x, Value *y) {
return (Value *)makeInt32(((Value_Int32 *)x)->i32 & ((Value_Int32 *)y)->i32);
}
Value *and_Int64(Value *x, Value *y) {
return (Value *)makeInt64(((Value_Int64 *)x)->i64 & ((Value_Int64 *)y)->i64);
}
Value *and_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_and_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mpz_and(retVal->i, ((Value_Integer *)x)->i, ((Value_Integer *)y)->i);
return (Value *)retVal;
}
/* or */
Value *or_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 | ((Value_Bits8 *)y)->ui8);
}
Value *or_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 |
((Value_Bits16 *)y)->ui16);
}
Value *or_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 |
((Value_Bits32 *)y)->ui32);
}
Value *or_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 |
((Value_Bits64 *)y)->ui64);
}
Value *or_Int8(Value *x, Value *y) {
return (Value *)makeInt8(((Value_Int8 *)x)->i8 | ((Value_Int8 *)y)->i8);
}
Value *or_Int16(Value *x, Value *y) {
return (Value *)makeInt16(((Value_Int16 *)x)->i16 | ((Value_Int16 *)y)->i16);
}
Value *or_Int32(Value *x, Value *y) {
return (Value *)makeInt32(((Value_Int32 *)x)->i32 | ((Value_Int32 *)y)->i32);
}
Value *or_Int64(Value *x, Value *y) {
return (Value *)makeInt64(((Value_Int64 *)x)->i64 | ((Value_Int64 *)y)->i64);
}
Value *or_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_or_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mpz_ior(retVal->i, ((Value_Integer *)x)->i, ((Value_Integer *)y)->i);
return (Value *)retVal;
}
/* xor */
Value *xor_Bits8(Value *x, Value *y) {
return (Value *)makeBits8(((Value_Bits8 *)x)->ui8 ^ ((Value_Bits8 *)y)->ui8);
}
Value *xor_Bits16(Value *x, Value *y) {
return (Value *)makeBits16(((Value_Bits16 *)x)->ui16 ^
((Value_Bits16 *)y)->ui16);
}
Value *xor_Bits32(Value *x, Value *y) {
return (Value *)makeBits32(((Value_Bits32 *)x)->ui32 ^
((Value_Bits32 *)y)->ui32);
}
Value *xor_Bits64(Value *x, Value *y) {
return (Value *)makeBits64(((Value_Bits64 *)x)->ui64 ^
((Value_Bits64 *)y)->ui64);
}
Value *xor_Int8(Value *x, Value *y) {
return (Value *)makeInt8(((Value_Int8 *)x)->i8 ^ ((Value_Int8 *)y)->i8);
}
Value *xor_Int16(Value *x, Value *y) {
return (Value *)makeInt16(((Value_Int16 *)x)->i16 ^ ((Value_Int16 *)y)->i16);
}
Value *xor_Int32(Value *x, Value *y) {
return (Value *)makeInt32(((Value_Int32 *)x)->i32 ^ ((Value_Int32 *)y)->i32);
}
Value *xor_Int64(Value *x, Value *y) {
return (Value *)makeInt64(((Value_Int64 *)x)->i64 ^ ((Value_Int64 *)y)->i64);
}
Value *xor_Integer(Value *x, Value *y) {
Value_Integer *retVal = makeInteger();
Value *idris2_xor_Integer(Value *x, Value *y) {
Value_Integer *retVal = idris2_mkInteger();
mpz_xor(retVal->i, ((Value_Integer *)x)->i, ((Value_Integer *)y)->i);
return (Value *)retVal;
}
/* lt */
Value *lt_Bits8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits8 *)x)->ui8 < ((Value_Bits8 *)y)->ui8);
}
Value *lt_Bits16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits16 *)x)->ui16 <
((Value_Bits16 *)y)->ui16);
}
Value *lt_Bits32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits32 *)x)->ui32 <
((Value_Bits32 *)y)->ui32);
}
Value *lt_Bits64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits64 *)x)->ui64 <
((Value_Bits64 *)y)->ui64);
}
Value *lt_Int8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int8 *)x)->i8 < ((Value_Int8 *)y)->i8);
}
Value *lt_Int16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int16 *)x)->i16 < ((Value_Int16 *)y)->i16);
}
Value *lt_Int32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int32 *)x)->i32 < ((Value_Int32 *)y)->i32);
}
Value *lt_Int64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int64 *)x)->i64 < ((Value_Int64 *)y)->i64);
}
Value *lt_Integer(Value *x, Value *y) {
return (Value *)makeBool(
mpz_cmp(((Value_Integer *)x)->i, ((Value_Integer *)y)->i) < 0);
}
Value *lt_double(Value *x, Value *y) {
return (Value *)makeBool(((Value_Double *)x)->d < ((Value_Double *)y)->d);
}
Value *lt_char(Value *x, Value *y) {
return (Value *)makeBool(((Value_Char *)x)->c < ((Value_Char *)y)->c);
}
Value *lt_string(Value *x, Value *y) {
return (Value *)makeBool(
strcmp(((Value_String *)x)->str, ((Value_String *)y)->str) < 0);
}
/* gt */
Value *gt_Bits8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits8 *)x)->ui8 > ((Value_Bits8 *)y)->ui8);
}
Value *gt_Bits16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits16 *)x)->ui16 >
((Value_Bits16 *)y)->ui16);
}
Value *gt_Bits32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits32 *)x)->ui32 >
((Value_Bits32 *)y)->ui32);
}
Value *gt_Bits64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits64 *)x)->ui64 >
((Value_Bits64 *)y)->ui64);
}
Value *gt_Int8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int8 *)x)->i8 > ((Value_Int8 *)y)->i8);
}
Value *gt_Int16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int16 *)x)->i16 > ((Value_Int16 *)y)->i16);
}
Value *gt_Int32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int32 *)x)->i32 > ((Value_Int32 *)y)->i32);
}
Value *gt_Int64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int64 *)x)->i64 > ((Value_Int64 *)y)->i64);
}
Value *gt_Integer(Value *x, Value *y) {
return (Value *)makeBool(
mpz_cmp(((Value_Integer *)x)->i, ((Value_Integer *)y)->i) > 0);
}
Value *gt_double(Value *x, Value *y) {
return (Value *)makeBool(((Value_Double *)x)->d > ((Value_Double *)y)->d);
}
Value *gt_char(Value *x, Value *y) {
return (Value *)makeBool(((Value_Char *)x)->c > ((Value_Char *)y)->c);
}
Value *gt_string(Value *x, Value *y) {
return (Value *)makeBool(
strcmp(((Value_String *)x)->str, ((Value_String *)y)->str) > 0);
}
/* eq */
Value *eq_Bits8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits8 *)x)->ui8 == ((Value_Bits8 *)y)->ui8);
}
Value *eq_Bits16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits16 *)x)->ui16 ==
((Value_Bits16 *)y)->ui16);
}
Value *eq_Bits32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits32 *)x)->ui32 ==
((Value_Bits32 *)y)->ui32);
}
Value *eq_Bits64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits64 *)x)->ui64 ==
((Value_Bits64 *)y)->ui64);
}
Value *eq_Int8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int8 *)x)->i8 == ((Value_Int8 *)y)->i8);
}
Value *eq_Int16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int16 *)x)->i16 == ((Value_Int16 *)y)->i16);
}
Value *eq_Int32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int32 *)x)->i32 == ((Value_Int32 *)y)->i32);
}
Value *eq_Int64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int64 *)x)->i64 == ((Value_Int64 *)y)->i64);
}
Value *eq_Integer(Value *x, Value *y) {
return (Value *)makeBool(
mpz_cmp(((Value_Integer *)x)->i, ((Value_Integer *)y)->i) == 0);
}
Value *eq_double(Value *x, Value *y) {
return (Value *)makeBool(((Value_Double *)x)->d == ((Value_Double *)y)->d);
}
Value *eq_char(Value *x, Value *y) {
return (Value *)makeBool(((Value_Char *)x)->c == ((Value_Char *)y)->c);
}
Value *eq_string(Value *x, Value *y) {
return (Value *)makeBool(
!strcmp(((Value_String *)x)->str, ((Value_String *)y)->str));
}
/* lte */
Value *lte_Bits8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits8 *)x)->ui8 <= ((Value_Bits8 *)y)->ui8);
}
Value *lte_Bits16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits16 *)x)->ui16 <=
((Value_Bits16 *)y)->ui16);
}
Value *lte_Bits32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits32 *)x)->ui32 <=
((Value_Bits32 *)y)->ui32);
}
Value *lte_Bits64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits64 *)x)->ui64 <=
((Value_Bits64 *)y)->ui64);
}
Value *lte_Int8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int8 *)x)->i8 <= ((Value_Int8 *)y)->i8);
}
Value *lte_Int16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int16 *)x)->i16 <= ((Value_Int16 *)y)->i16);
}
Value *lte_Int32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int32 *)x)->i32 <= ((Value_Int32 *)y)->i32);
}
Value *lte_Int64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int64 *)x)->i64 <= ((Value_Int64 *)y)->i64);
}
Value *lte_Integer(Value *x, Value *y) {
return (Value *)makeBool(
mpz_cmp(((Value_Integer *)x)->i, ((Value_Integer *)y)->i) <= 0);
}
Value *lte_double(Value *x, Value *y) {
return (Value *)makeBool(((Value_Double *)x)->d <= ((Value_Double *)y)->d);
}
Value *lte_char(Value *x, Value *y) {
return (Value *)makeBool(((Value_Char *)x)->c <= ((Value_Char *)y)->c);
}
Value *lte_string(Value *x, Value *y) {
return (Value *)makeBool(
strcmp(((Value_String *)x)->str, ((Value_String *)y)->str) <= 0);
}
/* gte */
Value *gte_Bits8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits8 *)x)->ui8 >= ((Value_Bits8 *)y)->ui8);
}
Value *gte_Bits16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits16 *)x)->ui16 >=
((Value_Bits16 *)y)->ui16);
}
Value *gte_Bits32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits32 *)x)->ui32 >=
((Value_Bits32 *)y)->ui32);
}
Value *gte_Bits64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Bits64 *)x)->ui64 >=
((Value_Bits64 *)y)->ui64);
}
Value *gte_Int8(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int8 *)x)->i8 >= ((Value_Int8 *)y)->i8);
}
Value *gte_Int16(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int16 *)x)->i16 >= ((Value_Int16 *)y)->i16);
}
Value *gte_Int32(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int32 *)x)->i32 >= ((Value_Int32 *)y)->i32);
}
Value *gte_Int64(Value *x, Value *y) {
return (Value *)makeBool(((Value_Int64 *)x)->i64 >= ((Value_Int64 *)y)->i64);
}
Value *gte_Integer(Value *x, Value *y) {
return (Value *)makeBool(
mpz_cmp(((Value_Integer *)x)->i, ((Value_Integer *)y)->i) >= 0);
}
Value *gte_double(Value *x, Value *y) {
return (Value *)makeBool(((Value_Double *)x)->d >= ((Value_Double *)y)->d);
}
Value *gte_char(Value *x, Value *y) {
return (Value *)makeBool(((Value_Char *)x)->c >= ((Value_Char *)y)->c);
}
Value *gte_string(Value *x, Value *y) {
return (Value *)makeBool(
strcmp(((Value_String *)x)->str, ((Value_String *)y)->str) >= 0);
}

View File

@ -4,201 +4,220 @@
#include <gmp.h>
#include <math.h>
double unpackDouble(Value *d);
Value *believe_me(Value *, Value *, Value *);
#define idris2_binop(ty, op, l, r) \
((Value *)idris2_mk##ty(idris2_vp_to_##ty(l) op idris2_vp_to_##ty(r)))
/* add */
Value *add_Bits8(Value *x, Value *y);
Value *add_Bits16(Value *x, Value *y);
Value *add_Bits32(Value *x, Value *y);
Value *add_Bits64(Value *x, Value *y);
Value *add_Int8(Value *x, Value *y);
Value *add_Int16(Value *x, Value *y);
Value *add_Int32(Value *x, Value *y);
Value *add_Int64(Value *x, Value *y);
Value *add_Integer(Value *x, Value *y);
Value *add_double(Value *x, Value *y);
#define idris2_add_Bits8(l, r) (idris2_binop(Bits8, +, l, r))
#define idris2_add_Bits16(l, r) (idris2_binop(Bits16, +, l, r))
#define idris2_add_Bits32(l, r) (idris2_binop(Bits32, +, l, r))
#define idris2_add_Bits64(l, r) (idris2_binop(Bits64, +, l, r))
#define idris2_add_Int8(l, r) (idris2_binop(Int8, +, l, r))
#define idris2_add_Int16(l, r) (idris2_binop(Int16, +, l, r))
#define idris2_add_Int32(l, r) (idris2_binop(Int32, +, l, r))
#define idris2_add_Int64(l, r) (idris2_binop(Int64, +, l, r))
Value *idris2_add_Integer(Value *x, Value *y);
#define idris2_add_Double(l, r) (idris2_binop(Double, +, l, r))
/* sub */
Value *sub_Bits8(Value *x, Value *y);
Value *sub_Bits16(Value *x, Value *y);
Value *sub_Bits32(Value *x, Value *y);
Value *sub_Bits64(Value *x, Value *y);
Value *sub_Int8(Value *x, Value *y);
Value *sub_Int16(Value *x, Value *y);
Value *sub_Int32(Value *x, Value *y);
Value *sub_Int64(Value *x, Value *y);
Value *sub_Integer(Value *x, Value *y);
Value *sub_double(Value *x, Value *y);
#define idris2_sub_Bits8(l, r) (idris2_binop(Bits8, -, l, r))
#define idris2_sub_Bits16(l, r) (idris2_binop(Bits16, -, l, r))
#define idris2_sub_Bits32(l, r) (idris2_binop(Bits32, -, l, r))
#define idris2_sub_Bits64(l, r) (idris2_binop(Bits64, -, l, r))
#define idris2_sub_Int8(l, r) (idris2_binop(Int8, -, l, r))
#define idris2_sub_Int16(l, r) (idris2_binop(Int16, -, l, r))
#define idris2_sub_Int32(l, r) (idris2_binop(Int32, -, l, r))
#define idris2_sub_Int64(l, r) (idris2_binop(Int64, -, l, r))
Value *idris2_sub_Integer(Value *x, Value *y);
#define idris2_sub_Double(l, r) (idris2_binop(Double, -, l, r))
/* negate */
Value *negate_Bits8(Value *x);
Value *negate_Bits16(Value *x);
Value *negate_Bits32(Value *x);
Value *negate_Bits64(Value *x);
Value *negate_Int8(Value *x);
Value *negate_Int16(Value *x);
Value *negate_Int32(Value *x);
Value *negate_Int64(Value *x);
Value *negate_Integer(Value *x);
Value *negate_double(Value *x);
#define idris2_nagate_Int8(x) (idris2_mkInt8(-(idris2_vp_to_Int8(x))))
#define idris2_nagate_Int16(x) (idris2_mkInt16(-(idris2_vp_to_Int16(x))))
#define idris2_nagate_Int32(x) (idris2_mkInt32(-(idris2_vp_to_Int32(x))))
#define idris2_nagate_Int64(x) (idris2_mkInt64(-(idris2_vp_to_Int64(x))))
Value *idris2_negate_Integer(Value *x);
#define idris2_nagate_Double(x) (idris2_mkDouble(-(idris2_vp_to_Double(x))))
/* mul */
Value *mul_Bits8(Value *x, Value *y);
Value *mul_Bits16(Value *x, Value *y);
Value *mul_Bits32(Value *x, Value *y);
Value *mul_Bits64(Value *x, Value *y);
Value *mul_Int8(Value *x, Value *y);
Value *mul_Int16(Value *x, Value *y);
Value *mul_Int32(Value *x, Value *y);
Value *mul_Int64(Value *x, Value *y);
Value *mul_Integer(Value *x, Value *y);
Value *mul_double(Value *x, Value *y);
#define idris2_mul_Bits8(l, r) (idris2_binop(Bits8, *, l, r))
#define idris2_mul_Bits16(l, r) (idris2_binop(Bits16, *, l, r))
#define idris2_mul_Bits32(l, r) (idris2_binop(Bits32, *, l, r))
#define idris2_mul_Bits64(l, r) (idris2_binop(Bits64, *, l, r))
#define idris2_mul_Int8(l, r) (idris2_binop(Int8, *, l, r))
#define idris2_mul_Int16(l, r) (idris2_binop(Int16, *, l, r))
#define idris2_mul_Int32(l, r) (idris2_binop(Int32, *, l, r))
#define idris2_mul_Int64(l, r) (idris2_binop(Int64, *, l, r))
Value *idris2_mul_Integer(Value *x, Value *y);
#define idris2_mul_Double(l, r) (idris2_binop(Double, *, l, r))
/* div */
Value *div_Bits8(Value *x, Value *y);
Value *div_Bits16(Value *x, Value *y);
Value *div_Bits32(Value *x, Value *y);
Value *div_Bits64(Value *x, Value *y);
Value *div_Int8(Value *x, Value *y);
Value *div_Int16(Value *x, Value *y);
Value *div_Int32(Value *x, Value *y);
Value *div_Int64(Value *x, Value *y);
Value *div_Integer(Value *x, Value *y);
Value *div_double(Value *x, Value *y);
#define idris2_div_Bits8(l, r) (idris2_binop(Bits8, /, l, r))
#define idris2_div_Bits16(l, r) (idris2_binop(Bits16, /, l, r))
#define idris2_div_Bits32(l, r) (idris2_binop(Bits32, /, l, r))
#define idris2_div_Bits64(l, r) (idris2_binop(Bits64, /, l, r))
Value *idris2_div_Int8(Value *x, Value *y);
Value *idris2_div_Int16(Value *x, Value *y);
Value *idris2_div_Int32(Value *x, Value *y);
Value *idris2_div_Int64(Value *x, Value *y);
Value *idris2_div_Integer(Value *x, Value *y);
#define idris2_div_Double(l, r) (idris2_binop(Double, /, l, r))
/* mod */
Value *mod_Bits8(Value *x, Value *y);
Value *mod_Bits16(Value *x, Value *y);
Value *mod_Bits32(Value *x, Value *y);
Value *mod_Bits64(Value *x, Value *y);
Value *mod_Int8(Value *x, Value *y);
Value *mod_Int16(Value *x, Value *y);
Value *mod_Int32(Value *x, Value *y);
Value *mod_Int64(Value *x, Value *y);
Value *mod_Integer(Value *x, Value *y);
#define idris2_mod_Bits8(l, r) (idris2_binop(Bits8, %, l, r))
#define idris2_mod_Bits16(l, r) (idris2_binop(Bits16, %, l, r))
#define idris2_mod_Bits32(l, r) (idris2_binop(Bits32, %, l, r))
#define idris2_mod_Bits64(l, r) (idris2_binop(Bits64, %, l, r))
Value *idris2_mod_Int8(Value *x, Value *y);
Value *idris2_mod_Int16(Value *x, Value *y);
Value *idris2_mod_Int32(Value *x, Value *y);
Value *idris2_mod_Int64(Value *x, Value *y);
Value *idris2_mod_Integer(Value *x, Value *y);
/* shiftl */
Value *shiftl_Bits8(Value *x, Value *y);
Value *shiftl_Bits16(Value *x, Value *y);
Value *shiftl_Bits32(Value *x, Value *y);
Value *shiftl_Bits64(Value *x, Value *y);
Value *shiftl_Int8(Value *x, Value *y);
Value *shiftl_Int16(Value *x, Value *y);
Value *shiftl_Int32(Value *x, Value *y);
Value *shiftl_Int64(Value *x, Value *y);
Value *shiftl_Integer(Value *x, Value *y);
#define idris2_shiftl_Bits8(l, r) (idris2_binop(Bits8, <<, l, r))
#define idris2_shiftl_Bits16(l, r) (idris2_binop(Bits16, <<, l, r))
#define idris2_shiftl_Bits32(l, r) (idris2_binop(Bits32, <<, l, r))
#define idris2_shiftl_Bits64(l, r) (idris2_binop(Bits64, <<, l, r))
#define idris2_shiftl_Int8(l, r) (idris2_binop(Int8, <<, l, r))
#define idris2_shiftl_Int16(l, r) (idris2_binop(Int16, <<, l, r))
#define idris2_shiftl_Int32(l, r) (idris2_binop(Int32, <<, l, r))
#define idris2_shiftl_Int64(l, r) (idris2_binop(Int64, <<, l, r))
Value *idris2_shiftl_Integer(Value *x, Value *y);
/* shiftr */
Value *shiftr_Bits8(Value *x, Value *y);
Value *shiftr_Bits16(Value *x, Value *y);
Value *shiftr_Bits32(Value *x, Value *y);
Value *shiftr_Bits64(Value *x, Value *y);
Value *shiftr_Int8(Value *x, Value *y);
Value *shiftr_Int16(Value *x, Value *y);
Value *shiftr_Int32(Value *x, Value *y);
Value *shiftr_Int64(Value *x, Value *y);
Value *shiftr_Integer(Value *x, Value *y);
#define idris2_shiftr_Bits8(l, r) (idris2_binop(Bits8, >>, l, r))
#define idris2_shiftr_Bits16(l, r) (idris2_binop(Bits16, >>, l, r))
#define idris2_shiftr_Bits32(l, r) (idris2_binop(Bits32, >>, l, r))
#define idris2_shiftr_Bits64(l, r) (idris2_binop(Bits64, >>, l, r))
#define idris2_shiftr_Int8(l, r) (idris2_binop(Int8, >>, l, r))
#define idris2_shiftr_Int16(l, r) (idris2_binop(Int16, >>, l, r))
#define idris2_shiftr_Int32(l, r) (idris2_binop(Int32, >>, l, r))
#define idris2_shiftr_Int64(l, r) (idris2_binop(Int64, >>, l, r))
Value *idris2_shiftr_Integer(Value *x, Value *y);
/* and */
Value *and_Bits8(Value *x, Value *y);
Value *and_Bits16(Value *x, Value *y);
Value *and_Bits32(Value *x, Value *y);
Value *and_Bits64(Value *x, Value *y);
Value *and_Int8(Value *x, Value *y);
Value *and_Int16(Value *x, Value *y);
Value *and_Int32(Value *x, Value *y);
Value *and_Int64(Value *x, Value *y);
Value *and_Integer(Value *x, Value *y);
#define idris2_and_Bits8(l, r) (idris2_binop(Bits8, &, l, r))
#define idris2_and_Bits16(l, r) (idris2_binop(Bits16, &, l, r))
#define idris2_and_Bits32(l, r) (idris2_binop(Bits32, &, l, r))
#define idris2_and_Bits64(l, r) (idris2_binop(Bits64, &, l, r))
#define idris2_and_Int8(l, r) (idris2_binop(Int8, &, l, r))
#define idris2_and_Int16(l, r) (idris2_binop(Int16, &, l, r))
#define idris2_and_Int32(l, r) (idris2_binop(Int32, &, l, r))
#define idris2_and_Int64(l, r) (idris2_binop(Int64, &, l, r))
Value *idris2_and_Integer(Value *x, Value *y);
/* or */
Value *or_Bits8(Value *x, Value *y);
Value *or_Bits16(Value *x, Value *y);
Value *or_Bits32(Value *x, Value *y);
Value *or_Bits64(Value *x, Value *y);
Value *or_Int8(Value *x, Value *y);
Value *or_Int16(Value *x, Value *y);
Value *or_Int32(Value *x, Value *y);
Value *or_Int64(Value *x, Value *y);
Value *or_Integer(Value *x, Value *y);
#define idris2_or_Bits8(l, r) (idris2_binop(Bits8, |, l, r))
#define idris2_or_Bits16(l, r) (idris2_binop(Bits16, |, l, r))
#define idris2_or_Bits32(l, r) (idris2_binop(Bits32, |, l, r))
#define idris2_or_Bits64(l, r) (idris2_binop(Bits64, |, l, r))
#define idris2_or_Int8(l, r) (idris2_binop(Int8, |, l, r))
#define idris2_or_Int16(l, r) (idris2_binop(Int16, |, l, r))
#define idris2_or_Int32(l, r) (idris2_binop(Int32, |, l, r))
#define idris2_or_Int64(l, r) (idris2_binop(Int64, |, l, r))
Value *idris2_or_Integer(Value *x, Value *y);
/* xor */
Value *xor_Bits8(Value *x, Value *y);
Value *xor_Bits16(Value *x, Value *y);
Value *xor_Bits32(Value *x, Value *y);
Value *xor_Bits64(Value *x, Value *y);
Value *xor_Int8(Value *x, Value *y);
Value *xor_Int16(Value *x, Value *y);
Value *xor_Int32(Value *x, Value *y);
Value *xor_Int64(Value *x, Value *y);
Value *xor_Integer(Value *x, Value *y);
#define idris2_xor_Bits8(l, r) (idris2_binop(Bits8, ^, l, r))
#define idris2_xor_Bits16(l, r) (idris2_binop(Bits16, ^, l, r))
#define idris2_xor_Bits32(l, r) (idris2_binop(Bits32, ^, l, r))
#define idris2_xor_Bits64(l, r) (idris2_binop(Bits64, ^, l, r))
#define idris2_xor_Int8(l, r) (idris2_binop(Int8, ^, l, r))
#define idris2_xor_Int16(l, r) (idris2_binop(Int16, ^, l, r))
#define idris2_xor_Int32(l, r) (idris2_binop(Int32, ^, l, r))
#define idris2_xor_Int64(l, r) (idris2_binop(Int64, ^, l, r))
Value *idris2_xor_Integer(Value *x, Value *y);
#define idris2_cmpop(ty, op, l, r) \
(idris2_mkBool((idris2_vp_to_##ty(l) op idris2_vp_to_##ty(r)) ? 1 : 0))
/* lt */
Value *lt_Bits8(Value *x, Value *y);
Value *lt_Bits16(Value *x, Value *y);
Value *lt_Bits32(Value *x, Value *y);
Value *lt_Bits64(Value *x, Value *y);
Value *lt_Int8(Value *x, Value *y);
Value *lt_Int16(Value *x, Value *y);
Value *lt_Int32(Value *x, Value *y);
Value *lt_Int64(Value *x, Value *y);
Value *lt_Integer(Value *x, Value *y);
Value *lt_double(Value *x, Value *y);
Value *lt_char(Value *x, Value *y);
Value *lt_string(Value *x, Value *y);
#define idris2_lt_Bits8(l, r) (idris2_cmpop(Bits8, <, l, r))
#define idris2_lt_Bits16(l, r) (idris2_cmpop(Bits16, <, l, r))
#define idris2_lt_Bits32(l, r) (idris2_cmpop(Bits32, <, l, r))
#define idris2_lt_Bits64(l, r) (idris2_cmpop(Bits64, <, l, r))
#define idris2_lt_Int8(l, r) (idris2_cmpop(Int8, <, l, r))
#define idris2_lt_Int16(l, r) (idris2_cmpop(Int16, <, l, r))
#define idris2_lt_Int32(l, r) (idris2_cmpop(Int32, <, l, r))
#define idris2_lt_Int64(l, r) (idris2_cmpop(Int64, <, l, r))
#define idris2_lt_Integer(l, r) \
(idris2_mkBool( \
mpz_cmp(((Value_Integer *)(l))->i, ((Value_Integer *)(r))->i) < 0))
#define idris2_lt_Double(l, r) (idris2_cmpop(Double, <, l, r))
#define idris2_lt_Char(l, r) (idris2_cmpop(Char, <, l, r))
#define idris2_lt_string(l, r) \
(idris2_mkBool( \
strcmp(((Value_String *)(l))->str, ((Value_String *)(r))->str) < 0))
/* gt */
Value *gt_Bits8(Value *x, Value *y);
Value *gt_Bits16(Value *x, Value *y);
Value *gt_Bits32(Value *x, Value *y);
Value *gt_Bits64(Value *x, Value *y);
Value *gt_Int8(Value *x, Value *y);
Value *gt_Int16(Value *x, Value *y);
Value *gt_Int32(Value *x, Value *y);
Value *gt_Int64(Value *x, Value *y);
Value *gt_Integer(Value *x, Value *y);
Value *gt_double(Value *x, Value *y);
Value *gt_char(Value *x, Value *y);
Value *gt_string(Value *x, Value *y);
#define idris2_gt_Bits8(l, r) (idris2_cmpop(Bits8, >, l, r))
#define idris2_gt_Bits16(l, r) (idris2_cmpop(Bits16, >, l, r))
#define idris2_gt_Bits32(l, r) (idris2_cmpop(Bits32, >, l, r))
#define idris2_gt_Bits64(l, r) (idris2_cmpop(Bits64, >, l, r))
#define idris2_gt_Int8(l, r) (idris2_cmpop(Int8, >, l, r))
#define idris2_gt_Int16(l, r) (idris2_cmpop(Int16, >, l, r))
#define idris2_gt_Int32(l, r) (idris2_cmpop(Int32, >, l, r))
#define idris2_gt_Int64(l, r) (idris2_cmpop(Int64, >, l, r))
#define idris2_gt_Integer(l, r) \
(idris2_mkBool( \
mpz_cmp(((Value_Integer *)(l))->i, ((Value_Integer *)(r))->i) > 0))
#define idris2_gt_Double(l, r) (idris2_cmpop(Double, >, l, r))
#define idris2_gt_Char(l, r) (idris2_cmpop(Char, >, l, r))
#define idris2_gt_string(l, r) \
(idris2_mkBool( \
strcmp(((Value_String *)(l))->str, ((Value_String *)(r))->str) > 0))
/* eq */
Value *eq_Bits8(Value *x, Value *y);
Value *eq_Bits16(Value *x, Value *y);
Value *eq_Bits32(Value *x, Value *y);
Value *eq_Bits64(Value *x, Value *y);
Value *eq_Int8(Value *x, Value *y);
Value *eq_Int16(Value *x, Value *y);
Value *eq_Int32(Value *x, Value *y);
Value *eq_Int64(Value *x, Value *y);
Value *eq_Integer(Value *x, Value *y);
Value *eq_double(Value *x, Value *y);
Value *eq_char(Value *x, Value *y);
Value *eq_string(Value *x, Value *y);
#define idris2_eq_Bits8(l, r) (idris2_cmpop(Bits8, ==, l, r))
#define idris2_eq_Bits16(l, r) (idris2_cmpop(Bits16, ==, l, r))
#define idris2_eq_Bits32(l, r) (idris2_cmpop(Bits32, ==, l, r))
#define idris2_eq_Bits64(l, r) (idris2_cmpop(Bits64, ==, l, r))
#define idris2_eq_Int8(l, r) (idris2_cmpop(Int8, ==, l, r))
#define idris2_eq_Int16(l, r) (idris2_cmpop(Int16, ==, l, r))
#define idris2_eq_Int32(l, r) (idris2_cmpop(Int32, ==, l, r))
#define idris2_eq_Int64(l, r) (idris2_cmpop(Int64, ==, l, r))
#define idris2_eq_Integer(l, r) \
(idris2_mkBool( \
mpz_cmp(((Value_Integer *)(l))->i, ((Value_Integer *)(r))->i) == 0))
#define idris2_eq_Double(l, r) (idris2_cmpop(Double, ==, l, r))
#define idris2_eq_Char(l, r) (idris2_cmpop(Char, ==, l, r))
#define idris2_eq_string(l, r) \
(idris2_mkBool( \
strcmp(((Value_String *)(l))->str, ((Value_String *)(r))->str) == 0))
/* lte */
Value *lte_Bits8(Value *x, Value *y);
Value *lte_Bits16(Value *x, Value *y);
Value *lte_Bits32(Value *x, Value *y);
Value *lte_Bits64(Value *x, Value *y);
Value *lte_Int8(Value *x, Value *y);
Value *lte_Int16(Value *x, Value *y);
Value *lte_Int32(Value *x, Value *y);
Value *lte_Int64(Value *x, Value *y);
Value *lte_Integer(Value *x, Value *y);
Value *lte_double(Value *x, Value *y);
Value *lte_char(Value *x, Value *y);
Value *lte_string(Value *x, Value *y);
#define idris2_lte_Bits8(l, r) (idris2_cmpop(Bits8, <=, l, r))
#define idris2_lte_Bits16(l, r) (idris2_cmpop(Bits16, <=, l, r))
#define idris2_lte_Bits32(l, r) (idris2_cmpop(Bits32, <=, l, r))
#define idris2_lte_Bits64(l, r) (idris2_cmpop(Bits64, <=, l, r))
#define idris2_lte_Int8(l, r) (idris2_cmpop(Int8, <=, l, r))
#define idris2_lte_Int16(l, r) (idris2_cmpop(Int16, <=, l, r))
#define idris2_lte_Int32(l, r) (idris2_cmpop(Int32, <=, l, r))
#define idris2_lte_Int64(l, r) (idris2_cmpop(Int64, <=, l, r))
#define idris2_lte_Integer(l, r) \
(idris2_mkBool( \
mpz_cmp(((Value_Integer *)(l))->i, ((Value_Integer *)(r))->i) <= 0))
#define idris2_lte_Double(l, r) (idris2_cmpop(Double, <=, l, r))
#define idris2_lte_Char(l, r) (idris2_cmpop(Char, <=, l, r))
#define idris2_lte_string(l, r) \
(idris2_mkBool( \
strcmp(((Value_String *)(l))->str, ((Value_String *)(r))->str) <= 0))
/* gte */
Value *gte_Bits8(Value *x, Value *y);
Value *gte_Bits16(Value *x, Value *y);
Value *gte_Bits32(Value *x, Value *y);
Value *gte_Bits64(Value *x, Value *y);
Value *gte_Int8(Value *x, Value *y);
Value *gte_Int16(Value *x, Value *y);
Value *gte_Int32(Value *x, Value *y);
Value *gte_Int64(Value *x, Value *y);
Value *gte_Integer(Value *x, Value *y);
Value *gte_double(Value *x, Value *y);
Value *gte_char(Value *x, Value *y);
Value *gte_string(Value *x, Value *y);
#define idris2_gte_Bits8(l, r) (idris2_cmpop(Bits8, >=, l, r))
#define idris2_gte_Bits16(l, r) (idris2_cmpop(Bits16, >=, l, r))
#define idris2_gte_Bits32(l, r) (idris2_cmpop(Bits32, >=, l, r))
#define idris2_gte_Bits64(l, r) (idris2_cmpop(Bits64, >=, l, r))
#define idris2_gte_Int8(l, r) (idris2_cmpop(Int8, >=, l, r))
#define idris2_gte_Int16(l, r) (idris2_cmpop(Int16, >=, l, r))
#define idris2_gte_Int32(l, r) (idris2_cmpop(Int32, >=, l, r))
#define idris2_gte_Int64(l, r) (idris2_cmpop(Int64, >=, l, r))
#define idris2_gte_Integer(l, r) \
(idris2_mkBool( \
mpz_cmp(((Value_Integer *)(l))->i, ((Value_Integer *)(r))->i) >= 0))
#define idris2_gte_Double(l, r) (idris2_cmpop(Double, >=, l, r))
#define idris2_gte_Char(l, r) (idris2_cmpop(Char, >=, l, r))
#define idris2_gte_string(l, r) \
(idris2_mkBool( \
strcmp(((Value_String *)(l))->str, ((Value_String *)(r))->str) >= 0))

View File

@ -2,8 +2,14 @@
#include "runtime.h"
Value *newValue(size_t size) {
#if defined(__STDC_VERSION__) && (__STDC_VERSION__ >= 201112L) /* C11 */
Value *retVal = (Value *)aligned_alloc(
sizeof(void *),
((size + sizeof(void *) - 1) / sizeof(void *)) * sizeof(void *));
#else
Value *retVal = (Value *)malloc(size);
IDRIS2_REFC_VERIFY(retVal, "malloc failed");
#endif
IDRIS2_REFC_VERIFY(retVal && !idris2_vp_is_unboxed(retVal), "malloc failed");
retVal->header.refCounter = 1;
retVal->header.tag = NO_TAG;
return retVal;
@ -40,92 +46,55 @@ Value_Closure *makeClosureFromArglist(fun_ptr_t f, Value_Arglist *arglist) {
return retVal;
}
Value_Double *makeDouble(double d) {
Value *idris2_mkDouble(double d) {
Value_Double *retVal = IDRIS2_NEW_VALUE(Value_Double);
retVal->header.tag = DOUBLE_TAG;
retVal->d = d;
return retVal;
return (Value *)retVal;
}
Value_Char *makeChar(char c) {
Value_Char *retVal = IDRIS2_NEW_VALUE(Value_Char);
retVal->header.tag = CHAR_TAG;
retVal->c = c;
return retVal;
}
Value_Bits8 *makeBits8(uint8_t i) {
Value_Bits8 *retVal = IDRIS2_NEW_VALUE(Value_Bits8);
retVal->header.tag = BITS8_TAG;
retVal->ui8 = i;
return retVal;
}
Value_Bits16 *makeBits16(uint16_t i) {
Value_Bits16 *retVal = IDRIS2_NEW_VALUE(Value_Bits16);
retVal->header.tag = BITS16_TAG;
retVal->ui16 = i;
return retVal;
}
Value_Bits32 *makeBits32(uint32_t i) {
Value *idris2_mkBits32_Boxed(uint32_t i) {
Value_Bits32 *retVal = IDRIS2_NEW_VALUE(Value_Bits32);
retVal->header.tag = BITS32_TAG;
retVal->ui32 = i;
return retVal;
return (Value *)retVal;
}
Value_Bits64 *makeBits64(uint64_t i) {
Value *idris2_mkBits64(uint64_t i) {
Value_Bits64 *retVal = IDRIS2_NEW_VALUE(Value_Bits64);
retVal->header.tag = BITS64_TAG;
retVal->ui64 = i;
return retVal;
return (Value *)retVal;
}
Value_Int8 *makeInt8(int8_t i) {
Value_Int8 *retVal = IDRIS2_NEW_VALUE(Value_Int8);
retVal->header.tag = INT8_TAG;
retVal->i8 = i;
return retVal;
}
Value_Int16 *makeInt16(int16_t i) {
Value_Int16 *retVal = IDRIS2_NEW_VALUE(Value_Int16);
retVal->header.tag = INT16_TAG;
retVal->i16 = i;
return retVal;
}
Value_Int32 *makeInt32(int32_t i) {
Value *idris2_mkInt32_Boxed(int32_t i) {
Value_Int32 *retVal = IDRIS2_NEW_VALUE(Value_Int32);
retVal->header.tag = INT32_TAG;
retVal->i32 = i;
return retVal;
return (Value *)retVal;
}
Value_Int64 *makeInt64(int64_t i) {
Value *idris2_mkInt64(int64_t i) {
Value_Int64 *retVal = IDRIS2_NEW_VALUE(Value_Int64);
retVal->header.tag = INT64_TAG;
retVal->i64 = i;
return retVal;
return (Value *)retVal;
}
Value_Int8 *makeBool(int p) { return makeInt8(p ? 1 : 0); }
Value_Integer *makeInteger() {
Value_Integer *idris2_mkInteger() {
Value_Integer *retVal = IDRIS2_NEW_VALUE(Value_Integer);
retVal->header.tag = INTEGER_TAG;
mpz_init(retVal->i);
return retVal;
}
Value_Integer *makeIntegerLiteral(char *i) {
Value_Integer *retVal = makeInteger();
Value_Integer *idris2_mkIntegerLiteral(char *i) {
Value_Integer *retVal = idris2_mkInteger();
mpz_set_str(retVal->i, i, 10);
return retVal;
}
Value_String *makeEmptyString(size_t l) {
Value_String *idris2_mkEmptyString(size_t l) {
Value_String *retVal = IDRIS2_NEW_VALUE(Value_String);
retVal->header.tag = STRING_TAG;
retVal->str = malloc(l);
@ -133,7 +102,7 @@ Value_String *makeEmptyString(size_t l) {
return retVal;
}
Value_String *makeString(char *s) {
Value_String *idris2_mkString(char *s) {
Value_String *retVal = IDRIS2_NEW_VALUE(Value_String);
int l = strlen(s);
retVal->header.tag = STRING_TAG;
@ -176,16 +145,17 @@ Value_Array *makeArray(int length) {
Value *newReference(Value *source) {
// note that we explicitly allow NULL as source (for erased arguments)
if (source) {
if (source && !idris2_vp_is_unboxed(source)) {
source->header.refCounter++;
}
return source;
}
void removeReference(Value *elem) {
if (!elem) {
if (!elem || idris2_vp_is_unboxed(elem)) {
return;
}
IDRIS2_REFC_VERIFY(elem->header.refCounter > 0, "refCounter %lld",
(long long)elem->header.refCounter);
// remove reference counter
@ -194,12 +164,8 @@ void removeReference(Value *elem) {
// recursively remove all references to all children
{
switch (elem->header.tag) {
case BITS8_TAG:
case BITS16_TAG:
case BITS32_TAG:
case BITS64_TAG:
case INT8_TAG:
case INT16_TAG:
case INT32_TAG:
case INT64_TAG:
/* nothing to delete, added for sake of completeness */
@ -211,9 +177,6 @@ void removeReference(Value *elem) {
case DOUBLE_TAG:
/* nothing to delete, added for sake of completeness */
break;
case CHAR_TAG:
/* nothing to delete, added for sake of completeness */
break;
case STRING_TAG:
free(((Value_String *)elem)->str);
break;

View File

@ -14,21 +14,53 @@ Value_Constructor *newConstructor(int total, int tag);
// copies arglist, no pointer bending
Value_Closure *makeClosureFromArglist(fun_ptr_t f, Value_Arglist *);
Value_Double *makeDouble(double d);
Value_Char *makeChar(char d);
Value_Bits8 *makeBits8(uint8_t i);
Value_Bits16 *makeBits16(uint16_t i);
Value_Bits32 *makeBits32(uint32_t i);
Value_Bits64 *makeBits64(uint64_t i);
Value_Int8 *makeInt8(int8_t i);
Value_Int16 *makeInt16(int16_t i);
Value_Int32 *makeInt32(int32_t i);
Value_Int64 *makeInt64(int64_t i);
Value_Int8 *makeBool(int p);
Value_Integer *makeInteger();
Value_Integer *makeIntegerLiteral(char *i);
Value_String *makeEmptyString(size_t l);
Value_String *makeString(char *);
Value *idris2_mkDouble(double d);
#define idris2_mkChar(x) \
((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1))
#define idris2_mkBits8(x) \
((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1))
#define idris2_mkBits16(x) \
((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1))
#if !defined(UINTPTR_WIDTH)
#define idris2_mkBits32(x) \
((idris2_vp_int_shift == 16) \
? (idris2_mkBits32_Boxed(x)) \
: ((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1)))
#define idris2_mkInt32(x) \
((idris2_vp_int_shift == 16) \
? (idris2_mkInt32_Boxed(x)) \
: ((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1)))
#elif UINTPTR_WIDTH >= 64
#define idris2_mkBits32(x) \
((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1))
#define idris2_mkInt32(x) \
((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1))
#elif UINTPTR_WIDTH >= 32
#define idris2_mkBits32(x) (idris2_mkBits32_Boxed(x))
#define idris2_mkInt32(x) (idris2_mkInt32_Boxed(x)))
#else
#error "unsupported uintptr_t width"
#endif
#define idris2_mkInt8(x) \
((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1))
#define idris2_mkInt16(x) \
((Value *)(((uintptr_t)(x) << idris2_vp_int_shift) + 1))
#define idris2_mkBool(x) (idris2_mkInt8(x))
Value *idris2_mkBits32_Boxed(uint32_t i);
Value *idris2_mkBits64(uint64_t i);
Value *idris2_mkInt32_Boxed(int32_t i);
Value *idris2_mkInt64(int64_t i);
Value_Integer *idris2_mkInteger();
Value_Integer *idris2_mkIntegerLiteral(char *i);
Value_String *idris2_mkEmptyString(size_t l);
Value_String *idris2_mkString(char *);
Value_Pointer *makePointer(void *);
Value_GCPointer *makeGCPointer(void *ptr_Raw, Value_Closure *onCollectFct);

View File

@ -29,7 +29,7 @@ static Value *osstring = NULL;
Value *idris2_System_Info_prim__os(void) {
if (osstring == NULL) {
osstring = (Value *)makeString(
osstring = (Value *)idris2_mkString(
#ifdef _WIN32
"windows"
#elif _WIN64
@ -63,7 +63,7 @@ static Value *codegenstring = NULL;
Value *idris2_System_Info_prim__codegen(void) {
if (codegenstring == NULL)
codegenstring = (Value *)makeString("refc");
codegenstring = (Value *)idris2_mkString("refc");
return newReference(codegenstring);
}
@ -82,7 +82,7 @@ Value *idris2_crash(Value *msg) {
Value *idris2_Data_IOArray_Prims_prim__newArray(Value *erased, Value *_length,
Value *v, Value *_word) {
int length = extractInt(_length);
int length = idris2_vp_to_Int64(_length);
Value_Array *a = makeArray(length);
for (int i = 0; i < length; i++) {
@ -93,11 +93,11 @@ Value *idris2_Data_IOArray_Prims_prim__newArray(Value *erased, Value *_length,
}
Value *idris2_Data_IOArray_Prims_prim__arraySet(Value *erased, Value *_array,
Value *_index, Value *v,
Value *index, Value *v,
Value *_word) {
Value_Array *a = (Value_Array *)_array;
removeReference(a->arr[((Value_Int64 *)_index)->i64]);
a->arr[((Value_Int64 *)_index)->i64] = newReference(v);
removeReference(a->arr[idris2_vp_to_Int64(index)]);
a->arr[idris2_vp_to_Int64(index)] = newReference(v);
return NULL;
}
@ -201,10 +201,10 @@ Value *System_Concurrency_Raw_prim__conditionWaitTimeout(Value *_condition,
Value *_world) {
Value_Condition *cond = (Value_Condition *)_condition;
Value_Mutex *mutex = (Value_Mutex *)_mutex;
Value_Int64 *timeout = (Value_Int64 *)_timeout;
int64_t timeout = idris2_vp_to_Int64(_timeout);
struct timespec t;
t.tv_sec = timeout->i64 / 1000000;
t.tv_nsec = timeout->i64 % 1000000;
t.tv_sec = timeout / 1000000;
t.tv_nsec = timeout % 1000000;
int r = pthread_cond_timedwait(cond->cond, mutex->mutex, &t);
IDRIS2_REFC_VERIFY(!r, "pthread_cond_timedwait failed: %s", strerror(r));
return NULL;

View File

@ -21,7 +21,7 @@ Value *idris2_crash(Value *msg);
Value *idris2_Data_IOArray_Prims_prim__newArray(Value *, Value *, Value *,
Value *);
#define idris2_Data_IOArray_Prims_prim__arrayGet(rased, array, i, word) \
(newReference(((Value_Array *)(array))->arr[((Value_Int64 *)i)->i64]))
(newReference(((Value_Array *)(array))->arr[idris2_vp_to_Int64(i)]))
Value *idris2_Data_IOArray_Prims_prim__arraySet(Value *, Value *, Value *,
Value *, Value *);

View File

@ -88,7 +88,8 @@ Value *apply_closure(Value *_clos, Value *arg) {
while (1) {
Value *retVal = f(newArgs);
deconstructArglist(newArgs);
if (!retVal || retVal->header.tag != COMPLETE_CLOSURE_TAG) {
if (!retVal || idris2_vp_is_unboxed(retVal) ||
retVal->header.tag != COMPLETE_CLOSURE_TAG) {
return retVal;
}
f = ((Value_Closure *)retVal)->f;
@ -112,36 +113,32 @@ Value *tailcall_apply_closure(Value *_clos, Value *arg) {
return (Value *)makeClosureFromArglist(f, newArgs);
}
int extractInt(Value *v) {
int idris2_extractInt(Value *v) {
if (idris2_vp_is_unboxed(v))
return (int)idris2_vp_to_Int32(v);
switch (v->header.tag) {
case BITS8_TAG:
return (int)((Value_Bits8 *)v)->ui8;
case BITS16_TAG:
return (int)((Value_Bits16 *)v)->ui16;
case BITS32_TAG:
return (int)((Value_Bits32 *)v)->ui32;
return (int)idris2_vp_to_Bits32(v);
case BITS64_TAG:
return (int)((Value_Bits64 *)v)->ui64;
case INT8_TAG:
return (int)((Value_Int8 *)v)->i8;
case INT16_TAG:
return (int)((Value_Int16 *)v)->i16;
return (int)idris2_vp_to_Bits64(v);
case INT32_TAG:
return (int)((Value_Int32 *)v)->i32;
return (int)idris2_vp_to_Bits32(v);
case INT64_TAG:
return (int)((Value_Int64 *)v)->i64;
return (int)idris2_vp_to_Int64(v);
case INTEGER_TAG:
return (int)mpz_get_si(((Value_Integer *)v)->i);
case DOUBLE_TAG:
return (int)((Value_Double *)v)->d;
case CHAR_TAG:
return (int)((Value_Char *)v)->c;
return (int)idris2_vp_to_Double(v);
default:
return -1;
}
}
Value *trampoline(Value *closure) {
if (!closure || idris2_vp_is_unboxed(closure))
return closure;
fun_ptr_t f = ((Value_Closure *)closure)->f;
Value_Arglist *_arglist = ((Value_Closure *)closure)->arglist;
Value_Arglist *arglist = (Value_Arglist *)newReference((Value *)_arglist);
@ -149,7 +146,8 @@ Value *trampoline(Value *closure) {
while (1) {
Value *retVal = f(arglist);
deconstructArglist(arglist);
if (!retVal || retVal->header.tag != COMPLETE_CLOSURE_TAG) {
if (!retVal || idris2_vp_is_unboxed(retVal) ||
retVal->header.tag != COMPLETE_CLOSURE_TAG) {
return retVal;
}
f = ((Value_Closure *)retVal)->f;

View File

@ -10,6 +10,6 @@ void removeReuseConstructor(Value_Constructor *constr);
Value *apply_closure(Value *, Value *arg);
void push_Arglist(Value_Arglist *arglist, Value *arg);
int extractInt(Value *);
int idris2_extractInt(Value *);
Value *trampoline(Value *closure);
Value *tailcall_apply_closure(Value *_clos, Value *arg);

View File

@ -1,18 +1,6 @@
#include "stringOps.h"
#include "refc_util.h"
Value *stringLength(Value *s) {
int length = strlen(((Value_String *)s)->str);
return (Value *)makeInt64(length);
}
Value *head(Value *str) {
Value_Char *c = IDRIS2_NEW_VALUE(Value_Char);
c->header.tag = CHAR_TAG;
c->c = ((Value_String *)str)->str[0];
return (Value *)c;
}
Value *tail(Value *input) {
Value_String *tailStr = IDRIS2_NEW_VALUE(Value_String);
tailStr->header.tag = STRING_TAG;
@ -50,14 +38,14 @@ Value *reverse(Value *str) {
Value *strIndex(Value *str, Value *i) {
char *s = ((Value_String *)str)->str;
int idx = ((Value_Int64 *)i)->i64;
return (Value *)makeChar(s[idx]);
int idx = idris2_vp_to_Int64(i);
return (Value *)idris2_mkChar(s[idx]);
}
Value *strCons(Value *c, Value *str) {
int l = strlen(((Value_String *)str)->str);
Value_String *retVal = makeEmptyString(l + 2);
retVal->str[0] = ((Value_Char *)c)->c;
Value_String *retVal = idris2_mkEmptyString(l + 2);
retVal->str[0] = idris2_vp_to_Char(c);
memcpy(retVal->str + 1, ((Value_String *)str)->str, l);
return (Value *)retVal;
}
@ -65,7 +53,7 @@ Value *strCons(Value *c, Value *str) {
Value *strAppend(Value *a, Value *b) {
int la = strlen(((Value_String *)a)->str);
int lb = strlen(((Value_String *)b)->str);
Value_String *retVal = makeEmptyString(la + lb + 1);
Value_String *retVal = idris2_mkEmptyString(la + lb + 1);
memcpy(retVal->str, ((Value_String *)a)->str, la);
memcpy(retVal->str + la, ((Value_String *)b)->str, lb);
return (Value *)retVal;
@ -73,15 +61,15 @@ Value *strAppend(Value *a, Value *b) {
Value *strSubstr(Value *start, Value *len, Value *s) {
char *input = ((Value_String *)s)->str;
int offset = extractInt(start); /* start and len is Nat. */
int l = extractInt(len);
int offset = idris2_vp_to_Int64(start); /* start and len was come from Nat. */
int l = idris2_vp_to_Int64(len);
int tailLen = strlen(input) - offset;
if (tailLen < l) {
l = tailLen;
}
Value_String *retVal = makeEmptyString(l + 1);
Value_String *retVal = idris2_mkEmptyString(l + 1);
memcpy(retVal->str, input + offset, l);
return (Value *)retVal;
@ -103,7 +91,7 @@ char *fastPack(Value *charList) {
int i = 0;
current = (Value_Constructor *)charList;
while (current != NULL) {
retVal[i++] = ((Value_Char *)current->args[0])->c;
retVal[i++] = idris2_vp_to_Char(current->args[0]);
current = (Value_Constructor *)current->args[1];
}
@ -116,20 +104,20 @@ Value *fastUnpack(char *str) {
}
Value_Constructor *retVal = newConstructor(2, 1);
retVal->args[0] = (Value *)makeChar(str[0]);
retVal->args[0] = idris2_mkChar(str[0]);
int i = 1;
Value_Constructor *current = retVal;
Value_Constructor *next;
while (str[i] != '\0') {
next = newConstructor(2, 1);
next->args[0] = (Value *)makeChar(str[i]);
next->args[0] = idris2_mkChar(str[i]);
current->args[1] = (Value *)next;
i++;
current = next;
}
current->args[1] = (Value *)NULL;
current->args[1] = NULL;
return (Value *)retVal;
}
@ -199,8 +187,8 @@ Value *onCollectStringIterator_arglist(Value_Arglist *arglist) {
Value *stringIteratorToString(void *a, char *str, Value *it_p,
Value_Closure *f) {
String_Iterator *it = ((Value_GCPointer *)it_p)->p->p;
Value *strVal = (Value *)makeString(it->str + it->pos);
return apply_closure(newReference((Value *)f), strVal);
return apply_closure(newReference((Value *)f),
(Value *)idris2_mkString(it->str + it->pos));
}
Value *stringIteratorNext(char *s, Value *it_p) {
@ -214,7 +202,7 @@ Value *stringIteratorNext(char *s, Value *it_p) {
it->pos++; // Ok to do this as StringIterator linear
Value_Constructor *retVal = newConstructor(2, 1);
retVal->args[0] = (Value *)makeChar(c);
retVal->args[0] = (Value *)idris2_mkChar(c);
retVal->args[1] = newReference(it_p);
return (Value *)retVal;

View File

@ -1,9 +1,12 @@
#pragma once
#include "cBackend.h"
#include "casts.h"
Value *stringLength(Value *);
Value *head(Value *str);
/* stringLength : String -> Int64!? WTH!. do you have over 4Gbytes text on
* memory!? */
#define stringLength(x) (idris2_mkInt64(strlen(((Value_String *)(x))->str)))
#define head(x) (idris2_cast_String_to_Char(x))
Value *tail(Value *str);
Value *reverse(Value *str);
Value *strIndex(Value *str, Value *i);

View File

@ -18,6 +18,7 @@ Show D where
show _ = "???"
main : IO ()
main = do
printLn $ map bw [Int8, Bits16, Char, String, Integer, (), Nat]
@ -39,6 +40,65 @@ main = do
'b' => 'b'
'c' => 'C'
_ => '?') $ unpack "abcdefg"
--
printLn $ map (\case
-128 => "-128"
-127 => "-127"
0 => "zero"
127 => "127"
x => "?" ++ show x
) $ [ the Int8 (-128), -127, 0, 127 ]
printLn $ map (\case
-128 => "-128"
-127 => "-127"
0 => "zero"
127 => "127"
255 => "255"
256 => "256"
x => "?" ++ show x
) $ [ the Bits8 0, 127, 255, 256, 257 ]
printLn $ map (\case
-32768 => "-32768"
0 => "zero"
32767 => "32767"
x => "?" ++ show x
) $ [ the Int16 (-32768), 0, 32767, 11 ]
printLn $ map (\case
32768 => "32768"
0 => "zero"
32767 => "32767"
0xffff => "65535"
0xff => "255" -- 0x100ff hits this
x => "?" ++ show x
) $ [ the Bits16 32768, 0, 0xffff, 0x100ff ]
printLn $ map (\case
-2147483648 => "-2147483648"
0 => "zero"
0x7fffffff => "0x7fffffff"
x => "?" ++ show x
) $ [ the Int32 (-2147483648), 0, 0x7fffffff, -1 ]
printLn $ map (\case
0x80000000 => "0x80000000"
21474834648 => "21474834648"
0 => "zero"
x => "?" ++ show x
) $ [ the Bits32 0x80000000, 0 ]
printLn $ map (\case
0 => "zero"
711 => "711"
-9223372036854775808 => "-9223372036854775808" -- FIXME: wont work
9223372036854775807 => "9223372036854775807" -- FIXME: wont work
x => "?" ++ show x
) $ [ the Int64 (-9223372036854775808), 0, 9223372036854775807, 711, 712 ]
printLn $ map (\case
0xffffffffffffffff => "0xffffffffffffffff"
0 => "zero"
32768 => "32678"
x => "?" ++ show x
) $ [ the Bits64 0xffffffffffffffff, 0, 32768 ]
--
printLn $ map show [D0, D1 99, D2, D3]
pure ()

View File

@ -2,4 +2,12 @@
["empty", "U", "ZZZ", "L", "L"]
["?", "?", "?", "1.0", "?", "?", "2.2"]
"AbC????"
["-128", "-127", "zero", "127"]
["zero", "127", "255", "zero", "?1"]
["-32768", "zero", "32767", "?11"]
["32768", "zero", "65535", "255"]
["-2147483648", "zero", "0x7fffffff", "?-1"]
["?2147483648", "zero"]
["zero", "zero", "?9223372036854775807", "711", "?712"]
["0xffffffffffffffff", "zero", "32678"]
["D0", "D1 99", "D2", "???"]

View File

@ -160,8 +160,8 @@ Value *Main_main(void)
Value *closure_81 = (Value*)makeClosureFromArglist(fPtr_81, arglist_80);
// end csegen_44()
Value * var_0 = trampoline(closure_81); // Prelude.Types:1121:1--1138:48
Value * var_1 = (Value*)makeIntegerLiteral("1"); // Prelude.Types:1121:1--1138:48
Value * var_2 = (Value*)makeIntegerLiteral("5"); // Prelude.Types:1121:1--1138:48
Value * var_1 = (Value*)idris2_mkIntegerLiteral("1"); // Prelude.Types:1121:1--1138:48
Value * var_2 = (Value*)idris2_mkIntegerLiteral("5"); // Prelude.Types:1121:1--1138:48
// start Prelude_Types_rangeFromTo_Range__dollara(var_0, var_1, var_2)
// Prelude.Types:1121:1--1138:48
Value_Arglist *arglist_82 = newArglist(0,3);
@ -223,7 +223,7 @@ Value *Main_main_11
Value *closure_90 = (Value*)makeClosureFromArglist(fPtr_90, arglist_89);
// end csegen_44()
Value * var_1 = trampoline(closure_90); // Prelude.Types:1121:1--1138:48
Value * var_2 = (Value*)makeIntegerLiteral("1"); // Prelude.Types:1121:1--1138:48
Value * var_2 = (Value*)idris2_mkIntegerLiteral("1"); // Prelude.Types:1121:1--1138:48
// start Prelude_Types_rangeFromTo_Range__dollara(var_1, var_2, var_0)
// Prelude.Types:1121:1--1138:48
Value_Arglist *arglist_91 = newArglist(0,3);
@ -251,7 +251,7 @@ Value *Main_main_10
, Value * var_0
)
{
Value * var_2 = (Value*)makeIntegerLiteral("0"); // Main:10:30--10:36
Value * var_2 = (Value*)idris2_mkIntegerLiteral("0"); // Main:10:30--10:36
// start Main_last(var_1, var_2) // Main:10:30--10:36
Value_Arglist *arglist_93 = newArglist(0,2);
arglist_93->args[0] = var_1;
@ -270,10 +270,10 @@ Value *Main_main_10
// Prelude.Show:110:1--112:50
// end Prelude_Show_show_Show_Integer(var_3) // Prelude.Show:110:1--112:50
Value * var_4 = trampoline(closure_96);
Value * var_5 = (Value*)makeString("\x0a""");
Value * var_5 = (Value*)idris2_mkString("\x0a""");
Value *primVar_97 = strAppend(var_4, var_5);
removeReference(var_5);
removeReference(var_4);
removeReference(var_5);
Value * var_6 = primVar_97; // Prelude.IO:98:22--98:34
// start Prelude_IO_prim__putStr(var_6, var_0) // Prelude.IO:98:22--98:34
Value_Arglist *arglist_98 = newArglist(0,2);

View File

@ -1,5 +1,5 @@
. ../../testutils.sh
idris2 --cg refc -o main Main.idr > /dev/null
idris2 --cg refc -o main Main.idr
awk -v RS= '/Value \*Main_last/' build/exec/main.c
awk -v RS= '/Value \*Main_main/' build/exec/main.c

View File

@ -75,8 +75,8 @@ Value *Main_insert
}
Value * var_19 = tmp_93;
Value *tmp_94 = NULL;
int tmp_95 = extractInt(var_19);
if (tmp_95 == 1) {
int64_t tmp_95 = idris2_extractInt(var_19);
if (tmp_95 == UINT8_C(1)) {
removeReference(var_19);
// start Main_insert(var_0, var_1, var_5) // Main:8:48--8:54
Value_Arglist *arglist_96 = newArglist(0,3);
@ -97,7 +97,7 @@ Value *Main_insert
constructor_92->args[1] = var_6;
constructor_92->args[2] = var_7;
tmp_94 = (Value*)constructor_92;
} else if (tmp_95 == 0) {
} else if (tmp_95 == UINT8_C(0)) {
removeReference(var_19);
// start Main_insert(var_0, var_1, var_7) // Main:9:52--9:58
Value_Arglist *arglist_98 = newArglist(0,3);