Add takeUntil Data.String.Parser parser

This commit is contained in:
Robert Wright 2021-10-22 16:33:14 +01:00
parent d531cc8dea
commit 12955bc5bc
3 changed files with 33 additions and 0 deletions

View File

@ -4,10 +4,12 @@ import public Control.Monad.Identity
import Control.Monad.Trans
import Data.String
import Data.String.Extra
import Data.Fin
import Data.List
import Data.List.Alternating
import Data.List1
import Data.SnocList
import Data.Vect
%default total
@ -241,6 +243,24 @@ covering
takeWhile1 : Monad m => (Char -> Bool) -> ParseT m String
takeWhile1 f = pack <$> some (satisfy f)
||| Takes from the input until the `stop` string is found.
||| Fails if the `stop` string cannot be found.
export
covering
takeUntil : Monad m => (stop : String) -> ParseT m String
takeUntil stop = do
let StrCons s top = strM stop
| StrNil => pure ""
takeUntil' s top [<]
where
takeUntil' : Monad m' => (s : Char) -> (top : String) -> (acc : SnocList String) -> ParseT m' String
takeUntil' s top acc = do
init <- takeWhile (/= s)
skip $ char s <?> "end of string reached - \{show stop} not found"
case !(succeeds $ string top) of
False => takeUntil' s top $ acc :< (init +> s)
True => pure $ concat {t = List} $ cast $ acc :< init
||| Parses zero or more space characters
export
covering

View File

@ -40,6 +40,13 @@ maybeParser = do res <- optional (string "abc")
ignore $ string "def"
pure $ isJust res
-- test takeUntil
takeUntilParser : Parser String
takeUntilParser = do ignore $ string "<!--"
res <- takeUntil "-->"
eos -- To check that takeUntil consumes the stop string itself
pure res
main : IO ()
main = do
res <- parseT parseStuff "abcdef"
@ -56,6 +63,9 @@ main = do
showRes $ pureParsing "63553"
s <- parseT (takeWhile isDigit) "887abc8993"
showRes s
showRes $ parse takeUntilParser "<!--XML Comment-->"
showRes $ parse takeUntilParser "<!--<- Complicated -- XML -- Comment ->-->"
showRes $ parse takeUntilParser "<!--Unclosed XML Comment"
res <- parseT optParser "123def"
showRes res
res <- parseT optParser "def"

View File

@ -6,6 +6,9 @@ Parse failed at position 0: Not good
['7', '6', '6', '7', '7', '5']
['6', '3', '5', '5', '3']
"887"
"XML Comment"
"<- Complicated -- XML -- Comment ->"
Parse failed at position 24: end of string reached - "-->" not found
"123"
""
True