Skip to content

Commit

Permalink
[ fix ] check indentation after = in declarations
Browse files Browse the repository at this point in the history
  • Loading branch information
dunhamsteve authored and gallais committed Jul 16, 2024
1 parent 182bcff commit 99c6657
Show file tree
Hide file tree
Showing 6 changed files with 24 additions and 5 deletions.
8 changes: 4 additions & 4 deletions libs/base/System/Clock.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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
Expand Down
1 change: 1 addition & 0 deletions src/Idris/Parser.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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)
Expand Down
2 changes: 1 addition & 1 deletion src/Idris/Pretty.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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
Expand Down
4 changes: 4 additions & 0 deletions tests/idris2/error/perror033/DeclIndent.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
namespace A
test : (x : Type -> Type) -> Type
test x =
x Type
11 changes: 11 additions & 0 deletions tests/idris2/error/perror033/expected
Original file line number Diff line number Diff line change
@@ -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 changes: 3 additions & 0 deletions tests/idris2/error/perror033/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
. ../../../testutils.sh

check DeclIndent.idr

0 comments on commit 99c6657

Please sign in to comment.