mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-10 05:07:15 +03:00
[ fix ] preserve spacing during update
This commit is contained in:
parent
3869233ce0
commit
53a7cf36a1
@ -56,22 +56,23 @@ doUpdates : {auto u : Ref UPD (List String)} ->
|
||||
Core (List SourcePart)
|
||||
doUpdates defs ups [] = pure []
|
||||
doUpdates defs ups (LBrace :: xs)
|
||||
= case dropSpace xs of
|
||||
= let (ws, nws) = spanSpace xs in map (LBrace :: ws ++) $
|
||||
case nws of
|
||||
Name n :: RBrace :: rest =>
|
||||
pure (LBrace :: Name n ::
|
||||
pure (Name n ::
|
||||
Whitespace " " :: Equal :: Whitespace " " ::
|
||||
!(doUpdates defs ups (Name n :: RBrace :: rest)))
|
||||
Name n :: Equal :: rest =>
|
||||
pure (LBrace :: Name n ::
|
||||
pure (Name n ::
|
||||
Whitespace " " :: Equal :: Whitespace " " ::
|
||||
!(doUpdates defs ups rest))
|
||||
_ => pure (LBrace :: !(doUpdates defs ups xs))
|
||||
_ => doUpdates defs ups xs
|
||||
where
|
||||
dropSpace : List SourcePart -> List SourcePart
|
||||
dropSpace [] = []
|
||||
dropSpace (RBrace :: xs) = RBrace :: xs
|
||||
dropSpace (Whitespace _ :: xs) = dropSpace xs
|
||||
dropSpace (x :: xs) = x :: dropSpace xs
|
||||
spanSpace : List SourcePart -> (List SourcePart, List SourcePart)
|
||||
spanSpace [] = ([], [])
|
||||
spanSpace (RBrace :: xs) = ([], RBrace :: xs)
|
||||
spanSpace (w@(Whitespace _) :: xs) = mapFst (w ::) (spanSpace xs)
|
||||
spanSpace (x :: xs) = map (x ::) (spanSpace xs)
|
||||
doUpdates defs ups (Name n :: xs)
|
||||
= case lookup n ups of
|
||||
Nothing => pure (Name n :: !(doUpdates defs ups xs))
|
||||
|
@ -52,6 +52,7 @@ idrisTests
|
||||
"interactive001", "interactive002", "interactive003", "interactive004",
|
||||
"interactive005", "interactive006", "interactive007", "interactive008",
|
||||
"interactive009", "interactive010", "interactive011", "interactive012",
|
||||
"interactive013",
|
||||
-- Literate
|
||||
"literate001", "literate002", "literate003", "literate004",
|
||||
"literate005", "literate006", "literate007", "literate008",
|
||||
|
4
tests/idris2/interactive013/Spacing.idr
Normal file
4
tests/idris2/interactive013/Spacing.idr
Normal file
@ -0,0 +1,4 @@
|
||||
module Spacing
|
||||
|
||||
id : {n : Nat} -> Nat
|
||||
id { n} = ?a
|
5
tests/idris2/interactive013/expected
Normal file
5
tests/idris2/interactive013/expected
Normal file
@ -0,0 +1,5 @@
|
||||
1/1: Building Spacing (Spacing.idr)
|
||||
Spacing> id { n = 0} = ?a_1
|
||||
id { n = (S k)} = ?a_2
|
||||
Spacing>
|
||||
Bye for now!
|
2
tests/idris2/interactive013/input
Normal file
2
tests/idris2/interactive013/input
Normal file
@ -0,0 +1,2 @@
|
||||
:cs 4 6 n
|
||||
:q
|
3
tests/idris2/interactive013/run
Executable file
3
tests/idris2/interactive013/run
Executable file
@ -0,0 +1,3 @@
|
||||
$1 --no-banner Spacing.idr < input
|
||||
|
||||
rm -rf build
|
Loading…
Reference in New Issue
Block a user