Idris2/tests/idris2/schemeeval005/Printf.idr
Edwin Brady 75716cd0d1
Fix casts in scheme evaluator (#2011)
* Fix casts in scheme evaluator

We really need test cases for all the primitives before we can use this
evaluator properly. Also test cases that run inside an environment,
which are a bit harder to construct.

* Add the cast fixes to racket support code

* More racket compile time evaluation fixes

We had the chez version of some primtives in the ct-support file. We
need a full set of tests for the primitives here too...
2021-10-16 14:19:26 +01:00

39 lines
1.1 KiB
Idris

module Printf
import Prelude
import Data.String
data Arg
= AInt Arg
| AOther Char Arg
| AEnd
buildArg : List Char -> Arg
buildArg fmt = case fmt of
'%' :: 'i' :: fmtTail => AInt (buildArg fmtTail)
c :: fmtTail => AOther c (buildArg fmtTail)
Nil => AEnd
argToType : Arg -> Type -> Type
argToType a result = case a of
AInt fmtTail => Int -> argToType fmtTail result
AOther _ fmtTail => argToType fmtTail result
AEnd => result
-- PrintfType "foo" result = result
-- PrintfType "%i\n" result = Int -> result
-- etc
PrintfType : String -> Type -> Type
PrintfType fmt result = argToType (buildArg (unpack fmt)) result
sprintf : (fmt : String) -> PrintfType fmt String
sprintf fmt = go "" (buildArg (unpack fmt)) where
go : String -> (arg : Arg) -> argToType arg String
go strTail arg = case arg of
AInt fmtTail => \i : Int => go (strTail ++ show i) fmtTail
AOther c fmtTail => go (strTail ++ singleton c) fmtTail
AEnd => strTail
test : ?result
test = sprintf "%i%i%i%i%i%i"