reinstate the good error messages

This commit is contained in:
Andre Videla 2023-11-04 19:53:28 +00:00 committed by André Videla
parent a73522ec66
commit 992fc62d86
9 changed files with 96 additions and 34 deletions

View File

@ -15,6 +15,7 @@ import Decidable.Equality
import Libraries.Data.NameMap
import Libraries.Text.PrettyPrint.Prettyprinter
import Libraries.Text.PrettyPrint.Prettyprinter.Util
import Libraries.Text.Bounded
import Libraries.Data.String.Extra
import Libraries.Data.SnocList.SizeOf
@ -172,11 +173,11 @@ data OperatorLHSInfo : tm -> Type where
-- Traditional operator wihtout binding, carries the lhs
NoBinder : (lhs : tm) -> OperatorLHSInfo tm
-- (nm : ty) ** fn x
BindType : (name : Name) -> (ty : tm) -> OperatorLHSInfo tm
BindType : (name : WithBounds Name) -> (ty : tm) -> OperatorLHSInfo tm
-- (nm := exp) ** fn nm
BindExpr : (name : Name) -> (expr : tm) -> OperatorLHSInfo tm
BindExpr : (name : WithBounds Name) -> (expr : tm) -> OperatorLHSInfo tm
-- (nm : ty := exp) ** fn nm
BindExplicitType : (name : Name) -> (type, expr : tm) -> OperatorLHSInfo tm
BindExplicitType : (name : WithBounds Name) -> (type, expr : tm) -> OperatorLHSInfo tm
export
Show (OperatorLHSInfo tm) where
@ -204,9 +205,9 @@ export
export
(.getBoundName) : OperatorLHSInfo tm -> Maybe Name
(.getBoundName) (NoBinder lhs) = Nothing
(.getBoundName) (BindType name ty) = Just name
(.getBoundName) (BindExpr name expr) = Just name
(.getBoundName) (BindExplicitType name type expr) = Just name
(.getBoundName) (BindType name ty) = Just name.val
(.getBoundName) (BindExpr name expr) = Just name.val
(.getBoundName) (BindExplicitType name type expr) = Just name.val
export
(.getBinder) : OperatorLHSInfo tm -> BindingModifier

View File

@ -766,20 +766,20 @@ mutual
= do desugaredLHS <- desugarB side ps lhs
desugaredRHS <- desugarTree side ps r
pure $ IApp loc (IApp loc (IVar opFC op) desugaredLHS)
(ILam loc top Explicit (Just name) desugaredLHS desugaredRHS)
(ILam loc top Explicit (Just name.val) desugaredLHS desugaredRHS)
-- (x := exp ** f x) ==>> (**) exp (\x : ? => f x)
desugarTree side ps (Infix loc opFC (op, Just (BindExpr name lhs)) _ r)
= do desugaredLHS <- desugarB side ps lhs
desugaredRHS <- desugarTree side ps r
pure $ IApp loc (IApp loc (IVar opFC op) desugaredLHS)
(ILam loc top Explicit (Just name) (Implicit opFC False) desugaredRHS)
(ILam loc top Explicit (Just name.val) (Implicit opFC False) desugaredRHS)
-- (x : ty := exp ** f x) ==>> (**) exp (\x : ty => f x)
desugarTree side ps (Infix loc opFC (op, Just (BindExplicitType name ty expr)) _ r)
= do desugaredLHS <- desugarB side ps expr
desugaredType <- desugarB side ps ty
desugaredRHS <- desugarTree side ps r
pure $ IApp loc (IApp loc (IVar opFC op) desugaredLHS)
(ILam loc top Explicit (Just name) desugaredType desugaredRHS)
(ILam loc top Explicit (Just name.val) desugaredType desugaredRHS)
desugarTree side ps (Infix loc opFC (op, Nothing) _ r)
= throw $ InternalError "illegal fixity: Parsed as infix but no binding information"

View File

@ -327,7 +327,7 @@ mutual
-- The different kinds of operator bindings `x : ty` for typebind
-- x := e and x : ty := e for autobind
opBinderTypes : OriginDesc -> IndentInfo -> Name -> Rule (OperatorLHSInfo PTerm)
opBinderTypes : OriginDesc -> IndentInfo -> WithBounds Name -> Rule (OperatorLHSInfo PTerm)
opBinderTypes fname indents boundName =
do decoratedSymbol fname ":"
ty <- typeExpr pdef fname indents
@ -343,14 +343,13 @@ mutual
opBinder : OriginDesc -> IndentInfo -> Rule (OperatorLHSInfo PTerm)
opBinder fname indents
= do boundName <- (UN . Basic <$> decoratedSimpleBinderName fname)
= do boundName <- bounds (UN . Basic <$> decoratedSimpleBinderName fname)
opBinderTypes fname indents boundName
autobindOp : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm
autobindOp q fname indents
= do binder <- bounds $ parens fname (opBinder fname indents)
continue indents
op <- bounds iOperator
autobindOp : WithBounds (OperatorLHSInfo PTerm) -> ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm
autobindOp binder q fname indents
= do op <- bounds iOperator
commit
e <- bounds (expr q fname indents)
pure (POp (boundToFC fname $ mergeBounds binder e)
(boundToFC fname op)
@ -387,7 +386,6 @@ mutual
opExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm
opExpr q fname indents = autobindOp q fname indents <|> opExprBase q fname indents
dpairType : OriginDesc -> WithBounds t -> IndentInfo -> Rule PTerm
dpairType fname start indents
= do loc <- bounds (do x <- UN . Basic <$> decoratedSimpleBinderName fname
@ -693,8 +691,11 @@ mutual
binderName = Basic <$> unqualifiedName
<|> symbol "_" $> Underscore
PiBindList : Type
PiBindList = List (RigCount, WithBounds (Maybe Name), PTerm)
pibindList : OriginDesc -> IndentInfo ->
Rule (List (RigCount, WithBounds (Maybe Name), PTerm))
Rule PiBindList
pibindList fname indents
= do params <- pibindListName fname indents
pure $ map (\(rig, n, ty) => (rig, map Just n, ty)) params
@ -704,14 +705,44 @@ mutual
= (decoratedSymbol fname "->" $> Explicit)
<|> (decoratedSymbol fname "=>" $> AutoImplicit)
explicitPi : OriginDesc -> IndentInfo -> Rule PTerm
explicitPi fname indents
= do b <- bounds $ parens fname $ pibindList fname indents
exp <- mustWorkBecause b.bounds "Cannot return a named argument"
explicitPi : WithBounds PiBindList -> OriginDesc -> IndentInfo -> Rule PTerm
explicitPi b fname indents
= do exp <- mustWorkBecause b.bounds "Cannot return a named argument"
$ bindSymbol fname
scope <- mustWork $ typeExpr pdef fname indents
pure (pibindAll fname exp b.val scope)
total
asOpInfo : PiBindList -> Maybe (OperatorLHSInfo PTerm)
asOpInfo [(rig, name, term)] = let True = rig == top
| _ => Nothing
Just n = sequence name
| _ => Nothing
in Just (BindType n term)
asOpInfo _ = Nothing
fromOpInfo : OperatorLHSInfo PTerm -> Maybe PiBindList
fromOpInfo (BindType name ty) = Just [(top, map Just name, ty)]
fromOpInfo _ = Nothing
piOrAutobind : OriginDesc -> IndentInfo -> Rule PTerm
piOrAutobind fname indents
= ((parens fname (Parser.choose (bounds $ opBinder fname indents) (bounds $ pibindList fname indents)))) >>=
\b : (Either (WithBounds (OperatorLHSInfo PTerm)) (WithBounds PiBindList)) =>
the (Rule PTerm) $ case b of
Left autobind => (autobindOp autobind pdef fname indents)
<|> let Just x = traverse fromOpInfo autobind
| _ => fail "not a binder"
in explicitPi x fname indents
Right bindList => (explicitPi bindList fname indents)
-- = (bounds $ pibindList fname indents) >>= (\bn : WithBounds PiBindList =>
-- let t = traverse asOpInfo bn
-- in the (Rule PTerm) $ case traverse asOpInfo bn of
-- Nothing => explicitPi bn fname indents
-- Just opInfo => autobindOp opInfo pdef fname indents
-- <|> explicitPi bn fname indents)
autoImplicitPi : OriginDesc -> IndentInfo -> Rule PTerm
autoImplicitPi fname indents
= do b <- bounds $ do
@ -1020,8 +1051,7 @@ mutual
<|> defaultImplicitPi fname indents
<|> forall_ fname indents
<|> implicitPi fname indents
<|> autobindOp pdef fname indents
<|> explicitPi fname indents
<|> piOrAutobind fname indents
<|> lam fname indents
typeExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm

View File

@ -333,15 +333,15 @@ mutual
prettyPrec d (PImplicit _) = "_"
prettyPrec d (PInfer _) = annotate Hole $ "?"
prettyPrec d (POp _ _ (BindType nm left) op right) =
group $ parens (prettyBinder nm <++> ":" <++> pretty left)
group $ parens (prettyBinder nm.val <++> ":" <++> pretty left)
<++> prettyOp op
<++> pretty right
prettyPrec d (POp _ _ (BindExpr nm left) op right) =
group $ parens (prettyBinder nm <++> ":=" <++> pretty left)
group $ parens (prettyBinder nm.val <++> ":=" <++> pretty left)
<++> prettyOp op
<++> pretty right
prettyPrec d (POp _ _ (BindExplicitType nm ty left) op right) =
group $ parens (prettyBinder nm <++> ":" <++> pretty ty <++> ":=" <++> pretty left)
group $ parens (prettyBinder nm.val <++> ":" <++> pretty ty <++> ":=" <++> pretty left)
<++> prettyOp op
<++> pretty right
prettyPrec d (POp _ _ (NoBinder x) op y) =

View File

@ -1,6 +1,6 @@
1/1: Building StrError1 (StrError1.idr)
Error: Couldn't parse any alternatives:
1: Expected '('.
1: Expected 'case', 'if', 'do', application or operator expression.
StrError1:2:24--2:25
1 | foo : String
@ -9,7 +9,7 @@ StrError1:2:24--2:25
... (42 others)
1/1: Building StrError2 (StrError2.idr)
Error: Couldn't parse any alternatives:
1: Expected '('.
1: Expected 'case', 'if', 'do', application or operator expression.
StrError2:2:19--2:20
1 | foo : String
@ -18,7 +18,7 @@ StrError2:2:19--2:20
... (31 others)
1/1: Building StrError3 (StrError3.idr)
Error: Couldn't parse any alternatives:
1: Expected '('.
1: Expected 'case', 'if', 'do', application or operator expression.
StrError3:2:7--2:8
1 | foo : String

View File

@ -1,6 +1,6 @@
1/1: Building ParseList (ParseList.idr)
Error: Couldn't parse any alternatives:
1: Expected '('.
1: Expected 'case', 'if', 'do', application or operator expression.
ParseList:8:5--8:6
4 | , 1

View File

@ -0,0 +1,10 @@
1/1: Building CongErr (CongErr.idr)
Error: While processing right hand side of fsprf. Can't solve constraint between: ?_ x and FS x.
CongErr:4:11--4:19
1 | import Data.Fin
2 |
3 | fsprf : x === y -> Fin.FS x = FS y
4 | fsprf p = cong _ p
^^^^^^^^

View File

@ -0,0 +1,21 @@
1/1: Building Issue735 (Issue735.idr)
Error: While processing left hand side of isCons. Can't match on (::) (Under-applied constructor).
Issue735:5:8--5:12
1 | module Issue735
2 |
3 | -- Not allowed to pattern-match on under-applied constructors
4 | isCons : (a -> List a -> List a) -> Bool
5 | isCons (::) = True
^^^^
Error: While processing left hand side of test. Can't match on A (Under-applied constructor).
Issue735:12:6--12:7
08 | interface A a where
09 |
10 | -- Not allowed to pattern-match on under-applied type constructors
11 | test : (kind : Type -> Type) -> Maybe Nat
12 | test A = Just 1
^

View File

@ -1,9 +1,9 @@
Main> Couldn't parse any alternatives:
1: Expected name.
1: Expected 'case', 'if', 'do', application or operator expression.
(Interactive):1:5--1:6
(Interactive):1:4--1:5
1 | :t (3 : Nat)
^
^
... (54 others)
Main> Expected string begin.