Refactor getFirstMatch function from tokeniser

This commit is contained in:
André Videla 2021-02-13 02:51:48 +00:00
parent 5cd7642991
commit 0705ba55c5

View File

@ -93,37 +93,41 @@ tokenise reject tokenizer line col acc str
(incol, []) => c + cast (length incol) (incol, []) => c + cast (length incol)
(incol, _) => cast (length incol) (incol, _) => cast (length incol)
-- get the next lexeme using the `Lexer` in argument, its position and the input
-- Returns the new position, the lexeme parsed and the rest of the input
-- If parsing the lexer fails, this returns `Nothing`
getNext : (lexer : Lexer) -> (line, col : Int) -> (input : List Char) -> Maybe (String, Int, Int, List Char)
getNext lexer line col str =
let Just (token, rest) = scan lexer [] str
| _ => Nothing
line' = line + cast (countNLs token)
col' = getCols token col
tokenStr = fastPack $ reverse token
in pure (tokenStr, line', col', rest)
getFirstMatch : Tokenizer a -> List Char -> getFirstMatch : Tokenizer a -> List Char ->
Either StopReason (List (WithBounds a), Int, Int, List Char) Either StopReason (List (WithBounds a), Int, Int, List Char)
getFirstMatch (Match lex fn) str getFirstMatch (Match lex fn) str
= let Just (tok, rest) = scan lex [] str = let Just (tok, line', col', rest) = getNext lex line col str
| _ => Left NoRuleApply | _ => Left NoRuleApply
line' = line + cast (countNLs tok) tok' = MkBounded (fn tok) False line col line' col'
col' = getCols tok col in Right ([tok'], line', col', rest)
tok' = MkBounded (fn (fastPack (reverse tok))) False line col line' col' in
Right ([tok'], line', col', rest)
getFirstMatch (Compose begin mapBegin tagger middleFn endFn mapEnd) str getFirstMatch (Compose begin mapBegin tagger middleFn endFn mapEnd) str
= do let Just (beginTok, rest) = scan begin [] str = let Just (beginTok', line', col' , rest) = getNext begin line col str
| _ => Left NoRuleApply | Nothing => Left NoRuleApply
let line' = line + cast (countNLs beginTok) tag = tagger beginTok'
let col' = getCols beginTok col middle = middleFn tag
let beginTok' = fastPack $ reverse beginTok end = endFn tag
let tag = tagger beginTok' beginTok'' = MkBounded (mapBegin beginTok') False line col line' col'
let beginTok'' = MkBounded (mapBegin beginTok') False line col line' col' (midToks, (reason, line'', col'', rest'')) =
let middle = middleFn tag
let end = endFn tag
let (midToks, (reason, line'', col'', rest'')) =
tokenise end middle line' col' [] rest tokenise end middle line' col' [] rest
case reason of in case reason of
reason@(ComposeNotClosing _ _) => Left reason (ComposeNotClosing _ _) => Left reason
_ => Right () _ => let Just (endTok', lineEnd, colEnd, restEnd) =
let Just (endTok, rest''') = scan end [] rest'' getNext end line'' col'' rest''
| _ => Left $ ComposeNotClosing (line, col) (line', col') | _ => Left $ ComposeNotClosing (line, col) (line', col')
let line''' = line'' + cast (countNLs endTok) endTok'' = MkBounded (mapEnd endTok') False line'' col'' lineEnd colEnd
let col''' = getCols endTok col'' in Right ([endTok''] ++ reverse midToks ++ [beginTok''], lineEnd, colEnd, restEnd)
let endTok' = fastPack $ reverse endTok
let endTok'' = MkBounded (mapEnd endTok') False line'' col'' line''' col'''
Right ([endTok''] ++ reverse midToks ++ [beginTok''], line''', col''', rest''')
getFirstMatch (Alt t1 t2) str getFirstMatch (Alt t1 t2) str
= case getFirstMatch t1 str of = case getFirstMatch t1 str of
Right result => Right result Right result => Right result