mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 12:42:02 +03:00
[ fix #3057 ] properly handle char literals in comments
This commit is contained in:
parent
a4ccb27c83
commit
e3214586a5
@ -150,10 +150,11 @@ mutual
|
|||||||
toEndComment : (k : Nat) -> Recognise (k /= 0)
|
toEndComment : (k : Nat) -> Recognise (k /= 0)
|
||||||
toEndComment Z = empty
|
toEndComment Z = empty
|
||||||
toEndComment (S k)
|
toEndComment (S k)
|
||||||
= some (pred (\c => c /= '-' && c /= '{' && c /= '"'))
|
= some (pred (\c => c /= '-' && c /= '{' && c /= '"' && c /= '\''))
|
||||||
<+> (eof <|> toEndComment (S k))
|
<+> (eof <|> toEndComment (S k))
|
||||||
<|> is '{' <+> singleBrace k
|
<|> is '{' <+> singleBrace k
|
||||||
<|> is '-' <+> singleDash k
|
<|> is '-' <+> singleDash k
|
||||||
|
<|> (charLit <|> pred (== '\'')) <+> toEndComment (S k)
|
||||||
<|> stringLit <+> toEndComment (S k)
|
<|> stringLit <+> toEndComment (S k)
|
||||||
|
|
||||||
||| After reading a single brace, we may either finish reading an
|
||| After reading a single brace, we may either finish reading an
|
||||||
|
@ -44,6 +44,7 @@ idrisTestsBasic = MkTestPool "Fundamental language features" [] Nothing
|
|||||||
"basic061", "basic062", "basic063", "basic064", "basic065",
|
"basic061", "basic062", "basic063", "basic064", "basic065",
|
||||||
"basic066", "basic067", "basic068", "basic069", "basic070",
|
"basic066", "basic067", "basic068", "basic069", "basic070",
|
||||||
"idiom001",
|
"idiom001",
|
||||||
|
"literals001",
|
||||||
"dotted001",
|
"dotted001",
|
||||||
"rewrite001",
|
"rewrite001",
|
||||||
"interpolation001", "interpolation002", "interpolation003",
|
"interpolation001", "interpolation002", "interpolation003",
|
||||||
|
4
tests/idris2/literals001/Test.idr
Normal file
4
tests/idris2/literals001/Test.idr
Normal file
@ -0,0 +1,4 @@
|
|||||||
|
{-
|
||||||
|
c : Char
|
||||||
|
c = '"'
|
||||||
|
-}
|
1
tests/idris2/literals001/expected
Normal file
1
tests/idris2/literals001/expected
Normal file
@ -0,0 +1 @@
|
|||||||
|
1/1: Building Test (Test.idr)
|
3
tests/idris2/literals001/run
Executable file
3
tests/idris2/literals001/run
Executable file
@ -0,0 +1,3 @@
|
|||||||
|
rm -rf build
|
||||||
|
|
||||||
|
$1 --no-color --console-width 0 --no-banner -c Test.idr
|
Loading…
Reference in New Issue
Block a user