mirror of
https://github.com/idris-lang/Idris2.git
synced 2025-01-04 22:24:19 +03:00
Add Data.String.Iterator implementation for JavaScript
This commit is contained in:
parent
0a203f8f52
commit
9add729ca3
@ -22,6 +22,7 @@ data StringIterator : String -> Type where [external]
|
||||
-- to avoid subverting the linearity guarantees of withString.
|
||||
%foreign
|
||||
"scheme:blodwen-string-iterator-new"
|
||||
"javascript:stringIterator:new"
|
||||
private
|
||||
fromString : (str : String) -> StringIterator str
|
||||
|
||||
@ -48,6 +49,7 @@ data UnconsResult : String -> Type where
|
||||
-- (e.g. byte offset into an UTF-8 string).
|
||||
%foreign
|
||||
"scheme:blodwen-string-iterator-next"
|
||||
"javascript:stringIterator:next"
|
||||
export
|
||||
uncons : (str : String) -> (1 it : StringIterator str) -> UnconsResult str
|
||||
|
||||
|
@ -73,6 +73,28 @@ addSupportToPreamble : {auto c : Ref ESs ESSt} -> String -> String -> Core Strin
|
||||
addSupportToPreamble name code =
|
||||
addToPreamble name name code
|
||||
|
||||
addStringIteratorToPreamble : {auto c : Ref ESs ESSt} -> Core String
|
||||
addStringIteratorToPreamble =
|
||||
do
|
||||
let defs = "
|
||||
function __prim_stringIteratorNew(str) {
|
||||
return str[Symbol.iterator]();
|
||||
}
|
||||
function __prim_stringIteratorNext(str, it) {
|
||||
const char = it.next();
|
||||
if (char.done) {
|
||||
return {h: 0}; // EOF
|
||||
} else {
|
||||
return {
|
||||
h: 1, // Character
|
||||
a1: char.value,
|
||||
a2: it
|
||||
};
|
||||
}
|
||||
}"
|
||||
let name = "stringIterator"
|
||||
let newName = esName name
|
||||
addToPreamble name newName defs
|
||||
|
||||
jsIdent : String -> String
|
||||
jsIdent s = concatMap okchar (unpack s)
|
||||
@ -311,6 +333,13 @@ makeForeign n x =
|
||||
lib_code <- readDataFile ("js/" ++ lib ++ ".js")
|
||||
addSupportToPreamble lib lib_code
|
||||
pure $ "const " ++ jsName n ++ " = " ++ lib ++ "_" ++ name ++ "\n"
|
||||
"stringIterator" =>
|
||||
do
|
||||
addStringIteratorToPreamble
|
||||
case def of
|
||||
"new" => pure $ "const " ++ jsName n ++ " = __prim_stringIteratorNew;\n"
|
||||
"next" => pure $ "const " ++ jsName n ++ " = __prim_stringIteratorNext;\n"
|
||||
_ => throw (InternalError $ "invalid string iterator function: " ++ def ++ ", supported functions are \"new\", \"next\"")
|
||||
|
||||
|
||||
_ => throw (InternalError $ "invalid foreign type : " ++ ty ++ ", supported types are \"lambda\", \"lambdaRequire\", \"support\"")
|
||||
|
Loading…
Reference in New Issue
Block a user