mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 04:33:45 +03:00
Add test case for lexer issue and fix other instance.
This commit is contained in:
parent
93b0f9c7b8
commit
90e46ef340
@ -155,7 +155,7 @@ tokenise pred line col acc tmap str
|
||||
|
||||
getCols : List Char -> Int -> Int
|
||||
getCols x c
|
||||
= case span (/= '\n') (reverse x) of
|
||||
= case span (/= '\n') x of
|
||||
(incol, []) => c + cast (length incol)
|
||||
(incol, _) => cast (length incol)
|
||||
|
||||
|
23
tests/contrib/lexer/Test.idr
Normal file
23
tests/contrib/lexer/Test.idr
Normal file
@ -0,0 +1,23 @@
|
||||
import Text.Lexer
|
||||
|
||||
data Kind = Ignore | Ident | Oper | Number
|
||||
|
||||
ignore : WithBounds (Token Kind) -> Bool
|
||||
ignore (MkBounded (Tok Ignore _) _ _) = True
|
||||
ignore _ = False
|
||||
|
||||
tmap : TokenMap (Token Kind)
|
||||
tmap = [
|
||||
(alpha <+> many alphaNum, Tok Ident),
|
||||
(some digit, Tok Number),
|
||||
(some symbol, Tok Oper),
|
||||
(spaces, Tok Ignore)
|
||||
]
|
||||
|
||||
show : WithBounds (Token Kind) -> String
|
||||
show t@(MkBounded (Tok _ v) _ _) = "\{show $ start t} \{show v}"
|
||||
|
||||
main : IO ()
|
||||
main = do
|
||||
let toks = filter (not . ignore) $ fst $ lex tmap "let x = 1\n y = 2\n in x + y"
|
||||
traverse_ (putStrLn . show) toks
|
11
tests/contrib/lexer/expected
Normal file
11
tests/contrib/lexer/expected
Normal file
@ -0,0 +1,11 @@
|
||||
(0, 0) "let"
|
||||
(0, 4) "x"
|
||||
(0, 6) "="
|
||||
(0, 8) "1"
|
||||
(1, 4) "y"
|
||||
(1, 6) "="
|
||||
(1, 8) "2"
|
||||
(2, 1) "in"
|
||||
(2, 4) "x"
|
||||
(2, 6) "+"
|
||||
(2, 8) "y"
|
3
tests/contrib/lexer/run
Normal file
3
tests/contrib/lexer/run
Normal file
@ -0,0 +1,3 @@
|
||||
$1 --no-banner --no-color --console-width 0 -p contrib -x main Test.idr
|
||||
|
||||
rm -rf build
|
Loading…
Reference in New Issue
Block a user