mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 01:01:59 +03:00
[ fix ] check indentation after = in declarations
This commit is contained in:
parent
182bcff0ff
commit
99c665774c
@ -53,10 +53,10 @@ Eq (Clock type) where
|
||||
public export
|
||||
Ord (Clock type) where
|
||||
compare (MkClock seconds1 nanoseconds1) (MkClock seconds2 nanoseconds2) =
|
||||
case compare seconds1 seconds2 of
|
||||
LT => LT
|
||||
GT => GT
|
||||
EQ => compare nanoseconds1 nanoseconds2
|
||||
case compare seconds1 seconds2 of
|
||||
LT => LT
|
||||
GT => GT
|
||||
EQ => compare nanoseconds1 nanoseconds2
|
||||
|
||||
public export
|
||||
Show (Clock type) where
|
||||
|
@ -1231,6 +1231,7 @@ mutual
|
||||
= do b <- bounds $ do
|
||||
decoratedSymbol fname "="
|
||||
mustWork $ do
|
||||
continue indents
|
||||
rhs <- typeExpr pdef fname indents
|
||||
ws <- option [] $ whereBlock fname col
|
||||
pure (rhs, ws)
|
||||
|
@ -313,7 +313,7 @@ mutual
|
||||
prettyPrec d (PDelay _ tm) = parenthesise (d > startPrec) $ "Delay" <++> prettyPrec appPrec tm
|
||||
prettyPrec d (PForce _ tm) = parenthesise (d > startPrec) $ "Force" <++> prettyPrec appPrec tm
|
||||
prettyPrec d (PAutoApp _ f a) =
|
||||
parenthesise (d > startPrec) $ group $ prettyPrec leftAppPrec f <++> "@" <+> braces (pretty a)
|
||||
parenthesise (d > startPrec) $ group $ prettyPrec leftAppPrec f <++> "@" <+> braces (pretty a)
|
||||
prettyPrec d (PNamedApp _ f n (PRef _ a)) =
|
||||
parenthesise (d > startPrec) $ group $
|
||||
if n == rawName a
|
||||
|
4
tests/idris2/error/perror033/DeclIndent.idr
Normal file
4
tests/idris2/error/perror033/DeclIndent.idr
Normal file
@ -0,0 +1,4 @@
|
||||
namespace A
|
||||
test : (x : Type -> Type) -> Type
|
||||
test x =
|
||||
x Type
|
11
tests/idris2/error/perror033/expected
Normal file
11
tests/idris2/error/perror033/expected
Normal file
@ -0,0 +1,11 @@
|
||||
1/1: Building DeclIndent (DeclIndent.idr)
|
||||
Error: Couldn't parse any alternatives:
|
||||
1: Unexpected end of expression.
|
||||
|
||||
DeclIndent:4:2--4:3
|
||||
1 | namespace A
|
||||
2 | test : (x : Type -> Type) -> Type
|
||||
3 | test x =
|
||||
4 | x Type
|
||||
^
|
||||
... (2 others)
|
3
tests/idris2/error/perror033/run
Normal file
3
tests/idris2/error/perror033/run
Normal file
@ -0,0 +1,3 @@
|
||||
. ../../../testutils.sh
|
||||
|
||||
check DeclIndent.idr
|
Loading…
Reference in New Issue
Block a user