address comment

This commit is contained in:
Andy Lok 2021-02-25 18:43:42 +08:00
parent 775d7b4bdb
commit 8c23f5944b
4 changed files with 21 additions and 20 deletions

View File

@ -424,7 +424,6 @@ mutual
-- merge neighbouring StrLiteral
mergeStrLit : List PStr -> List PStr
mergeStrLit [] = []
mergeStrLit [x] = [x]
mergeStrLit ((StrLiteral fc str1)::(StrLiteral _ str2)::xs)
= (StrLiteral fc (str1 ++ str2)) :: xs
mergeStrLit (x::xs) = x :: mergeStrLit xs
@ -451,12 +450,12 @@ mutual
trimLast _ (initLines `snoc` []) | Snoc [] initLines _ = pure lines
trimLast _ (initLines `snoc` [StrLiteral fc str]) | Snoc [(StrLiteral _ _)] initLines _
= if any (not . isSpace) (fastUnpack str)
then throw $ BadMultiline fc "Unexpected character in the last line"
then throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters"
else pure initLines
trimLast _ (initLines `snoc` xs) | Snoc xs initLines _
= let fc = fromMaybe fc $ findBy (\case StrInterp fc _ => Just fc;
StrLiteral _ _ => Nothing) xs in
throw $ BadMultiline fc "Unexpected interpolation in the last line"
throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters"
dropLastNL : List PStr -> List PStr
dropLastNL pstrs with (snocList pstrs)
@ -470,14 +469,14 @@ mutual
trimLeft indent [(StrLiteral fc str)]
= let (trimed, rest) = splitAt indent (fastUnpack str) in
if any (not . isSpace) trimed
then throw $ BadMultiline fc "Indentation not enough"
then throw $ BadMultiline fc "Line is less indented than the closing delimiter"
else pure $ [StrLiteral fc (fastPack rest)]
trimLeft indent ((StrLiteral fc str)::xs)
= let (trimed, rest) = splitAt indent (fastUnpack str) in
if any (not . isSpace) trimed || length trimed < indent
then throw $ BadMultiline fc "Indentation not enough"
then throw $ BadMultiline fc "Line is less indented than the closing delimiter"
else pure $ (StrLiteral fc (fastPack rest))::xs
trimLeft indent xs = throw $ BadMultiline fc "Indentation not enough"
trimLeft indent xs = throw $ BadMultiline fc "Line is less indented than the closing delimiter"
expandDo : {auto s : Ref Syn SyntaxInfo} ->
{auto c : Ref Ctxt Defs} ->

View File

@ -795,6 +795,8 @@ mutual
toPStr : (WithBounds $ Either PTerm (List1 String)) -> Either String PStr
toPStr x = case x.val of
Right (str:::[]) => Right $ StrLiteral (boundToFC fname x) str
-- FIXME: This throws the error at the next line after the line wrap.
-- But it's supposed to point to the string begin quote.
Right (_:::strs) => Left "Multi-line string is expected to begin with \"\"\""
Left tm => Right $ StrInterp (boundToFC fname x) tm
@ -811,15 +813,15 @@ mutual
PMultiline (boundToFC fname b) (fromInteger $ cast col) xs
where
toLines : List (WithBounds $ Either PTerm (List1 String)) -> List PStr -> List (List PStr) -> List (List PStr)
toLines [] line acc = reverse (reverse line::acc)
toLines [] line acc = acc `snoc` line
toLines (x::xs) line acc
= case x.val of
Left tm => toLines xs ((StrInterp (boundToFC fname x) tm)::line) acc
Right (str:::[]) => toLines xs ((StrLiteral (boundToFC fname x) str)::line) acc
Left tm => toLines xs (line `snoc` (StrInterp (boundToFC fname x) tm)) acc
Right (str:::[]) => toLines xs (line `snoc` (StrLiteral (boundToFC fname x) str)) acc
-- FIXME: calculate the precise FC so as to improve error report for invalid indentation.
Right (str:::strs@(_::_)) => toLines xs [StrLiteral (boundToFC fname x) (last strs)]
((map (\str => [StrLiteral (boundToFC fname x) str]) (List.drop 1 $ reverse strs))
++ (List.reverse ((StrLiteral (boundToFC fname x) str)::line)::acc))
((acc `snoc` (line `snoc` (StrLiteral (boundToFC fname x) str))) ++
((\str => [StrLiteral (boundToFC fname x) str]) <$> (init strs)))
visOption : Rule Visibility
visOption

View File

@ -499,9 +499,9 @@ mutual
= "rewrite " ++ show rule
export
showPStr : PStr -> String
showPStr (StrLiteral _ str) = show str
showPStr (StrInterp _ tm) = show tm
Show PStr where
show (StrLiteral _ str) = show str
show (StrInterp _ tm) = show tm
showUpdate : PFieldUpdate -> String
showUpdate (PSetField p v) = showSep "." p ++ " = " ++ show v
@ -590,8 +590,8 @@ mutual
showPrec d (PSectionR _ x op) = "(" ++ showPrec d x ++ " " ++ showPrec d op ++ ")"
showPrec d (PEq fc l r) = showPrec d l ++ " = " ++ showPrec d r
showPrec d (PBracketed _ tm) = "(" ++ showPrec d tm ++ ")"
showPrec d (PString _ xs) = join " ++ " $ showPStr <$> xs
showPrec d (PMultiline _ indent xs) = "multiline (" ++ (join " ++ " $ showPStr <$> concat xs) ++ ")"
showPrec d (PString _ xs) = join " ++ " $ show <$> xs
showPrec d (PMultiline _ indent xs) = "multiline (" ++ (join " ++ " $ show <$> concat xs) ++ ")"
showPrec d (PDoBlock _ ns ds)
= "do " ++ showSep " ; " (map showDo ds)
showPrec d (PBang _ tm) = "!" ++ showPrec d tm

View File

@ -31,7 +31,7 @@ StrError4.idr:2:7--2:8
^
1/1: Building StrError5 (StrError5.idr)
Error: While processing multi-line string. Indentation not enough.
Error: While processing multi-line string. Line is less indented than the closing delimiter.
StrError5.idr:3:1--4:5
3 | a
@ -62,7 +62,7 @@ StrError8.idr:2:7--2:8
^
1/1: Building StrError9 (StrError9.idr)
Error: While processing multi-line string. Unexpected interpolation in the last line.
Error: While processing multi-line string. Closing delimiter of multiline strings cannot be preceded by non-whitespace characters.
StrError9.idr:3:1--3:11
1 | foo : String
@ -71,7 +71,7 @@ StrError9.idr:3:1--3:11
^^^^^^^^^^
1/1: Building StrError10 (StrError10.idr)
Error: While processing multi-line string. Unexpected character in the last line.
Error: While processing multi-line string. Closing delimiter of multiline strings cannot be preceded by non-whitespace characters.
StrError10.idr:3:1--3:2
1 | foo : String
@ -89,7 +89,7 @@ StrError11.idr:4:1--4:2
^
1/1: Building StrError12 (StrError12.idr)
Error: While processing multi-line string. Indentation not enough.
Error: While processing multi-line string. Line is less indented than the closing delimiter.
StrError12.idr:3:1--3:3
1 | foo : String