mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-26 05:01:34 +03:00
[ error ] Improve locality of parse errors in implicit arguments
This commit is contained in:
parent
bff18428b4
commit
56a9bc6be4
@ -244,6 +244,7 @@ mutual
|
|||||||
braceArgs : OriginDesc -> IndentInfo -> Rule (List ArgType)
|
braceArgs : OriginDesc -> IndentInfo -> Rule (List ArgType)
|
||||||
braceArgs fname indents
|
braceArgs fname indents
|
||||||
= do start <- bounds (decoratedSymbol fname "{")
|
= do start <- bounds (decoratedSymbol fname "{")
|
||||||
|
mustWork $ do
|
||||||
list <- sepBy (decoratedSymbol fname ",")
|
list <- sepBy (decoratedSymbol fname ",")
|
||||||
$ do x <- bounds (UN . Basic <$> decoratedSimpleNamedArg fname)
|
$ do x <- bounds (UN . Basic <$> decoratedSimpleNamedArg fname)
|
||||||
let fc = boundToFC fname x
|
let fc = boundToFC fname x
|
||||||
@ -1782,13 +1783,13 @@ topDecl fname indents
|
|||||||
pure [d]
|
pure [d]
|
||||||
<|> do ds <- claims fname indents
|
<|> do ds <- claims fname indents
|
||||||
pure (forget ds)
|
pure (forget ds)
|
||||||
|
<|> do d <- implDecl fname indents
|
||||||
|
pure [d]
|
||||||
<|> do d <- definition fname indents
|
<|> do d <- definition fname indents
|
||||||
pure [d]
|
pure [d]
|
||||||
<|> fixDecl fname indents
|
<|> fixDecl fname indents
|
||||||
<|> do d <- ifaceDecl fname indents
|
<|> do d <- ifaceDecl fname indents
|
||||||
pure [d]
|
pure [d]
|
||||||
<|> do d <- implDecl fname indents
|
|
||||||
pure [d]
|
|
||||||
<|> do d <- recordDecl fname indents
|
<|> do d <- recordDecl fname indents
|
||||||
pure [d]
|
pure [d]
|
||||||
<|> do d <- namespaceDecl fname indents
|
<|> do d <- namespaceDecl fname indents
|
||||||
|
@ -93,7 +93,8 @@ idrisTestsError = MkTestPool "Error messages" [] Nothing
|
|||||||
"perror001", "perror002", "perror003", "perror004", "perror005",
|
"perror001", "perror002", "perror003", "perror004", "perror005",
|
||||||
"perror006", "perror007", "perror008", "perror009", "perror010",
|
"perror006", "perror007", "perror008", "perror009", "perror010",
|
||||||
"perror011", "perror012", "perror013", "perror014", "perror015",
|
"perror011", "perror012", "perror013", "perror014", "perror015",
|
||||||
"perror016", "perror017", "perror018", "perror019", "perror020"]
|
"perror016", "perror017", "perror018", "perror019", "perror020",
|
||||||
|
"perror021"]
|
||||||
|
|
||||||
idrisTestsInteractive : TestPool
|
idrisTestsInteractive : TestPool
|
||||||
idrisTestsInteractive = MkTestPool "Interactive editing" [] Nothing
|
idrisTestsInteractive = MkTestPool "Interactive editing" [] Nothing
|
||||||
|
6
tests/idris2/perror021/Implicit.idr
Normal file
6
tests/idris2/perror021/Implicit.idr
Normal file
@ -0,0 +1,6 @@
|
|||||||
|
import Data.Vect
|
||||||
|
|
||||||
|
myReverse : Vect n el -> Vect n el
|
||||||
|
myReverse [] = []
|
||||||
|
myReverse {n = S k} (x :: xs) =
|
||||||
|
replace {p=\n => Vect n el} (plusCommutative k 1) (myReverse xs ++ [x])
|
11
tests/idris2/perror021/expected
Normal file
11
tests/idris2/perror021/expected
Normal file
@ -0,0 +1,11 @@
|
|||||||
|
1/1: Building Implicit (Implicit.idr)
|
||||||
|
Error: Expected '}'.
|
||||||
|
|
||||||
|
Implicit:6:15--6:17
|
||||||
|
2 |
|
||||||
|
3 | myReverse : Vect n el -> Vect n el
|
||||||
|
4 | myReverse [] = []
|
||||||
|
5 | myReverse {n = S k} (x :: xs) =
|
||||||
|
6 | replace {p=\n => Vect n el} (plusCommutative k 1) (myReverse xs ++ [x])
|
||||||
|
^^
|
||||||
|
|
4
tests/idris2/perror021/run
Normal file
4
tests/idris2/perror021/run
Normal file
@ -0,0 +1,4 @@
|
|||||||
|
rm -rf build
|
||||||
|
|
||||||
|
$1 --no-banner --no-color --console-width 0 --check Implicit.idr
|
||||||
|
|
Loading…
Reference in New Issue
Block a user