Idris2/libs/contrib/Data/String/Parser.idr

146 lines
5.0 KiB
Idris
Raw Normal View History

2020-07-04 15:45:29 +03:00
||| A simple parser combinator library for strings. Inspired by attoparsec zepto.
module Data.String.Parser
import Control.Monad.Identity
2020-07-05 11:22:23 +03:00
import Control.Monad.Trans
2020-07-04 15:45:29 +03:00
import Data.Strings
%default partial
2020-07-05 17:39:34 +03:00
||| The input state, pos is position in the string and maxPos is the length of the input string.
public export
2020-07-04 15:45:29 +03:00
record State where
constructor S
input : String
pos : Int
2020-07-05 01:59:57 +03:00
maxPos : Int
2020-07-04 15:45:29 +03:00
2020-07-05 15:22:49 +03:00
Show State where
show s = "(" ++ show s.input ++", " ++ show s.pos ++ ", " ++ show s.maxPos ++")"
2020-07-05 17:39:34 +03:00
||| Result of applying a parser
public export
2020-07-04 15:45:29 +03:00
data Result a = Fail Int String | OK a State
2020-07-04 15:45:29 +03:00
public export
record ParseT (m : Type -> Type) (a : Type) where
constructor P
runParser : State -> m (Result a)
public export
Parser : Type -> Type
Parser a = ParseT Identity a
public export
implementation Monad m => Functor (ParseT m) where
map f p = P $ \state =>
do res <- p.runParser state
case res of
OK r state' => pure (OK (f r) state')
Fail i err => pure (Fail i err)
public export
Monad m => Applicative (ParseT m) where
pure x = P $ \s => pure (OK x s)
f <*> x = P $ \s => case !(x.runParser s) of
OK r s' => case !(f.runParser s') of
OK f' fs => pure (OK (f' r) fs)
Fail i err => pure (Fail i err)
Fail i err => pure (Fail i err)
public export
Monad m => Monad (ParseT m) where
m >>= k = P $ \state =>
do res <- m.runParser state
case res of
OK a state' => (k a).runParser state'
Fail i err => pure (Fail i err)
public export
Monad m => Alternative (ParseT m) where
empty = P $ \s => pure $ Fail (s.pos) "no alternative left"
a <|> b = P $ \s => case !(a.runParser s) of
OK r s' => pure $ OK r s'
Fail _ _ => b.runParser s
2020-07-05 11:22:23 +03:00
public export
MonadTrans ParseT where
lift x = P $ \s => do res <-x
pure $ OK res s
2020-07-04 15:45:29 +03:00
2020-07-05 17:39:34 +03:00
||| Run a parser in a monad
||| Returns a tuple of the result and final position on success.
||| Returns an error message on failure.
export
2020-07-05 17:39:34 +03:00
parseT : Monad m => ParseT m a -> String -> m (Either String (a, Int))
parseT p str = do res <- p.runParser (S str 0 (strLength str))
2020-07-04 15:45:29 +03:00
case res of
2020-07-05 17:39:34 +03:00
OK r s => pure $ Right (r, s.pos)
2020-07-05 11:22:23 +03:00
Fail i err => pure $ Left $ fastAppend ["Parse failed at position ", show i, ": ", err]
2020-07-04 15:45:29 +03:00
2020-07-05 17:39:34 +03:00
||| Run a parser in a pure function
||| Returns a tuple of the result and final position on success.
||| Returns an error message on failure.
export
2020-07-05 17:39:34 +03:00
parse : Parser a -> String -> Either String (a, Int)
2020-07-04 15:45:29 +03:00
parse p str = runIdentity $ parseT p str
2020-07-05 17:39:34 +03:00
||| Succeeds if the next char satisfies the predicate `f`
export
2020-07-04 15:45:29 +03:00
satisfy : Monad m => (Char -> Bool) -> ParseT m Char
2020-07-05 01:59:57 +03:00
satisfy f = P $ \s => do if s.pos < s.maxPos
then do
let ch = strIndex s.input s.pos
if f ch
then pure $ OK ch (S s.input (s.pos + 1) s.maxPos)
else pure $ Fail (s.pos) "satisfy"
else pure $ Fail (s.pos) "satisfy"
2020-07-05 17:39:34 +03:00
||| Succeeds if the string `str` follows.
export
2020-07-04 15:45:29 +03:00
string : Monad m => String -> ParseT m ()
string str = P $ \s => do let len = strLength str
2020-07-05 15:22:49 +03:00
if s.pos+len <= s.maxPos
then do let head = strSubstr s.pos len s.input
2020-07-05 01:59:57 +03:00
if head == str
then pure $ OK () (S s.input (s.pos + len) s.maxPos)
else pure $ Fail (s.pos) ("string " ++ show str)
else pure $ Fail (s.pos) ("string " ++ show str)
2020-07-05 17:39:34 +03:00
||| Succeeds if the next char is `c`
export
2020-07-05 01:59:57 +03:00
char : Monad m => Char -> ParseT m ()
char c = do _ <- satisfy (== c)
pure ()
mutual
2020-07-05 17:39:34 +03:00
||| Succeeds if `p` succeeds, will continue to match `p` until it fails
||| and accumulate the results in a list
export
2020-07-05 01:59:57 +03:00
some : Monad m => ParseT m a -> ParseT m (List a)
some p = pure (!p :: !(many p))
2020-07-05 17:39:34 +03:00
||| Always succeeds, will accumulate the results of `p` in a list until it fails.
export
2020-07-05 01:59:57 +03:00
many : Monad m => ParseT m a -> ParseT m (List a)
many p = some p <|> pure []
2020-07-04 15:45:29 +03:00
2020-07-05 17:39:34 +03:00
||| Always succeeds, applies the predicate `f` on chars until it fails and creates a string
||| from the results.
export
takeWhile : Monad m => (Char -> Bool) -> ParseT m String
takeWhile f = do ls <- many (satisfy f)
2020-07-05 15:22:49 +03:00
pure $ pack ls
2020-07-05 17:39:34 +03:00
||| Returns the result of the parser `p` or `def` if it fails.
2020-07-05 15:22:49 +03:00
export
option : Monad m => a -> ParseT m a -> ParseT m a
2020-07-05 17:00:29 +03:00
option def p = p <|> pure def
2020-07-05 17:39:34 +03:00
||| Returns a Maybe that contains the result of `p` if it succeeds or `Nothing` if it fails.
2020-07-05 17:00:29 +03:00
export
optional : Monad m => ParseT m a -> ParseT m (Maybe a)
2020-07-05 17:39:34 +03:00
optional p = (p >>= \res => pure $ Just res) <|> pure Nothing