From 1d7d07a667196d2c562b6194d7792e51c21634a6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Wed, 18 Oct 2023 12:06:36 +0100 Subject: [PATCH 01/57] add test file --- src/Idris/Doc/Keywords.idr | 25 ++++++++++++++++++-- src/Parser/Lexer/Source.idr | 2 +- tests/Main.idr | 4 ++++ tests/idris2/operators/operators001/Test.idr | 6 +++++ tests/idris2/operators/operators001/expected | 1 + tests/idris2/operators/operators001/run | 3 +++ 6 files changed, 38 insertions(+), 3 deletions(-) create mode 100644 tests/idris2/operators/operators001/Test.idr create mode 100644 tests/idris2/operators/operators001/expected create mode 100755 tests/idris2/operators/operators001/run diff --git a/src/Idris/Doc/Keywords.idr b/src/Idris/Doc/Keywords.idr index 65084e008..1544b5ab4 100644 --- a/src/Idris/Doc/Keywords.idr +++ b/src/Idris/Doc/Keywords.idr @@ -458,6 +458,27 @@ namespaceblock = vcat $ You can use `export` or `public export` to control whether a function declared in a namespace is available outside of it. """] +autobindDoc : Doc IdrisDocAnn +autobindDoc = """ + Autobind + + `autobind` is a modifier for operator precedence and fixity declaration. + It tells the parser that this operator behaves like a binding operator, + allowing you to give a name to the left-hand-side of the operator and + use it on the right-hand-side. + + A typical example of an autobind operator is `(**)` the type constructor + for dependent pairs. It is used like this: `(x : Nat ** Vect x String)` + + If you declare a new operator to be autobind you can use it the same + way. + + Start by defining `autobind infixr 1 =@`, and then you can use it + like so: `(n : Nat =@ f n)`. + + `autobind` only applies to `infixr` operators and cannot be used as + operator sections. + """ rewriteeq : Doc IdrisDocAnn rewriteeq = vcat $ @@ -487,8 +508,7 @@ withabstraction = vcat $ If we additionally need to remember that the link between the patterns and the intermediate computation we can use the `proof` keyword to retain an equality proof. - """, "", - """ + In the following example we want to implement a `filter` function that not only returns values that satisfy the input predicate but also proofs that they do. The `with (p x)` construct introduces a value of type `Bool` @@ -564,6 +584,7 @@ keywordsDoc = :: "else" ::= ifthenelse :: "forall" ::= forallquantifier :: "rewrite" ::= rewriteeq + :: "autobind" ::= autobindDoc :: "using" ::= "" :: "interface" ::= interfacemechanism :: "implementation" ::= interfacemechanism diff --git a/src/Parser/Lexer/Source.idr b/src/Parser/Lexer/Source.idr index bc4c3d60c..3bfaa700e 100644 --- a/src/Parser/Lexer/Source.idr +++ b/src/Parser/Lexer/Source.idr @@ -245,7 +245,7 @@ keywords : List String keywords = ["data", "module", "where", "let", "in", "do", "record", "auto", "default", "implicit", "failing", "mutual", "namespace", "parameters", "with", "proof", "impossible", "case", "of", - "if", "then", "else", "forall", "rewrite", + "if", "then", "else", "forall", "rewrite", "autobind", "using", "interface", "implementation", "open", "import", "public", "export", "private"] ++ fixityKeywords ++ diff --git a/tests/Main.idr b/tests/Main.idr index 0ee17265e..52016ca93 100644 --- a/tests/Main.idr +++ b/tests/Main.idr @@ -112,6 +112,9 @@ idrisTestsReflection = testsInDir "idris2/reflection" "Quotation and Reflection" idrisTestsWith : IO TestPool idrisTestsWith = testsInDir "idris2/with" "With abstraction" +idrisTestsOperators : IO TestPool +idrisTestsOperators = testsInDir "idris2/operators" "Operator and fixities" + idrisTestsIPKG : IO TestPool idrisTestsIPKG = testsInDir "idris2/pkg" "Package and .ipkg files" @@ -216,6 +219,7 @@ main = runner $ , !idrisTestsSchemeEval , !idrisTestsReflection , !idrisTestsWith + , !idrisTestsOperators , !idrisTestsDebug , !idrisTestsIPKG , testPaths "idris2/misc" idrisTests diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr new file mode 100644 index 000000000..c113aa3d8 --- /dev/null +++ b/tests/idris2/operators/operators001/Test.idr @@ -0,0 +1,6 @@ + +autobind infixr 0 =@ + +(=@) : (a : Type) -> (a -> Type) -> Type +(=@) a f = (1 x : a) -> f x + diff --git a/tests/idris2/operators/operators001/expected b/tests/idris2/operators/operators001/expected new file mode 100644 index 000000000..31e5db7f3 --- /dev/null +++ b/tests/idris2/operators/operators001/expected @@ -0,0 +1 @@ +1/1: Building Test (Test.idr) diff --git a/tests/idris2/operators/operators001/run b/tests/idris2/operators/operators001/run new file mode 100755 index 000000000..d35387bb3 --- /dev/null +++ b/tests/idris2/operators/operators001/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr From b8d81b768e9bfbaace919cfc92235a7c612d4315 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Wed, 18 Oct 2023 13:04:42 +0100 Subject: [PATCH 02/57] add autobind keyword --- src/Core/Binary.idr | 2 +- src/Idris/Desugar.idr | 4 ++-- src/Idris/Desugar/Mutual.idr | 4 ++-- src/Idris/Parser.idr | 3 ++- src/Idris/Syntax.idr | 14 ++++++++------ src/Idris/Syntax/TTC.idr | 4 +++- src/Idris/Syntax/Traversals.idr | 4 ++-- tests/idris2/operators/operators001/Test.idr | 2 +- 8 files changed, 21 insertions(+), 16 deletions(-) diff --git a/src/Core/Binary.idr b/src/Core/Binary.idr index bd9a4e4ce..6f985cde4 100644 --- a/src/Core/Binary.idr +++ b/src/Core/Binary.idr @@ -29,7 +29,7 @@ import public Libraries.Utils.Binary ||| version number if you're changing the version more than once in the same day. export ttcVersion : Int -ttcVersion = 2023_09_08_00 +ttcVersion = 2023_11_12_00 export checkTTCVersion : String -> Int -> Int -> Core () diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index f9196dd34..d512552ab 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -1109,7 +1109,7 @@ mutual NS ns (DN str (MN ("__mk" ++ str) 0)) mkConName n = DN (show n) (MN ("__mk" ++ show n) 0) - desugarDecl ps (PFixity fc vis fix prec opName) + desugarDecl ps (PFixity fc vis binding fix prec opName) = do ctx <- get Ctxt -- We update the context of fixities by adding a namespaced fixity -- given by the current namespace and its fixity name. @@ -1118,7 +1118,7 @@ mutual let updatedNS = NS (mkNestedNamespace (Just ctx.currentNS) (show fix)) (UN $ Basic $ nameRoot opName) - update Syn { fixities $= addName updatedNS (MkFixityInfo fc vis fix prec) } + update Syn { fixities $= addName updatedNS (MkFixityInfo fc vis binding fix prec) } pure [] desugarDecl ps d@(PFail fc mmsg ds) = do -- save the state: the content of a failing block should be discarded diff --git a/src/Idris/Desugar/Mutual.idr b/src/Idris/Desugar/Mutual.idr index b05449cee..c13ca8c41 100644 --- a/src/Idris/Desugar/Mutual.idr +++ b/src/Idris/Desugar/Mutual.idr @@ -29,7 +29,7 @@ getDecl AsType (PRecord fc doc vis mbtot (MkPRecord n ps _ _ _)) mkRecType : List (Name, RigCount, PiInfo PTerm, PTerm) -> PTerm mkRecType [] = PType fc mkRecType ((n, c, p, t) :: ts) = PPi fc c p (Just n) t (mkRecType ts) -getDecl AsType d@(PFixity _ _ _ _ _) = Just d +getDecl AsType d@(PFixity _ _ _ _ _ _) = Just d getDecl AsType d@(PDirective _ _) = Just d getDecl AsType d = Nothing @@ -37,7 +37,7 @@ getDecl AsDef (PClaim _ _ _ _ _) = Nothing getDecl AsDef d@(PData _ _ _ _ (MkPLater _ _ _)) = Just d getDecl AsDef (PInterface _ _ _ _ _ _ _ _ _) = Nothing getDecl AsDef d@(PRecord _ _ _ _ (MkPRecordLater _ _)) = Just d -getDecl AsDef (PFixity _ _ _ _ _) = Nothing +getDecl AsDef (PFixity _ _ _ _ _ _) = Nothing getDecl AsDef (PDirective _ _) = Nothing getDecl AsDef d = Just d diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 545da2e70..8d47e5890 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -1828,13 +1828,14 @@ definition fname indents fixDecl : OriginDesc -> IndentInfo -> Rule (List PDecl) fixDecl fname indents = do vis <- exportVisibility fname + binding <- map isJust (optional (keyword "autobind")) b <- bounds (do fixity <- decorate fname Keyword $ fix commit prec <- decorate fname Keyword $ intLit ops <- sepBy1 (decoratedSymbol fname ",") iOperator pure (fixity, prec, ops)) (fixity, prec, ops) <- pure b.val - pure (map (PFixity (boundToFC fname b) vis fixity (fromInteger prec)) (forget ops)) + pure (map (PFixity (boundToFC fname b) vis binding fixity (fromInteger prec)) (forget ops)) directiveDecl : OriginDesc -> IndentInfo -> Rule PDecl directiveDecl fname indents diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index ed365477c..6a77dc490 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -53,17 +53,19 @@ record FixityInfo where constructor MkFixityInfo fc : FC vis : Visibility + isAutobind : Bool fix : Fixity precedence : Nat export Show FixityInfo where - show fx = "fc: \{show fx.fc}, visibility: \{show fx.vis}, fixity: \{show fx.fix}, precedence: \{show fx.precedence}" + show fx = "fc: \{show fx.fc}, visibility: \{show fx.vis}, isAutobind: \{show fx.isAutobind}, fixity: \{show fx.fix}, precedence: \{show fx.precedence}" export Eq FixityInfo where x == y = x.fc == y.fc && x.vis == y.vis + && x.isAutobind == y.isAutobind && x.fix == y.fix && x.precedence == y.precedence @@ -470,7 +472,7 @@ mutual PFail : FC -> Maybe String -> List (PDecl' nm) -> PDecl' nm PMutual : FC -> List (PDecl' nm) -> PDecl' nm - PFixity : FC -> Visibility -> Fixity -> Nat -> OpStr -> PDecl' nm + PFixity : FC -> Visibility -> (isAutobind : Bool) -> Fixity -> Nat -> OpStr -> PDecl' nm PNamespace : FC -> Namespace -> List (PDecl' nm) -> PDecl' nm PTransform : FC -> String -> PTerm' nm -> PTerm' nm -> PDecl' nm PRunElabDecl : FC -> PTerm' nm -> PDecl' nm @@ -489,7 +491,7 @@ mutual getPDeclLoc (PRecord fc _ _ _ _) = fc getPDeclLoc (PMutual fc _) = fc getPDeclLoc (PFail fc _ _) = fc - getPDeclLoc (PFixity fc _ _ _ _) = fc + getPDeclLoc (PFixity fc _ _ _ _ _) = fc getPDeclLoc (PNamespace fc _ _) = fc getPDeclLoc (PTransform fc _ _ _) = fc getPDeclLoc (PRunElabDecl fc _) = fc @@ -999,9 +1001,9 @@ initSyntax initFixities : ANameMap FixityInfo initFixities = fromList - [ (UN $ Basic "-", MkFixityInfo EmptyFC Export Prefix 10) - , (UN $ Basic "negate", MkFixityInfo EmptyFC Export Prefix 10) -- for documentation purposes - , (UN $ Basic "=", MkFixityInfo EmptyFC Export Infix 0) + [ (UN $ Basic "-", MkFixityInfo EmptyFC Export False Prefix 10) + , (UN $ Basic "negate", MkFixityInfo EmptyFC Export False Prefix 10) -- for documentation purposes + , (UN $ Basic "=", MkFixityInfo EmptyFC Export False Infix 0) ] initDocStrings : ANameMap String diff --git a/src/Idris/Syntax/TTC.idr b/src/Idris/Syntax/TTC.idr index 0ae8a2af3..47e1b63ad 100644 --- a/src/Idris/Syntax/TTC.idr +++ b/src/Idris/Syntax/TTC.idr @@ -86,14 +86,16 @@ TTC FixityInfo where toBuf b fx = do toBuf b fx.fc toBuf b fx.vis + toBuf b fx.isAutobind toBuf b fx.fix toBuf b fx.precedence fromBuf b = do fc <- fromBuf b vis <- fromBuf b + isAutobind <- fromBuf b fix <- fromBuf b prec <- fromBuf b - pure $ MkFixityInfo fc vis fix prec + pure $ MkFixityInfo fc vis isAutobind fix prec export diff --git a/src/Idris/Syntax/Traversals.idr b/src/Idris/Syntax/Traversals.idr index 9db5d8fa2..fd390ceea 100644 --- a/src/Idris/Syntax/Traversals.idr +++ b/src/Idris/Syntax/Traversals.idr @@ -265,7 +265,7 @@ mapPTermM f = goPTerm where pure $ PRecord fc doc v tot (MkPRecordLater n !(go4TupledPTerms nts)) goPDecl (PFail fc msg ps) = PFail fc msg <$> goPDecls ps goPDecl (PMutual fc ps) = PMutual fc <$> goPDecls ps - goPDecl p@(PFixity _ _ _ _ _) = pure p + goPDecl p@(PFixity _ _ _ _ _ _) = pure p goPDecl (PNamespace fc strs ps) = PNamespace fc strs <$> goPDecls ps goPDecl (PTransform fc n a b) = PTransform fc n <$> goPTerm a <*> goPTerm b goPDecl (PRunElabDecl fc a) = PRunElabDecl fc <$> goPTerm a @@ -533,7 +533,7 @@ mapPTerm f = goPTerm where = PRecord fc doc v tot (MkPRecordLater n (go4TupledPTerms nts)) goPDecl (PFail fc msg ps) = PFail fc msg $ goPDecl <$> ps goPDecl (PMutual fc ps) = PMutual fc $ goPDecl <$> ps - goPDecl p@(PFixity _ _ _ _ _) = p + goPDecl p@(PFixity _ _ _ _ _ _) = p goPDecl (PNamespace fc strs ps) = PNamespace fc strs $ goPDecl <$> ps goPDecl (PTransform fc n a b) = PTransform fc n (goPTerm a) (goPTerm b) goPDecl (PRunElabDecl fc a) = PRunElabDecl fc $ goPTerm a diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index c113aa3d8..053dd0f1e 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -1,6 +1,6 @@ autobind infixr 0 =@ -(=@) : (a : Type) -> (a -> Type) -> Type +0 (=@) : (a : Type) -> (a -> Type) -> Type (=@) a f = (1 x : a) -> f x From 0ed65eb5870adbbf5d91ddb268f2742b96636b5c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 21 Oct 2023 14:28:17 +0100 Subject: [PATCH 03/57] desugar autobind properly --- src/Idris/Desugar.idr | 18 ++++++--- src/Idris/Parser.idr | 41 +++++++++++++++++--- src/Idris/Pretty.idr | 6 ++- src/Idris/Resugar.idr | 6 +-- src/Idris/Syntax.idr | 11 ++++-- src/Idris/Syntax/Traversals.idr | 15 ++++--- src/Idris/Syntax/Views.idr | 3 +- tests/idris2/operators/operators001/Test.idr | 8 ++++ 8 files changed, 84 insertions(+), 24 deletions(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index d512552ab..040674b17 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -160,7 +160,7 @@ checkConflictingFixities isPrefix exprFC opn toTokList : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> PTerm -> Core (List (Tok OpStr PTerm)) -toTokList (POp fc opFC opn l r) +toTokList (POp fc opFC _ opn l r) = do precInfo <- checkConflictingFixities False fc opn rtoks <- toTokList r pure (Expr l :: Op fc opFC opn precInfo :: rtoks) @@ -308,8 +308,16 @@ mutual [apply (IVar fc (UN $ Basic "===")) [l', r'], apply (IVar fc (UN $ Basic "~=~")) [l', r']] desugarB side ps (PBracketed fc e) = desugarB side ps e - desugarB side ps (POp fc opFC op l r) - = do ts <- toTokList (POp fc opFC op l r) + -- desugarB side ps (PDPair fc opFC (PRef namefc n@(UN _)) ty r) + -- = do ty' <- desugarB side ps ty + -- r' <- desugarB side ps r + -- pure $ apply (IVar opFC dpairname) + -- [ty', ILam namefc top Explicit (Just n) ty' r'] + desugarB side ps (POp fc opFC (Just nm) op l r) + = desugarB side ps (POp fc opFC Nothing op l + $ PLam (virtualiseFC opFC) top Explicit nm (PImplicit opFC) r) + desugarB side ps (POp fc opFC Nothing op l r) + = do ts <- toTokList (POp fc opFC Nothing op l r) desugarTree side ps !(parseOps ts) desugarB side ps (PPrefixOp fc opFC op arg) = do ts <- toTokList (PPrefixOp fc opFC op arg) @@ -322,12 +330,12 @@ mutual [] => desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) - (POp fc opFC op (PRef fc (MN "arg" 0)) arg)) + (POp fc opFC Nothing op (PRef fc (MN "arg" 0)) arg)) (prec :: _) => desugarB side ps (PPrefixOp fc opFC op arg) desugarB side ps (PSectionR fc opFC arg op) = desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) - (POp fc opFC op arg (PRef fc (MN "arg" 0)))) + (POp fc opFC Nothing op arg (PRef fc (MN "arg" 0)))) desugarB side ps (PSearch fc depth) = pure $ ISearch fc depth desugarB side ps (PPrimVal fc (BI x)) = case !fromIntegerName of diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 8d47e5890..8ea4b1c00 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -325,15 +325,41 @@ mutual decoratedSymbol fname "]" pure (map (\ n => (boundToFC fname n, n.val)) $ forget ns) - opExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm - opExpr q fname indents + opBinder : OriginDesc -> IndentInfo -> Rule (PTerm, PTerm) + opBinder fname indents + = do loc <- bounds (do x <- UN . Basic <$> decoratedSimpleBinderName fname + decoratedSymbol fname ":" + ty <- typeExpr pdef fname indents + pure (x, ty)) + (nm, ty) <- pure loc.val + pure (PRef (boundToFC fname loc) nm, ty) + + + autobindOp : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm + autobindOp q fname indents + = do bindRes <- bounds $ parens fname (opBinder fname indents) + (bindName, bindExpr) <- pure $ bindRes.val + continue indents + op <- bounds iOperator + e <- bounds (expr q fname indents) + pure (POp (boundToFC fname $ mergeBounds bindRes e) + (boundToFC fname op) + (Just $ bindName) + op.val + bindExpr + e.val) + + opExprBase : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm + opExprBase q fname indents = do l <- bounds (appExpr q fname indents) (if eqOK q - then do r <- bounds (continue indents *> decoratedSymbol fname "=" *> opExpr q fname indents) + then do r <- bounds (continue indents + *> decoratedSymbol fname "=" + *> opExprBase q fname indents) pure $ let fc = boundToFC fname (mergeBounds l r) opFC = virtualiseFC fc -- already been highlighted: we don't care - in POp fc opFC (UN $ Basic "=") l.val r.val + in POp fc opFC Nothing (UN $ Basic "=") l.val r.val else fail "= not allowed") <|> (do b <- bounds $ do @@ -346,9 +372,13 @@ mutual (op, r) <- pure b.val let fc = boundToFC fname (mergeBounds l b) let opFC = boundToFC fname op - pure (POp fc opFC op.val l.val r)) + pure (POp fc opFC Nothing op.val l.val r)) <|> pure l.val + opExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm + opExpr q fname indents = opExprBase q fname indents + <|> autobindOp q fname indents + dpairType : OriginDesc -> WithBounds t -> IndentInfo -> Rule PTerm dpairType fname start indents = do loc <- bounds (do x <- UN . Basic <$> decoratedSimpleBinderName fname @@ -981,6 +1011,7 @@ mutual <|> defaultImplicitPi fname indents <|> forall_ fname indents <|> implicitPi fname indents + <|> autobindOp pdef fname indents <|> explicitPi fname indents <|> lam fname indents diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index 6c4cffa8f..296579516 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -332,7 +332,11 @@ mutual prettyPrec d (PDotted _ p) = dot <+> prettyPrec d p prettyPrec d (PImplicit _) = "_" prettyPrec d (PInfer _) = annotate Hole $ "?" - prettyPrec d (POp _ _ op x y) = + prettyPrec d (POp _ _ (Just ab) op left right) = + group $ parens (pretty ab <++> ":" <++> pretty left) + <++> prettyOp op + <++> pretty right + prettyPrec d (POp _ _ Nothing op x y) = parenthesise (d >= App) $ group $ pretty x <++> prettyOp op diff --git a/src/Idris/Resugar.idr b/src/Idris/Resugar.idr index 47b02b54b..b2e87cc61 100644 --- a/src/Idris/Resugar.idr +++ b/src/Idris/Resugar.idr @@ -46,7 +46,7 @@ mkOp tm@(PApp fc (PApp _ (PRef opFC kn) x) y) -- to know if the name is an operator or not, it's enough to check -- that the fixity context contains the name `(++)` let rootName = UN (Basic (nameRoot raw)) - let asOp = POp fc opFC kn (unbracketApp x) (unbracketApp y) + let asOp = POp fc opFC Nothing kn (unbracketApp x) (unbracketApp y) if not (null (lookupName rootName (infixes syn))) then pure asOp else case dropNS raw of @@ -590,8 +590,8 @@ cleanPTerm ptm cleanNode : IPTerm -> Core IPTerm cleanNode (PRef fc nm) = PRef fc <$> cleanKindedName nm - cleanNode (POp fc opFC op x y) = - (\ op => POp fc opFC op x y) <$> cleanKindedName op + cleanNode (POp fc opFC abi op x y) = + (\ op => POp fc opFC abi op x y) <$> cleanKindedName op cleanNode (PPrefixOp fc opFC op x) = (\ op => PPrefixOp fc opFC op x) <$> cleanKindedName op cleanNode (PSectionL fc opFC op x) = diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index 6a77dc490..763d4dc24 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -138,7 +138,10 @@ mutual -- Operators - POp : (full, opFC : FC) -> OpStr' nm -> PTerm' nm -> PTerm' nm -> PTerm' nm + POp : (full, opFC : FC) -> + -- If the operator is autobind, we expect a pterm for the bound name + (isAutobind : Maybe (PTerm' nm)) -> + OpStr' nm -> PTerm' nm -> PTerm' nm -> PTerm' nm PPrefixOp : (full, opFC : FC) -> OpStr' nm -> PTerm' nm -> PTerm' nm PSectionL : (full, opFC : FC) -> OpStr' nm -> PTerm' nm -> PTerm' nm PSectionR : (full, opFC : FC) -> PTerm' nm -> OpStr' nm -> PTerm' nm @@ -204,7 +207,7 @@ mutual getPTermLoc (PDotted fc _) = fc getPTermLoc (PImplicit fc) = fc getPTermLoc (PInfer fc) = fc - getPTermLoc (POp fc _ _ _ _) = fc + getPTermLoc (POp fc _ _ _ _ _) = fc getPTermLoc (PPrefixOp fc _ _ _) = fc getPTermLoc (PSectionL fc _ _ _) = fc getPTermLoc (PSectionR fc _ _ _) = fc @@ -808,7 +811,9 @@ parameters {0 nm : Type} (toName : nm -> Name) showPTermPrec d (PDotted _ p) = "." ++ showPTermPrec d p showPTermPrec _ (PImplicit _) = "_" showPTermPrec _ (PInfer _) = "?" - showPTermPrec d (POp _ _ op x y) = showPTermPrec d x ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d y + showPTermPrec d (POp _ _ Nothing op x y) = showPTermPrec d x ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d y + showPTermPrec d (POp _ _ (Just nm) op x y) + = "(" ++ showPTermPrec d nm ++ " : " ++ showPTermPrec d x ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d y ++ ")" showPTermPrec d (PPrefixOp _ _ op x) = showOpPrec d op ++ showPTermPrec d x showPTermPrec d (PSectionL _ _ op x) = "(" ++ showOpPrec d op ++ " " ++ showPTermPrec d x ++ ")" showPTermPrec d (PSectionR _ _ x op) = "(" ++ showPTermPrec d x ++ " " ++ showOpPrec d op ++ ")" diff --git a/src/Idris/Syntax/Traversals.idr b/src/Idris/Syntax/Traversals.idr index fd390ceea..f0202f54e 100644 --- a/src/Idris/Syntax/Traversals.idr +++ b/src/Idris/Syntax/Traversals.idr @@ -99,9 +99,12 @@ mapPTermM f = goPTerm where >>= f goPTerm t@(PImplicit _) = f t goPTerm t@(PInfer _) = f t - goPTerm (POp fc opFC x y z) = - POp fc opFC x <$> goPTerm y - <*> goPTerm z + goPTerm (POp fc opFC autobindInfo x y z) = + POp fc opFC + <$> traverseOpt f autobindInfo + <*> pure x + <*> goPTerm y + <*> goPTerm z >>= f goPTerm (PPrefixOp fc opFC x y) = PPrefixOp fc opFC x <$> goPTerm y @@ -434,8 +437,8 @@ mapPTerm f = goPTerm where = f $ PDotted fc $ goPTerm x goPTerm t@(PImplicit _) = f t goPTerm t@(PInfer _) = f t - goPTerm (POp fc opFC x y z) - = f $ POp fc opFC x (goPTerm y) (goPTerm z) + goPTerm (POp fc opFC autoBindInfo opName y z) + = f $ POp fc opFC (map f autoBindInfo) opName (goPTerm y) (goPTerm z) goPTerm (PPrefixOp fc opFC x y) = f $ PPrefixOp fc opFC x $ goPTerm y goPTerm (PSectionL fc opFC x y) @@ -618,7 +621,7 @@ substFC fc = mapPTerm $ \case PDotted _ x => PDotted fc x PImplicit _ => PImplicit fc PInfer _ => PInfer fc - POp _ _ x y z => POp fc fc x y z + POp _ _ ab nm l r => POp fc fc ab nm l r PPrefixOp _ _ x y => PPrefixOp fc fc x y PSectionL _ _ x y => PSectionL fc fc x y PSectionR _ _ x y => PSectionR fc fc x y diff --git a/src/Idris/Syntax/Views.idr b/src/Idris/Syntax/Views.idr index d9984aeaf..73bd5a710 100644 --- a/src/Idris/Syntax/Views.idr +++ b/src/Idris/Syntax/Views.idr @@ -26,7 +26,8 @@ getFnArgs embed fts = go fts [] where go (PAutoApp fc f t) = go f . (Auto fc t ::) go (PNamedApp fc f n t) = go f . (Named fc n t ::) go (PBracketed fc f) = go f - go (POp fc opFC op l r) = (PRef opFC op,) . (Explicit fc l ::) . (Explicit fc r ::) + -- we don't care about the binder info here + go (POp fc opFC _ op l r) = (PRef opFC op,) . (Explicit fc l ::) . (Explicit fc r ::) go (PEq fc l r) = (PRef fc $ embed eqName,) . (Explicit fc l ::) . (Explicit fc r ::) -- ambiguous, picking the type constructor here go (PPair fc l r) = (PRef fc $ embed pairname,) . (Explicit fc l ::) . (Explicit fc r ::) diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index 053dd0f1e..b04c839a8 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -1,6 +1,14 @@ + autobind infixr 0 =@ 0 (=@) : (a : Type) -> (a -> Type) -> Type (=@) a f = (1 x : a) -> f x + +data S : {ty : Type} -> (x : ty) -> Type where + MkS : (x : ty) =@ S x + -- same as + -- MkS : (1 x : ty) -> S x + + From 046e08d173fd0d7f72ee00c3bc19bb68e2296895 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 21 Oct 2023 21:00:12 +0100 Subject: [PATCH 04/57] differenciate between type bind and expr bind --- src/Idris/Desugar.idr | 30 +++++++----- src/Idris/Parser.idr | 40 +++++++++------ src/Idris/Pretty.idr | 14 ++++-- src/Idris/Resugar.idr | 6 +-- src/Idris/Syntax.idr | 51 +++++++++++++++++--- src/Idris/Syntax/Traversals.idr | 33 ++++++++++--- src/Idris/Syntax/Views.idr | 2 +- tests/idris2/operators/operators001/Test.idr | 21 +++++++- 8 files changed, 146 insertions(+), 51 deletions(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 040674b17..9d19e59b2 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -160,7 +160,7 @@ checkConflictingFixities isPrefix exprFC opn toTokList : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> PTerm -> Core (List (Tok OpStr PTerm)) -toTokList (POp fc opFC _ opn l r) +toTokList (POp fc opFC (NotAutobind l) opn r) = do precInfo <- checkConflictingFixities False fc opn rtoks <- toTokList r pure (Expr l :: Op fc opFC opn precInfo :: rtoks) @@ -308,16 +308,20 @@ mutual [apply (IVar fc (UN $ Basic "===")) [l', r'], apply (IVar fc (UN $ Basic "~=~")) [l', r']] desugarB side ps (PBracketed fc e) = desugarB side ps e - -- desugarB side ps (PDPair fc opFC (PRef namefc n@(UN _)) ty r) - -- = do ty' <- desugarB side ps ty - -- r' <- desugarB side ps r - -- pure $ apply (IVar opFC dpairname) - -- [ty', ILam namefc top Explicit (Just n) ty' r'] - desugarB side ps (POp fc opFC (Just nm) op l r) - = desugarB side ps (POp fc opFC Nothing op l - $ PLam (virtualiseFC opFC) top Explicit nm (PImplicit opFC) r) - desugarB side ps (POp fc opFC Nothing op l r) - = do ts <- toTokList (POp fc opFC Nothing op l r) + -- a bindtype operator desugars (x : a ** b x) into ((**) a (\x : a => b x)) + desugarB side ps (POp fc opFC (BindType nm l) op r) + = desugarB side ps (POp fc opFC (NotAutobind l) op + $ PLam (virtualiseFC opFC) top Explicit nm l r) + -- a bindexpr operator desugars (x := a ** b x) into ((**) a (\x : ? => b x)) + desugarB side ps (POp fc opFC (BindExpr nm l) op r) + = desugarB side ps (POp fc opFC (NotAutobind l) op + $ PLam (virtualiseFC opFC) top Explicit nm (PImplicit $ virtualiseFC opFC) r) + -- an explicit bind operator desugars (x : ty := a ** b x) into ((**) a (\x : ty => b x)) + desugarB side ps (POp fc opFC (BindExplicitType nm ty l) op r) + = desugarB side ps (POp fc opFC (NotAutobind l) op + $ PLam (virtualiseFC opFC) top Explicit nm ty r) + desugarB side ps (POp fc opFC (NotAutobind l) op r) + = do ts <- toTokList (POp fc opFC (NotAutobind l) op r) desugarTree side ps !(parseOps ts) desugarB side ps (PPrefixOp fc opFC op arg) = do ts <- toTokList (PPrefixOp fc opFC op arg) @@ -330,12 +334,12 @@ mutual [] => desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) - (POp fc opFC Nothing op (PRef fc (MN "arg" 0)) arg)) + (POp fc opFC (NotAutobind (PRef fc (MN "arg" 0))) op arg)) (prec :: _) => desugarB side ps (PPrefixOp fc opFC op arg) desugarB side ps (PSectionR fc opFC arg op) = desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) - (POp fc opFC Nothing op arg (PRef fc (MN "arg" 0)))) + (POp fc opFC (NotAutobind arg) op (PRef fc (MN "arg" 0)))) desugarB side ps (PSearch fc depth) = pure $ ISearch fc depth desugarB side ps (PPrimVal fc (BI x)) = case !fromIntegerName of diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 8ea4b1c00..6ff4a78a5 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -325,28 +325,38 @@ mutual decoratedSymbol fname "]" pure (map (\ n => (boundToFC fname n, n.val)) $ forget ns) - opBinder : OriginDesc -> IndentInfo -> Rule (PTerm, PTerm) - opBinder fname indents - = do loc <- bounds (do x <- UN . Basic <$> decoratedSimpleBinderName fname - decoratedSymbol fname ":" - ty <- typeExpr pdef fname indents - pure (x, ty)) - (nm, ty) <- pure loc.val - pure (PRef (boundToFC fname loc) nm, ty) + -- The different kinds of operator bindings `x : ty` for typebind + -- x := e and x : ty := e for autobind + opBinderTypes : OriginDesc -> IndentInfo -> PTerm -> Rule (OperatorLHSInfo PTerm) + opBinderTypes fname indents boundName = + do decoratedSymbol fname ":" + ty <- typeExpr pdef fname indents + pure (BindType boundName ty) + <|> do decoratedSymbol fname ":=" + exp <- expr pdef fname indents + pure (BindExpr boundName exp) + <|> do decoratedSymbol fname ":" + ty <- typeExpr pdef fname indents + decoratedSymbol fname ":=" + exp <- expr pdef fname indents + pure (BindExplicitType boundName ty exp) + opBinder : OriginDesc -> IndentInfo -> Rule (OperatorLHSInfo PTerm) + opBinder fname indents + = do x <- bounds (UN . Basic <$> decoratedSimpleBinderName fname) + let boundName = PRef (boundToFC fname x) x.val + opBinderTypes fname indents boundName autobindOp : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm autobindOp q fname indents - = do bindRes <- bounds $ parens fname (opBinder fname indents) - (bindName, bindExpr) <- pure $ bindRes.val + = do binder <- bounds $ parens fname (opBinder fname indents) continue indents op <- bounds iOperator e <- bounds (expr q fname indents) - pure (POp (boundToFC fname $ mergeBounds bindRes e) + pure (POp (boundToFC fname $ mergeBounds binder e) (boundToFC fname op) - (Just $ bindName) + binder.val op.val - bindExpr e.val) opExprBase : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm @@ -359,7 +369,7 @@ mutual pure $ let fc = boundToFC fname (mergeBounds l r) opFC = virtualiseFC fc -- already been highlighted: we don't care - in POp fc opFC Nothing (UN $ Basic "=") l.val r.val + in POp fc opFC (NotAutobind l.val) (UN $ Basic "=") r.val else fail "= not allowed") <|> (do b <- bounds $ do @@ -372,7 +382,7 @@ mutual (op, r) <- pure b.val let fc = boundToFC fname (mergeBounds l b) let opFC = boundToFC fname op - pure (POp fc opFC Nothing op.val l.val r)) + pure (POp fc opFC (NotAutobind l.val) op.val r)) <|> pure l.val opExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index 296579516..8f299368d 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -332,11 +332,19 @@ mutual prettyPrec d (PDotted _ p) = dot <+> prettyPrec d p prettyPrec d (PImplicit _) = "_" prettyPrec d (PInfer _) = annotate Hole $ "?" - prettyPrec d (POp _ _ (Just ab) op left right) = - group $ parens (pretty ab <++> ":" <++> pretty left) + prettyPrec d (POp _ _ (BindType nm left) op right) = + group $ parens (pretty nm <++> ":" <++> pretty left) <++> prettyOp op <++> pretty right - prettyPrec d (POp _ _ Nothing op x y) = + prettyPrec d (POp _ _ (BindExpr nm left) op right) = + group $ parens (pretty nm <++> ":=" <++> pretty left) + <++> prettyOp op + <++> pretty right + prettyPrec d (POp _ _ (BindExplicitType nm ty left) op right) = + group $ parens (pretty nm <++> ":" <++> pretty ty <++> ":=" <++> pretty left) + <++> prettyOp op + <++> pretty right + prettyPrec d (POp _ _ (NotAutobind x) op y) = parenthesise (d >= App) $ group $ pretty x <++> prettyOp op diff --git a/src/Idris/Resugar.idr b/src/Idris/Resugar.idr index b2e87cc61..fdc3fe03a 100644 --- a/src/Idris/Resugar.idr +++ b/src/Idris/Resugar.idr @@ -46,7 +46,7 @@ mkOp tm@(PApp fc (PApp _ (PRef opFC kn) x) y) -- to know if the name is an operator or not, it's enough to check -- that the fixity context contains the name `(++)` let rootName = UN (Basic (nameRoot raw)) - let asOp = POp fc opFC Nothing kn (unbracketApp x) (unbracketApp y) + let asOp = POp fc opFC (NotAutobind (unbracketApp x)) kn (unbracketApp y) if not (null (lookupName rootName (infixes syn))) then pure asOp else case dropNS raw of @@ -590,8 +590,8 @@ cleanPTerm ptm cleanNode : IPTerm -> Core IPTerm cleanNode (PRef fc nm) = PRef fc <$> cleanKindedName nm - cleanNode (POp fc opFC abi op x y) = - (\ op => POp fc opFC abi op x y) <$> cleanKindedName op + cleanNode (POp fc opFC abi op y) = + (\ op => POp fc opFC abi op y) <$> cleanKindedName op cleanNode (PPrefixOp fc opFC op x) = (\ op => PPrefixOp fc opFC op x) <$> cleanKindedName op cleanNode (PSectionL fc opFC op x) = diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index 763d4dc24..686ed9d52 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -81,6 +81,39 @@ public export data HidingDirective = HideName Name | HideFixity Fixity Name +-- Left-hand-side information for operators, carries autobind information +-- an operator can either be +-- - not autobind, a regular operator +-- - binding types, such that (nm : ty ** fn nm) desugars into (ty ** \(nm : ty) => fn nm) +-- - binding expressing with an inferred type such that +-- (nm := exp ** fn nm) desugars into (exp ** \(nm : ?) => fn nm) +-- - binding both types and expression such that +-- (nm : ty := exp ** fn nm) desugars into (exp ** \(nm : ty) => fn nm) +public export +data OperatorLHSInfo : tm -> Type where + -- Traditional operator wihtout binding, carries the lhs + NotAutobind : (lhs : tm) -> OperatorLHSInfo tm + -- (nm : ty) ** fn x + BindType : (name, ty : tm) -> OperatorLHSInfo tm + -- (nm := exp) ** fn nm + BindExpr : (name, expr : tm) -> OperatorLHSInfo tm + -- (nm : ty := exp) ** fn nm + BindExplicitType : (name, type, expr : tm) -> OperatorLHSInfo tm + +export +Functor OperatorLHSInfo where + map f (NotAutobind lhs) = NotAutobind $ f lhs + map f (BindType nm lhs) = BindType (f nm) (f lhs) + map f (BindExpr nm lhs) = BindExpr (f nm) (f lhs) + map f (BindExplicitType nm ty lhs) = BindExplicitType (f nm) (f ty) (f lhs) + +export +(.getLhs) : OperatorLHSInfo tm -> tm +(.getLhs) (NotAutobind lhs) = lhs +(.getLhs) (BindExpr _ lhs) = lhs +(.getLhs) (BindType _ lhs) = lhs +(.getLhs) (BindExplicitType _ _ lhs) = lhs + mutual ||| Source language as produced by the parser @@ -139,9 +172,8 @@ mutual -- Operators POp : (full, opFC : FC) -> - -- If the operator is autobind, we expect a pterm for the bound name - (isAutobind : Maybe (PTerm' nm)) -> - OpStr' nm -> PTerm' nm -> PTerm' nm -> PTerm' nm + (lhsInfo : OperatorLHSInfo (PTerm' nm)) -> + OpStr' nm -> (rhs : PTerm' nm) -> PTerm' nm PPrefixOp : (full, opFC : FC) -> OpStr' nm -> PTerm' nm -> PTerm' nm PSectionL : (full, opFC : FC) -> OpStr' nm -> PTerm' nm -> PTerm' nm PSectionR : (full, opFC : FC) -> PTerm' nm -> OpStr' nm -> PTerm' nm @@ -207,7 +239,7 @@ mutual getPTermLoc (PDotted fc _) = fc getPTermLoc (PImplicit fc) = fc getPTermLoc (PInfer fc) = fc - getPTermLoc (POp fc _ _ _ _ _) = fc + getPTermLoc (POp fc _ _ _ _) = fc getPTermLoc (PPrefixOp fc _ _ _) = fc getPTermLoc (PSectionL fc _ _ _) = fc getPTermLoc (PSectionR fc _ _ _) = fc @@ -811,9 +843,14 @@ parameters {0 nm : Type} (toName : nm -> Name) showPTermPrec d (PDotted _ p) = "." ++ showPTermPrec d p showPTermPrec _ (PImplicit _) = "_" showPTermPrec _ (PInfer _) = "?" - showPTermPrec d (POp _ _ Nothing op x y) = showPTermPrec d x ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d y - showPTermPrec d (POp _ _ (Just nm) op x y) - = "(" ++ showPTermPrec d nm ++ " : " ++ showPTermPrec d x ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d y ++ ")" + showPTermPrec d (POp _ _ (NotAutobind left) op right) + = showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right + showPTermPrec d (POp _ _ (BindType nm left) op right) + = "(" ++ showPTermPrec d nm ++ " : " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + showPTermPrec d (POp _ _ (BindExpr nm left) op right) + = "(" ++ showPTermPrec d nm ++ " := " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + showPTermPrec d (POp _ _ (BindExplicitType nm ty left) op right) + = "(" ++ showPTermPrec d nm ++ " : " ++ showPTermPrec d ty ++ ":=" ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (PPrefixOp _ _ op x) = showOpPrec d op ++ showPTermPrec d x showPTermPrec d (PSectionL _ _ op x) = "(" ++ showOpPrec d op ++ " " ++ showPTermPrec d x ++ ")" showPTermPrec d (PSectionR _ _ x op) = "(" ++ showPTermPrec d x ++ " " ++ showOpPrec d op ++ ")" diff --git a/src/Idris/Syntax/Traversals.idr b/src/Idris/Syntax/Traversals.idr index f0202f54e..1e0edfefe 100644 --- a/src/Idris/Syntax/Traversals.idr +++ b/src/Idris/Syntax/Traversals.idr @@ -99,12 +99,29 @@ mapPTermM f = goPTerm where >>= f goPTerm t@(PImplicit _) = f t goPTerm t@(PInfer _) = f t - goPTerm (POp fc opFC autobindInfo x y z) = + goPTerm (POp fc opFC (NotAutobind left) op right) = POp fc opFC - <$> traverseOpt f autobindInfo - <*> pure x - <*> goPTerm y - <*> goPTerm z + <$> (NotAutobind <$> goPTerm left) + <*> pure op + <*> goPTerm right + >>= f + goPTerm (POp fc opFC (BindType nm left) op right) = + POp fc opFC + <$> (BindType nm <$> goPTerm left) + <*> pure op + <*> goPTerm right + >>= f + goPTerm (POp fc opFC (BindExpr nm left) op right) = + POp fc opFC + <$> (BindExpr nm <$> goPTerm left) + <*> pure op + <*> goPTerm right + >>= f + goPTerm (POp fc opFC (BindExplicitType nm ty left) op right) = + POp fc opFC + <$> (BindExplicitType nm <$> goPTerm ty <*> goPTerm left) + <*> pure op + <*> goPTerm right >>= f goPTerm (PPrefixOp fc opFC x y) = PPrefixOp fc opFC x <$> goPTerm y @@ -437,8 +454,8 @@ mapPTerm f = goPTerm where = f $ PDotted fc $ goPTerm x goPTerm t@(PImplicit _) = f t goPTerm t@(PInfer _) = f t - goPTerm (POp fc opFC autoBindInfo opName y z) - = f $ POp fc opFC (map f autoBindInfo) opName (goPTerm y) (goPTerm z) + goPTerm (POp fc opFC autoBindInfo opName z) + = f $ POp fc opFC (map f autoBindInfo) opName (goPTerm z) goPTerm (PPrefixOp fc opFC x y) = f $ PPrefixOp fc opFC x $ goPTerm y goPTerm (PSectionL fc opFC x y) @@ -621,7 +638,7 @@ substFC fc = mapPTerm $ \case PDotted _ x => PDotted fc x PImplicit _ => PImplicit fc PInfer _ => PInfer fc - POp _ _ ab nm l r => POp fc fc ab nm l r + POp _ _ ab nm r => POp fc fc ab nm r PPrefixOp _ _ x y => PPrefixOp fc fc x y PSectionL _ _ x y => PSectionL fc fc x y PSectionR _ _ x y => PSectionR fc fc x y diff --git a/src/Idris/Syntax/Views.idr b/src/Idris/Syntax/Views.idr index 73bd5a710..34e9de748 100644 --- a/src/Idris/Syntax/Views.idr +++ b/src/Idris/Syntax/Views.idr @@ -27,7 +27,7 @@ getFnArgs embed fts = go fts [] where go (PNamedApp fc f n t) = go f . (Named fc n t ::) go (PBracketed fc f) = go f -- we don't care about the binder info here - go (POp fc opFC _ op l r) = (PRef opFC op,) . (Explicit fc l ::) . (Explicit fc r ::) + go (POp fc opFC leftSide op r) = (PRef opFC op,) . (Explicit fc leftSide.getLhs ::) . (Explicit fc r ::) go (PEq fc l r) = (PRef fc $ embed eqName,) . (Explicit fc l ::) . (Explicit fc r ::) -- ambiguous, picking the type constructor here go (PPair fc l r) = (PRef fc $ embed pairname,) . (Explicit fc l ::) . (Explicit fc r ::) diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index b04c839a8..dc7159e30 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -1,6 +1,8 @@ -autobind infixr 0 =@ +import Data.Vect + +typebind infixr 0 =@ 0 (=@) : (a : Type) -> (a -> Type) -> Type (=@) a f = (1 x : a) -> f x @@ -11,4 +13,21 @@ data S : {ty : Type} -> (x : ty) -> Type where -- same as -- MkS : (1 x : ty) -> S x +typebind infixr 0 *> +-- (*>) : (ty : Type) -> (ty -> Type) -> Type +-- (*>) = DPair +-- +-- testCompose : (x : Nat) *> (y : Nat) *> Vect (n + m) String +-- testCompose = (1 ** 2 ** ["hello", "world", "!"]) + +autobind infixr 7 `MyLet` + +MyLet : (val) -> (val -> rest) -> rest +MyLet arg fn = fn arg + +program : Nat +program = (n := 3) `MyLet` 2 + n + +program2 : Nat +program2 = (n : Nat := 3) `MyLet` 2 + n From 6ce4ec2ebf9166b9b5b898403da8650e3bbdcca1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 21 Oct 2023 22:26:10 +0100 Subject: [PATCH 05/57] add both typebind and autobind --- src/Idris/Doc/Keywords.idr | 42 ++++++++++++++++---- src/Idris/Parser.idr | 16 +++++--- src/Idris/Syntax.idr | 31 +++++++++++---- src/Idris/Syntax/TTC.idr | 18 +++++++-- src/Parser/Lexer/Source.idr | 2 +- tests/idris2/operators/operators001/Test.idr | 2 - 6 files changed, 86 insertions(+), 25 deletions(-) diff --git a/src/Idris/Doc/Keywords.idr b/src/Idris/Doc/Keywords.idr index 1544b5ab4..b79d64485 100644 --- a/src/Idris/Doc/Keywords.idr +++ b/src/Idris/Doc/Keywords.idr @@ -467,17 +467,44 @@ autobindDoc = """ allowing you to give a name to the left-hand-side of the operator and use it on the right-hand-side. - A typical example of an autobind operator is `(**)` the type constructor - for dependent pairs. It is used like this: `(x : Nat ** Vect x String)` + `autobind` differs from `typebind` in the syntax it allows and the way + it desugars. Delcaring an operator as `autobind infixr 0 =|>` allows + you to use it with the syntax `(x := e =|> f x)` or `(x : t := e =|> f x)` - If you declare a new operator to be autobind you can use it the same - way. - - Start by defining `autobind infixr 1 =@`, and then you can use it - like so: `(n : Nat =@ f n)`. + You will need to use `autobind` instead of `typebind` whenever the + type of the lambda of the function called by the operator has a type + that is not equal to the argument given on the left side of the operator `autobind` only applies to `infixr` operators and cannot be used as operator sections. + + `autobind` operators are desugared as a lambda: + `(x := expr =|> fn x)` -> `(expr =@ (\x : ? =|> fn x))` + `(x : ty := expr =|> fn x)` -> `(expr =@ (\x : ty =|> fn x))` + """ +typebindDoc : Doc IdrisDocAnn +typebindDoc = """ + Typebind + + `typebind` is a modifier for operator precedence and fixity declaration. + It tells the parser that this operator behaves like a binding operator, + allowing you to give a name to the left-hand-side of the operator and + use it on the right-hand-side. + + A typical example of a typebind operator is `(**)` the type constructor + for dependent pairs. It is used like this: `(x : Nat ** Vect x String)` + + If you declare a new operator to be typebind you can use it the same + way. + + Start by defining `typebind infixr 1 =@`, and then you can use it + like so: `(n : Nat =@ f n)` or `(n : Nat) =@ f n` + + `typebind` only applies to `infixr` operators and cannot be used as + operator sections. + + `typebind` operators are desugared as a lambda: + `(x : ty =@ fn x)` -> `(ty =@ (\x : ty =@ fn x))` """ rewriteeq : Doc IdrisDocAnn @@ -584,6 +611,7 @@ keywordsDoc = :: "else" ::= ifthenelse :: "forall" ::= forallquantifier :: "rewrite" ::= rewriteeq + :: "typebind" ::= autobindDoc :: "autobind" ::= autobindDoc :: "using" ::= "" :: "interface" ::= interfacemechanism diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 6ff4a78a5..556223201 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -331,15 +331,15 @@ mutual opBinderTypes fname indents boundName = do decoratedSymbol fname ":" ty <- typeExpr pdef fname indents - pure (BindType boundName ty) + decoratedSymbol fname ":=" + exp <- expr pdef fname indents + pure (BindExplicitType boundName ty exp) <|> do decoratedSymbol fname ":=" exp <- expr pdef fname indents pure (BindExpr boundName exp) <|> do decoratedSymbol fname ":" ty <- typeExpr pdef fname indents - decoratedSymbol fname ":=" - exp <- expr pdef fname indents - pure (BindExplicitType boundName ty exp) + pure (BindType boundName ty) opBinder : OriginDesc -> IndentInfo -> Rule (OperatorLHSInfo PTerm) opBinder fname indents @@ -1866,10 +1866,16 @@ definition fname indents = do nd <- bounds (clause 0 Nothing fname indents) pure (PDef (boundToFC fname nd) [nd.val]) +operatorBindingKeyword : EmptyRule BindingModifier +operatorBindingKeyword + = (keyword "autobind" >> pure Autobind) + <|> (keyword "typebind" >> pure Typebind) + <|> pure NotBinding + fixDecl : OriginDesc -> IndentInfo -> Rule (List PDecl) fixDecl fname indents = do vis <- exportVisibility fname - binding <- map isJust (optional (keyword "autobind")) + binding <- operatorBindingKeyword b <- bounds (do fixity <- decorate fname Keyword $ fix commit prec <- decorate fname Keyword $ intLit diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index 686ed9d52..e12fa374d 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -47,25 +47,42 @@ Eq Fixity where Prefix == Prefix = True _ == _ = False + +public export +data BindingModifier = NotBinding | Autobind | Typebind + +export +Eq BindingModifier where + NotBinding == NotBinding = True + Autobind == Autobind = True + Typebind == Typebind = True + _ == _ = False + +export +Show BindingModifier where + show NotBinding = "not binding" + show Typebind = "typebind" + show Autobind = "autobind" + -- A record to hold all the information about a fixity public export record FixityInfo where constructor MkFixityInfo fc : FC vis : Visibility - isAutobind : Bool + bindingInfo : BindingModifier fix : Fixity precedence : Nat export Show FixityInfo where - show fx = "fc: \{show fx.fc}, visibility: \{show fx.vis}, isAutobind: \{show fx.isAutobind}, fixity: \{show fx.fix}, precedence: \{show fx.precedence}" + show fx = "fc: \{show fx.fc}, visibility: \{show fx.vis}, binding: \{show fx.bindingInfo}, fixity: \{show fx.fix}, precedence: \{show fx.precedence}" export Eq FixityInfo where x == y = x.fc == y.fc && x.vis == y.vis - && x.isAutobind == y.isAutobind + && x.bindingInfo == y.bindingInfo && x.fix == y.fix && x.precedence == y.precedence @@ -507,7 +524,7 @@ mutual PFail : FC -> Maybe String -> List (PDecl' nm) -> PDecl' nm PMutual : FC -> List (PDecl' nm) -> PDecl' nm - PFixity : FC -> Visibility -> (isAutobind : Bool) -> Fixity -> Nat -> OpStr -> PDecl' nm + PFixity : FC -> Visibility -> BindingModifier -> Fixity -> Nat -> OpStr -> PDecl' nm PNamespace : FC -> Namespace -> List (PDecl' nm) -> PDecl' nm PTransform : FC -> String -> PTerm' nm -> PTerm' nm -> PDecl' nm PRunElabDecl : FC -> PTerm' nm -> PDecl' nm @@ -1043,9 +1060,9 @@ initSyntax initFixities : ANameMap FixityInfo initFixities = fromList - [ (UN $ Basic "-", MkFixityInfo EmptyFC Export False Prefix 10) - , (UN $ Basic "negate", MkFixityInfo EmptyFC Export False Prefix 10) -- for documentation purposes - , (UN $ Basic "=", MkFixityInfo EmptyFC Export False Infix 0) + [ (UN $ Basic "-", MkFixityInfo EmptyFC Export NotBinding Prefix 10) + , (UN $ Basic "negate", MkFixityInfo EmptyFC Export NotBinding Prefix 10) -- for documentation purposes + , (UN $ Basic "=", MkFixityInfo EmptyFC Export NotBinding Infix 0) ] initDocStrings : ANameMap String diff --git a/src/Idris/Syntax/TTC.idr b/src/Idris/Syntax/TTC.idr index 47e1b63ad..a8b144d27 100644 --- a/src/Idris/Syntax/TTC.idr +++ b/src/Idris/Syntax/TTC.idr @@ -81,21 +81,33 @@ TTC Import where nameAs <- fromBuf b pure (MkImport loc reexport path nameAs) +export +TTC BindingModifier where + toBuf b NotBinding = tag 0 + toBuf b Typebind = tag 1 + toBuf b Autobind = tag 2 + fromBuf b + = case !getTag of + 0 => pure NotBinding + 1 => pure Typebind + 2 => pure Autobind + _ => corrupt "binding" + export TTC FixityInfo where toBuf b fx = do toBuf b fx.fc toBuf b fx.vis - toBuf b fx.isAutobind + toBuf b fx.bindingInfo toBuf b fx.fix toBuf b fx.precedence fromBuf b = do fc <- fromBuf b vis <- fromBuf b - isAutobind <- fromBuf b + binding <- fromBuf b fix <- fromBuf b prec <- fromBuf b - pure $ MkFixityInfo fc vis isAutobind fix prec + pure $ MkFixityInfo fc vis binding fix prec export diff --git a/src/Parser/Lexer/Source.idr b/src/Parser/Lexer/Source.idr index 3bfaa700e..b3acb22f0 100644 --- a/src/Parser/Lexer/Source.idr +++ b/src/Parser/Lexer/Source.idr @@ -245,7 +245,7 @@ keywords : List String keywords = ["data", "module", "where", "let", "in", "do", "record", "auto", "default", "implicit", "failing", "mutual", "namespace", "parameters", "with", "proof", "impossible", "case", "of", - "if", "then", "else", "forall", "rewrite", "autobind", + "if", "then", "else", "forall", "rewrite", "typebind", "autobind", "using", "interface", "implementation", "open", "import", "public", "export", "private"] ++ fixityKeywords ++ diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index dc7159e30..035e5a0af 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -10,8 +10,6 @@ typebind infixr 0 =@ data S : {ty : Type} -> (x : ty) -> Type where MkS : (x : ty) =@ S x - -- same as - -- MkS : (1 x : ty) -> S x typebind infixr 0 *> From 1adb780e3fbbd74ad5c9e3977a4a19e6aeee2d4b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sun, 22 Oct 2023 13:42:13 +0100 Subject: [PATCH 06/57] add tests to chaining and combing operators --- tests/idris2/operators/operators001/Test.idr | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index 035e5a0af..aa31184ae 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -3,13 +3,20 @@ import Data.Vect typebind infixr 0 =@ +infixr 0 -@ 0 (=@) : (a : Type) -> (a -> Type) -> Type (=@) a f = (1 x : a) -> f x +(-@) : (a, b : Type) -> Type +(-@) a b = (1 _ : a) -> b data S : {ty : Type} -> (x : ty) -> Type where MkS : (x : ty) =@ S x + Mk2 : (x : ty) =@ (y : ty) =@ S (x, y) + Mk3 : (x : ty) =@ ty -@ S x + Mk4 : ty -@ (x : ty) =@ S x + -- Chain : (x : ty =@ y : ty =@ S (x, y)) typebind infixr 0 *> From 02a6751796f6cc9aa13109fb0325eb52a26b186e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 24 Oct 2023 17:11:27 +0100 Subject: [PATCH 07/57] add shunting test --- tests/idris2/operators/operators001/Test.idr | 22 +++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index aa31184ae..e7dcfc9d4 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -1,13 +1,19 @@ +module Test import Data.Vect typebind infixr 0 =@ infixr 0 -@ +typebind infix 1 =@@ + 0 (=@) : (a : Type) -> (a -> Type) -> Type (=@) a f = (1 x : a) -> f x +0 (=@@) : (a : Type) -> (a -> Type) -> Type +(=@@) a f = (1 x : a) -> f x + (-@) : (a, b : Type) -> Type (-@) a b = (1 _ : a) -> b @@ -16,7 +22,21 @@ data S : {ty : Type} -> (x : ty) -> Type where Mk2 : (x : ty) =@ (y : ty) =@ S (x, y) Mk3 : (x : ty) =@ ty -@ S x Mk4 : ty -@ (x : ty) =@ S x - -- Chain : (x : ty =@ y : ty =@ S (x, y)) + +map : (x : a) =@@ b -@ (y : List a) =@ List b + +map2 : ((x : a) =@ b) -@ (y : List a) =@ List b + +map3 : (x : a) =@ b -@ (y : List a) =@ List b + +map4 : (x : a) =@ (b -@ (y : List a) =@ List b) + +test : Test.map === Test.map2 + +failing + test2 : Test.map === Test.map3 + +test3 : Test.map3 === Test.map4 typebind infixr 0 *> From 891f84107760e2667a289ecbc268b158d0c9bcac Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 24 Oct 2023 16:58:52 +0100 Subject: [PATCH 08/57] correctly rearrange binding operators according to shunting rules --- src/Idris/Desugar.idr | 90 ++++++++++++++++++++++---------- src/Idris/Parser.idr | 5 +- src/Idris/Pretty.idr | 6 +-- src/Idris/Syntax.idr | 27 ++++++---- src/Libraries/Utils/Shunting.idr | 7 ++- 5 files changed, 89 insertions(+), 46 deletions(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 9d19e59b2..d77c54033 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -126,12 +126,14 @@ checkConflictingFixities isPrefix exprFC opn then throw (GenericMsg exprFC "Unknown operator '\{op}'") else pure (NonAssoc 1) -- Backticks are non associative by default + -- (True, ((fxName, fx) :: _), _) => do -- in the prefix case, remove conflicts with infix (-) let extraFixities = pre ++ (filter (\(nm, _) => not $ nameRoot nm == "-") inf) unless (isCompatible fx extraFixities) $ warnConflict fxName extraFixities pure (mkPrec fx.fix fx.precedence) - -- Could not find any prefix operator fixities, there may be infix ones + -- Could not find any prefix operator fixities, there may still be conflicts with + -- the infix ones. (True, [] , _) => throw (GenericMsg exprFC $ "'\{op}' is not a prefix operator") (False, _, ((fxName, fx) :: _)) => do @@ -145,7 +147,9 @@ checkConflictingFixities isPrefix exprFC opn -- Fixities are compatible with all others of the same name that share the same fixity and precedence isCompatible : FixityInfo -> (fixities : List (Name, FixityInfo)) -> Bool isCompatible fx - = all (\fx' => fx.fix == fx'.fix && fx.precedence == fx'.precedence) . map snd + = all (\fx' => fx.fix == fx'.fix + && fx.precedence == fx'.precedence + && fx.bindingInfo == fx'.bindingInfo ) . map snd -- Emits a warning using the fixity that we picked and the list of all conflicting fixities warnConflict : (picked : Name) -> (conflicts : List (Name, FixityInfo)) -> Core () @@ -157,17 +161,30 @@ checkConflictingFixities isPrefix exprFC opn For example: %hide \{show fxName} """ + toTokList : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> - PTerm -> Core (List (Tok OpStr PTerm)) -toTokList (POp fc opFC (NotAutobind l) opn r) + PTerm -> Core (List (Tok (OpStr, OperatorLHSInfo PTerm) PTerm)) + -- -- a bindtype operator desugars (x : a ** b x) into ((**) a (\x : a => b x)) +-- toTokList (POp fc opFC (BindType nm l) opn r) + -- = desugarB side ps (POp fc opFC (NotAutobind l) op + -- $ PLam (virtualiseFC opFC) top Explicit nm l r) + -- -- a bindexpr operator desugars (x := a ** b x) into ((**) a (\x : ? => b x)) + -- desugarB side ps (POp fc opFC (BindExpr nm l) op r) + -- = desugarB side ps (POp fc opFC (NotAutobind l) op + -- $ PLam (virtualiseFC opFC) top Explicit nm (PImplicit $ virtualiseFC opFC) r) + -- -- an explicit bind operator desugars (x : ty := a ** b x) into ((**) a (\x : ty => b x)) + -- desugarB side ps (POp fc opFC (BindExplicitType nm ty l) op r) + -- = desugarB side ps (POp fc opFC (NotAutobind l) op + -- $ PLam (virtualiseFC opFC) top Explicit nm ty r) +toTokList (POp fc opFC (l) opn r) = do precInfo <- checkConflictingFixities False fc opn rtoks <- toTokList r - pure (Expr l :: Op fc opFC opn precInfo :: rtoks) + pure (Expr l.getLhs :: Op fc opFC (opn, l) precInfo :: rtoks) toTokList (PPrefixOp fc opFC opn arg) = do precInfo <- checkConflictingFixities True fc opn rtoks <- toTokList arg - pure (Op fc opFC opn precInfo :: rtoks) + pure (Op fc opFC (opn, ?bidningForPrefix) precInfo :: rtoks) toTokList t = pure [Expr t] record BangData where @@ -308,20 +325,8 @@ mutual [apply (IVar fc (UN $ Basic "===")) [l', r'], apply (IVar fc (UN $ Basic "~=~")) [l', r']] desugarB side ps (PBracketed fc e) = desugarB side ps e - -- a bindtype operator desugars (x : a ** b x) into ((**) a (\x : a => b x)) - desugarB side ps (POp fc opFC (BindType nm l) op r) - = desugarB side ps (POp fc opFC (NotAutobind l) op - $ PLam (virtualiseFC opFC) top Explicit nm l r) - -- a bindexpr operator desugars (x := a ** b x) into ((**) a (\x : ? => b x)) - desugarB side ps (POp fc opFC (BindExpr nm l) op r) - = desugarB side ps (POp fc opFC (NotAutobind l) op - $ PLam (virtualiseFC opFC) top Explicit nm (PImplicit $ virtualiseFC opFC) r) - -- an explicit bind operator desugars (x : ty := a ** b x) into ((**) a (\x : ty => b x)) - desugarB side ps (POp fc opFC (BindExplicitType nm ty l) op r) - = desugarB side ps (POp fc opFC (NotAutobind l) op - $ PLam (virtualiseFC opFC) top Explicit nm ty r) - desugarB side ps (POp fc opFC (NotAutobind l) op r) - = do ts <- toTokList (POp fc opFC (NotAutobind l) op r) + desugarB side ps (POp fc opFC (l) op r) + = do ts <- toTokList (POp fc opFC (l) op r) desugarTree side ps !(parseOps ts) desugarB side ps (PPrefixOp fc opFC op arg) = do ts <- toTokList (PPrefixOp fc opFC op arg) @@ -726,27 +731,57 @@ mutual rule' <- desugarDo side ps ns rule pure $ IRewrite fc rule' rest' + -- Handle special cases for operator overloads + -- Maybe overloading ** should be in here desugarTree : {auto s : Ref Syn SyntaxInfo} -> {auto b : Ref Bang BangData} -> {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> {auto m : Ref MD Metadata} -> {auto o : Ref ROpts REPLOpts} -> - Side -> List Name -> Tree OpStr PTerm -> Core RawImp - desugarTree side ps (Infix loc eqFC (UN $ Basic "=") l r) -- special case since '=' is special syntax + Side -> List Name -> Tree (OpStr, OperatorLHSInfo PTerm) PTerm -> Core RawImp + desugarTree side ps (Infix loc eqFC (UN $ Basic "=", _) l r) -- special case since '=' is special syntax = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IAlternative loc FirstSuccess [apply (IVar eqFC eqName) [l', r'], apply (IVar eqFC heqName) [l', r']]) - desugarTree side ps (Infix loc _ (UN $ Basic "$") l r) -- special case since '$' is special syntax + desugarTree side ps (Infix loc _ (UN $ Basic "$", _) l r) -- special case since '$' is special syntax = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc l' r') - desugarTree side ps (Infix loc opFC op l r) + desugarTree side ps (Infix loc opFC (op, NotAutobind lhs) l r) = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc (IApp loc (IVar opFC op) l') r') + -- -- a bindtype operator desugars (x : a ** b x) into ((**) a (\x : a => b x)) + -- desugarB side ps (POp fc opFC (BindType nm l) op r) + -- = desugarB side ps (POp fc opFC (NotAutobind l) op + -- $ PLam (virtualiseFC opFC) top Explicit nm l r) + -- -- a bindexpr operator desugars (x := a ** b x) into ((**) a (\x : ? => b x)) + -- desugarB side ps (POp fc opFC (BindExpr nm l) op r) + -- = desugarB side ps (POp fc opFC (NotAutobind l) op + -- $ PLam (virtualiseFC opFC) top Explicit nm (PImplicit $ virtualiseFC opFC) r) + -- -- an explicit bind operator desugars (x : ty := a ** b x) into ((**) a (\x : ty => b x)) + -- desugarB side ps (POp fc opFC (BindExplicitType nm ty l) op r) + -- = desugarB side ps (POp fc opFC (NotAutobind l) op + -- $ PLam (virtualiseFC opFC) top Explicit nm ty r) + desugarTree side ps (Infix loc opFC (op, BindType name lhs) l 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) desugaredLHS desugaredRHS) + desugarTree side ps (Infix loc opFC (op, BindExpr name lhs) l 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) + desugarTree side ps (Infix loc opFC (op, BindExplicitType name ty expr) l 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) -- negation is a special case, since we can't have an operator with -- two meanings otherwise @@ -754,7 +789,7 @@ mutual -- Note: In case of negated signed integer literals, we apply the -- negation directly. Otherwise, the literal might be -- truncated to 0 before being passed on to `negate`. - desugarTree side ps (Pre loc opFC (UN $ Basic "-") $ Leaf $ PPrimVal fc c) + desugarTree side ps (Pre loc opFC (UN $ Basic "-", _) $ Leaf $ PPrimVal fc c) = let newFC = fromMaybe EmptyFC (mergeFC loc fc) continue = desugarTree side ps . Leaf . PPrimVal newFC in case c of @@ -770,11 +805,11 @@ mutual _ => do arg' <- desugarTree side ps (Leaf $ PPrimVal fc c) pure (IApp loc (IVar opFC (UN $ Basic "negate")) arg') - desugarTree side ps (Pre loc opFC (UN $ Basic "-") arg) + desugarTree side ps (Pre loc opFC (UN $ Basic "-", _) arg) = do arg' <- desugarTree side ps arg pure (IApp loc (IVar opFC (UN $ Basic "negate")) arg') - desugarTree side ps (Pre loc opFC op arg) + desugarTree side ps (Pre loc opFC (op, _) arg) = do arg' <- desugarTree side ps arg pure (IApp loc (IVar opFC op) arg') desugarTree side ps (Leaf t) = desugarB side ps t @@ -1242,4 +1277,5 @@ mutual {auto u : Ref UST UState} -> {auto o : Ref ROpts REPLOpts} -> Side -> List Name -> PTerm -> Core RawImp + desugar s ps tm = desugarDo s ps Nothing tm diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 556223201..144d247a5 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -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 -> PTerm -> Rule (OperatorLHSInfo PTerm) + opBinderTypes : OriginDesc -> IndentInfo -> Name -> Rule (OperatorLHSInfo PTerm) opBinderTypes fname indents boundName = do decoratedSymbol fname ":" ty <- typeExpr pdef fname indents @@ -343,8 +343,7 @@ mutual opBinder : OriginDesc -> IndentInfo -> Rule (OperatorLHSInfo PTerm) opBinder fname indents - = do x <- bounds (UN . Basic <$> decoratedSimpleBinderName fname) - let boundName = PRef (boundToFC fname x) x.val + = do boundName <- (UN . Basic <$> decoratedSimpleBinderName fname) opBinderTypes fname indents boundName autobindOp : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index 8f299368d..7a54b7a26 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -333,15 +333,15 @@ mutual prettyPrec d (PImplicit _) = "_" prettyPrec d (PInfer _) = annotate Hole $ "?" prettyPrec d (POp _ _ (BindType nm left) op right) = - group $ parens (pretty nm <++> ":" <++> pretty left) + group $ parens (prettyBinder nm <++> ":" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (BindExpr nm left) op right) = - group $ parens (pretty nm <++> ":=" <++> pretty left) + group $ parens (prettyBinder nm <++> ":=" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (BindExplicitType nm ty left) op right) = - group $ parens (pretty nm <++> ":" <++> pretty ty <++> ":=" <++> pretty left) + group $ parens (prettyBinder nm <++> ":" <++> pretty ty <++> ":=" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (NotAutobind x) op y) = diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index e12fa374d..c1e850c5d 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -111,18 +111,27 @@ data OperatorLHSInfo : tm -> Type where -- Traditional operator wihtout binding, carries the lhs NotAutobind : (lhs : tm) -> OperatorLHSInfo tm -- (nm : ty) ** fn x - BindType : (name, ty : tm) -> OperatorLHSInfo tm + BindType : (name : Name) -> (ty : tm) -> OperatorLHSInfo tm -- (nm := exp) ** fn nm - BindExpr : (name, expr : tm) -> OperatorLHSInfo tm + BindExpr : (name : Name) -> (expr : tm) -> OperatorLHSInfo tm -- (nm : ty := exp) ** fn nm - BindExplicitType : (name, type, expr : tm) -> OperatorLHSInfo tm + BindExplicitType : (name : Name) -> (type, expr : tm) -> OperatorLHSInfo tm + +export +Show (OperatorLHSInfo tm) where + show (NotAutobind lhs) = "Regular operator" + show (BindType name ty) = "Type-binding operator" + show (BindExpr name expr) = "Automatically-binding operator" + show (BindExplicitType name type expr) = "Automatically-binding operator (explicit)" + +%name OperatorLHSInfo opInfo export Functor OperatorLHSInfo where map f (NotAutobind lhs) = NotAutobind $ f lhs - map f (BindType nm lhs) = BindType (f nm) (f lhs) - map f (BindExpr nm lhs) = BindExpr (f nm) (f lhs) - map f (BindExplicitType nm ty lhs) = BindExplicitType (f nm) (f ty) (f lhs) + map f (BindType nm lhs) = BindType nm (f lhs) + map f (BindExpr nm lhs) = BindExpr nm (f lhs) + map f (BindExplicitType nm ty lhs) = BindExplicitType nm (f ty) (f lhs) export (.getLhs) : OperatorLHSInfo tm -> tm @@ -863,11 +872,11 @@ parameters {0 nm : Type} (toName : nm -> Name) showPTermPrec d (POp _ _ (NotAutobind left) op right) = showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right showPTermPrec d (POp _ _ (BindType nm left) op right) - = "(" ++ showPTermPrec d nm ++ " : " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ show nm ++ " : " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (POp _ _ (BindExpr nm left) op right) - = "(" ++ showPTermPrec d nm ++ " := " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ show nm ++ " := " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (POp _ _ (BindExplicitType nm ty left) op right) - = "(" ++ showPTermPrec d nm ++ " : " ++ showPTermPrec d ty ++ ":=" ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ show nm ++ " : " ++ showPTermPrec d ty ++ ":=" ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (PPrefixOp _ _ op x) = showOpPrec d op ++ showPTermPrec d x showPTermPrec d (PSectionL _ _ op x) = "(" ++ showOpPrec d op ++ " " ++ showPTermPrec d x ++ ")" showPTermPrec d (PSectionR _ _ x op) = "(" ++ showPTermPrec d x ++ " " ++ showOpPrec d op ++ ")" diff --git a/src/Libraries/Utils/Shunting.idr b/src/Libraries/Utils/Shunting.idr index 424893cc1..e60f97936 100644 --- a/src/Libraries/Utils/Shunting.idr +++ b/src/Libraries/Utils/Shunting.idr @@ -22,10 +22,9 @@ data OpPrec -- Tokens are either operators or already parsed expressions in some -- higher level language public export -data Tok op a - = ||| The second FC is for the operator alone - Op FC FC op OpPrec - | Expr a +data Tok : (op, a : Type) -> Type where + Op : (expressionLoc : FC) -> (operatorLoc : FC) -> (operatorInfo : op) -> OpPrec -> Tok op a + Expr : a -> Tok op a -- The result of shunting is a parse tree with the precedences made explicit -- in the tree. From 4c7010871fe2353eac208354e081a2545f0c9321 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 24 Oct 2023 17:28:55 +0100 Subject: [PATCH 09/57] cleanup hole --- src/Idris/Desugar.idr | 32 +++++++++++--------------------- 1 file changed, 11 insertions(+), 21 deletions(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index d77c54033..610f70cc2 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -164,7 +164,7 @@ checkConflictingFixities isPrefix exprFC opn toTokList : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> - PTerm -> Core (List (Tok (OpStr, OperatorLHSInfo PTerm) PTerm)) + PTerm -> Core (List (Tok (OpStr, Maybe (OperatorLHSInfo PTerm)) PTerm)) -- -- a bindtype operator desugars (x : a ** b x) into ((**) a (\x : a => b x)) -- toTokList (POp fc opFC (BindType nm l) opn r) -- = desugarB side ps (POp fc opFC (NotAutobind l) op @@ -177,14 +177,14 @@ toTokList : {auto s : Ref Syn SyntaxInfo} -> -- desugarB side ps (POp fc opFC (BindExplicitType nm ty l) op r) -- = desugarB side ps (POp fc opFC (NotAutobind l) op -- $ PLam (virtualiseFC opFC) top Explicit nm ty r) -toTokList (POp fc opFC (l) opn r) +toTokList (POp fc opFC l opn r) = do precInfo <- checkConflictingFixities False fc opn rtoks <- toTokList r - pure (Expr l.getLhs :: Op fc opFC (opn, l) precInfo :: rtoks) + pure (Expr l.getLhs :: Op fc opFC (opn, Just l) precInfo :: rtoks) toTokList (PPrefixOp fc opFC opn arg) = do precInfo <- checkConflictingFixities True fc opn rtoks <- toTokList arg - pure (Op fc opFC (opn, ?bidningForPrefix) precInfo :: rtoks) + pure (Op fc opFC (opn, Nothing) precInfo :: rtoks) toTokList t = pure [Expr t] record BangData where @@ -739,7 +739,7 @@ mutual {auto u : Ref UST UState} -> {auto m : Ref MD Metadata} -> {auto o : Ref ROpts REPLOpts} -> - Side -> List Name -> Tree (OpStr, OperatorLHSInfo PTerm) PTerm -> Core RawImp + Side -> List Name -> Tree (OpStr, Maybe $ OperatorLHSInfo PTerm) PTerm -> Core RawImp desugarTree side ps (Infix loc eqFC (UN $ Basic "=", _) l r) -- special case since '=' is special syntax = do l' <- desugarTree side ps l r' <- desugarTree side ps r @@ -750,38 +750,28 @@ mutual = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc l' r') - desugarTree side ps (Infix loc opFC (op, NotAutobind lhs) l r) + desugarTree side ps (Infix loc opFC (op, Just (NotAutobind lhs)) l r) = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc (IApp loc (IVar opFC op) l') r') - -- -- a bindtype operator desugars (x : a ** b x) into ((**) a (\x : a => b x)) - -- desugarB side ps (POp fc opFC (BindType nm l) op r) - -- = desugarB side ps (POp fc opFC (NotAutobind l) op - -- $ PLam (virtualiseFC opFC) top Explicit nm l r) - -- -- a bindexpr operator desugars (x := a ** b x) into ((**) a (\x : ? => b x)) - -- desugarB side ps (POp fc opFC (BindExpr nm l) op r) - -- = desugarB side ps (POp fc opFC (NotAutobind l) op - -- $ PLam (virtualiseFC opFC) top Explicit nm (PImplicit $ virtualiseFC opFC) r) - -- -- an explicit bind operator desugars (x : ty := a ** b x) into ((**) a (\x : ty => b x)) - -- desugarB side ps (POp fc opFC (BindExplicitType nm ty l) op r) - -- = desugarB side ps (POp fc opFC (NotAutobind l) op - -- $ PLam (virtualiseFC opFC) top Explicit nm ty r) - desugarTree side ps (Infix loc opFC (op, BindType name lhs) l r) + desugarTree side ps (Infix loc opFC (op, Just (BindType 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) desugaredLHS desugaredRHS) - desugarTree side ps (Infix loc opFC (op, BindExpr name lhs) l r) + 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) - desugarTree side ps (Infix loc opFC (op, BindExplicitType name ty expr) l r) + 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) + desugarTree side ps (Infix loc opFC (op, Nothing) _ r) + = throw $ InternalError "illegal fixity: Parsed as infix but no binding information" -- negation is a special case, since we can't have an operator with -- two meanings otherwise From 653ed454ef6c914119ef94b6dd0acd74634e7805 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 24 Oct 2023 18:19:31 +0100 Subject: [PATCH 10/57] add basics error messages for operator misuse --- src/Idris/Desugar.idr | 59 +++++++++++++++++++++++++++---------------- 1 file changed, 37 insertions(+), 22 deletions(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 610f70cc2..8b69efa91 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -113,7 +113,7 @@ mkPrec Prefix = Prefix checkConflictingFixities : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> (isPrefix : Bool) -> - FC -> Name -> Core OpPrec + FC -> Name -> Core (OpPrec, BindingModifier) checkConflictingFixities isPrefix exprFC opn = do syn <- get Syn let op = nameRoot opn @@ -124,14 +124,14 @@ checkConflictingFixities isPrefix exprFC opn -- characters, if not, it must be a backticked expression. (_, [], []) => if any isOpChar (fastUnpack op) then throw (GenericMsg exprFC "Unknown operator '\{op}'") - else pure (NonAssoc 1) -- Backticks are non associative by default + else pure (NonAssoc 1, NotBinding) -- Backticks are non associative by default -- (True, ((fxName, fx) :: _), _) => do -- in the prefix case, remove conflicts with infix (-) let extraFixities = pre ++ (filter (\(nm, _) => not $ nameRoot nm == "-") inf) unless (isCompatible fx extraFixities) $ warnConflict fxName extraFixities - pure (mkPrec fx.fix fx.precedence) + pure (mkPrec fx.fix fx.precedence, NotBinding) -- Could not find any prefix operator fixities, there may still be conflicts with -- the infix ones. (True, [] , _) => throw (GenericMsg exprFC $ "'\{op}' is not a prefix operator") @@ -140,16 +140,17 @@ checkConflictingFixities isPrefix exprFC opn -- In the infix case, remove conflicts with prefix (-) let extraFixities = (filter (\(nm, _) => not $ nm == UN (Basic "-")) pre) ++ inf unless (isCompatible fx extraFixities) $ warnConflict fxName extraFixities - pure (mkPrec fx.fix fx.precedence) + pure (mkPrec fx.fix fx.precedence, fx.bindingInfo) -- Could not find any infix operator fixities, there may be prefix ones (False, _, []) => throw (GenericMsg exprFC $ "'\{op}' is not an infix operator") where - -- Fixities are compatible with all others of the same name that share the same fixity and precedence + -- Fixities are compatible with all others of the same name that share the same + -- fixity, precedence, and binding information isCompatible : FixityInfo -> (fixities : List (Name, FixityInfo)) -> Bool isCompatible fx = all (\fx' => fx.fix == fx'.fix && fx.precedence == fx'.precedence - && fx.bindingInfo == fx'.bindingInfo ) . map snd + && fx.bindingInfo == fx'.bindingInfo) . map snd -- Emits a warning using the fixity that we picked and the list of all conflicting fixities warnConflict : (picked : Name) -> (conflicts : List (Name, FixityInfo)) -> Core () @@ -161,28 +162,38 @@ checkConflictingFixities isPrefix exprFC opn For example: %hide \{show fxName} """ +checkConflictingBinding : FC -> (foundFixity : BindingModifier) -> (usage : OperatorLHSInfo a) -> Core () +checkConflictingBinding fc NotBinding (NotAutobind lhs) = pure () +checkConflictingBinding fc NotBinding (BindType name ty) + = throw $ GenericMsg fc "Operator \{show name} is regular but is used in a type-binding position" +checkConflictingBinding fc NotBinding (BindExpr name expr) + = throw $ GenericMsg fc "Operator \{show name} is regular but is used in an automatically-binding position" +checkConflictingBinding fc NotBinding (BindExplicitType name type expr) + = throw $ GenericMsg fc "Operator \{show name} is regular but is used in an automatically-binding position" +checkConflictingBinding fc Autobind (NotAutobind lhs) + = throw $ GenericMsg fc "Operator is automatically-binding but is used as a regular operator" +checkConflictingBinding fc Autobind (BindType name ty) + = throw $ GenericMsg fc "Operator is automatically-binding but is used in a type-binding position" +checkConflictingBinding fc Autobind (BindExpr name expr) = pure () +checkConflictingBinding fc Autobind (BindExplicitType name type expr) = pure () +checkConflictingBinding fc Typebind (NotAutobind lhs) + = throw $ GenericMsg fc "Operator is type-binding but is used in as a regular operator" +checkConflictingBinding fc Typebind (BindType name ty) = pure() +checkConflictingBinding fc Typebind (BindExpr name expr) + = throw $ GenericMsg fc "Operator is type-binding but is used in an automatically-binding position" +checkConflictingBinding fc Typebind (BindExplicitType name type expr) + = throw $ GenericMsg fc "Operator is type-binding but is used in an automatically-binding position" toTokList : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> PTerm -> Core (List (Tok (OpStr, Maybe (OperatorLHSInfo PTerm)) PTerm)) - -- -- a bindtype operator desugars (x : a ** b x) into ((**) a (\x : a => b x)) --- toTokList (POp fc opFC (BindType nm l) opn r) - -- = desugarB side ps (POp fc opFC (NotAutobind l) op - -- $ PLam (virtualiseFC opFC) top Explicit nm l r) - -- -- a bindexpr operator desugars (x := a ** b x) into ((**) a (\x : ? => b x)) - -- desugarB side ps (POp fc opFC (BindExpr nm l) op r) - -- = desugarB side ps (POp fc opFC (NotAutobind l) op - -- $ PLam (virtualiseFC opFC) top Explicit nm (PImplicit $ virtualiseFC opFC) r) - -- -- an explicit bind operator desugars (x : ty := a ** b x) into ((**) a (\x : ty => b x)) - -- desugarB side ps (POp fc opFC (BindExplicitType nm ty l) op r) - -- = desugarB side ps (POp fc opFC (NotAutobind l) op - -- $ PLam (virtualiseFC opFC) top Explicit nm ty r) toTokList (POp fc opFC l opn r) - = do precInfo <- checkConflictingFixities False fc opn + = do (precInfo, bindInfo) <- checkConflictingFixities False fc opn + checkConflictingBinding opFC bindInfo l rtoks <- toTokList r pure (Expr l.getLhs :: Op fc opFC (opn, Just l) precInfo :: rtoks) toTokList (PPrefixOp fc opFC opn arg) - = do precInfo <- checkConflictingFixities True fc opn + = do (precInfo, _) <- checkConflictingFixities True fc opn rtoks <- toTokList arg pure (Op fc opFC (opn, Nothing) precInfo :: rtoks) toTokList t = pure [Expr t] @@ -325,8 +336,8 @@ mutual [apply (IVar fc (UN $ Basic "===")) [l', r'], apply (IVar fc (UN $ Basic "~=~")) [l', r']] desugarB side ps (PBracketed fc e) = desugarB side ps e - desugarB side ps (POp fc opFC (l) op r) - = do ts <- toTokList (POp fc opFC (l) op r) + desugarB side ps (POp fc opFC l op r) + = do ts <- toTokList (POp fc opFC l op r) desugarTree side ps !(parseOps ts) desugarB side ps (PPrefixOp fc opFC op arg) = do ts <- toTokList (PPrefixOp fc opFC op arg) @@ -750,20 +761,24 @@ mutual = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc l' r') + -- normal operators desugarTree side ps (Infix loc opFC (op, Just (NotAutobind lhs)) l r) = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc (IApp loc (IVar opFC op) l') r') + -- (x : ty ** f x) ==>> (**) exp (\x : ty => f x) desugarTree side ps (Infix loc opFC (op, Just (BindType 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) 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) + -- (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 From fd40777e0daa1472e5b12c96f99f7000dda373f1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Wed, 25 Oct 2023 12:01:52 +0100 Subject: [PATCH 11/57] add error message tests --- .../idris2/operators/operators002/Error2.idr | 12 ++++++ .../idris2/operators/operators002/Errors.idr | 9 +++++ .../idris2/operators/operators002/Errors2.idr | 8 ++++ .../idris2/operators/operators002/Errors3.idr | 8 ++++ .../idris2/operators/operators002/Errors4.idr | 9 +++++ .../idris2/operators/operators002/Errors5.idr | 10 +++++ tests/idris2/operators/operators002/Test.idr | 38 +++++++++++++++++++ tests/idris2/operators/operators002/expected | 28 ++++++++++++++ tests/idris2/operators/operators002/run | 8 ++++ 9 files changed, 130 insertions(+) create mode 100644 tests/idris2/operators/operators002/Error2.idr create mode 100644 tests/idris2/operators/operators002/Errors.idr create mode 100644 tests/idris2/operators/operators002/Errors2.idr create mode 100644 tests/idris2/operators/operators002/Errors3.idr create mode 100644 tests/idris2/operators/operators002/Errors4.idr create mode 100644 tests/idris2/operators/operators002/Errors5.idr create mode 100644 tests/idris2/operators/operators002/Test.idr create mode 100644 tests/idris2/operators/operators002/expected create mode 100755 tests/idris2/operators/operators002/run diff --git a/tests/idris2/operators/operators002/Error2.idr b/tests/idris2/operators/operators002/Error2.idr new file mode 100644 index 000000000..e3b6a420d --- /dev/null +++ b/tests/idris2/operators/operators002/Error2.idr @@ -0,0 +1,12 @@ + +typebind infixr 0 =@ + +0 (=@) : (a : Type) -> (a -> Type) -> Type +(=@) a f = (1 x : a) -> f x + + +data S : {ty : Type} -> (x : ty) -> Type where + MkS : (x : ty) =@ S x + +wrongId : {0 a : Type} -> a =@ a +wrongId x = x diff --git a/tests/idris2/operators/operators002/Errors.idr b/tests/idris2/operators/operators002/Errors.idr new file mode 100644 index 000000000..4ee059bf5 --- /dev/null +++ b/tests/idris2/operators/operators002/Errors.idr @@ -0,0 +1,9 @@ + +typebind infixr 0 =@ + +0 (=@) : (a : Type) -> (a -> Type) -> Type +(=@) a f = (1 x : a) -> f x + +data S : {ty : Type} -> (x : ty) -> Type where + MkS : (x := ty) =@ S x + diff --git a/tests/idris2/operators/operators002/Errors2.idr b/tests/idris2/operators/operators002/Errors2.idr new file mode 100644 index 000000000..f7b6e7b22 --- /dev/null +++ b/tests/idris2/operators/operators002/Errors2.idr @@ -0,0 +1,8 @@ + +autobind infixr 0 =@ + +0 (=@) : (a : Type) -> (a -> Type) -> Type +(=@) a f = (1 x : a) -> f x + +wrongId : {0 a : Type} -> a =@ a +wrongId x = x diff --git a/tests/idris2/operators/operators002/Errors3.idr b/tests/idris2/operators/operators002/Errors3.idr new file mode 100644 index 000000000..0e8994ec5 --- /dev/null +++ b/tests/idris2/operators/operators002/Errors3.idr @@ -0,0 +1,8 @@ + +typebind infixr 0 =@ + +0 (=@) : (a : Type) -> (a -> Type) -> Type +(=@) a f = (1 x : a) -> f x + +wrongId : {0 a : Type} -> a =@ a +wrongId x = x diff --git a/tests/idris2/operators/operators002/Errors4.idr b/tests/idris2/operators/operators002/Errors4.idr new file mode 100644 index 000000000..d6a52f928 --- /dev/null +++ b/tests/idris2/operators/operators002/Errors4.idr @@ -0,0 +1,9 @@ + +infixr 0 =@ + +0 (=@) : (a : Type) -> (a -> Type) -> Type +(=@) a f = (1 x : a) -> f x + + +data S : {ty : Type} -> (x : ty) -> Type where + MkS : (x : ty) =@ S x diff --git a/tests/idris2/operators/operators002/Errors5.idr b/tests/idris2/operators/operators002/Errors5.idr new file mode 100644 index 000000000..fd7225832 --- /dev/null +++ b/tests/idris2/operators/operators002/Errors5.idr @@ -0,0 +1,10 @@ + + +infixr 0 =@ + +0 (=@) : (a : Type) -> (a -> Type) -> Type +(=@) a f = (1 x : a) -> f x + + +data S : {ty : Type} -> (x : ty) -> Type where + MkS : (x := ty) =@ S x diff --git a/tests/idris2/operators/operators002/Test.idr b/tests/idris2/operators/operators002/Test.idr new file mode 100644 index 000000000..b50c78f4c --- /dev/null +++ b/tests/idris2/operators/operators002/Test.idr @@ -0,0 +1,38 @@ + + +import Data.Vect + +typebind infixr 0 =@ +infixr 0 -@ + +0 (=@) : (a : Type) -> (a -> Type) -> Type +(=@) a f = (1 x : a) -> f x + +(-@) : (a, b : Type) -> Type +(-@) a b = (1 _ : a) -> b + +data S : {ty : Type} -> (x : ty) -> Type where + MkS : (x : ty) =@ S x + Mk2 : (x : ty) =@ (y : ty) =@ S (x, y) + Mk3 : (x : ty) =@ ty -@ S x + Mk4 : ty -@ (x : ty) =@ S x + Chain : (x : ty =@ y : ty =@ S (x, y)) + +typebind infixr 0 *> + +-- (*>) : (ty : Type) -> (ty -> Type) -> Type +-- (*>) = DPair +-- +-- testCompose : (x : Nat) *> (y : Nat) *> Vect (n + m) String +-- testCompose = (1 ** 2 ** ["hello", "world", "!"]) + +autobind infixr 7 `MyLet` + +MyLet : (val) -> (val -> rest) -> rest +MyLet arg fn = fn arg + +program : Nat +program = (n := 3) `MyLet` 2 + n + +program2 : Nat +program2 = (n : Nat := 3) `MyLet` 2 + n diff --git a/tests/idris2/operators/operators002/expected b/tests/idris2/operators/operators002/expected new file mode 100644 index 000000000..5cfceb8fd --- /dev/null +++ b/tests/idris2/operators/operators002/expected @@ -0,0 +1,28 @@ +1/1: Building Errors (Errors.idr) +Operator (=@) is a type-binding operator but is used in +automatically binding position. +Either: +- Write (x : ty) =@ S x to use the operator as type-binding +- Change the fixity defined at Errors:(2:1) : autobind infixr 0 =@ +- Import a module with an autobind fixity +1/1: Building Errors2 (Errors2.idr) +Operator (=@) is an automatically binding operator (autobind) but is used +as regular operator. +Either: +- Change the fixity defined at Errors2:(2:1) to : infixr 0 =@ +- Remove the fixity and import a module that exports a compatible fixity +1/1: Building Errors3 (Errors3.idr) +Operator (=@) is a type-binding operator but is used as a regular operator. +Either: +- Change the fixity defined at Errors3:(2:1) to : typebind infixr 0 =@ +- Remove the fixity and import a module that exports a compatible fixity +1/1: Building Errors4 (Errors4.idr) +Operator (=@) is regular but is used in a type-binding position. +Either: +- Change the fixity defined at Errors4:(2:1) to : typebind infixr 0 =@ +- Remove the fixity and import a module that exports a compatible fixity +1/1: Building Errors5 (Errors5.idr) +Operator (=@) is regular but is used in an automatically-binding position. +Either: +- Change the fixity defined at Errors4:(2:1) to : autobind infixr 0 =@ +- Remove the fixity and import a module that exports a compatible fixity diff --git a/tests/idris2/operators/operators002/run b/tests/idris2/operators/operators002/run new file mode 100755 index 000000000..13baec981 --- /dev/null +++ b/tests/idris2/operators/operators002/run @@ -0,0 +1,8 @@ +. ../../../testutils.sh + +check Errors.idr +check Errors2.idr +check Errors3.idr +check Errors4.idr +check Errors5.idr + From 0f46e2c2fa1da9f2ed530f3b50672cfb5bca2ea7 Mon Sep 17 00:00:00 2001 From: Andre Videla Date: Thu, 26 Oct 2023 02:13:16 +0100 Subject: [PATCH 12/57] Display dynamic errors for binding operators Whenever the binding structure of an operator does not match the fixity in scope, the error displays what was the expected syntax and what was given. Additionally, it provides code samples of possible fixes for the current error, dynamically generated from the information at the error's site. --- src/Core/Context.idr | 4 ++ src/Core/Core.idr | 17 ++++- src/Core/TT.idr | 114 ++++++++++++++++++++++++++++++++ src/Idris/Desugar.idr | 53 +++++++-------- src/Idris/Error.idr | 67 +++++++++++++++++++ src/Idris/Parser.idr | 4 +- src/Idris/Pretty.idr | 2 +- src/Idris/Resugar.idr | 2 +- src/Idris/Syntax.idr | 101 +--------------------------- src/Idris/Syntax/Traversals.idr | 4 +- 10 files changed, 229 insertions(+), 139 deletions(-) diff --git a/src/Core/Context.idr b/src/Core/Context.idr index 04ae3e8b9..004f3ddcd 100644 --- a/src/Core/Context.idr +++ b/src/Core/Context.idr @@ -820,6 +820,8 @@ HasNames Error where full gam (InRHS fc n err) = InRHS fc <$> full gam n <*> full gam err full gam (MaybeMisspelling err xs) = MaybeMisspelling <$> full gam err <*> pure xs full gam (WarningAsError wrn) = WarningAsError <$> full gam wrn + full gam (OperatorBindingMismatch {print} fc expected actual opName rhs) + = OperatorBindingMismatch {print} fc expected actual <$> full gam opName <*> pure rhs resolved gam (Fatal err) = Fatal <$> resolved gam err resolved _ (CantConvert fc gam rho s t) @@ -911,6 +913,8 @@ HasNames Error where resolved gam (InRHS fc n err) = InRHS fc <$> resolved gam n <*> resolved gam err resolved gam (MaybeMisspelling err xs) = MaybeMisspelling <$> resolved gam err <*> pure xs resolved gam (WarningAsError wrn) = WarningAsError <$> resolved gam wrn + resolved gam (OperatorBindingMismatch {print} fc expected actual opName rhs) + = OperatorBindingMismatch {print} fc expected actual <$> resolved gam opName <*> pure rhs export HasNames Totality where diff --git a/src/Core/Core.idr b/src/Core/Core.idr index 3fe662fab..e4f8fcf91 100644 --- a/src/Core/Core.idr +++ b/src/Core/Core.idr @@ -10,6 +10,7 @@ import Data.Vect import Libraries.Data.IMaybe import Libraries.Text.PrettyPrint.Prettyprinter import Libraries.Text.PrettyPrint.Prettyprinter.Util +import Libraries.Text.PrettyPrint.Prettyprinter.Doc import Libraries.Data.Tap import public Data.IORef @@ -162,6 +163,9 @@ data Error : Type where FC -> Env Term vars -> Term vars -> (description : String) -> Error RunElabFail : Error -> Error GenericMsg : FC -> String -> Error + OperatorBindingMismatch : {a : Type} -> {print : a -> Doc ()} -> + FC -> (expectedFixity : Maybe FixityInfo) -> (use_site : OperatorLHSInfo a) -> + (opName : Name) -> (rhs : a) -> Error TTCError : TTCErrorMsg -> Error FileErr : String -> FileError -> Error CantFindPackage : String -> Error @@ -390,10 +394,13 @@ Show Error where show err show (MaybeMisspelling err ns) - = show err ++ "\nDid you mean" ++ case ns of - (n ::: []) => ": " ++ n ++ "?" - _ => " any of: " ++ showSep ", " (map show (forget ns)) ++ "?" + = show err ++ "\nDid you mean" ++ case ns of + (n ::: []) => ": " ++ n ++ "?" + _ => " any of: " ++ showSep ", " (map show (forget ns)) ++ "?" show (WarningAsError w) = show w + show (OperatorBindingMismatch fc expected actual opName rhs) + = show fc ++ ": Operator " ++ show opName ++ " is " ++ show expected + ++ " but used as a " ++ show actual ++ " operator" export getWarningLoc : Warning -> FC @@ -483,6 +490,7 @@ getErrorLoc (InLHS _ _ err) = getErrorLoc err getErrorLoc (InRHS _ _ err) = getErrorLoc err getErrorLoc (MaybeMisspelling err _) = getErrorLoc err getErrorLoc (WarningAsError warn) = Just (getWarningLoc warn) +getErrorLoc (OperatorBindingMismatch fc _ _ _ _) = Just fc export killWarningLoc : Warning -> Warning @@ -571,6 +579,9 @@ killErrorLoc (InLHS fc x err) = InLHS emptyFC x (killErrorLoc err) killErrorLoc (InRHS fc x err) = InRHS emptyFC x (killErrorLoc err) killErrorLoc (MaybeMisspelling err xs) = MaybeMisspelling (killErrorLoc err) xs killErrorLoc (WarningAsError wrn) = WarningAsError (killWarningLoc wrn) +killErrorLoc (OperatorBindingMismatch {print} fc expected actual opName rhs) + = OperatorBindingMismatch {print} emptyFC expected actual opName rhs + -- Core is a wrapper around IO that is specialised for efficiency. export diff --git a/src/Core/TT.idr b/src/Core/TT.idr index 427ee4638..59e0e4502 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -101,6 +101,120 @@ Ord Visibility where compare Public Private = GT compare Public Export = GT +public export +data Fixity = InfixL | InfixR | Infix | Prefix + +export +Show Fixity where + show InfixL = "infixl" + show InfixR = "infixr" + show Infix = "infix" + show Prefix = "prefix" + +export +Eq Fixity where + InfixL == InfixL = True + InfixR == InfixR = True + Infix == Infix = True + Prefix == Prefix = True + _ == _ = False + + +public export +data BindingModifier = NotBinding | Autobind | Typebind + +export +Eq BindingModifier where + NotBinding == NotBinding = True + Autobind == Autobind = True + Typebind == Typebind = True + _ == _ = False + +export +Show BindingModifier where + show NotBinding = "regular" + show Typebind = "typebind" + show Autobind = "autobind" + +-- A record to hold all the information about a fixity +public export +record FixityInfo where + constructor MkFixityInfo + fc : FC + vis : Visibility + bindingInfo : BindingModifier + fix : Fixity + precedence : Nat + +export +Show FixityInfo where + show fx = "fc: \{show fx.fc}, visibility: \{show fx.vis}, binding: \{show fx.bindingInfo}, fixity: \{show fx.fix}, precedence: \{show fx.precedence}" + + +export +Eq FixityInfo where + x == y = x.fc == y.fc + && x.vis == y.vis + && x.bindingInfo == y.bindingInfo + && x.fix == y.fix + && x.precedence == y.precedence + +-- Left-hand-side information for operators, carries autobind information +-- an operator can either be +-- - not autobind, a regular operator +-- - binding types, such that (nm : ty ** fn nm) desugars into (ty ** \(nm : ty) => fn nm) +-- - binding expressing with an inferred type such that +-- (nm := exp ** fn nm) desugars into (exp ** \(nm : ?) => fn nm) +-- - binding both types and expression such that +-- (nm : ty := exp ** fn nm) desugars into (exp ** \(nm : ty) => fn nm) +public export +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 + -- (nm := exp) ** fn nm + BindExpr : (name : Name) -> (expr : tm) -> OperatorLHSInfo tm + -- (nm : ty := exp) ** fn nm + BindExplicitType : (name : Name) -> (type, expr : tm) -> OperatorLHSInfo tm + +export +Show (OperatorLHSInfo tm) where + show (NoBinder lhs) = "regular" + show (BindType name ty) = "type-binding (typebind)" + show (BindExpr name expr) = "automatically-binding (autobind)" + show (BindExplicitType name type expr) = "automatically-binding (autobind)" + +%name OperatorLHSInfo opInfo + +export +Functor OperatorLHSInfo where + map f (NoBinder lhs) = NoBinder $ f lhs + map f (BindType nm lhs) = BindType nm (f lhs) + map f (BindExpr nm lhs) = BindExpr nm (f lhs) + map f (BindExplicitType nm ty lhs) = BindExplicitType nm (f ty) (f lhs) + +export +(.getLhs) : OperatorLHSInfo tm -> tm +(.getLhs) (NoBinder lhs) = lhs +(.getLhs) (BindExpr _ lhs) = lhs +(.getLhs) (BindType _ lhs) = lhs +(.getLhs) (BindExplicitType _ _ lhs) = lhs + +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 + +export +(.getBinder) : OperatorLHSInfo tm -> BindingModifier +(.getBinder) (NoBinder lhs) = NotBinding +(.getBinder) (BindType name ty) = Typebind +(.getBinder) (BindExpr name expr) = Autobind +(.getBinder) (BindExplicitType name type expr) = Autobind + public export data TotalReq = Total | CoveringOnly | PartialOK %name TotalReq treq diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 8b69efa91..8d22a3d97 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -113,7 +113,7 @@ mkPrec Prefix = Prefix checkConflictingFixities : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> (isPrefix : Bool) -> - FC -> Name -> Core (OpPrec, BindingModifier) + FC -> Name -> Core (OpPrec, Maybe FixityInfo) checkConflictingFixities isPrefix exprFC opn = do syn <- get Syn let op = nameRoot opn @@ -124,14 +124,14 @@ checkConflictingFixities isPrefix exprFC opn -- characters, if not, it must be a backticked expression. (_, [], []) => if any isOpChar (fastUnpack op) then throw (GenericMsg exprFC "Unknown operator '\{op}'") - else pure (NonAssoc 1, NotBinding) -- Backticks are non associative by default + else pure (NonAssoc 1, Nothing) -- Backticks are non associative by default -- (True, ((fxName, fx) :: _), _) => do -- in the prefix case, remove conflicts with infix (-) let extraFixities = pre ++ (filter (\(nm, _) => not $ nameRoot nm == "-") inf) unless (isCompatible fx extraFixities) $ warnConflict fxName extraFixities - pure (mkPrec fx.fix fx.precedence, NotBinding) + pure (mkPrec fx.fix fx.precedence, Just fx) -- Could not find any prefix operator fixities, there may still be conflicts with -- the infix ones. (True, [] , _) => throw (GenericMsg exprFC $ "'\{op}' is not a prefix operator") @@ -140,7 +140,7 @@ checkConflictingFixities isPrefix exprFC opn -- In the infix case, remove conflicts with prefix (-) let extraFixities = (filter (\(nm, _) => not $ nm == UN (Basic "-")) pre) ++ inf unless (isCompatible fx extraFixities) $ warnConflict fxName extraFixities - pure (mkPrec fx.fix fx.precedence, fx.bindingInfo) + pure (mkPrec fx.fix fx.precedence, Just fx) -- Could not find any infix operator fixities, there may be prefix ones (False, _, []) => throw (GenericMsg exprFC $ "'\{op}' is not an infix operator") where @@ -162,34 +162,27 @@ checkConflictingFixities isPrefix exprFC opn For example: %hide \{show fxName} """ -checkConflictingBinding : FC -> (foundFixity : BindingModifier) -> (usage : OperatorLHSInfo a) -> Core () -checkConflictingBinding fc NotBinding (NotAutobind lhs) = pure () -checkConflictingBinding fc NotBinding (BindType name ty) - = throw $ GenericMsg fc "Operator \{show name} is regular but is used in a type-binding position" -checkConflictingBinding fc NotBinding (BindExpr name expr) - = throw $ GenericMsg fc "Operator \{show name} is regular but is used in an automatically-binding position" -checkConflictingBinding fc NotBinding (BindExplicitType name type expr) - = throw $ GenericMsg fc "Operator \{show name} is regular but is used in an automatically-binding position" -checkConflictingBinding fc Autobind (NotAutobind lhs) - = throw $ GenericMsg fc "Operator is automatically-binding but is used as a regular operator" -checkConflictingBinding fc Autobind (BindType name ty) - = throw $ GenericMsg fc "Operator is automatically-binding but is used in a type-binding position" -checkConflictingBinding fc Autobind (BindExpr name expr) = pure () -checkConflictingBinding fc Autobind (BindExplicitType name type expr) = pure () -checkConflictingBinding fc Typebind (NotAutobind lhs) - = throw $ GenericMsg fc "Operator is type-binding but is used in as a regular operator" -checkConflictingBinding fc Typebind (BindType name ty) = pure() -checkConflictingBinding fc Typebind (BindExpr name expr) - = throw $ GenericMsg fc "Operator is type-binding but is used in an automatically-binding position" -checkConflictingBinding fc Typebind (BindExplicitType name type expr) - = throw $ GenericMsg fc "Operator is type-binding but is used in an automatically-binding position" +checkConflictingBinding : FC -> Name -> (foundFixity : Maybe FixityInfo) -> (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () +checkConflictingBinding fc opName foundFixity use_site rhs + = if isCompatible foundFixity use_site + then pure () + else throw $ OperatorBindingMismatch {print=byShow} fc foundFixity use_site opName rhs + where + isCompatible : Maybe FixityInfo -> OperatorLHSInfo PTerm -> Bool + isCompatible Nothing (NoBinder lhs) = True + isCompatible Nothing _ = False + isCompatible (Just fixInfo) (NoBinder lhs) = fixInfo.bindingInfo == NotBinding + isCompatible (Just fixInfo) (BindType name ty) = fixInfo.bindingInfo == Typebind + isCompatible (Just fixInfo) (BindExpr name expr) = fixInfo.bindingInfo == Autobind + isCompatible (Just fixInfo) (BindExplicitType name type expr) + = fixInfo.bindingInfo == Autobind toTokList : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> PTerm -> Core (List (Tok (OpStr, Maybe (OperatorLHSInfo PTerm)) PTerm)) toTokList (POp fc opFC l opn r) - = do (precInfo, bindInfo) <- checkConflictingFixities False fc opn - checkConflictingBinding opFC bindInfo l + = do (precInfo, fixInfo) <- checkConflictingFixities False fc opn + checkConflictingBinding opFC opn fixInfo l r rtoks <- toTokList r pure (Expr l.getLhs :: Op fc opFC (opn, Just l) precInfo :: rtoks) toTokList (PPrefixOp fc opFC opn arg) @@ -350,12 +343,12 @@ mutual [] => desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) - (POp fc opFC (NotAutobind (PRef fc (MN "arg" 0))) op arg)) + (POp fc opFC (NoBinder (PRef fc (MN "arg" 0))) op arg)) (prec :: _) => desugarB side ps (PPrefixOp fc opFC op arg) desugarB side ps (PSectionR fc opFC arg op) = desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) - (POp fc opFC (NotAutobind arg) op (PRef fc (MN "arg" 0)))) + (POp fc opFC (NoBinder arg) op (PRef fc (MN "arg" 0)))) desugarB side ps (PSearch fc depth) = pure $ ISearch fc depth desugarB side ps (PPrimVal fc (BI x)) = case !fromIntegerName of @@ -762,7 +755,7 @@ mutual r' <- desugarTree side ps r pure (IApp loc l' r') -- normal operators - desugarTree side ps (Infix loc opFC (op, Just (NotAutobind lhs)) l r) + desugarTree side ps (Infix loc opFC (op, Just (NoBinder lhs)) l r) = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc (IApp loc (IVar opFC op) l') r') diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index 5289d466d..38bbbed28 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -619,6 +619,73 @@ perrorRaw (BadRunElab fc env script desc) perrorRaw (RunElabFail e) = pure $ reflow "Error during reflection" <+> colon <++> !(perrorRaw e) perrorRaw (GenericMsg fc str) = pure $ pretty0 str <+> line <+> !(ploc fc) +perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) + = pure $ "Operator" <++> pretty0 !(getFullName opName) <++> "is" + <++> printBindingInfo expected + <++> "operator, but is used as" <++> printBindingModifier actual.getBinder + <++> "operator." + <+> line <+> !(ploc fc) + <+> "Explanation: regular, typebind and autobind operators all use a slightly different" + <++> "syntax, typebind looks like this: '(name : type)" <++> pretty0 opName + <++> "expr', autobind looks like this: '(name := expr)" <++> pretty0 opName + <++> "expr'." + <+> line <+> line + <+> "Possible solutions:" <+> line + <+> indent 1 (vsep (map ("-" <++>) + (expressionDiagnositc ++ [fixityDiagnostic, moduleDiagnostic]))) + where + moduleDiagnostic : Doc IdrisAnn + moduleDiagnostic = case expected of + Nothing => "Import a module that exports a suitable fixity." + (Just a) => "Hide or remove the fixity at" <++> byShow a.fc + <++> "and import a module that exports a compatible fixity." + infixOpName : Doc IdrisAnn + infixOpName = case expected of + Nothing => enclose "`" "`" (byShow opName) + _ => byShow opName + + displayFixityInfo : FixityInfo -> BindingModifier -> Doc IdrisAnn + displayFixityInfo (MkFixityInfo fc1 vis _ fix precedence) NotBinding + = byShow vis <++> byShow fix <++> byShow precedence <++> pretty0 opName + displayFixityInfo (MkFixityInfo _ vis _ fix precedence) usedBinder + = byShow vis <++> byShow usedBinder <++> byShow fix <++> byShow precedence <++> pretty0 opName + expressionDiagnositc : List (Doc IdrisAnn) + expressionDiagnositc = case expected of + Nothing => [] + (Just e) => let sentence = "Write the expression using" <++> byShow e.bindingInfo <++> "syntax:" + in pure $ sentence <++> enclose "'" "'" (case e.bindingInfo of + NotBinding => + reAnnotate (const Code) (p actual.getLhs) + <++> infixOpName <++> reAnnotate (const Code) (p rhs) + Autobind => + parens (maybe "_" pretty0 actual.getBoundName <++> ":=" + <++> reAnnotate (const Code) (p actual.getLhs)) + <++> infixOpName <++> reAnnotate (const Code) (p rhs) + Typebind => + parens (maybe "_" pretty0 actual.getBoundName <++> ":" + <++> reAnnotate (const Code) (p actual.getLhs)) + <++> infixOpName <++> reAnnotate (const Code) (p rhs) + ) <+> dot + + + fixityDiagnostic : Doc IdrisAnn + fixityDiagnostic = case expected of + Nothing => "Define a new fixity:" <++> "infixr 0" <++> infixOpName + (Just fix) => + "Change the fixity defined at" <++> pretty0 fix.fc <++> "to" + <++> enclose "'" "'" (displayFixityInfo fix actual.getBinder) + <+> dot + + printBindingModifier : BindingModifier -> Doc IdrisAnn + printBindingModifier NotBinding = "a regular" + printBindingModifier Typebind = "a type-binding (typebind)" + printBindingModifier Autobind = "an automatically-binding (autobind)" + + printBindingInfo : Maybe FixityInfo -> Doc IdrisAnn + printBindingInfo Nothing = "a regular" + printBindingInfo (Just x) = printBindingModifier x.bindingInfo + + perrorRaw (TTCError msg) = pure $ errorDesc (reflow "Error in TTC file" <+> colon <++> byShow msg) <++> parens "the most likely case is that the ./build directory in your current project contains files from a previous build of idris2 or the idris2 executable is from a different build than the installed .ttc files" diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 144d247a5..858381af6 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -368,7 +368,7 @@ mutual pure $ let fc = boundToFC fname (mergeBounds l r) opFC = virtualiseFC fc -- already been highlighted: we don't care - in POp fc opFC (NotAutobind l.val) (UN $ Basic "=") r.val + in POp fc opFC (NoBinder l.val) (UN $ Basic "=") r.val else fail "= not allowed") <|> (do b <- bounds $ do @@ -381,7 +381,7 @@ mutual (op, r) <- pure b.val let fc = boundToFC fname (mergeBounds l b) let opFC = boundToFC fname op - pure (POp fc opFC (NotAutobind l.val) op.val r)) + pure (POp fc opFC (NoBinder l.val) op.val r)) <|> pure l.val opExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index 7a54b7a26..80f17bf62 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -344,7 +344,7 @@ mutual group $ parens (prettyBinder nm <++> ":" <++> pretty ty <++> ":=" <++> pretty left) <++> prettyOp op <++> pretty right - prettyPrec d (POp _ _ (NotAutobind x) op y) = + prettyPrec d (POp _ _ (NoBinder x) op y) = parenthesise (d >= App) $ group $ pretty x <++> prettyOp op diff --git a/src/Idris/Resugar.idr b/src/Idris/Resugar.idr index fdc3fe03a..2f0607ba0 100644 --- a/src/Idris/Resugar.idr +++ b/src/Idris/Resugar.idr @@ -46,7 +46,7 @@ mkOp tm@(PApp fc (PApp _ (PRef opFC kn) x) y) -- to know if the name is an operator or not, it's enough to check -- that the fixity context contains the name `(++)` let rootName = UN (Basic (nameRoot raw)) - let asOp = POp fc opFC (NotAutobind (unbracketApp x)) kn (unbracketApp y) + let asOp = POp fc opFC (NoBinder (unbracketApp x)) kn (unbracketApp y) if not (null (lookupName rootName (infixes syn))) then pure asOp else case dropNS raw of diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index c1e850c5d..70a9ef3ec 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -29,63 +29,6 @@ import Parser.Lexer.Source %default covering -public export -data Fixity = InfixL | InfixR | Infix | Prefix - -export -Show Fixity where - show InfixL = "infixl" - show InfixR = "infixr" - show Infix = "infix" - show Prefix = "prefix" - -export -Eq Fixity where - InfixL == InfixL = True - InfixR == InfixR = True - Infix == Infix = True - Prefix == Prefix = True - _ == _ = False - - -public export -data BindingModifier = NotBinding | Autobind | Typebind - -export -Eq BindingModifier where - NotBinding == NotBinding = True - Autobind == Autobind = True - Typebind == Typebind = True - _ == _ = False - -export -Show BindingModifier where - show NotBinding = "not binding" - show Typebind = "typebind" - show Autobind = "autobind" - --- A record to hold all the information about a fixity -public export -record FixityInfo where - constructor MkFixityInfo - fc : FC - vis : Visibility - bindingInfo : BindingModifier - fix : Fixity - precedence : Nat - -export -Show FixityInfo where - show fx = "fc: \{show fx.fc}, visibility: \{show fx.vis}, binding: \{show fx.bindingInfo}, fixity: \{show fx.fix}, precedence: \{show fx.precedence}" - -export -Eq FixityInfo where - x == y = x.fc == y.fc - && x.vis == y.vis - && x.bindingInfo == y.bindingInfo - && x.fix == y.fix - && x.precedence == y.precedence - public export OpStr' : Type -> Type OpStr' nm = nm @@ -98,48 +41,6 @@ public export data HidingDirective = HideName Name | HideFixity Fixity Name --- Left-hand-side information for operators, carries autobind information --- an operator can either be --- - not autobind, a regular operator --- - binding types, such that (nm : ty ** fn nm) desugars into (ty ** \(nm : ty) => fn nm) --- - binding expressing with an inferred type such that --- (nm := exp ** fn nm) desugars into (exp ** \(nm : ?) => fn nm) --- - binding both types and expression such that --- (nm : ty := exp ** fn nm) desugars into (exp ** \(nm : ty) => fn nm) -public export -data OperatorLHSInfo : tm -> Type where - -- Traditional operator wihtout binding, carries the lhs - NotAutobind : (lhs : tm) -> OperatorLHSInfo tm - -- (nm : ty) ** fn x - BindType : (name : Name) -> (ty : tm) -> OperatorLHSInfo tm - -- (nm := exp) ** fn nm - BindExpr : (name : Name) -> (expr : tm) -> OperatorLHSInfo tm - -- (nm : ty := exp) ** fn nm - BindExplicitType : (name : Name) -> (type, expr : tm) -> OperatorLHSInfo tm - -export -Show (OperatorLHSInfo tm) where - show (NotAutobind lhs) = "Regular operator" - show (BindType name ty) = "Type-binding operator" - show (BindExpr name expr) = "Automatically-binding operator" - show (BindExplicitType name type expr) = "Automatically-binding operator (explicit)" - -%name OperatorLHSInfo opInfo - -export -Functor OperatorLHSInfo where - map f (NotAutobind lhs) = NotAutobind $ f lhs - map f (BindType nm lhs) = BindType nm (f lhs) - map f (BindExpr nm lhs) = BindExpr nm (f lhs) - map f (BindExplicitType nm ty lhs) = BindExplicitType nm (f ty) (f lhs) - -export -(.getLhs) : OperatorLHSInfo tm -> tm -(.getLhs) (NotAutobind lhs) = lhs -(.getLhs) (BindExpr _ lhs) = lhs -(.getLhs) (BindType _ lhs) = lhs -(.getLhs) (BindExplicitType _ _ lhs) = lhs - mutual ||| Source language as produced by the parser @@ -869,7 +770,7 @@ parameters {0 nm : Type} (toName : nm -> Name) showPTermPrec d (PDotted _ p) = "." ++ showPTermPrec d p showPTermPrec _ (PImplicit _) = "_" showPTermPrec _ (PInfer _) = "?" - showPTermPrec d (POp _ _ (NotAutobind left) op right) + showPTermPrec d (POp _ _ (NoBinder left) op right) = showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right showPTermPrec d (POp _ _ (BindType nm left) op right) = "(" ++ show nm ++ " : " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" diff --git a/src/Idris/Syntax/Traversals.idr b/src/Idris/Syntax/Traversals.idr index 1e0edfefe..b2aefb1e1 100644 --- a/src/Idris/Syntax/Traversals.idr +++ b/src/Idris/Syntax/Traversals.idr @@ -99,9 +99,9 @@ mapPTermM f = goPTerm where >>= f goPTerm t@(PImplicit _) = f t goPTerm t@(PInfer _) = f t - goPTerm (POp fc opFC (NotAutobind left) op right) = + goPTerm (POp fc opFC (NoBinder left) op right) = POp fc opFC - <$> (NotAutobind <$> goPTerm left) + <$> (NoBinder <$> goPTerm left) <*> pure op <*> goPTerm right >>= f From 07c42c22a56a7aa8956b561d109aaa6469522818 Mon Sep 17 00:00:00 2001 From: Andre Videla Date: Thu, 26 Oct 2023 02:24:36 +0100 Subject: [PATCH 13/57] update error messages --- tests/idris2/error/perror007/expected | 6 +- tests/idris2/error/perror014/expected | 2 +- .../interactive/interactive028/expected | 6 +- tests/idris2/literate017/project-output.ipkg | 47 +++++++++ tests/idris2/operators/operators002/expected | 98 ++++++++++++++---- tests/idris2/pkg010/expected | 6 ++ tests/idris2/pkg013/expected | 1 + tests/idris2/pkg014/expected | 1 + tests/idris2/pkg017/expected | 22 ++++ tests/node/node004/test.buf | Bin 0 -> 100 bytes tests/node/node017/testdir/test.txt | 1 + tests/node/node018/testout.txt | 2 + 12 files changed, 162 insertions(+), 30 deletions(-) create mode 100644 tests/idris2/literate017/project-output.ipkg create mode 100644 tests/idris2/pkg010/expected create mode 100644 tests/idris2/pkg013/expected create mode 100644 tests/idris2/pkg014/expected create mode 100644 tests/idris2/pkg017/expected create mode 100644 tests/node/node004/test.buf create mode 100644 tests/node/node017/testdir/test.txt create mode 100644 tests/node/node018/testout.txt diff --git a/tests/idris2/error/perror007/expected b/tests/idris2/error/perror007/expected index 4b8f8536d..50c05993d 100644 --- a/tests/idris2/error/perror007/expected +++ b/tests/idris2/error/perror007/expected @@ -1,6 +1,6 @@ 1/1: Building StrError1 (StrError1.idr) Error: Couldn't parse any alternatives: -1: Expected 'case', 'if', 'do', application or operator expression. +1: Expected '('. 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 'case', 'if', 'do', application or operator expression. +1: Expected '('. 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 'case', 'if', 'do', application or operator expression. +1: Expected '('. StrError3:2:7--2:8 1 | foo : String diff --git a/tests/idris2/error/perror014/expected b/tests/idris2/error/perror014/expected index 0bf8c4db7..199fd3290 100644 --- a/tests/idris2/error/perror014/expected +++ b/tests/idris2/error/perror014/expected @@ -1,6 +1,6 @@ 1/1: Building ParseList (ParseList.idr) Error: Couldn't parse any alternatives: -1: Expected 'case', 'if', 'do', application or operator expression. +1: Expected '('. ParseList:8:5--8:6 4 | , 1 diff --git a/tests/idris2/interactive/interactive028/expected b/tests/idris2/interactive/interactive028/expected index 529623d7f..0c511e535 100644 --- a/tests/idris2/interactive/interactive028/expected +++ b/tests/idris2/interactive/interactive028/expected @@ -1,9 +1,9 @@ Main> Couldn't parse any alternatives: -1: Expected 'case', 'if', 'do', application or operator expression. +1: Expected name. -(Interactive):1:4--1:5 +(Interactive):1:5--1:6 1 | :t (3 : Nat) - ^ + ^ ... (54 others) Main> Expected string begin. diff --git a/tests/idris2/literate017/project-output.ipkg b/tests/idris2/literate017/project-output.ipkg new file mode 100644 index 000000000..78ec8933a --- /dev/null +++ b/tests/idris2/literate017/project-output.ipkg @@ -0,0 +1,47 @@ +package project +-- version = +-- authors = +-- maintainers = +-- license = +-- brief = +-- readme = +-- homepage = +-- sourceloc = +-- bugtracker = + +-- the Idris2 version required (e.g. langversion >= 0.5.1) +-- langversion + +-- packages to add to search path +-- depends = + +-- modules to install +modules = A.A, A.AB, B.B, B.BA, Main + +-- main file (i.e. file to load at REPL) +-- main = + +-- name of executable +-- executable = +-- opts = +sourcedir = "src/" +-- builddir = +-- outputdir = + +-- script to run before building +-- prebuild = + +-- script to run after building +-- postbuild = + +-- script to run after building, before installing +-- preinstall = + +-- script to run after installing +-- postinstall = + +-- script to run before cleaning +-- preclean = + +-- script to run after cleaning +-- postclean = \ No newline at end of file diff --git a/tests/idris2/operators/operators002/expected b/tests/idris2/operators/operators002/expected index 5cfceb8fd..8451e1d0b 100644 --- a/tests/idris2/operators/operators002/expected +++ b/tests/idris2/operators/operators002/expected @@ -1,28 +1,80 @@ 1/1: Building Errors (Errors.idr) -Operator (=@) is a type-binding operator but is used in -automatically binding position. -Either: -- Write (x : ty) =@ S x to use the operator as type-binding -- Change the fixity defined at Errors:(2:1) : autobind infixr 0 =@ -- Import a module with an autobind fixity +Error: Operator =@ is a type-binding (typebind) operator, but is used as an automatically-binding (autobind) operator. + +Errors:8:19--8:21 + 4 | 0 (=@) : (a : Type) -> (a -> Type) -> Type + 5 | (=@) a f = (1 x : a) -> f x + 6 | + 7 | data S : {ty : Type} -> (x : ty) -> Type where + 8 | MkS : (x := ty) =@ S x + ^^ +Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) =@ expr', autobind looks like this: '(name := expr) =@ expr'. + +Possible solutions: + - Write the expression using typebind syntax: '(x : ty) =@ S x'. + - Change the fixity defined at Errors:2:10--2:21 to 'export autobind infixr 0 =@'. + - Hide or remove the fixity at Errors:2:10--2:21 and import a module that exports a compatible fixity. 1/1: Building Errors2 (Errors2.idr) -Operator (=@) is an automatically binding operator (autobind) but is used -as regular operator. -Either: -- Change the fixity defined at Errors2:(2:1) to : infixr 0 =@ -- Remove the fixity and import a module that exports a compatible fixity +Error: Operator =@ is an automatically-binding (autobind) operator, but is used as a regular operator. + +Errors2:7:29--7:31 + 3 | + 4 | 0 (=@) : (a : Type) -> (a -> Type) -> Type + 5 | (=@) a f = (1 x : a) -> f x + 6 | + 7 | wrongId : {0 a : Type} -> a =@ a + ^^ +Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) =@ expr', autobind looks like this: '(name := expr) =@ expr'. + +Possible solutions: + - Write the expression using autobind syntax: '(_ := a) =@ a'. + - Change the fixity defined at Errors2:2:10--2:21 to 'export infixr 0 =@'. + - Hide or remove the fixity at Errors2:2:10--2:21 and import a module that exports a compatible fixity. 1/1: Building Errors3 (Errors3.idr) -Operator (=@) is a type-binding operator but is used as a regular operator. -Either: -- Change the fixity defined at Errors3:(2:1) to : typebind infixr 0 =@ -- Remove the fixity and import a module that exports a compatible fixity +Error: Operator =@ is a type-binding (typebind) operator, but is used as a regular operator. + +Errors3:7:29--7:31 + 3 | + 4 | 0 (=@) : (a : Type) -> (a -> Type) -> Type + 5 | (=@) a f = (1 x : a) -> f x + 6 | + 7 | wrongId : {0 a : Type} -> a =@ a + ^^ +Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) =@ expr', autobind looks like this: '(name := expr) =@ expr'. + +Possible solutions: + - Write the expression using typebind syntax: '(_ : a) =@ a'. + - Change the fixity defined at Errors3:2:10--2:21 to 'export infixr 0 =@'. + - Hide or remove the fixity at Errors3:2:10--2:21 and import a module that exports a compatible fixity. 1/1: Building Errors4 (Errors4.idr) -Operator (=@) is regular but is used in a type-binding position. -Either: -- Change the fixity defined at Errors4:(2:1) to : typebind infixr 0 =@ -- Remove the fixity and import a module that exports a compatible fixity +Error: Operator =@ is a regular operator, but is used as a type-binding (typebind) operator. + +Errors4:9:18--9:20 + 5 | (=@) a f = (1 x : a) -> f x + 6 | + 7 | + 8 | data S : {ty : Type} -> (x : ty) -> Type where + 9 | MkS : (x : ty) =@ S x + ^^ +Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) =@ expr', autobind looks like this: '(name := expr) =@ expr'. + +Possible solutions: + - Write the expression using regular syntax: 'ty =@ S x'. + - Change the fixity defined at Errors4:2:1--2:12 to 'export typebind infixr 0 =@'. + - Hide or remove the fixity at Errors4:2:1--2:12 and import a module that exports a compatible fixity. 1/1: Building Errors5 (Errors5.idr) -Operator (=@) is regular but is used in an automatically-binding position. -Either: -- Change the fixity defined at Errors4:(2:1) to : autobind infixr 0 =@ -- Remove the fixity and import a module that exports a compatible fixity +Error: Operator =@ is a regular operator, but is used as an automatically-binding (autobind) operator. + +Errors5:10:19--10:21 + 06 | (=@) a f = (1 x : a) -> f x + 07 | + 08 | + 09 | data S : {ty : Type} -> (x : ty) -> Type where + 10 | MkS : (x := ty) =@ S x + ^^ +Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) =@ expr', autobind looks like this: '(name := expr) =@ expr'. + +Possible solutions: + - Write the expression using regular syntax: 'ty =@ S x'. + - Change the fixity defined at Errors5:3:1--3:12 to 'export autobind infixr 0 =@'. + - Hide or remove the fixity at Errors5:3:1--3:12 and import a module that exports a compatible fixity. diff --git a/tests/idris2/pkg010/expected b/tests/idris2/pkg010/expected new file mode 100644 index 000000000..acc095d89 --- /dev/null +++ b/tests/idris2/pkg010/expected @@ -0,0 +1,6 @@ +1/1: Building Main (Main.idr) +Installing /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/build/ttc/Main.ttc to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 +Installing /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/build/ttc/Main.ttm to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 +Installing /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/build/ttc/Main.ttc to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 +Installing /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/build/ttc/Main.ttm to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 +Installing package file for testpkg to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 diff --git a/tests/idris2/pkg013/expected b/tests/idris2/pkg013/expected new file mode 100644 index 000000000..2ecdeeb6a --- /dev/null +++ b/tests/idris2/pkg013/expected @@ -0,0 +1 @@ +LOG package.depends:10: all depends: ["/home/andre/Programming/Idris/Compiler/tests/idris2/pkg013/depends/bar-0.7.2", "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg013/depends/foo-0.1.0"] diff --git a/tests/idris2/pkg014/expected b/tests/idris2/pkg014/expected new file mode 100644 index 000000000..26c79cd81 --- /dev/null +++ b/tests/idris2/pkg014/expected @@ -0,0 +1 @@ +LOG package.depends:10: all depends: ["/home/andre/Programming/Idris/Compiler/tests/idris2/pkg014/depends/bar-0.1.0", "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg014/depends/baz-0.2.0", "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg014/depends/foo-0.2.0"] diff --git a/tests/idris2/pkg017/expected b/tests/idris2/pkg017/expected new file mode 100644 index 000000000..139b610ff --- /dev/null +++ b/tests/idris2/pkg017/expected @@ -0,0 +1,22 @@ +(:protocol-version 2 1) +(:return (:ok "Current working directory is \"/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/b1\"") 1) +(:write-string "1/1: Building B1 (src/B1.idr)" 2) +(:return (:ok ()) 2) +(:return (:ok "1" ((0 1 ((:decor :data))))) 4) +(:return (:ok (("A.i" (:filename "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/prefix/idris2-0.6.0/a1-0/A.idr") (:start 2 0) (:end 3 7)))) 5) +(:return (:ok "Current working directory is \"/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/b2\"") 6) +(:write-string "1/1: Building B2 (src/B2.idr)" 7) +(:return (:ok ()) 7) +(:return (:ok "2" ((0 1 ((:decor :data))))) 8) +(:return (:ok (("A.i" (:filename "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/prefix/idris2-0.6.0/a2-0/A.idr") (:start 2 0) (:end 3 7)))) 9) +he file is done, aborting +(:protocol-version 2 1) +(:return (:ok "Current working directory is \"/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/b2\"") 1) +(:return (:ok ()) 2) +(:return (:ok "2" ((0 1 ((:decor :data))))) 4) +(:return (:ok (("A.i" (:filename "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/prefix/idris2-0.6.0/a2-0/A.idr") (:start 2 0) (:end 3 7)))) 5) +(:return (:ok "Current working directory is \"/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/b1\"") 6) +(:return (:ok ()) 7) +(:return (:ok "1" ((0 1 ((:decor :data))))) 8) +(:return (:ok (("A.i" (:filename "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/prefix/idris2-0.6.0/a1-0/A.idr") (:start 2 0) (:end 3 7)))) 9) +he file is done, aborting diff --git a/tests/node/node004/test.buf b/tests/node/node004/test.buf new file mode 100644 index 0000000000000000000000000000000000000000..c0bfaac24e0e259fb7e91d3a87ed08ce848ae498 GIT binary patch literal 100 scmZRGV_;x#1OkR?k#+75XN5a3FnFZq Date: Wed, 1 Nov 2023 23:12:27 +0000 Subject: [PATCH 14/57] fix repl documentation --- src/Idris/Doc/Keywords.idr | 13 +++++++------ 1 file changed, 7 insertions(+), 6 deletions(-) diff --git a/src/Idris/Doc/Keywords.idr b/src/Idris/Doc/Keywords.idr index b79d64485..0f1d4a408 100644 --- a/src/Idris/Doc/Keywords.idr +++ b/src/Idris/Doc/Keywords.idr @@ -469,18 +469,19 @@ autobindDoc = """ `autobind` differs from `typebind` in the syntax it allows and the way it desugars. Delcaring an operator as `autobind infixr 0 =|>` allows - you to use it with the syntax `(x := e =|> f x)` or `(x : t := e =|> f x)` + you to use it with the syntax `(x := e) =|> f x`, `(x : t := e) =|> f x`. + You will need to use `autobind` instead of `typebind` whenever the type of the lambda of the function called by the operator has a type - that is not equal to the argument given on the left side of the operator + that is not equal to the argument given on the left side of the operator. `autobind` only applies to `infixr` operators and cannot be used as operator sections. `autobind` operators are desugared as a lambda: - `(x := expr =|> fn x)` -> `(expr =@ (\x : ? =|> fn x))` - `(x : ty := expr =|> fn x)` -> `(expr =@ (\x : ty =|> fn x))` + `(x := expr) =|> fn x` -> `(expr =@ (\x : ? =|> fn x))` + `(x : ty := expr) =|> fn x` -> `(expr =@ (\x : ty =|> fn x))` """ typebindDoc : Doc IdrisDocAnn typebindDoc = """ @@ -498,13 +499,13 @@ typebindDoc = """ way. Start by defining `typebind infixr 1 =@`, and then you can use it - like so: `(n : Nat =@ f n)` or `(n : Nat) =@ f n` + like so: `(n : Nat) =@ f n` `typebind` only applies to `infixr` operators and cannot be used as operator sections. `typebind` operators are desugared as a lambda: - `(x : ty =@ fn x)` -> `(ty =@ (\x : ty =@ fn x))` + `(x : ty) =@ fn x` -> `(ty =@ (\x : ty =@ fn x))` """ rewriteeq : Doc IdrisDocAnn From a7a25914b72b7c03a3b20058c0a46bb524696e71 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 11 Nov 2023 19:15:47 +0000 Subject: [PATCH 15/57] Allow binding operators on LHS When a binding operator is found on the LHS, we don't check if the first argument is a binder. --- src/Idris/Desugar.idr | 36 +++++++++++--------- src/Idris/Parser.idr | 10 +++--- tests/idris2/operators/operators001/Test.idr | 21 ++++++++++++ 3 files changed, 45 insertions(+), 22 deletions(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 8d22a3d97..5d1fce6a1 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -162,8 +162,9 @@ checkConflictingFixities isPrefix exprFC opn For example: %hide \{show fxName} """ -checkConflictingBinding : FC -> Name -> (foundFixity : Maybe FixityInfo) -> (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () -checkConflictingBinding fc opName foundFixity use_site rhs +checkConflictingBinding : Side -> FC -> Name -> (foundFixity : Maybe FixityInfo) -> (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () +checkConflictingBinding LHS fc opName foundFixity use_site rhs = pure () -- don't check if we're on the LHS +checkConflictingBinding side fc opName foundFixity use_site rhs = if isCompatible foundFixity use_site then pure () else throw $ OperatorBindingMismatch {print=byShow} fc foundFixity use_site opName rhs @@ -177,19 +178,20 @@ checkConflictingBinding fc opName foundFixity use_site rhs isCompatible (Just fixInfo) (BindExplicitType name type expr) = fixInfo.bindingInfo == Autobind -toTokList : {auto s : Ref Syn SyntaxInfo} -> - {auto c : Ref Ctxt Defs} -> - PTerm -> Core (List (Tok (OpStr, Maybe (OperatorLHSInfo PTerm)) PTerm)) -toTokList (POp fc opFC l opn r) - = do (precInfo, fixInfo) <- checkConflictingFixities False fc opn - checkConflictingBinding opFC opn fixInfo l r - rtoks <- toTokList r - pure (Expr l.getLhs :: Op fc opFC (opn, Just l) precInfo :: rtoks) -toTokList (PPrefixOp fc opFC opn arg) - = do (precInfo, _) <- checkConflictingFixities True fc opn - rtoks <- toTokList arg - pure (Op fc opFC (opn, Nothing) precInfo :: rtoks) -toTokList t = pure [Expr t] +parameters (side : Side) + toTokList : {auto s : Ref Syn SyntaxInfo} -> + {auto c : Ref Ctxt Defs} -> + PTerm -> Core (List (Tok (OpStr, Maybe (OperatorLHSInfo PTerm)) PTerm)) + toTokList (POp fc opFC l opn r) + = do (precInfo, fixInfo) <- checkConflictingFixities False fc opn + checkConflictingBinding side opFC opn fixInfo l r + rtoks <- toTokList r + pure (Expr l.getLhs :: Op fc opFC (opn, Just l) precInfo :: rtoks) + toTokList (PPrefixOp fc opFC opn arg) + = do (precInfo, _) <- checkConflictingFixities True fc opn + rtoks <- toTokList arg + pure (Op fc opFC (opn, Nothing) precInfo :: rtoks) + toTokList t = pure [Expr t] record BangData where constructor MkBangData @@ -330,10 +332,10 @@ mutual apply (IVar fc (UN $ Basic "~=~")) [l', r']] desugarB side ps (PBracketed fc e) = desugarB side ps e desugarB side ps (POp fc opFC l op r) - = do ts <- toTokList (POp fc opFC l op r) + = do ts <- toTokList side (POp fc opFC l op r) desugarTree side ps !(parseOps ts) desugarB side ps (PPrefixOp fc opFC op arg) - = do ts <- toTokList (PPrefixOp fc opFC op arg) + = do ts <- toTokList side (PPrefixOp fc opFC op arg) desugarTree side ps !(parseOps ts) desugarB side ps (PSectionL fc opFC op arg) = do syn <- get Syn diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 858381af6..328ef09dc 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -105,14 +105,14 @@ pnoeq = { eqOK := False } export pdef : ParseOpts -pdef = MkParseOpts True True +pdef = MkParseOpts {eqOK = True, withOK = True} pnowith : ParseOpts -pnowith = MkParseOpts True False +pnowith = MkParseOpts {eqOK = True, withOK = False} export plhs : ParseOpts -plhs = MkParseOpts False False +plhs = MkParseOpts {eqOK = False, withOK = False} %hide Prelude.(>>) %hide Prelude.(>>=) @@ -385,8 +385,8 @@ mutual <|> pure l.val opExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm - opExpr q fname indents = opExprBase q fname indents - <|> autobindOp q fname indents + opExpr q fname indents = autobindOp q fname indents <|> opExprBase q fname indents + dpairType : OriginDesc -> WithBounds t -> IndentInfo -> Rule PTerm dpairType fname start indents diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index e7dcfc9d4..a09045124 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -56,3 +56,24 @@ program = (n := 3) `MyLet` 2 + n program2 : Nat program2 = (n : Nat := 3) `MyLet` 2 + n + +typebind infixr 6 |> + +record Container where + constructor (|>) + shape : Type + position : shape -> Type + +typebind infixr 7 @@ + +record (@@) (x : Type) (y : x -> Type) where + constructor PairUp + fst : x + snd : y fst + +compose : Container -> Container -> Container +compose (a |> a') (b |> b') = + (x : (y : a) @@ a' y -> b) |> + (y : a' x.fst) @@ + b' (x.snd y) + From f079781c17659905a5270df4e3e95c25b634cf9a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sun, 12 Nov 2023 00:43:14 +0000 Subject: [PATCH 16/57] fix tests --- tests/idris2/operators/operators001/Test.idr | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index a09045124..3b866339a 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -57,14 +57,14 @@ program = (n := 3) `MyLet` 2 + n program2 : Nat program2 = (n : Nat := 3) `MyLet` 2 + n -typebind infixr 6 |> +typebind infixr 0 |> record Container where constructor (|>) shape : Type position : shape -> Type -typebind infixr 7 @@ +typebind infixr 0 @@ record (@@) (x : Type) (y : x -> Type) where constructor PairUp @@ -73,7 +73,7 @@ record (@@) (x : Type) (y : x -> Type) where compose : Container -> Container -> Container compose (a |> a') (b |> b') = - (x : (y : a) @@ a' y -> b) |> + (x : (y : a) @@ (a' y -> b)) |> (y : a' x.fst) @@ b' (x.snd y) From a73522ec663cbdd654a582c8694f7dc57c960cf7 Mon Sep 17 00:00:00 2001 From: Andre Videla Date: Mon, 27 Nov 2023 19:23:19 +0000 Subject: [PATCH 17/57] cleanup tests --- tests/idris2/literate017/project-output.ipkg | 47 ------------------- tests/idris2/pkg010/expected | 6 --- tests/idris2/pkg013/expected | 1 - tests/idris2/pkg014/expected | 1 - tests/idris2/pkg017/expected | 22 --------- tests/node/node004/test.buf | Bin 100 -> 0 bytes tests/node/node017/testdir/test.txt | 1 - tests/node/node018/testout.txt | 2 - 8 files changed, 80 deletions(-) delete mode 100644 tests/idris2/literate017/project-output.ipkg delete mode 100644 tests/idris2/pkg010/expected delete mode 100644 tests/idris2/pkg013/expected delete mode 100644 tests/idris2/pkg014/expected delete mode 100644 tests/idris2/pkg017/expected delete mode 100644 tests/node/node004/test.buf delete mode 100644 tests/node/node017/testdir/test.txt delete mode 100644 tests/node/node018/testout.txt diff --git a/tests/idris2/literate017/project-output.ipkg b/tests/idris2/literate017/project-output.ipkg deleted file mode 100644 index 78ec8933a..000000000 --- a/tests/idris2/literate017/project-output.ipkg +++ /dev/null @@ -1,47 +0,0 @@ -package project --- version = --- authors = --- maintainers = --- license = --- brief = --- readme = --- homepage = --- sourceloc = --- bugtracker = - --- the Idris2 version required (e.g. langversion >= 0.5.1) --- langversion - --- packages to add to search path --- depends = - --- modules to install -modules = A.A, A.AB, B.B, B.BA, Main - --- main file (i.e. file to load at REPL) --- main = - --- name of executable --- executable = --- opts = -sourcedir = "src/" --- builddir = --- outputdir = - --- script to run before building --- prebuild = - --- script to run after building --- postbuild = - --- script to run after building, before installing --- preinstall = - --- script to run after installing --- postinstall = - --- script to run before cleaning --- preclean = - --- script to run after cleaning --- postclean = \ No newline at end of file diff --git a/tests/idris2/pkg010/expected b/tests/idris2/pkg010/expected deleted file mode 100644 index acc095d89..000000000 --- a/tests/idris2/pkg010/expected +++ /dev/null @@ -1,6 +0,0 @@ -1/1: Building Main (Main.idr) -Installing /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/build/ttc/Main.ttc to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 -Installing /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/build/ttc/Main.ttm to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 -Installing /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/build/ttc/Main.ttc to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 -Installing /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/build/ttc/Main.ttm to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 -Installing package file for testpkg to /home/andre/Programming/Idris/Compiler/tests/idris2/pkg010/currently/nonexistent/dir/idris2-0.6.0/testpkg-0 diff --git a/tests/idris2/pkg013/expected b/tests/idris2/pkg013/expected deleted file mode 100644 index 2ecdeeb6a..000000000 --- a/tests/idris2/pkg013/expected +++ /dev/null @@ -1 +0,0 @@ -LOG package.depends:10: all depends: ["/home/andre/Programming/Idris/Compiler/tests/idris2/pkg013/depends/bar-0.7.2", "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg013/depends/foo-0.1.0"] diff --git a/tests/idris2/pkg014/expected b/tests/idris2/pkg014/expected deleted file mode 100644 index 26c79cd81..000000000 --- a/tests/idris2/pkg014/expected +++ /dev/null @@ -1 +0,0 @@ -LOG package.depends:10: all depends: ["/home/andre/Programming/Idris/Compiler/tests/idris2/pkg014/depends/bar-0.1.0", "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg014/depends/baz-0.2.0", "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg014/depends/foo-0.2.0"] diff --git a/tests/idris2/pkg017/expected b/tests/idris2/pkg017/expected deleted file mode 100644 index 139b610ff..000000000 --- a/tests/idris2/pkg017/expected +++ /dev/null @@ -1,22 +0,0 @@ -(:protocol-version 2 1) -(:return (:ok "Current working directory is \"/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/b1\"") 1) -(:write-string "1/1: Building B1 (src/B1.idr)" 2) -(:return (:ok ()) 2) -(:return (:ok "1" ((0 1 ((:decor :data))))) 4) -(:return (:ok (("A.i" (:filename "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/prefix/idris2-0.6.0/a1-0/A.idr") (:start 2 0) (:end 3 7)))) 5) -(:return (:ok "Current working directory is \"/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/b2\"") 6) -(:write-string "1/1: Building B2 (src/B2.idr)" 7) -(:return (:ok ()) 7) -(:return (:ok "2" ((0 1 ((:decor :data))))) 8) -(:return (:ok (("A.i" (:filename "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/prefix/idris2-0.6.0/a2-0/A.idr") (:start 2 0) (:end 3 7)))) 9) -he file is done, aborting -(:protocol-version 2 1) -(:return (:ok "Current working directory is \"/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/b2\"") 1) -(:return (:ok ()) 2) -(:return (:ok "2" ((0 1 ((:decor :data))))) 4) -(:return (:ok (("A.i" (:filename "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/prefix/idris2-0.6.0/a2-0/A.idr") (:start 2 0) (:end 3 7)))) 5) -(:return (:ok "Current working directory is \"/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/b1\"") 6) -(:return (:ok ()) 7) -(:return (:ok "1" ((0 1 ((:decor :data))))) 8) -(:return (:ok (("A.i" (:filename "/home/andre/Programming/Idris/Compiler/tests/idris2/pkg017/prefix/idris2-0.6.0/a1-0/A.idr") (:start 2 0) (:end 3 7)))) 9) -he file is done, aborting diff --git a/tests/node/node004/test.buf b/tests/node/node004/test.buf deleted file mode 100644 index c0bfaac24e0e259fb7e91d3a87ed08ce848ae498..0000000000000000000000000000000000000000 GIT binary patch literal 0 HcmV?d00001 literal 100 scmZRGV_;x#1OkR?k#+75XN5a3FnFZq Date: Sat, 4 Nov 2023 19:53:28 +0000 Subject: [PATCH 18/57] reinstate the good error messages --- src/Core/TT.idr | 13 ++-- src/Idris/Desugar.idr | 6 +- src/Idris/Parser.idr | 60 ++++++++++++++----- src/Idris/Pretty.idr | 6 +- tests/idris2/error/perror007/expected | 6 +- tests/idris2/error/perror014/expected | 2 +- tests/idris2/error007/expected | 10 ++++ tests/idris2/error014/expected | 21 +++++++ .../interactive/interactive028/expected | 6 +- 9 files changed, 96 insertions(+), 34 deletions(-) create mode 100644 tests/idris2/error007/expected create mode 100644 tests/idris2/error014/expected diff --git a/src/Core/TT.idr b/src/Core/TT.idr index 59e0e4502..b3e1039cf 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -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 diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 5d1fce6a1..3f9fd2ed4 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -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" diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 328ef09dc..db4dec674 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -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 diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index 80f17bf62..ca91e4ded 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -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) = diff --git a/tests/idris2/error/perror007/expected b/tests/idris2/error/perror007/expected index 50c05993d..4b8f8536d 100644 --- a/tests/idris2/error/perror007/expected +++ b/tests/idris2/error/perror007/expected @@ -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 diff --git a/tests/idris2/error/perror014/expected b/tests/idris2/error/perror014/expected index 199fd3290..0bf8c4db7 100644 --- a/tests/idris2/error/perror014/expected +++ b/tests/idris2/error/perror014/expected @@ -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 diff --git a/tests/idris2/error007/expected b/tests/idris2/error007/expected new file mode 100644 index 000000000..2e9fd2b47 --- /dev/null +++ b/tests/idris2/error007/expected @@ -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 + ^^^^^^^^ + diff --git a/tests/idris2/error014/expected b/tests/idris2/error014/expected new file mode 100644 index 000000000..03ed7e741 --- /dev/null +++ b/tests/idris2/error014/expected @@ -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 + ^ + diff --git a/tests/idris2/interactive/interactive028/expected b/tests/idris2/interactive/interactive028/expected index 0c511e535..529623d7f 100644 --- a/tests/idris2/interactive/interactive028/expected +++ b/tests/idris2/interactive/interactive028/expected @@ -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. From 1633b7ec1cdbd4cc6f65c724c12e9e137a6cded1 Mon Sep 17 00:00:00 2001 From: Andre Videla Date: Wed, 6 Dec 2023 13:54:36 +0000 Subject: [PATCH 19/57] reverse parsing of pi --- src/Idris/Parser.idr | 39 +++++++++++++-------------------------- 1 file changed, 13 insertions(+), 26 deletions(-) diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index db4dec674..9abeb7c69 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -346,9 +346,11 @@ mutual = do boundName <- bounds (UN . Basic <$> decoratedSimpleBinderName fname) opBinderTypes fname indents boundName - autobindOp : WithBounds (OperatorLHSInfo PTerm) -> ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm - autobindOp binder q fname indents - = do op <- bounds iOperator + autobindOp : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm + autobindOp q fname indents + = do binder <- bounds $ parens fname (opBinder fname indents) + continue indents + op <- bounds iOperator commit e <- bounds (expr q fname indents) pure (POp (boundToFC fname $ mergeBounds binder e) @@ -384,7 +386,8 @@ mutual <|> pure l.val opExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm - opExpr q fname indents = autobindOp q fname indents <|> opExprBase q fname indents + opExpr q fname indents = autobindOp q fname indents + <|> opExprBase q fname indents dpairType : OriginDesc -> WithBounds t -> IndentInfo -> Rule PTerm dpairType fname start indents @@ -705,9 +708,10 @@ mutual = (decoratedSymbol fname "->" $> Explicit) <|> (decoratedSymbol fname "=>" $> AutoImplicit) - explicitPi : WithBounds PiBindList -> OriginDesc -> IndentInfo -> Rule PTerm - explicitPi b fname indents - = do exp <- mustWorkBecause b.bounds "Cannot return a named argument" + 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" $ bindSymbol fname scope <- mustWork $ typeExpr pdef fname indents pure (pibindAll fname exp b.val scope) @@ -725,24 +729,6 @@ mutual 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 @@ -1051,7 +1037,8 @@ mutual <|> defaultImplicitPi fname indents <|> forall_ fname indents <|> implicitPi fname indents - <|> piOrAutobind fname indents + <|> autobindOp pdef fname indents + <|> explicitPi fname indents <|> lam fname indents typeExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm From 38fbad17f8321d1350f94308e4cd2daf0f7a44d8 Mon Sep 17 00:00:00 2001 From: Andre Videla Date: Wed, 6 Dec 2023 14:51:18 +0000 Subject: [PATCH 20/57] update test --- tests/idris2/error007/expected | 10 ---------- tests/idris2/error014/expected | 21 --------------------- 2 files changed, 31 deletions(-) delete mode 100644 tests/idris2/error007/expected delete mode 100644 tests/idris2/error014/expected diff --git a/tests/idris2/error007/expected b/tests/idris2/error007/expected deleted file mode 100644 index 2e9fd2b47..000000000 --- a/tests/idris2/error007/expected +++ /dev/null @@ -1,10 +0,0 @@ -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 - ^^^^^^^^ - diff --git a/tests/idris2/error014/expected b/tests/idris2/error014/expected deleted file mode 100644 index 03ed7e741..000000000 --- a/tests/idris2/error014/expected +++ /dev/null @@ -1,21 +0,0 @@ -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 - ^ - From 3d2c85425760c3119e65d8fc0b6b0499844aaea8 Mon Sep 17 00:00:00 2001 From: Andre Videla Date: Wed, 6 Dec 2023 16:25:39 +0000 Subject: [PATCH 21/57] update docs --- CHANGELOG.md | 3 + docs/source/reference/operators.rst | 176 ++++++++++++++++++++++++++++ 2 files changed, 179 insertions(+) create mode 100644 docs/source/reference/operators.rst diff --git a/CHANGELOG.md b/CHANGELOG.md index 5585c41c9..e8dedf7f0 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -34,6 +34,9 @@ This CHANGELOG describes the history of already-released versions. Please see [C referred to by another definitions, and in which function currently elaborator is. These features together give an ability to inspect whether particular expressions are recursive (including mutual recursion). +* Autobind and Typebind modifier on operators allow the user to + customise the syntax of operator to look more like a binder. + See [#3113](https://github.com/idris-lang/Idris2/issues/3113). ### REPL/CLI changes diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst new file mode 100644 index 000000000..dfb89c568 --- /dev/null +++ b/docs/source/reference/operators.rst @@ -0,0 +1,176 @@ +.. _operators: + +********* +Operators +********* + +Idris2 does not have syntax blocs (like in Idris1) or mixfix operators (like in Agda). +Instead, it expands on the ability of infix operators to enable library designers +to write DSL while keeping error messages under control. + +Because operators are not linked to function definitions, they are also part of the +file namespacing and follow the same rules as other defintions. + +.. warning:: + Operators can and will make some code less legible. Use with taste and caution. + This document is meant to be mainly used by library authors and advanced users. + If you are on the fence about using operators, err on the side of caution and + avoid them. + +Basics +====== + +Before we jump into the fancy features, let us explain how infix operators work +for most users. + +When you see an expression + + +.. code-block:: idris + + 1 + n + +It means that there is a function ``(+)`` and a *fixity* declaration +in scope. A fixity for this operator looks like this + +.. code-block:: idris + + infixl 8 + + +It starts with a fixity keyword, you have the choice to use either ``infixl``, +``infixr``, ``infix`` or ``prefix``. + +``infixl`` means the operator is left-associative, so that successive applications +of the operator will bracket to the left: ``n + m + 3 + x = (((n + m) + 3) + x)```. +Similarly, ``infixr`` is right-associative, and ``infix`` is non-associative, so the +brackets are mandatory. + +The number afterwards indicate the *precedence* of the operator, that is, if it should +be bracketed before, or after, other operators used in the same expression. For example, +we want ``*`` to *take precedence* over ``+`` , because of this, we define it like this: + +.. code-block:: idris + + infixl 9 * + +This way, the expression ``n + m * x`` is correctly interpreted as ``n + (m * x)``. + + +Fixity & Precedence Namespacing +=============================== + +Just like other top-level declarations in the language, fixities can be exported +with the ``export`` access modifier, or kept private with ``private``. + +A ``private`` fixity will remain in scope for the rest of the file but will not be +visible to users that import the module. Because fixities and operators are +separate, this does not mean you cannot use the functions that have this operator +name, it merely means that you cannot use it in infix position. But you can use +it as a regular function application using brackets. Let us see what this +looks like + +.. code-block:: idris + + module A + + private infixl &&& 8 + + -- a binary function making use of our fixity definition + export + (&&&) : ... + +.. code-block:: idris + + module B + + import A + + main : IO () + main = do print (a &&& b) -- won't work + print ((&&&) a b) -- ok + +Private record fixity pattern +----------------------------- + +Private fixity declarations are useful in conjuction with records. When +you declare a record with operators as fields, it is helpful to write +them in infix position. However, the compiler will also synthesize a +projection for the field that takes a value of the record as the first +argument, making it unsuitable for use as a binary operator. + +.. code-block:: idris + + + infixl 7 <@> + record SomeRelation (a : Type) where + (<@>) : a -> a -> Type + compose : (x, y, z : a) -> x <@> y -> y <@> z -> x <@> z + + lteRel : SomeRelation Nat + lteRel = ... + + natRel : Nat -> Nat -> Type + natRel x y = (<@>) lteRel x y + +What we really want to write is ``natRel x y = <@> x y`` but +``(<@>)`` now has type ``SomeRelation a -> a -> a -> Type``. + +The solution is to define a private field with a private fixity +and a public one which relies on proof search to find the appropriate +argument: + +.. code-block:: idris + + private infixl 7 + export infixl 7 <@> + + record SomeRelation (a : Type) where + () : a -> a -> Type + compose : (x, y, z : a) -> x y -> y z -> x z + + export + (<@>) : (rel : SomeRelation a) => a -> a -> Type + x <@> y = () rel x y + + %hint + lteRel : SomeRelation Nat + lteRel = ... + + natRel : Nat -> Nat -> Type + natRel x y = x <@> y + +We define ``(<@>)`` as a projection function with an implicit parameter +allowing it to be used as a binary operator once again. + +Private Local definition +------------------------ + +Private fixity definitions are useful when redefining an operator fixity +in a module. This happens when multiple DSLs are imported as the same time +and you do not want to expose conflicting fixity declarations: + +.. code-block:: idris + + module Coproduct + + import Product + + -- mark this as private since we don't want to clash + -- with the Prelude + when importing the module + private infixr 5 + + + data (+) : a -> a -> Type where + ... + + distrib1 : {x, y, z : a} -> x + y + z -> (x + y) + z + +Here ``distrib1`` makes explicit use of the operator being defined as +right-associative. + +Typebind Operators +================== + +Autobind Operators +================== + + From c3651509e0a3b8b4ca1e8bb4fb0d60bf02b146fd Mon Sep 17 00:00:00 2001 From: Andre Videla Date: Wed, 6 Dec 2023 22:20:21 +0000 Subject: [PATCH 22/57] add operator documentation --- docs/source/reference/builtins.rst | 4 +- docs/source/reference/index.rst | 1 + docs/source/reference/operators.rst | 180 ++++++++++++++++++++++++++-- 3 files changed, 175 insertions(+), 10 deletions(-) diff --git a/docs/source/reference/builtins.rst b/docs/source/reference/builtins.rst index 013317685..6274ab9f3 100644 --- a/docs/source/reference/builtins.rst +++ b/docs/source/reference/builtins.rst @@ -47,7 +47,7 @@ To ensure that a type is optimized to an ``Integer``, use ``%builtin Natural`` i data MyNat = Succ MyNat | Zero - + %builtin Natural MyNat Casting between natural numbers and integer @@ -106,7 +106,7 @@ that a function is correct. It is your responsibility to make sure this is corre data ComplexNat = Zero | Succ ComplexNat - + integerToMaybeNat : Integer -> Maybe ComplexNat integerToMaybeNat _ = ... diff --git a/docs/source/reference/index.rst b/docs/source/reference/index.rst index d1217d09c..6c615141d 100644 --- a/docs/source/reference/index.rst +++ b/docs/source/reference/index.rst @@ -22,6 +22,7 @@ This is a placeholder, to get set up with readthedocs. records literate overloadedlit + operators strings pragmas builtins diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index dfb89c568..ff6f3b767 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -5,8 +5,8 @@ Operators ********* Idris2 does not have syntax blocs (like in Idris1) or mixfix operators (like in Agda). -Instead, it expands on the ability of infix operators to enable library designers -to write DSL while keeping error messages under control. +Instead, it expands on the abilities of infix operators to enable library designers +to write Domain Specific Languages (DSLs) while keeping error messages under control. Because operators are not linked to function definitions, they are also part of the file namespacing and follow the same rules as other defintions. @@ -43,9 +43,9 @@ It starts with a fixity keyword, you have the choice to use either ``infixl``, ``infixl`` means the operator is left-associative, so that successive applications of the operator will bracket to the left: ``n + m + 3 + x = (((n + m) + 3) + x)```. Similarly, ``infixr`` is right-associative, and ``infix`` is non-associative, so the -brackets are mandatory. +brackets are mandatory. Here, we chose for ``+`` to be left-associative, hence ``infixl``. -The number afterwards indicate the *precedence* of the operator, that is, if it should +The number after the fixity indicate the *precedence level* of the operator, that is, if it should be bracketed before, or after, other operators used in the same expression. For example, we want ``*`` to *take precedence* over ``+`` , because of this, we define it like this: @@ -55,9 +55,71 @@ we want ``*`` to *take precedence* over ``+`` , because of this, we define it li This way, the expression ``n + m * x`` is correctly interpreted as ``n + (m * x)``. +Fixities declarations are optional and only change how a file is parsed, but you can +use any function defined with operator symbols with parenthesis around it: + +.. code-block:: idris + + -- those two are the same + n + 3 + (+) n 3 + +Because fixities are separated from the function definitions, a single operator +can have 0 or multiple fixity definitions. In the next section we explain how to +deal with multiple fixity definitions. Fixity & Precedence Namespacing =============================== +Sometimes it could be that you need to import two libraries that export +conflicting fixities. If that is the case, the compiler will emit a warning +and pick one of the fixities to parse the file. If that happens, you should +hide the fixity definitions that you do not wish to use. For this, use the +``%hide`` directive, just like you would to hide a function definition, but +use the fixity and the operator to hide at the end. Let's work through an +example: + +.. code-block:: idris + + module A + export infixl 8 - + +.. code-block:: idris + + module B + export infixr 5 - + +.. code-block:: idris + + module C + + import A + import B + + test : Int + test = 1 - 3 - 10 + +This program will raise a warning on the last line of module ``C`` because +there are two conflicting fixities in scope, should we parse the expression +as ``(1 - 3) - 10`` or as ``1 - (3 - 10)``? In those cases, you can hide +the extra fixity you do not wish to use by using ``%hide``: + +.. code-block:: idris + + module C + + import A + import B + + %hide A.infixl.(-) + + test : Int + test = 1 - 3 - 10 -- all good, no error + +In which case the program will be parsed as ``1 - (3 - 10)`` and not emit +any errors. + +Export modifiers on fixities +---------------------------- Just like other top-level declarations in the language, fixities can be exported with the ``export`` access modifier, or kept private with ``private``. @@ -89,26 +151,33 @@ looks like main = do print (a &&& b) -- won't work print ((&&&) a b) -- ok +In what follows we have two examples of programs that benefit from +declaring a fixity ``private`` rather than ``export``. + Private record fixity pattern ----------------------------- Private fixity declarations are useful in conjuction with records. When you declare a record with operators as fields, it is helpful to write them in infix position. However, the compiler will also synthesize a -projection for the field that takes a value of the record as the first -argument, making it unsuitable for use as a binary operator. +projection function for the field that takes as first argument the +a value of the record to project from. This makes using the operator +as a binary infix operator impossible, since it now has 3 arguments. .. code-block:: idris infixl 7 <@> + record SomeRelation (a : Type) where (<@>) : a -> a -> Type - compose : (x, y, z : a) -> x <@> y -> y <@> z -> x <@> z + -- we use the field here in infix position + compose : {x, y, z : a} -> x <@> y -> y <@> z -> x <@> z lteRel : SomeRelation Nat lteRel = ... + -- we want to use <@> in infix position here as well but we cannot natRel : Nat -> Nat -> Type natRel x y = (<@>) lteRel x y @@ -126,7 +195,7 @@ argument: record SomeRelation (a : Type) where () : a -> a -> Type - compose : (x, y, z : a) -> x y -> y z -> x z + compose : {x, y, z : a} -> x y -> y z -> x z export (<@>) : (rel : SomeRelation a) => a -> a -> Type @@ -170,7 +239,102 @@ right-associative. Typebind Operators ================== +In dependently-typed programming, we have the ability define types which +first argument is a type and the second is a lambda using the first argument +as it's type. A typical example of this is the dependent linear arrow: + + .. code-block:: idris + + infixr 0 =@ + 0 (=@) : (x : Type) -> (x -> Type) -> Type + (=@) x f = (1 v : x) -> f v + + +However, we cannot use as is because the second argument is +a lambda, and writing out any value using this operator will look a bit awkward: + +.. code-block:: idris + + linearSingleton : Nat =@ (\x => Singleton x) + linearSingleton = ... + +What we really want to write, is something like the dependent arrow ``->`` but +for linear types: + +.. code-block:: idris + + linearSingleton : (x : Nat) =@ Singleton x + linearSingleton = ... + +The above syntax is allowed if the operator is declared as ``typebind``. For +this, simply add the ``typebind`` keyword in front of the fixity declaration. + +.. code-block:: idris + + typebind infixr 0 =@ + +``typebind`` is a mofidier like ``export`` and both can be used at the same time. + + +An operator defined as ``typebind`` cannot be used in regular position anymore, +writing ``Nat =@ (\x => Singleton x)`` will raise an error. + +This new syntax is purely syntax sugar and converts any instance of +``(name : type) op expr`` into ``type op (\name : type => expr)`` + +Because of its left-to-right binding structure, typebind operators can +only ever be ``infixr`` with precedence 0. + + Autobind Operators ================== +Typebind operators allow to bind a *type* on the left side of an operator, but +sometimes, there is no dependency between the first argument, and the expression +on the right side of the operator. For those case, we use ``autobind``. +An example of this is a DSL for a dependently-typed programming language +where the constructor for ``Pi`` takes terms on the left and lambdas on the right: + +.. code-block:: idris + + VPi : Value -> (Value -> Value) -> Value + + sig : Value + sig = VPi VStar (\fstTy -> VPi + (VPi fstTy (const VStar)) (\sndTy -> VPi + fstTy (\val1 -> VPi + (sndTy `vapp` val1) (\val2 -> + VSigma fstTy sndTy))))) + +We would like to use a custom operator to build values using ``VPi``, but its +signature does not follow the pattern that ``typebind`` uses. Instead, we use +``autobind`` to tell the compiler that the type of the lambda is not given +by the first argument. For this we use ``:=`` instead of ``:``: + +.. code-block:: idris + + autobind infixr 0 =>> + (=>>) : Value -> (Value -> Value) -> Value + (=>>) = VPi + + + sig : Value + sig = + (fstTy := VStar) =>> + (sndTy := (_ := fstTy) =>> VStar) =>> + (val1 := fstTy) =>> + (val2 := sndTy `vapp` val1) =>> + VSgima fstTy sndTy + +This new syntax is much closer to what the code is meant to look like for users +accustomed to dependently-typed programming languages. + +More technically, any ``autobind`` operator is called with the syntax +``(name := expr) op body`` and is desugared into ``expr op (\name : ? => body)``. +If you want, or need, to give the type explicitly, you can still do so by using +the full syntax: ``(name : type := expr) op body`` which is desugared into +``expr op (\name : type => body)``. + +Like ``typebind``, ``autobind`` operators cannot be used as regular operators anymore +, additionally an ``autobind`` operator cannot use the ``typebind`` syntax either. From 63c167637c01d1791091d1b74c165a6ceb2837a9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 12 Dec 2023 15:23:50 +0000 Subject: [PATCH 23/57] wip more error messages --- src/Core/Context.idr | 2 ++ src/Core/Core.idr | 4 +++ src/Core/TT.idr | 8 +++++ src/Idris/Desugar.idr | 27 +++++++++++++- src/Idris/Error.idr | 1 + tests/idris2/operators/operators001/Test.idr | 15 ++++---- .../idris2/operators/operators003/Error1.idr | 3 ++ .../idris2/operators/operators003/Error2.idr | 3 ++ .../idris2/operators/operators003/Error3.idr | 3 ++ tests/idris2/operators/operators003/expected | 36 +++++++++++++++++++ tests/idris2/operators/operators003/run | 5 +++ 11 files changed, 99 insertions(+), 8 deletions(-) create mode 100644 tests/idris2/operators/operators003/Error1.idr create mode 100644 tests/idris2/operators/operators003/Error2.idr create mode 100644 tests/idris2/operators/operators003/Error3.idr create mode 100644 tests/idris2/operators/operators003/expected create mode 100755 tests/idris2/operators/operators003/run diff --git a/src/Core/Context.idr b/src/Core/Context.idr index 004f3ddcd..eae47cc22 100644 --- a/src/Core/Context.idr +++ b/src/Core/Context.idr @@ -796,6 +796,7 @@ HasNames Error where full gam (BadRunElab fc rho s desc) = BadRunElab fc <$> full gam rho <*> full gam s <*> pure desc full gam (RunElabFail e) = RunElabFail <$> full gam e full gam (GenericMsg fc x) = pure (GenericMsg fc x) + full gam (GenericMsgSol fc x y) = pure (GenericMsgSol fc x y) full gam (TTCError x) = pure (TTCError x) full gam (FileErr x y) = pure (FileErr x y) full gam (CantFindPackage x) = pure (CantFindPackage x) @@ -889,6 +890,7 @@ HasNames Error where resolved gam (BadRunElab fc rho s desc) = BadRunElab fc <$> resolved gam rho <*> resolved gam s <*> pure desc resolved gam (RunElabFail e) = RunElabFail <$> resolved gam e resolved gam (GenericMsg fc x) = pure (GenericMsg fc x) + resolved gam (GenericMsgSol fc x y) = pure (GenericMsgSol fc x y) resolved gam (TTCError x) = pure (TTCError x) resolved gam (FileErr x y) = pure (FileErr x y) resolved gam (CantFindPackage x) = pure (CantFindPackage x) diff --git a/src/Core/Core.idr b/src/Core/Core.idr index e4f8fcf91..d12c8f357 100644 --- a/src/Core/Core.idr +++ b/src/Core/Core.idr @@ -163,6 +163,7 @@ data Error : Type where FC -> Env Term vars -> Term vars -> (description : String) -> Error RunElabFail : Error -> Error GenericMsg : FC -> String -> Error + GenericMsgSol : FC -> (message : String) -> (solutions : List String) -> Error OperatorBindingMismatch : {a : Type} -> {print : a -> Doc ()} -> FC -> (expectedFixity : Maybe FixityInfo) -> (use_site : OperatorLHSInfo a) -> (opName : Name) -> (rhs : a) -> Error @@ -354,6 +355,7 @@ Show Error where show (BadRunElab fc env script desc) = show fc ++ ":Bad elaborator script " ++ show script ++ " (" ++ desc ++ ")" show (RunElabFail e) = "Error during reflection: " ++ show e show (GenericMsg fc str) = show fc ++ ":" ++ str + show (GenericMsgSol fc msg sols) = show fc ++ ":" ++ msg ++ " Solutions: " ++ show sols show (TTCError msg) = "Error in TTC file: " ++ show msg show (FileErr fname err) = "File error (" ++ fname ++ "): " ++ show err show (CantFindPackage fname) = "Can't find package " ++ fname @@ -466,6 +468,7 @@ getErrorLoc (BadImplicit loc _) = Just loc getErrorLoc (BadRunElab loc _ _ _) = Just loc getErrorLoc (RunElabFail e) = getErrorLoc e getErrorLoc (GenericMsg loc _) = Just loc +getErrorLoc (GenericMsgSol loc _ _) = Just loc getErrorLoc (TTCError _) = Nothing getErrorLoc (FileErr _ _) = Nothing getErrorLoc (CantFindPackage _) = Nothing @@ -555,6 +558,7 @@ killErrorLoc (BadImplicit fc x) = BadImplicit emptyFC x killErrorLoc (BadRunElab fc x y description) = BadRunElab emptyFC x y description killErrorLoc (RunElabFail e) = RunElabFail $ killErrorLoc e killErrorLoc (GenericMsg fc x) = GenericMsg emptyFC x +killErrorLoc (GenericMsgSol fc x y) = GenericMsgSol emptyFC x y killErrorLoc (TTCError x) = TTCError x killErrorLoc (FileErr x y) = FileErr x y killErrorLoc (CantFindPackage x) = CantFindPackage x diff --git a/src/Core/TT.idr b/src/Core/TT.idr index b3e1039cf..65ee1f8c4 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -112,6 +112,10 @@ Show Fixity where show Infix = "infix" show Prefix = "prefix" +export +Interpolation Fixity where + interpolate = show + export Eq Fixity where InfixL == InfixL = True @@ -137,6 +141,10 @@ Show BindingModifier where show Typebind = "typebind" show Autobind = "autobind" +export +Interpolation BindingModifier where + interpolate = show + -- A record to hold all the information about a fixity public export record FixityInfo where diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 3f9fd2ed4..51049cc7e 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -178,6 +178,25 @@ checkConflictingBinding side fc opName foundFixity use_site rhs isCompatible (Just fixInfo) (BindExplicitType name type expr) = fixInfo.bindingInfo == Autobind +checkValidFixity : BindingModifier -> Fixity -> Nat -> Bool + +-- If the fixity declaration is not binding, there are no restrictions +checkValidFixity NotBinding _ _ = True + +-- If the fixity declaration is not binding, either typebind or autobind +-- the fixity can only be `infixr` with precedence `0`. We might want +-- to revisit that in the future, but as of right now we want to be +-- conservative and avoid abuse. +-- having multiple precedence levels would mean that if +-- =@ has higher precedence than -@ +-- the expression (x : a) =@ b -@ (y : List a) =@ List b +-- would be bracketed as ((x : a) =@ b) -@ (y : List a) =@ List b +-- instead of (x : a) =@ (b -@ ((y : List a) =@ List b)) +checkValidFixity _ InfixR 0 = True + +-- If it's binding but not infixr precedence 0, raise an error +checkValidFixity _ _ _ = False + parameters (side : Side) toTokList : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> @@ -1157,7 +1176,13 @@ mutual mkConName n = DN (show n) (MN ("__mk" ++ show n) 0) desugarDecl ps (PFixity fc vis binding fix prec opName) - = do ctx <- get Ctxt + = do unless (checkValidFixity binding fix prec) + (throw $ GenericMsgSol fc + "invalid fixity, \{binding} operator must be infixr 0." + ["write `\{binding} infixr 0 \{show opName}`" + , "remove the binding keyword `\{fix} \{show prec} \{show opName}`" + ]) + ctx <- get Ctxt -- We update the context of fixities by adding a namespaced fixity -- given by the current namespace and its fixity name. -- This allows fixities to be stored along with the namespace at their diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index 38bbbed28..fa0dcbd87 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -110,6 +110,7 @@ Eq Error where BadRunElab fc1 rho1 s1 d1 == BadRunElab fc2 rho2 s2 d2 = fc1 == fc2 && d1 == d2 RunElabFail e1 == RunElabFail e2 = e1 == e2 GenericMsg fc1 x1 == GenericMsg fc2 x2 = fc1 == fc2 && x1 == x2 + GenericMsgSol fc1 x1 y1 == GenericMsgSol fc2 x2 y2 = fc1 == fc2 && x1 == x2 && y1 == y2 TTCError x1 == TTCError x2 = x1 == x2 FileErr x1 y1 == FileErr x2 y2 = x1 == x2 && y1 == y2 CantFindPackage x1 == CantFindPackage x2 = x1 == x2 diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index 3b866339a..e3d5b8f5a 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -6,7 +6,7 @@ import Data.Vect typebind infixr 0 =@ infixr 0 -@ -typebind infix 1 =@@ +-- typebind infixr 1 =@@ 0 (=@) : (a : Type) -> (a -> Type) -> Type (=@) a f = (1 x : a) -> f x @@ -23,7 +23,7 @@ data S : {ty : Type} -> (x : ty) -> Type where Mk3 : (x : ty) =@ ty -@ S x Mk4 : ty -@ (x : ty) =@ S x -map : (x : a) =@@ b -@ (y : List a) =@ List b +-- map : (x : a) =@@ b -@ (y : List a) =@ List b map2 : ((x : a) =@ b) -@ (y : List a) =@ List b @@ -31,10 +31,11 @@ map3 : (x : a) =@ b -@ (y : List a) =@ List b map4 : (x : a) =@ (b -@ (y : List a) =@ List b) -test : Test.map === Test.map2 - -failing - test2 : Test.map === Test.map3 +-- this could be possible if we allowed binding operatprs with binding operators +-- with higher precedences +-- test : Test.map === Test.map2 +-- failing +-- test2 : Test.map === Test.map3 test3 : Test.map3 === Test.map4 @@ -46,7 +47,7 @@ typebind infixr 0 *> -- testCompose : (x : Nat) *> (y : Nat) *> Vect (n + m) String -- testCompose = (1 ** 2 ** ["hello", "world", "!"]) -autobind infixr 7 `MyLet` +autobind infixr 0 `MyLet` MyLet : (val) -> (val -> rest) -> rest MyLet arg fn = fn arg diff --git a/tests/idris2/operators/operators003/Error1.idr b/tests/idris2/operators/operators003/Error1.idr new file mode 100644 index 000000000..8860d8899 --- /dev/null +++ b/tests/idris2/operators/operators003/Error1.idr @@ -0,0 +1,3 @@ +module Error1 + +typebind infixl 0 =@ diff --git a/tests/idris2/operators/operators003/Error2.idr b/tests/idris2/operators/operators003/Error2.idr new file mode 100644 index 000000000..de22cfe71 --- /dev/null +++ b/tests/idris2/operators/operators003/Error2.idr @@ -0,0 +1,3 @@ +module Error2 + +typebind infixr 3 =@ diff --git a/tests/idris2/operators/operators003/Error3.idr b/tests/idris2/operators/operators003/Error3.idr new file mode 100644 index 000000000..f91d022ab --- /dev/null +++ b/tests/idris2/operators/operators003/Error3.idr @@ -0,0 +1,3 @@ +module Error3 + +typebind infixl 3 =@ diff --git a/tests/idris2/operators/operators003/expected b/tests/idris2/operators/operators003/expected new file mode 100644 index 000000000..5bae83e6a --- /dev/null +++ b/tests/idris2/operators/operators003/expected @@ -0,0 +1,36 @@ +1/1: Building Error1 (Error1.idr) +Error: invalid fixity, typebind operator must be infixr 0. +Possible solutions: +- write `typebind infixr 0 =@` +- remove the binding keyword `infixl 0 =@` + +Error1:3:10--3:21 + 1 | module Error1 + 2 | + 3 | typebind infixl 0 =@ + ^^^^^^^^^^^ + +1/1: Building Error2 (Error2.idr) +Error: invalid fixity, typebind operator must be infixr 0. +Possible solutions: +- write `typebind infixr 0 =@` +- remove the binding keyword `infixr 3 =@` + +Error2:3:10--3:21 + 1 | module Error2 + 2 | + 3 | typebind infixr 3 =@ + ^^^^^^^^^^^ + +1/1: Building Error3 (Error3.idr) +Error: invalid fixity, typebind operator must be infixr 0. +Possible solutions: +- write `typebind infixr 0 =@` +- remove the binding keyword `infixl 3 =@` + +Error3:3:10--3:21 + 1 | module Error3 + 2 | + 3 | typebind infixl 3 =@ + ^^^^^^^^^^^ + diff --git a/tests/idris2/operators/operators003/run b/tests/idris2/operators/operators003/run new file mode 100755 index 000000000..e87f60f6f --- /dev/null +++ b/tests/idris2/operators/operators003/run @@ -0,0 +1,5 @@ +. ../../../testutils.sh + +check Error1.idr +check Error2.idr +check Error3.idr From 4cb8dc507bea38b1eadfc7f2d8ca344e2efa059c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Thu, 14 Dec 2023 21:34:16 +0000 Subject: [PATCH 24/57] update error message for infixr --- src/Idris/Desugar.idr | 6 ++--- src/Idris/Error.idr | 5 ++++ tests/idris2/operators/operators003/expected | 24 ++++++++++---------- 3 files changed, 20 insertions(+), 15 deletions(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 51049cc7e..b91ba7c84 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -1178,9 +1178,9 @@ mutual desugarDecl ps (PFixity fc vis binding fix prec opName) = do unless (checkValidFixity binding fix prec) (throw $ GenericMsgSol fc - "invalid fixity, \{binding} operator must be infixr 0." - ["write `\{binding} infixr 0 \{show opName}`" - , "remove the binding keyword `\{fix} \{show prec} \{show opName}`" + "Invalid fixity, \{binding} operator must be infixr 0." + [ "Make it `infixr 0`: `\{binding} infixr 0 \{show opName}`" + , "Remove the binding keyword: `\{fix} \{show prec} \{show opName}`" ]) ctx <- get Ctxt -- We update the context of fixities by adding a namespaced fixity diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index fa0dcbd87..5880b2c22 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -620,6 +620,11 @@ perrorRaw (BadRunElab fc env script desc) perrorRaw (RunElabFail e) = pure $ reflow "Error during reflection" <+> colon <++> !(perrorRaw e) perrorRaw (GenericMsg fc str) = pure $ pretty0 str <+> line <+> !(ploc fc) +perrorRaw (GenericMsgSol fc header solutions) + = pure $ pretty0 header <+> line <+> !(ploc fc) + <+> line + <+> "Possible solutions:" <+> line + <+> indent 1 (vsep (map (\s => "-" <++> pretty0 s) solutions)) perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) = pure $ "Operator" <++> pretty0 !(getFullName opName) <++> "is" <++> printBindingInfo expected diff --git a/tests/idris2/operators/operators003/expected b/tests/idris2/operators/operators003/expected index 5bae83e6a..1b262ce82 100644 --- a/tests/idris2/operators/operators003/expected +++ b/tests/idris2/operators/operators003/expected @@ -1,8 +1,5 @@ 1/1: Building Error1 (Error1.idr) -Error: invalid fixity, typebind operator must be infixr 0. -Possible solutions: -- write `typebind infixr 0 =@` -- remove the binding keyword `infixl 0 =@` +Error: Invalid fixity, typebind operator must be infixr 0. Error1:3:10--3:21 1 | module Error1 @@ -10,11 +7,11 @@ Error1:3:10--3:21 3 | typebind infixl 0 =@ ^^^^^^^^^^^ -1/1: Building Error2 (Error2.idr) -Error: invalid fixity, typebind operator must be infixr 0. Possible solutions: -- write `typebind infixr 0 =@` -- remove the binding keyword `infixr 3 =@` + - Make it `infixr 0`: `typebind infixr 0 =@` + - Remove the binding keyword: `infixl 0 =@` +1/1: Building Error2 (Error2.idr) +Error: Invalid fixity, typebind operator must be infixr 0. Error2:3:10--3:21 1 | module Error2 @@ -22,11 +19,11 @@ Error2:3:10--3:21 3 | typebind infixr 3 =@ ^^^^^^^^^^^ -1/1: Building Error3 (Error3.idr) -Error: invalid fixity, typebind operator must be infixr 0. Possible solutions: -- write `typebind infixr 0 =@` -- remove the binding keyword `infixl 3 =@` + - Make it `infixr 0`: `typebind infixr 0 =@` + - Remove the binding keyword: `infixr 3 =@` +1/1: Building Error3 (Error3.idr) +Error: Invalid fixity, typebind operator must be infixr 0. Error3:3:10--3:21 1 | module Error3 @@ -34,3 +31,6 @@ Error3:3:10--3:21 3 | typebind infixl 3 =@ ^^^^^^^^^^^ +Possible solutions: + - Make it `infixr 0`: `typebind infixr 0 =@` + - Remove the binding keyword: `infixl 3 =@` From eb9d4d65a5d52a305bc501487d2366b469f43fd3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Thu, 14 Dec 2023 22:21:28 +0000 Subject: [PATCH 25/57] fix typo in doc --- docs/source/reference/operators.rst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index ff6f3b767..9cb7bfb57 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -273,7 +273,7 @@ this, simply add the ``typebind`` keyword in front of the fixity declaration. typebind infixr 0 =@ -``typebind`` is a mofidier like ``export`` and both can be used at the same time. +``typebind`` is a modifier like ``export`` and both can be used at the same time. An operator defined as ``typebind`` cannot be used in regular position anymore, From b5895394d1d8152aab97fd4063d370343579a34e Mon Sep 17 00:00:00 2001 From: Andre Videla Date: Fri, 15 Dec 2023 18:24:33 +0000 Subject: [PATCH 26/57] update comment documentation strings --- src/Core/TT.idr | 14 +++++++------- src/Idris/Desugar.idr | 6 +++--- src/Idris/Parser.idr | 10 +++++----- 3 files changed, 15 insertions(+), 15 deletions(-) diff --git a/src/Core/TT.idr b/src/Core/TT.idr index 65ee1f8c4..6be2a8031 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -171,25 +171,25 @@ Eq FixityInfo where -- Left-hand-side information for operators, carries autobind information -- an operator can either be -- - not autobind, a regular operator --- - binding types, such that (nm : ty ** fn nm) desugars into (ty ** \(nm : ty) => fn nm) +-- - binding types, such that `(nm : ty) =@ fn nm` desugars into `(=@) ty (\(nm : ty) => fn nm)` -- - binding expressing with an inferred type such that --- (nm := exp ** fn nm) desugars into (exp ** \(nm : ?) => fn nm) +-- `(nm := exp) =@ fn nm` desugars into `(=@) exp (\(nm : ?) => fn nm)` -- - binding both types and expression such that --- (nm : ty := exp ** fn nm) desugars into (exp ** \(nm : ty) => fn nm) +-- `(nm : ty := exp) =@ fn nm` desugars into `(=@) exp (\(nm : ty) => fn nm)` public export data OperatorLHSInfo : tm -> Type where -- Traditional operator wihtout binding, carries the lhs NoBinder : (lhs : tm) -> OperatorLHSInfo tm - -- (nm : ty) ** fn x + -- (nm : ty) =@ fn x BindType : (name : WithBounds Name) -> (ty : tm) -> OperatorLHSInfo tm - -- (nm := exp) ** fn nm + -- (nm := exp) =@ fn nm BindExpr : (name : WithBounds Name) -> (expr : tm) -> OperatorLHSInfo tm - -- (nm : ty := exp) ** fn nm + -- (nm : ty := exp) =@ fn nm BindExplicitType : (name : WithBounds Name) -> (type, expr : tm) -> OperatorLHSInfo tm export Show (OperatorLHSInfo tm) where - show (NoBinder lhs) = "regular" + show (NoBinder lhs) = "regular" show (BindType name ty) = "type-binding (typebind)" show (BindExpr name expr) = "automatically-binding (autobind)" show (BindExplicitType name type expr) = "automatically-binding (autobind)" diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index b91ba7c84..c5efa8eab 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -780,19 +780,19 @@ mutual = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc (IApp loc (IVar opFC op) l') r') - -- (x : ty ** f x) ==>> (**) exp (\x : ty => f x) + -- (x : ty) =@ f x ==>> (=@) exp (\x : ty => f x) desugarTree side ps (Infix loc opFC (op, Just (BindType 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.val) desugaredLHS desugaredRHS) - -- (x := exp ** f x) ==>> (**) exp (\x : ? => f x) + -- (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.val) (Implicit opFC False) desugaredRHS) - -- (x : ty := exp ** f x) ==>> (**) exp (\x : ty => f x) + -- (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 diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 9abeb7c69..14ab15a5d 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -1882,16 +1882,16 @@ definition fname indents = do nd <- bounds (clause 0 Nothing fname indents) pure (PDef (boundToFC fname nd) [nd.val]) -operatorBindingKeyword : EmptyRule BindingModifier -operatorBindingKeyword - = (keyword "autobind" >> pure Autobind) - <|> (keyword "typebind" >> pure Typebind) +operatorBindingKeyword : OriginDesc -> EmptyRule BindingModifier +operatorBindingKeyword fname + = (decoratedKeyword fname "autobind" >> pure Autobind) + <|> (decoratedKeyword fname "typebind" >> pure Typebind) <|> pure NotBinding fixDecl : OriginDesc -> IndentInfo -> Rule (List PDecl) fixDecl fname indents = do vis <- exportVisibility fname - binding <- operatorBindingKeyword + binding <- operatorBindingKeyword fname b <- bounds (do fixity <- decorate fname Keyword $ fix commit prec <- decorate fname Keyword $ intLit From e583e73a7ce97f2381de417c9d636fc39a9f015d Mon Sep 17 00:00:00 2001 From: Mathew Polzin Date: Mon, 1 Jan 2024 21:47:36 -0600 Subject: [PATCH 27/57] move CHANGELOG entry to CHANGELOG_NEXT --- CHANGELOG.md | 3 --- CHANGELOG_NEXT.md | 4 ++++ 2 files changed, 4 insertions(+), 3 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index e8dedf7f0..5585c41c9 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -34,9 +34,6 @@ This CHANGELOG describes the history of already-released versions. Please see [C referred to by another definitions, and in which function currently elaborator is. These features together give an ability to inspect whether particular expressions are recursive (including mutual recursion). -* Autobind and Typebind modifier on operators allow the user to - customise the syntax of operator to look more like a binder. - See [#3113](https://github.com/idris-lang/Idris2/issues/3113). ### REPL/CLI changes diff --git a/CHANGELOG_NEXT.md b/CHANGELOG_NEXT.md index 997f2b100..992ab72d8 100644 --- a/CHANGELOG_NEXT.md +++ b/CHANGELOG_NEXT.md @@ -26,6 +26,10 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO ### Language changes +* Autobind and Typebind modifier on operators allow the user to + customise the syntax of operator to look more like a binder. + See [#3113](https://github.com/idris-lang/Idris2/issues/3113). + ### Backend changes #### RefC From 210f9d9c1500642b1e36dc8e0623f95dcbbb04c1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 6 Jan 2024 22:11:52 +0000 Subject: [PATCH 28/57] fix error message for non-associative fixity --- src/Idris/Desugar.idr | 8 ++++++-- src/Libraries/Utils/Shunting.idr | 8 ++++---- tests/idris2/operators/operators004/Test.idr | 9 +++++++++ tests/idris2/operators/operators004/expected | 11 +++++++++++ tests/idris2/operators/operators004/run | 3 +++ 5 files changed, 33 insertions(+), 6 deletions(-) create mode 100644 tests/idris2/operators/operators004/Test.idr create mode 100644 tests/idris2/operators/operators004/expected create mode 100755 tests/idris2/operators/operators004/run diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index c5efa8eab..bc6ea83ce 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -105,6 +105,10 @@ mkPrec InfixR = AssocR mkPrec Infix = NonAssoc mkPrec Prefix = Prefix +-- This is used to print the error message for fixities +[showFst] Show a => Show (a, b) where + show (x, y) = show x + -- Check that an operator does not have any conflicting fixities in scope. -- Each operator can have its fixity defined multiple times across multiple -- modules as long as the fixities are consistent. If they aren't, the fixity @@ -352,10 +356,10 @@ mutual desugarB side ps (PBracketed fc e) = desugarB side ps e desugarB side ps (POp fc opFC l op r) = do ts <- toTokList side (POp fc opFC l op r) - desugarTree side ps !(parseOps ts) + desugarTree side ps !(parseOps @{showFst} ts) desugarB side ps (PPrefixOp fc opFC op arg) = do ts <- toTokList side (PPrefixOp fc opFC op arg) - desugarTree side ps !(parseOps ts) + desugarTree side ps !(parseOps @{showFst} ts) desugarB side ps (PSectionL fc opFC op arg) = do syn <- get Syn -- It might actually be a prefix argument rather than a section diff --git a/src/Libraries/Utils/Shunting.idr b/src/Libraries/Utils/Shunting.idr index e60f97936..a3c3cb314 100644 --- a/src/Libraries/Utils/Shunting.idr +++ b/src/Libraries/Utils/Shunting.idr @@ -90,13 +90,13 @@ higher : Show op => FC -> op -> OpPrec -> op -> OpPrec -> Core Bool higher loc opx op opy (Prefix p) = pure False higher loc opx (NonAssoc x) opy oy = if x == getPrec oy - then throw (GenericMsg loc ("Operator '" ++ show opx ++ - "' is non-associative")) + then throw (GenericMsg loc ("Operator " ++ show opx ++ + " is non-associative")) else pure (x > getPrec oy) higher loc opx ox opy (NonAssoc y) = if getPrec ox == y - then throw (GenericMsg loc ("Operator '" ++ show opy ++ - "' is non-associative")) + then throw (GenericMsg loc ("Operator " ++ show opy ++ + " is non-associative")) else pure (getPrec ox > y) higher loc opl l opr r = pure $ (getPrec l > getPrec r) || diff --git a/tests/idris2/operators/operators004/Test.idr b/tests/idris2/operators/operators004/Test.idr new file mode 100644 index 000000000..485981d8a --- /dev/null +++ b/tests/idris2/operators/operators004/Test.idr @@ -0,0 +1,9 @@ + +infix 5 -:- + +(-:-) : a -> List a -> List a +(-:-) = (::) + +test : List Nat +test = 4 -:- 3 -:- [] + diff --git a/tests/idris2/operators/operators004/expected b/tests/idris2/operators/operators004/expected new file mode 100644 index 000000000..6a40255e4 --- /dev/null +++ b/tests/idris2/operators/operators004/expected @@ -0,0 +1,11 @@ +1/1: Building Test (Test.idr) +Error: Operator -:- is non-associative + +Test:8:8--8:22 + 4 | (-:-) : a -> List a -> List a + 5 | (-:-) = (::) + 6 | + 7 | test : List Nat + 8 | test = 4 -:- 3 -:- [] + ^^^^^^^^^^^^^^ + diff --git a/tests/idris2/operators/operators004/run b/tests/idris2/operators/operators004/run new file mode 100755 index 000000000..d35387bb3 --- /dev/null +++ b/tests/idris2/operators/operators004/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr From 91f66bc7b5c841a3ac9061168edad0a3e61790b1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sun, 7 Jan 2024 10:19:21 +0000 Subject: [PATCH 29/57] add location of fixity to change in error message --- src/Core/Core.idr | 7 ++- src/Core/TT.idr | 8 ++++ src/Idris/Desugar.idr | 46 +++++++++++--------- src/Idris/Error.idr | 20 ++++----- src/Libraries/Utils/Shunting.idr | 24 ++++++---- tests/idris2/operators/operators004/expected | 3 ++ 6 files changed, 68 insertions(+), 40 deletions(-) diff --git a/src/Core/Core.idr b/src/Core/Core.idr index d12c8f357..0fc0eaee3 100644 --- a/src/Core/Core.idr +++ b/src/Core/Core.idr @@ -165,7 +165,7 @@ data Error : Type where GenericMsg : FC -> String -> Error GenericMsgSol : FC -> (message : String) -> (solutions : List String) -> Error OperatorBindingMismatch : {a : Type} -> {print : a -> Doc ()} -> - FC -> (expectedFixity : Maybe FixityInfo) -> (use_site : OperatorLHSInfo a) -> + FC -> (expectedFixity : BacktickOrOperatorFixity) -> (use_site : OperatorLHSInfo a) -> (opName : Name) -> (rhs : a) -> Error TTCError : TTCErrorMsg -> Error FileErr : String -> FileError -> Error @@ -400,9 +400,12 @@ Show Error where (n ::: []) => ": " ++ n ++ "?" _ => " any of: " ++ showSep ", " (map show (forget ns)) ++ "?" show (WarningAsError w) = show w - show (OperatorBindingMismatch fc expected actual opName rhs) + show (OperatorBindingMismatch fc (DeclaredFixity expected) actual opName rhs) = show fc ++ ": Operator " ++ show opName ++ " is " ++ show expected ++ " but used as a " ++ show actual ++ " operator" + show (OperatorBindingMismatch fc Backticked actual opName rhs) + = show fc ++ ": Operator " ++ show opName ++ " has no declared fixity" + ++ " but used as a " ++ show actual ++ " operator" export getWarningLoc : Warning -> FC diff --git a/src/Core/TT.idr b/src/Core/TT.idr index 6be2a8031..6c1dde8a6 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -168,6 +168,14 @@ Eq FixityInfo where && x.fix == y.fix && x.precedence == y.precedence +||| Whenever we read an operator from the parser, we don't know if it's a backticked expression with no fixity +||| declaration, or if it has a fixity declaration. If it does not have a declaration, we represent this state +||| with `Backticked`. +||| Note that a backticked expression can have a fixity declaration, in which case it is represented with +||| `DeclaredFixity`. +public export +data BacktickOrOperatorFixity = Backticked | DeclaredFixity FixityInfo + -- Left-hand-side information for operators, carries autobind information -- an operator can either be -- - not autobind, a regular operator diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index bc6ea83ce..f834bf1d2 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -106,8 +106,12 @@ mkPrec Infix = NonAssoc mkPrec Prefix = Prefix -- This is used to print the error message for fixities -[showFst] Show a => Show (a, b) where - show (x, y) = show x +[interpName] Interpolation ((Name, BacktickOrOperatorFixity), b) where + interpolate ((x, _), _) = show x + +[showWithLoc] Show ((Name, BacktickOrOperatorFixity), b) where + show ((x, DeclaredFixity y), _) = show x ++ " at " ++ show y.fc + show ((x, Backticked), _) = show x -- Check that an operator does not have any conflicting fixities in scope. -- Each operator can have its fixity defined multiple times across multiple @@ -117,7 +121,7 @@ mkPrec Prefix = Prefix checkConflictingFixities : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> (isPrefix : Bool) -> - FC -> Name -> Core (OpPrec, Maybe FixityInfo) + FC -> Name -> Core (OpPrec, BacktickOrOperatorFixity) checkConflictingFixities isPrefix exprFC opn = do syn <- get Syn let op = nameRoot opn @@ -128,14 +132,14 @@ checkConflictingFixities isPrefix exprFC opn -- characters, if not, it must be a backticked expression. (_, [], []) => if any isOpChar (fastUnpack op) then throw (GenericMsg exprFC "Unknown operator '\{op}'") - else pure (NonAssoc 1, Nothing) -- Backticks are non associative by default + else pure (NonAssoc 1, Backticked) -- Backticks are non associative by default -- (True, ((fxName, fx) :: _), _) => do -- in the prefix case, remove conflicts with infix (-) let extraFixities = pre ++ (filter (\(nm, _) => not $ nameRoot nm == "-") inf) unless (isCompatible fx extraFixities) $ warnConflict fxName extraFixities - pure (mkPrec fx.fix fx.precedence, Just fx) + pure (mkPrec fx.fix fx.precedence, DeclaredFixity fx) -- Could not find any prefix operator fixities, there may still be conflicts with -- the infix ones. (True, [] , _) => throw (GenericMsg exprFC $ "'\{op}' is not a prefix operator") @@ -144,7 +148,7 @@ checkConflictingFixities isPrefix exprFC opn -- In the infix case, remove conflicts with prefix (-) let extraFixities = (filter (\(nm, _) => not $ nm == UN (Basic "-")) pre) ++ inf unless (isCompatible fx extraFixities) $ warnConflict fxName extraFixities - pure (mkPrec fx.fix fx.precedence, Just fx) + pure (mkPrec fx.fix fx.precedence, DeclaredFixity fx) -- Could not find any infix operator fixities, there may be prefix ones (False, _, []) => throw (GenericMsg exprFC $ "'\{op}' is not an infix operator") where @@ -166,20 +170,20 @@ checkConflictingFixities isPrefix exprFC opn For example: %hide \{show fxName} """ -checkConflictingBinding : Side -> FC -> Name -> (foundFixity : Maybe FixityInfo) -> (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () +checkConflictingBinding : Side -> FC -> Name -> (foundFixity : BacktickOrOperatorFixity) -> (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () checkConflictingBinding LHS fc opName foundFixity use_site rhs = pure () -- don't check if we're on the LHS checkConflictingBinding side fc opName foundFixity use_site rhs = if isCompatible foundFixity use_site then pure () else throw $ OperatorBindingMismatch {print=byShow} fc foundFixity use_site opName rhs where - isCompatible : Maybe FixityInfo -> OperatorLHSInfo PTerm -> Bool - isCompatible Nothing (NoBinder lhs) = True - isCompatible Nothing _ = False - isCompatible (Just fixInfo) (NoBinder lhs) = fixInfo.bindingInfo == NotBinding - isCompatible (Just fixInfo) (BindType name ty) = fixInfo.bindingInfo == Typebind - isCompatible (Just fixInfo) (BindExpr name expr) = fixInfo.bindingInfo == Autobind - isCompatible (Just fixInfo) (BindExplicitType name type expr) + isCompatible : BacktickOrOperatorFixity -> OperatorLHSInfo PTerm -> Bool + isCompatible Backticked (NoBinder lhs) = True + isCompatible Backticked _ = False + isCompatible (DeclaredFixity fixInfo) (NoBinder lhs) = fixInfo.bindingInfo == NotBinding + isCompatible (DeclaredFixity fixInfo) (BindType name ty) = fixInfo.bindingInfo == Typebind + isCompatible (DeclaredFixity fixInfo) (BindExpr name expr) = fixInfo.bindingInfo == Autobind + isCompatible (DeclaredFixity fixInfo) (BindExplicitType name type expr) = fixInfo.bindingInfo == Autobind checkValidFixity : BindingModifier -> Fixity -> Nat -> Bool @@ -204,16 +208,16 @@ checkValidFixity _ _ _ = False parameters (side : Side) toTokList : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> - PTerm -> Core (List (Tok (OpStr, Maybe (OperatorLHSInfo PTerm)) PTerm)) + PTerm -> Core (List (Tok ((OpStr, BacktickOrOperatorFixity), Maybe (OperatorLHSInfo PTerm)) PTerm)) toTokList (POp fc opFC l opn r) = do (precInfo, fixInfo) <- checkConflictingFixities False fc opn checkConflictingBinding side opFC opn fixInfo l r rtoks <- toTokList r - pure (Expr l.getLhs :: Op fc opFC (opn, Just l) precInfo :: rtoks) + pure (Expr l.getLhs :: Op fc opFC ((opn, fixInfo), Just l) precInfo :: rtoks) toTokList (PPrefixOp fc opFC opn arg) - = do (precInfo, _) <- checkConflictingFixities True fc opn + = do (precInfo, fixInfo) <- checkConflictingFixities True fc opn rtoks <- toTokList arg - pure (Op fc opFC (opn, Nothing) precInfo :: rtoks) + pure (Op fc opFC ((opn, fixInfo), Nothing) precInfo :: rtoks) toTokList t = pure [Expr t] record BangData where @@ -356,10 +360,12 @@ mutual desugarB side ps (PBracketed fc e) = desugarB side ps e desugarB side ps (POp fc opFC l op r) = do ts <- toTokList side (POp fc opFC l op r) - desugarTree side ps !(parseOps @{showFst} ts) + tree <- parseOps @{interpName} @{showWithLoc} ts + desugarTree side ps (mapFst (\((x, _), y) => (x, y)) tree) desugarB side ps (PPrefixOp fc opFC op arg) = do ts <- toTokList side (PPrefixOp fc opFC op arg) - desugarTree side ps !(parseOps @{showFst} ts) + tree <- parseOps @{interpName} @{showWithLoc} ts + desugarTree side ps (mapFst (\((x, _), y) => (x, y)) tree) desugarB side ps (PSectionL fc opFC op arg) = do syn <- get Syn -- It might actually be a prefix argument rather than a section diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index 5880b2c22..d72af4df0 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -642,12 +642,12 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) where moduleDiagnostic : Doc IdrisAnn moduleDiagnostic = case expected of - Nothing => "Import a module that exports a suitable fixity." - (Just a) => "Hide or remove the fixity at" <++> byShow a.fc + Backticked => "Import a module that exports a suitable fixity." + (DeclaredFixity a) => "Hide or remove the fixity at" <++> byShow a.fc <++> "and import a module that exports a compatible fixity." infixOpName : Doc IdrisAnn infixOpName = case expected of - Nothing => enclose "`" "`" (byShow opName) + Backticked => enclose "`" "`" (byShow opName) _ => byShow opName displayFixityInfo : FixityInfo -> BindingModifier -> Doc IdrisAnn @@ -657,8 +657,8 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) = byShow vis <++> byShow usedBinder <++> byShow fix <++> byShow precedence <++> pretty0 opName expressionDiagnositc : List (Doc IdrisAnn) expressionDiagnositc = case expected of - Nothing => [] - (Just e) => let sentence = "Write the expression using" <++> byShow e.bindingInfo <++> "syntax:" + Backticked => [] + (DeclaredFixity e) => let sentence = "Write the expression using" <++> byShow e.bindingInfo <++> "syntax:" in pure $ sentence <++> enclose "'" "'" (case e.bindingInfo of NotBinding => reAnnotate (const Code) (p actual.getLhs) @@ -676,8 +676,8 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) fixityDiagnostic : Doc IdrisAnn fixityDiagnostic = case expected of - Nothing => "Define a new fixity:" <++> "infixr 0" <++> infixOpName - (Just fix) => + Backticked => "Define a new fixity:" <++> "infixr 0" <++> infixOpName + (DeclaredFixity fix) => "Change the fixity defined at" <++> pretty0 fix.fc <++> "to" <++> enclose "'" "'" (displayFixityInfo fix actual.getBinder) <+> dot @@ -687,9 +687,9 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) printBindingModifier Typebind = "a type-binding (typebind)" printBindingModifier Autobind = "an automatically-binding (autobind)" - printBindingInfo : Maybe FixityInfo -> Doc IdrisAnn - printBindingInfo Nothing = "a regular" - printBindingInfo (Just x) = printBindingModifier x.bindingInfo + printBindingInfo : BacktickOrOperatorFixity -> Doc IdrisAnn + printBindingInfo Backticked = "a regular" + printBindingInfo (DeclaredFixity x) = printBindingModifier x.bindingInfo perrorRaw (TTCError msg) diff --git a/src/Libraries/Utils/Shunting.idr b/src/Libraries/Utils/Shunting.idr index a3c3cb314..f6cc37904 100644 --- a/src/Libraries/Utils/Shunting.idr +++ b/src/Libraries/Utils/Shunting.idr @@ -40,6 +40,12 @@ data Tree op a = Infix FC FC op (Tree op a) (Tree op a) | Pre FC FC op (Tree op a) | Leaf a +public export +Bifunctor Tree where + bimap f g (Infix fc fc1 x y z) = Infix fc fc1 (f x) (bimap f g y) (bimap f g z) + bimap f g (Pre fc fc1 x y) = Pre fc fc1 (f x) (bimap f g y) + bimap f g (Leaf x) = Leaf (g x) + export (Show op, Show a) => Show (Tree op a) where show (Infix _ _ op l r) = "(" ++ show op ++ " " ++ show l ++ " " ++ show r ++ ")" @@ -86,23 +92,25 @@ isLAssoc _ = False -- Return whether the first operator should be applied before the second, -- assuming -higher : Show op => FC -> op -> OpPrec -> op -> OpPrec -> Core Bool +higher : Interpolation op => (showLoc : Show op) => FC -> op -> OpPrec -> op -> OpPrec -> Core Bool higher loc opx op opy (Prefix p) = pure False higher loc opx (NonAssoc x) opy oy = if x == getPrec oy - then throw (GenericMsg loc ("Operator " ++ show opx ++ - " is non-associative")) + then throw (GenericMsgSol loc ("Operator \{opx} is non-associative") + ["Add brackets around \{opy}" + , "Change the fixity of \{show opy} to `infixl` or `infixr`"]) else pure (x > getPrec oy) higher loc opx ox opy (NonAssoc y) = if getPrec ox == y - then throw (GenericMsg loc ("Operator " ++ show opy ++ - " is non-associative")) + then throw (GenericMsgSol loc ("Operator \{opy} is non-associative") + ["Add brackets around \{opy}" + , "Change the fixity of \{show opy} to `infixl` or `infixr`"]) else pure (getPrec ox > y) higher loc opl l opr r = pure $ (getPrec l > getPrec r) || ((getPrec l == getPrec r) && isLAssoc l) -processStack : Show op => {auto o : Ref Out (List (Tree op a))} -> +processStack : Interpolation op => (showLoc : Show op) => {auto o : Ref Out (List (Tree op a))} -> List (FC, FC, op, OpPrec) -> op -> OpPrec -> Core (List (FC, FC, op, OpPrec)) processStack [] op prec = pure [] @@ -112,7 +120,7 @@ processStack (x@(loc, opFC, opx, sprec) :: xs) opy prec processStack xs opy prec else pure (x :: xs) -shunt : Show op => {auto o : Ref Out (List (Tree op a))} -> +shunt : Interpolation op => (showLoc : Show op) => {auto o : Ref Out (List (Tree op a))} -> (opstk : List (FC, FC, op, OpPrec)) -> List (Tok op a) -> Core (Tree op a) shunt stk (Expr x :: rest) @@ -131,7 +139,7 @@ shunt stk [] mkOp (loc, opFC, op, prec) = Op loc opFC op prec export -parseOps : Show op => List (Tok op a) -> Core (Tree op a) +parseOps : Interpolation op => (showLoc : Show op) => List (Tok op a) -> Core (Tree op a) parseOps toks = do o <- newRef {t = List (Tree op a)} Out [] shunt [] toks diff --git a/tests/idris2/operators/operators004/expected b/tests/idris2/operators/operators004/expected index 6a40255e4..70d253093 100644 --- a/tests/idris2/operators/operators004/expected +++ b/tests/idris2/operators/operators004/expected @@ -9,3 +9,6 @@ Test:8:8--8:22 8 | test = 4 -:- 3 -:- [] ^^^^^^^^^^^^^^ +Possible solutions: + - Add brackets around -:- + - Change the fixity of -:- at Test:2:1--2:12 to `infixl` or `infixr` From 1d3668d9a65e0b1bda874a5ebd4b42ef31c403c5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sun, 7 Jan 2024 10:42:33 +0000 Subject: [PATCH 30/57] Allow underscore as a valid name for binder operator --- src/Idris/Parser.idr | 3 ++- tests/idris2/operators/operators005/Test.idr | 19 +++++++++++++++++++ tests/idris2/operators/operators005/expected | 1 + tests/idris2/operators/operators005/run | 3 +++ 4 files changed, 25 insertions(+), 1 deletion(-) create mode 100644 tests/idris2/operators/operators005/Test.idr create mode 100644 tests/idris2/operators/operators005/expected create mode 100755 tests/idris2/operators/operators005/run diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 14ab15a5d..70cc616fc 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -343,7 +343,8 @@ mutual opBinder : OriginDesc -> IndentInfo -> Rule (OperatorLHSInfo PTerm) opBinder fname indents - = do boundName <- bounds (UN . Basic <$> decoratedSimpleBinderName fname) + = do boundName <- bounds (UN <$> (Basic <$> decoratedSimpleBinderName fname + <|> symbol "_" $> Underscore)) opBinderTypes fname indents boundName autobindOp : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm diff --git a/tests/idris2/operators/operators005/Test.idr b/tests/idris2/operators/operators005/Test.idr new file mode 100644 index 000000000..a96b27435 --- /dev/null +++ b/tests/idris2/operators/operators005/Test.idr @@ -0,0 +1,19 @@ + +import Data.String + +typebind infixr 0 :- +autobind infixr 0 `for` + +record Container where + constructor (:-) + a1 : Type + a2 : a1 -> Type + +const : Type -> Type -> Container +const a b = (_ : a) :- b + +test : Maybe (List Double) +test = (_ := ["1", "two", "3"]) `for` Just 3 + +test2 : Maybe (List Double) +test2 = (_ : String := ["1", "two", "3"]) `for` Just 3 diff --git a/tests/idris2/operators/operators005/expected b/tests/idris2/operators/operators005/expected new file mode 100644 index 000000000..31e5db7f3 --- /dev/null +++ b/tests/idris2/operators/operators005/expected @@ -0,0 +1 @@ +1/1: Building Test (Test.idr) diff --git a/tests/idris2/operators/operators005/run b/tests/idris2/operators/operators005/run new file mode 100755 index 000000000..d35387bb3 --- /dev/null +++ b/tests/idris2/operators/operators005/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr From 4eb9c978065cad1f592d4052e4dc87a6ff658e37 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Thu, 11 Jan 2024 23:11:34 +0000 Subject: [PATCH 31/57] Allow patterns in operator binders --- src/Core/TT.idr | 22 +++---- src/Idris/Desugar.idr | 64 +++++++++---------- src/Idris/Error.idr | 19 +++--- src/Idris/Parser.idr | 18 +----- src/Idris/Pretty.idr | 6 +- src/Idris/Syntax.idr | 6 +- .../interactive/interactive028/expected | 2 +- tests/idris2/operators/operators006/Test.idr | 10 +++ tests/idris2/operators/operators006/expected | 1 + tests/idris2/operators/operators006/run | 3 + 10 files changed, 76 insertions(+), 75 deletions(-) create mode 100644 tests/idris2/operators/operators006/Test.idr create mode 100644 tests/idris2/operators/operators006/expected create mode 100755 tests/idris2/operators/operators006/run diff --git a/src/Core/TT.idr b/src/Core/TT.idr index 6c1dde8a6..68bc6c6dd 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -189,11 +189,11 @@ data OperatorLHSInfo : tm -> Type where -- Traditional operator wihtout binding, carries the lhs NoBinder : (lhs : tm) -> OperatorLHSInfo tm -- (nm : ty) =@ fn x - BindType : (name : WithBounds Name) -> (ty : tm) -> OperatorLHSInfo tm + BindType : (name : WithBounds tm) -> (ty : tm) -> OperatorLHSInfo tm -- (nm := exp) =@ fn nm - BindExpr : (name : WithBounds Name) -> (expr : tm) -> OperatorLHSInfo tm + BindExpr : (name : WithBounds tm) -> (expr : tm) -> OperatorLHSInfo tm -- (nm : ty := exp) =@ fn nm - BindExplicitType : (name : WithBounds Name) -> (type, expr : tm) -> OperatorLHSInfo tm + BindExplicitType : (name : WithBounds tm) -> (type, expr : tm) -> OperatorLHSInfo tm export Show (OperatorLHSInfo tm) where @@ -207,9 +207,9 @@ Show (OperatorLHSInfo tm) where export Functor OperatorLHSInfo where map f (NoBinder lhs) = NoBinder $ f lhs - map f (BindType nm lhs) = BindType nm (f lhs) - map f (BindExpr nm lhs) = BindExpr nm (f lhs) - map f (BindExplicitType nm ty lhs) = BindExplicitType nm (f ty) (f lhs) + map f (BindType nm lhs) = BindType (map f nm) (f lhs) + map f (BindExpr nm lhs) = BindExpr (map f nm) (f lhs) + map f (BindExplicitType nm ty lhs) = BindExplicitType (map f nm) (f ty) (f lhs) export (.getLhs) : OperatorLHSInfo tm -> tm @@ -219,11 +219,11 @@ export (.getLhs) (BindExplicitType _ _ lhs) = lhs export -(.getBoundName) : OperatorLHSInfo tm -> Maybe Name -(.getBoundName) (NoBinder lhs) = Nothing -(.getBoundName) (BindType name ty) = Just name.val -(.getBoundName) (BindExpr name expr) = Just name.val -(.getBoundName) (BindExplicitType name type expr) = Just name.val +(.getBoundPat) : OperatorLHSInfo tm -> Maybe tm +(.getBoundPat) (NoBinder lhs) = Nothing +(.getBoundPat) (BindType name ty) = Just name.val +(.getBoundPat) (BindExpr name expr) = Just name.val +(.getBoundPat) (BindExplicitType name type expr) = Just name.val export (.getBinder) : OperatorLHSInfo tm -> BindingModifier diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index f834bf1d2..824bfe6fc 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -175,7 +175,7 @@ checkConflictingBinding LHS fc opName foundFixity use_site rhs = pure () -- don' checkConflictingBinding side fc opName foundFixity use_site rhs = if isCompatible foundFixity use_site then pure () - else throw $ OperatorBindingMismatch {print=byShow} fc foundFixity use_site opName rhs + else throw $ OperatorBindingMismatch {print = byShow} fc foundFixity use_site opName rhs where isCompatible : BacktickOrOperatorFixity -> OperatorLHSInfo PTerm -> Bool isCompatible Backticked (NoBinder lhs) = True @@ -361,11 +361,13 @@ mutual desugarB side ps (POp fc opFC l op r) = do ts <- toTokList side (POp fc opFC l op r) tree <- parseOps @{interpName} @{showWithLoc} ts - desugarTree side ps (mapFst (\((x, _), y) => (x, y)) tree) + unop <- desugarTree side ps (mapFst (\((x, _), y) => (x, y)) tree) + desugarB side ps unop desugarB side ps (PPrefixOp fc opFC op arg) = do ts <- toTokList side (PPrefixOp fc opFC op arg) tree <- parseOps @{interpName} @{showWithLoc} ts - desugarTree side ps (mapFst (\((x, _), y) => (x, y)) tree) + unop <- desugarTree side ps (mapFst (\((x, _), y) => (x, y)) tree) + desugarB side ps unop desugarB side ps (PSectionL fc opFC op arg) = do syn <- get Syn -- It might actually be a prefix argument rather than a section @@ -766,49 +768,45 @@ mutual rule' <- desugarDo side ps ns rule pure $ IRewrite fc rule' rest' - -- Handle special cases for operator overloads - -- Maybe overloading ** should be in here + -- Replace all operator by function application desugarTree : {auto s : Ref Syn SyntaxInfo} -> {auto b : Ref Bang BangData} -> {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> {auto m : Ref MD Metadata} -> {auto o : Ref ROpts REPLOpts} -> - Side -> List Name -> Tree (OpStr, Maybe $ OperatorLHSInfo PTerm) PTerm -> Core RawImp + Side -> List Name -> Tree (OpStr, Maybe $ OperatorLHSInfo PTerm) PTerm -> + Core PTerm desugarTree side ps (Infix loc eqFC (UN $ Basic "=", _) l r) -- special case since '=' is special syntax - = do l' <- desugarTree side ps l - r' <- desugarTree side ps r - pure (IAlternative loc FirstSuccess - [apply (IVar eqFC eqName) [l', r'], - apply (IVar eqFC heqName) [l', r']]) + = pure $ PEq loc !(desugarTree side ps l) !(desugarTree side ps r) desugarTree side ps (Infix loc _ (UN $ Basic "$", _) l r) -- special case since '$' is special syntax = do l' <- desugarTree side ps l r' <- desugarTree side ps r - pure (IApp loc l' r') + pure (PApp loc l' r') -- normal operators desugarTree side ps (Infix loc opFC (op, Just (NoBinder lhs)) l r) = do l' <- desugarTree side ps l r' <- desugarTree side ps r - pure (IApp loc (IApp loc (IVar opFC op) l') r') + pure (PApp loc (PApp loc (PRef opFC op) l') r') -- (x : ty) =@ f x ==>> (=@) exp (\x : ty => f x) - desugarTree side ps (Infix loc opFC (op, Just (BindType 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.val) desugaredLHS desugaredRHS) + desugarTree side ps (Infix loc opFC (op, Just (BindType pat lhs)) l r) + = do l' <- desugarTree side ps l + body <- desugarTree side ps r + pure $ PApp loc (PApp loc (PRef opFC op) l') + (PLam loc top Explicit pat.val l' body) -- (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.val) (Implicit opFC False) desugaredRHS) + desugarTree side ps (Infix loc opFC (op, Just (BindExpr pat lhs)) l r) + = do l' <- desugarTree side ps l + body <- desugarTree side ps r + pure $ PApp loc (PApp loc (PRef opFC op) l') + (PLam loc top Explicit pat.val (PInfer opFC) body) + -- (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.val) desugaredType desugaredRHS) + desugarTree side ps (Infix loc opFC (op, Just (BindExplicitType pat ty expr)) l r) + = do l' <- desugarTree side ps l + body <- desugarTree side ps r + pure $ PApp loc (PApp loc (PRef opFC op) l') + (PLam loc top Explicit pat.val ty body) desugarTree side ps (Infix loc opFC (op, Nothing) _ r) = throw $ InternalError "illegal fixity: Parsed as infix but no binding information" @@ -832,16 +830,16 @@ mutual -- not a signed integer literal. proceed by desugaring -- and applying to `negate`. _ => do arg' <- desugarTree side ps (Leaf $ PPrimVal fc c) - pure (IApp loc (IVar opFC (UN $ Basic "negate")) arg') + pure (PApp loc (PRef opFC (UN $ Basic "negate")) arg') desugarTree side ps (Pre loc opFC (UN $ Basic "-", _) arg) = do arg' <- desugarTree side ps arg - pure (IApp loc (IVar opFC (UN $ Basic "negate")) arg') + pure (PApp loc (PRef opFC (UN $ Basic "negate")) arg') desugarTree side ps (Pre loc opFC (op, _) arg) = do arg' <- desugarTree side ps arg - pure (IApp loc (IVar opFC op) arg') - desugarTree side ps (Leaf t) = desugarB side ps t + pure (PApp loc (PRef opFC op) arg') + desugarTree side ps (Leaf t) = pure t desugarType : {auto s : Ref Syn SyntaxInfo} -> {auto c : Ref Ctxt Defs} -> diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index d72af4df0..0426135b1 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -655,22 +655,25 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) = byShow vis <++> byShow fix <++> byShow precedence <++> pretty0 opName displayFixityInfo (MkFixityInfo _ vis _ fix precedence) usedBinder = byShow vis <++> byShow usedBinder <++> byShow fix <++> byShow precedence <++> pretty0 opName + + printE : ? -> Doc IdrisAnn + printE x = reAnnotate (const Code) (p x) + expressionDiagnositc : List (Doc IdrisAnn) expressionDiagnositc = case expected of Backticked => [] (DeclaredFixity e) => let sentence = "Write the expression using" <++> byShow e.bindingInfo <++> "syntax:" in pure $ sentence <++> enclose "'" "'" (case e.bindingInfo of NotBinding => - reAnnotate (const Code) (p actual.getLhs) - <++> infixOpName <++> reAnnotate (const Code) (p rhs) + printE actual.getLhs <++> infixOpName <++> printE rhs Autobind => - parens (maybe "_" pretty0 actual.getBoundName <++> ":=" - <++> reAnnotate (const Code) (p actual.getLhs)) - <++> infixOpName <++> reAnnotate (const Code) (p rhs) + parens (maybe "_" printE actual.getBoundPat <++> ":=" + <++> printE actual.getLhs) + <++> infixOpName <++> printE rhs Typebind => - parens (maybe "_" pretty0 actual.getBoundName <++> ":" - <++> reAnnotate (const Code) (p actual.getLhs)) - <++> infixOpName <++> reAnnotate (const Code) (p rhs) + parens (maybe "_" printE actual.getBoundPat <++> ":" + <++> printE actual.getLhs) + <++> infixOpName <++> printE rhs ) <+> dot diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 70cc616fc..1e1979111 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -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 -> WithBounds Name -> Rule (OperatorLHSInfo PTerm) + opBinderTypes : OriginDesc -> IndentInfo -> WithBounds PTerm -> Rule (OperatorLHSInfo PTerm) opBinderTypes fname indents boundName = do decoratedSymbol fname ":" ty <- typeExpr pdef fname indents @@ -343,8 +343,7 @@ mutual opBinder : OriginDesc -> IndentInfo -> Rule (OperatorLHSInfo PTerm) opBinder fname indents - = do boundName <- bounds (UN <$> (Basic <$> decoratedSimpleBinderName fname - <|> symbol "_" $> Underscore)) + = do boundName <- bounds (expr plhs fname indents) opBinderTypes fname indents boundName autobindOp : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm @@ -717,19 +716,6 @@ mutual 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 - autoImplicitPi : OriginDesc -> IndentInfo -> Rule PTerm autoImplicitPi fname indents = do b <- bounds $ do diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index ca91e4ded..3280d50c9 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -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.val <++> ":" <++> pretty left) + group $ parens (prettyPrec d nm.val <++> ":" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (BindExpr nm left) op right) = - group $ parens (prettyBinder nm.val <++> ":=" <++> pretty left) + group $ parens (prettyPrec d nm.val <++> ":=" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (BindExplicitType nm ty left) op right) = - group $ parens (prettyBinder nm.val <++> ":" <++> pretty ty <++> ":=" <++> pretty left) + group $ parens (prettyPrec d nm.val <++> ":" <++> pretty ty <++> ":=" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (NoBinder x) op y) = diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index 70a9ef3ec..7c076d7de 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -773,11 +773,11 @@ parameters {0 nm : Type} (toName : nm -> Name) showPTermPrec d (POp _ _ (NoBinder left) op right) = showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right showPTermPrec d (POp _ _ (BindType nm left) op right) - = "(" ++ show nm ++ " : " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ showPTermPrec d nm.val ++ " : " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (POp _ _ (BindExpr nm left) op right) - = "(" ++ show nm ++ " := " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ showPTermPrec d nm.val ++ " := " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (POp _ _ (BindExplicitType nm ty left) op right) - = "(" ++ show nm ++ " : " ++ showPTermPrec d ty ++ ":=" ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ showPTermPrec d nm.val ++ " : " ++ showPTermPrec d ty ++ ":=" ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (PPrefixOp _ _ op x) = showOpPrec d op ++ showPTermPrec d x showPTermPrec d (PSectionL _ _ op x) = "(" ++ showOpPrec d op ++ " " ++ showPTermPrec d x ++ ")" showPTermPrec d (PSectionR _ _ x op) = "(" ++ showPTermPrec d x ++ " " ++ showOpPrec d op ++ ")" diff --git a/tests/idris2/interactive/interactive028/expected b/tests/idris2/interactive/interactive028/expected index 529623d7f..6b78dac3e 100644 --- a/tests/idris2/interactive/interactive028/expected +++ b/tests/idris2/interactive/interactive028/expected @@ -4,7 +4,7 @@ Main> Couldn't parse any alternatives: (Interactive):1:4--1:5 1 | :t (3 : Nat) ^ -... (54 others) +... (56 others) Main> Expected string begin. (Interactive):1:5--1:7 diff --git a/tests/idris2/operators/operators006/Test.idr b/tests/idris2/operators/operators006/Test.idr new file mode 100644 index 000000000..b9aa1639c --- /dev/null +++ b/tests/idris2/operators/operators006/Test.idr @@ -0,0 +1,10 @@ + +import Data.Fin + +autobind infixr 0 `bind` + +bind : Monad m => m a -> (a -> m b) -> m b +bind = (>>=) + +both : Maybe (Nat, Nat) -> Maybe Nat +both m = (MkPair x y := m) `bind` Just (x + y) diff --git a/tests/idris2/operators/operators006/expected b/tests/idris2/operators/operators006/expected new file mode 100644 index 000000000..31e5db7f3 --- /dev/null +++ b/tests/idris2/operators/operators006/expected @@ -0,0 +1 @@ +1/1: Building Test (Test.idr) diff --git a/tests/idris2/operators/operators006/run b/tests/idris2/operators/operators006/run new file mode 100755 index 000000000..d35387bb3 --- /dev/null +++ b/tests/idris2/operators/operators006/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr From ebb8b55d4564c60856a494e7a052596dc1f0e0d7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Mon, 15 Jan 2024 16:19:54 +0000 Subject: [PATCH 32/57] Udpate IDE test This change is correct because `=` is indeed a function and should be colored as such. --- src/Idris/Desugar.idr | 2 +- tests/ideMode/ideMode005/expected | 6 ++++++ 2 files changed, 7 insertions(+), 1 deletion(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 824bfe6fc..79e25dcbd 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -778,7 +778,7 @@ mutual Side -> List Name -> Tree (OpStr, Maybe $ OperatorLHSInfo PTerm) PTerm -> Core PTerm desugarTree side ps (Infix loc eqFC (UN $ Basic "=", _) l r) -- special case since '=' is special syntax - = pure $ PEq loc !(desugarTree side ps l) !(desugarTree side ps r) + = pure $ PEq loc !(desugarTree side ps l) !(desugarTree side ps r) desugarTree side ps (Infix loc _ (UN $ Basic "$", _) l r) -- special case since '$' is special syntax = do l' <- desugarTree side ps l r' <- desugarTree side ps r diff --git a/tests/ideMode/ideMode005/expected b/tests/ideMode/ideMode005/expected index e69de29bb..299fc037f 100644 --- a/tests/ideMode/ideMode005/expected +++ b/tests/ideMode/ideMode005/expected @@ -0,0 +1,6 @@ +24a25 +> 0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 40) (:end 2 49)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +49a50 +> 0000dd(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 7) (:end 17 47)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +40a41 +> 0000d3(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 8 27) (:end 8 44)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) From 31ea83039c1ee9627b603bce5267bd5b236752bf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 20 Jan 2024 16:59:57 +0000 Subject: [PATCH 33/57] first attempt at suggesting different operators --- src/Core/Context.idr | 10 ++++++---- src/Core/Core.idr | 12 ++++++------ src/Idris/Desugar.idr | 13 +++++++++++-- src/Idris/Error.idr | 11 +++++++++-- tests/idris2/operators/operators007/Test.idr | 11 +++++++++++ tests/idris2/operators/operators007/expected | 17 +++++++++++++++++ tests/idris2/operators/operators007/run | 3 +++ 7 files changed, 63 insertions(+), 14 deletions(-) create mode 100644 tests/idris2/operators/operators007/Test.idr create mode 100644 tests/idris2/operators/operators007/expected create mode 100755 tests/idris2/operators/operators007/run diff --git a/src/Core/Context.idr b/src/Core/Context.idr index eae47cc22..d7d52015b 100644 --- a/src/Core/Context.idr +++ b/src/Core/Context.idr @@ -821,8 +821,9 @@ HasNames Error where full gam (InRHS fc n err) = InRHS fc <$> full gam n <*> full gam err full gam (MaybeMisspelling err xs) = MaybeMisspelling <$> full gam err <*> pure xs full gam (WarningAsError wrn) = WarningAsError <$> full gam wrn - full gam (OperatorBindingMismatch {print} fc expected actual opName rhs) - = OperatorBindingMismatch {print} fc expected actual <$> full gam opName <*> pure rhs + full gam (OperatorBindingMismatch {print} fc expected actual opName rhs candidates) + = OperatorBindingMismatch {print} fc expected actual + <$> full gam opName <*> pure rhs <*> pure candidates resolved gam (Fatal err) = Fatal <$> resolved gam err resolved _ (CantConvert fc gam rho s t) @@ -915,8 +916,9 @@ HasNames Error where resolved gam (InRHS fc n err) = InRHS fc <$> resolved gam n <*> resolved gam err resolved gam (MaybeMisspelling err xs) = MaybeMisspelling <$> resolved gam err <*> pure xs resolved gam (WarningAsError wrn) = WarningAsError <$> resolved gam wrn - resolved gam (OperatorBindingMismatch {print} fc expected actual opName rhs) - = OperatorBindingMismatch {print} fc expected actual <$> resolved gam opName <*> pure rhs + resolved gam (OperatorBindingMismatch {print} fc expected actual opName rhs candidates) + = OperatorBindingMismatch {print} fc expected actual + <$> resolved gam opName <*> pure rhs <*> pure candidates export HasNames Totality where diff --git a/src/Core/Core.idr b/src/Core/Core.idr index 0fc0eaee3..8eaa16b60 100644 --- a/src/Core/Core.idr +++ b/src/Core/Core.idr @@ -166,7 +166,7 @@ data Error : Type where GenericMsgSol : FC -> (message : String) -> (solutions : List String) -> Error OperatorBindingMismatch : {a : Type} -> {print : a -> Doc ()} -> FC -> (expectedFixity : BacktickOrOperatorFixity) -> (use_site : OperatorLHSInfo a) -> - (opName : Name) -> (rhs : a) -> Error + (opName : Name) -> (rhs : a) -> (candidates : List String) -> Error TTCError : TTCErrorMsg -> Error FileErr : String -> FileError -> Error CantFindPackage : String -> Error @@ -400,10 +400,10 @@ Show Error where (n ::: []) => ": " ++ n ++ "?" _ => " any of: " ++ showSep ", " (map show (forget ns)) ++ "?" show (WarningAsError w) = show w - show (OperatorBindingMismatch fc (DeclaredFixity expected) actual opName rhs) + show (OperatorBindingMismatch fc (DeclaredFixity expected) actual opName rhs _) = show fc ++ ": Operator " ++ show opName ++ " is " ++ show expected ++ " but used as a " ++ show actual ++ " operator" - show (OperatorBindingMismatch fc Backticked actual opName rhs) + show (OperatorBindingMismatch fc Backticked actual opName rhs _) = show fc ++ ": Operator " ++ show opName ++ " has no declared fixity" ++ " but used as a " ++ show actual ++ " operator" @@ -496,7 +496,7 @@ getErrorLoc (InLHS _ _ err) = getErrorLoc err getErrorLoc (InRHS _ _ err) = getErrorLoc err getErrorLoc (MaybeMisspelling err _) = getErrorLoc err getErrorLoc (WarningAsError warn) = Just (getWarningLoc warn) -getErrorLoc (OperatorBindingMismatch fc _ _ _ _) = Just fc +getErrorLoc (OperatorBindingMismatch fc _ _ _ _ _) = Just fc export killWarningLoc : Warning -> Warning @@ -586,8 +586,8 @@ killErrorLoc (InLHS fc x err) = InLHS emptyFC x (killErrorLoc err) killErrorLoc (InRHS fc x err) = InRHS emptyFC x (killErrorLoc err) killErrorLoc (MaybeMisspelling err xs) = MaybeMisspelling (killErrorLoc err) xs killErrorLoc (WarningAsError wrn) = WarningAsError (killWarningLoc wrn) -killErrorLoc (OperatorBindingMismatch {print} fc expected actual opName rhs) - = OperatorBindingMismatch {print} emptyFC expected actual opName rhs +killErrorLoc (OperatorBindingMismatch {print} fc expected actual opName rhs candidates) + = OperatorBindingMismatch {print} emptyFC expected actual opName rhs candidates -- Core is a wrapper around IO that is specialised for efficiency. diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 79e25dcbd..21782b898 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -170,12 +170,15 @@ checkConflictingFixities isPrefix exprFC opn For example: %hide \{show fxName} """ -checkConflictingBinding : Side -> FC -> Name -> (foundFixity : BacktickOrOperatorFixity) -> (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () +checkConflictingBinding : Ref Ctxt Defs => + Side -> FC -> Name -> (foundFixity : BacktickOrOperatorFixity) -> + (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () checkConflictingBinding LHS fc opName foundFixity use_site rhs = pure () -- don't check if we're on the LHS checkConflictingBinding side fc opName foundFixity use_site rhs = if isCompatible foundFixity use_site then pure () - else throw $ OperatorBindingMismatch {print = byShow} fc foundFixity use_site opName rhs + else throw $ OperatorBindingMismatch + {print = byShow} fc foundFixity use_site opName rhs !candidates where isCompatible : BacktickOrOperatorFixity -> OperatorLHSInfo PTerm -> Bool isCompatible Backticked (NoBinder lhs) = True @@ -186,6 +189,12 @@ checkConflictingBinding side fc opName foundFixity use_site rhs isCompatible (DeclaredFixity fixInfo) (BindExplicitType name type expr) = fixInfo.bindingInfo == Autobind + candidates : Core (List String) + candidates = do Just (nm, cs) <- getSimilarNames opName + | Nothing => pure [] + ns <- currentNS <$> get Ctxt + pure (showSimilarNames ns opName nm cs) + checkValidFixity : BindingModifier -> Fixity -> Nat -> Bool -- If the fixity declaration is not binding, there are no restrictions diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index 0426135b1..e855cb9e3 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -625,7 +625,7 @@ perrorRaw (GenericMsgSol fc header solutions) <+> line <+> "Possible solutions:" <+> line <+> indent 1 (vsep (map (\s => "-" <++> pretty0 s) solutions)) -perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) +perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs candidates) = pure $ "Operator" <++> pretty0 !(getFullName opName) <++> "is" <++> printBindingInfo expected <++> "operator, but is used as" <++> printBindingModifier actual.getBinder @@ -638,8 +638,15 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs) <+> line <+> line <+> "Possible solutions:" <+> line <+> indent 1 (vsep (map ("-" <++>) - (expressionDiagnositc ++ [fixityDiagnostic, moduleDiagnostic]))) + (expressionDiagnositc ++ [fixityDiagnostic, moduleDiagnostic] ++ spellingCandidates))) where + spellingCandidates : List (Doc IdrisAnn) + spellingCandidates = if null candidates + then [] + else ["did you mean one of:" <++> hcat (punctuate ", " + (map byShow candidates))] + + moduleDiagnostic : Doc IdrisAnn moduleDiagnostic = case expected of Backticked => "Import a module that exports a suitable fixity." diff --git a/tests/idris2/operators/operators007/Test.idr b/tests/idris2/operators/operators007/Test.idr new file mode 100644 index 000000000..61049346d --- /dev/null +++ b/tests/idris2/operators/operators007/Test.idr @@ -0,0 +1,11 @@ + +import Data.Fin + +autobind infixr 0 >> +infixr 0 !> + +(>>) : Monad m => m a -> (a -> m b) -> m b +(>>) = (>>=) + +both : Maybe (Nat, Nat) -> Maybe Nat +both m = (MkPair x y := m) !> Just (x + y) diff --git a/tests/idris2/operators/operators007/expected b/tests/idris2/operators/operators007/expected new file mode 100644 index 000000000..85f117972 --- /dev/null +++ b/tests/idris2/operators/operators007/expected @@ -0,0 +1,17 @@ +1/1: Building Test (Test.idr) +Error: Operator !> is a regular operator, but is used as an automatically-binding (autobind) operator. + +Test:11:28--11:30 + 07 | (>>) : Monad m => m a -> (a -> m b) -> m b + 08 | (>>) = (>>=) + 09 | + 10 | both : Maybe (Nat, Nat) -> Maybe Nat + 11 | both m = (MkPair x y := m) !> Just (x + y) + ^^ +Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) !> expr', autobind looks like this: '(name := expr) !> expr'. + +Possible solutions: + - Write the expression using regular syntax: 'm !> Just (x + y)'. + - Change the fixity defined at Test:5:1--5:12 to 'export autobind infixr 0 !>'. + - Hide or remove the fixity at Test:5:1--5:12 and import a module that exports a compatible fixity. + - Did you mean '>>' ? diff --git a/tests/idris2/operators/operators007/run b/tests/idris2/operators/operators007/run new file mode 100755 index 000000000..d35387bb3 --- /dev/null +++ b/tests/idris2/operators/operators007/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr From e0f5ee999633f9c0510cdbcacdfd2df28d3c9fb6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 20 Jan 2024 19:47:23 +0000 Subject: [PATCH 34/57] implement compatible operator suggestions --- src/Core/Context.idr | 8 +++- src/Idris/Desugar.idr | 18 ++++++-- src/Idris/Error.idr | 9 ++-- src/Idris/Syntax.idr | 7 ++++ tests/idris2/operators/operators007/Test.idr | 3 +- tests/idris2/operators/operators007/Test2.idr | 14 +++++++ tests/idris2/operators/operators007/expected | 41 +++++++++++++------ tests/idris2/operators/operators007/run | 2 + 8 files changed, 79 insertions(+), 23 deletions(-) create mode 100644 tests/idris2/operators/operators007/Test2.idr diff --git a/src/Core/Context.idr b/src/Core/Context.idr index d7d52015b..4c388606d 100644 --- a/src/Core/Context.idr +++ b/src/Core/Context.idr @@ -1131,7 +1131,10 @@ getFieldNames ctxt recNS -- Find similar looking names in the context export getSimilarNames : {auto c : Ref Ctxt Defs} -> - Name -> Core (Maybe (String, List (Name, Visibility, Nat))) + -- Predicate run to customise the behavior of looking for similar names + -- Sometime we might want to hide names that we know make no sense. + {default Nothing keepPredicate : Maybe ((Name, GlobalDef) -> Core Bool)} -> + Name -> Core (Maybe (String, List (Name, Visibility, Nat))) getSimilarNames nm = case show <$> userNameRoot nm of Nothing => pure Nothing Just str => if length str <= 1 then pure (Just (str, [])) else @@ -1145,6 +1148,9 @@ getSimilarNames nm = case show <$> userNameRoot nm of | False => pure Nothing Just def <- lookupCtxtExact nm (gamma defs) | Nothing => pure Nothing -- should be impossible + let predicate = fromMaybe (\_ => pure True) keepPredicate + True <- predicate (nm, def) + | False => pure Nothing pure (Just (collapseDefault $ visibility def, dist)) kept <- NameMap.mapMaybeM @{CORE} test (resolvedAs (gamma defs)) pure $ Just (str, toList kept) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 21782b898..6cc08edd1 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -42,6 +42,7 @@ import Data.Maybe import Data.List import Data.List.Views import Data.String +import Debug.Trace -- Convert high level Idris declarations (PDecl from Idris.Syntax) into -- TTImp, recording any high level syntax info on the way (e.g. infix @@ -123,9 +124,8 @@ checkConflictingFixities : {auto s : Ref Syn SyntaxInfo} -> (isPrefix : Bool) -> FC -> Name -> Core (OpPrec, BacktickOrOperatorFixity) checkConflictingFixities isPrefix exprFC opn - = do syn <- get Syn - let op = nameRoot opn - let foundFixities : List (Name, FixityInfo) = lookupName (UN (Basic op)) (fixities syn) + = do let op = nameRoot opn + foundFixities <- getFixityInfo op let (pre, inf) = partition ((== Prefix) . fix . snd) foundFixities case (isPrefix, pre, inf) of -- If we do not find any fixity for this operator we check that it uses operator @@ -171,6 +171,7 @@ checkConflictingFixities isPrefix exprFC opn """ checkConflictingBinding : Ref Ctxt Defs => + Ref Syn SyntaxInfo => Side -> FC -> Name -> (foundFixity : BacktickOrOperatorFixity) -> (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () checkConflictingBinding LHS fc opName foundFixity use_site rhs = pure () -- don't check if we're on the LHS @@ -189,8 +190,17 @@ checkConflictingBinding side fc opName foundFixity use_site rhs isCompatible (DeclaredFixity fixInfo) (BindExplicitType name type expr) = fixInfo.bindingInfo == Autobind + keepCompatibleBinding : BindingModifier -> (Name, GlobalDef) -> Core Bool + keepCompatibleBinding compatibleBinder (name, def) = do + fixities <- getFixityInfo (nameRoot name) + let compatible = any (\(_, fx) => fx.bindingInfo == use_site.getBinder) fixities + pure compatible + candidates : Core (List String) - candidates = do Just (nm, cs) <- getSimilarNames opName + candidates = do let DeclaredFixity fxInfo = foundFixity + | _ => pure [] -- if there is no declared fixity we can't know what's + -- supposed to go there. + Just (nm, cs) <- getSimilarNames {keepPredicate = Just (keepCompatibleBinding fxInfo.bindingInfo)} opName | Nothing => pure [] ns <- currentNS <$> get Ctxt pure (showSimilarNames ns opName nm cs) diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index e855cb9e3..a20be6fce 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -641,10 +641,11 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs candi (expressionDiagnositc ++ [fixityDiagnostic, moduleDiagnostic] ++ spellingCandidates))) where spellingCandidates : List (Doc IdrisAnn) - spellingCandidates = if null candidates - then [] - else ["did you mean one of:" <++> hcat (punctuate ", " - (map byShow candidates))] + spellingCandidates = case candidates of + [] => [] + [x] => ["Did you mean" <++> enclose "'" "'" (pretty0 x) <++> "?"] + xs => ["Did you mean one of:" <++> hcat (punctuate ", " + (map (enclose "'" "'" . pretty0) xs)) <++> "?"] moduleDiagnostic : Doc IdrisAnn diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index 7c076d7de..7226a7b28 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -1005,6 +1005,13 @@ removeFixity : {auto s : Ref Syn SyntaxInfo} -> Fixity -> Name -> Core () removeFixity _ key = update Syn ({fixities $= removeExact key }) +||| Return all fixity declarations for an operator name +export +getFixityInfo : {auto s : Ref Syn SyntaxInfo} -> String -> Core (List (Name, FixityInfo)) +getFixityInfo nm = do + syn <- get Syn + pure $ lookupName (UN $ Basic nm) (fixities syn) + export covering Show PTypeDecl where diff --git a/tests/idris2/operators/operators007/Test.idr b/tests/idris2/operators/operators007/Test.idr index 61049346d..b15090ce2 100644 --- a/tests/idris2/operators/operators007/Test.idr +++ b/tests/idris2/operators/operators007/Test.idr @@ -2,10 +2,9 @@ import Data.Fin autobind infixr 0 >> -infixr 0 !> (>>) : Monad m => m a -> (a -> m b) -> m b (>>) = (>>=) both : Maybe (Nat, Nat) -> Maybe Nat -both m = (MkPair x y := m) !> Just (x + y) +both m = (MkPair x y := m) >>= Just (x + y) diff --git a/tests/idris2/operators/operators007/Test2.idr b/tests/idris2/operators/operators007/Test2.idr new file mode 100644 index 000000000..021923a00 --- /dev/null +++ b/tests/idris2/operators/operators007/Test2.idr @@ -0,0 +1,14 @@ + +import Data.Fin + +autobind infixr 0 >> +autobind infixr 0 >= + +(>>) : Monad m => m a -> (a -> m b) -> m b +(>>) = (>>=) + +(>=) : Monad m => m a -> (a -> m b) -> m b +(>=) = (>>=) + +both : Maybe (Nat, Nat) -> Maybe Nat +both m = (MkPair x y := m) >>= Just (x + y) diff --git a/tests/idris2/operators/operators007/expected b/tests/idris2/operators/operators007/expected index 85f117972..d006e178e 100644 --- a/tests/idris2/operators/operators007/expected +++ b/tests/idris2/operators/operators007/expected @@ -1,17 +1,34 @@ 1/1: Building Test (Test.idr) -Error: Operator !> is a regular operator, but is used as an automatically-binding (autobind) operator. +Error: Operator >>= is a regular operator, but is used as an automatically-binding (autobind) operator. -Test:11:28--11:30 - 07 | (>>) : Monad m => m a -> (a -> m b) -> m b - 08 | (>>) = (>>=) - 09 | - 10 | both : Maybe (Nat, Nat) -> Maybe Nat - 11 | both m = (MkPair x y := m) !> Just (x + y) - ^^ -Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) !> expr', autobind looks like this: '(name := expr) !> expr'. +Test:10:28--10:31 + 06 | (>>) : Monad m => m a -> (a -> m b) -> m b + 07 | (>>) = (>>=) + 08 | + 09 | both : Maybe (Nat, Nat) -> Maybe Nat + 10 | both m = (MkPair x y := m) >>= Just (x + y) + ^^^ +Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) >>= expr', autobind looks like this: '(name := expr) >>= expr'. Possible solutions: - - Write the expression using regular syntax: 'm !> Just (x + y)'. - - Change the fixity defined at Test:5:1--5:12 to 'export autobind infixr 0 !>'. - - Hide or remove the fixity at Test:5:1--5:12 and import a module that exports a compatible fixity. + - Write the expression using regular syntax: 'm >>= Just (x + y)'. + - Change the fixity defined at Prelude.Ops:20:1--20:37 to 'export autobind infixl 1 >>='. + - Hide or remove the fixity at Prelude.Ops:20:1--20:37 and import a module that exports a compatible fixity. - Did you mean '>>' ? +1/1: Building Test2 (Test2.idr) +Error: Operator >>= is a regular operator, but is used as an automatically-binding (autobind) operator. + +Test2:14:28--14:31 + 10 | (>=) : Monad m => m a -> (a -> m b) -> m b + 11 | (>=) = (>>=) + 12 | + 13 | both : Maybe (Nat, Nat) -> Maybe Nat + 14 | both m = (MkPair x y := m) >>= Just (x + y) + ^^^ +Explanation: regular, typebind and autobind operators all use a slightly different syntax, typebind looks like this: '(name : type) >>= expr', autobind looks like this: '(name := expr) >>= expr'. + +Possible solutions: + - Write the expression using regular syntax: 'm >>= Just (x + y)'. + - Change the fixity defined at Prelude.Ops:20:1--20:37 to 'export autobind infixl 1 >>='. + - Hide or remove the fixity at Prelude.Ops:20:1--20:37 and import a module that exports a compatible fixity. + - Did you mean one of: '>=', '>>' ? diff --git a/tests/idris2/operators/operators007/run b/tests/idris2/operators/operators007/run index d35387bb3..09019ff63 100755 --- a/tests/idris2/operators/operators007/run +++ b/tests/idris2/operators/operators007/run @@ -1,3 +1,5 @@ . ../../../testutils.sh check Test.idr + +check Test2.idr From c6fbd27c6c5d178a66f0de8e14619a477f03cfea Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:32:11 +0900 Subject: [PATCH 35/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index 9cb7bfb57..48a5e4284 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -20,7 +20,7 @@ file namespacing and follow the same rules as other defintions. Basics ====== -Before we jump into the fancy features, let us explain how infix operators work +Before we jump into the fancy features, let us explain how operators work for most users. When you see an expression From bcb225568a3a469f2ebede3c6d9537553fecd0cb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:32:21 +0900 Subject: [PATCH 36/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index 48a5e4284..ecf9ea09c 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -47,7 +47,7 @@ brackets are mandatory. Here, we chose for ``+`` to be left-associative, hence ` The number after the fixity indicate the *precedence level* of the operator, that is, if it should be bracketed before, or after, other operators used in the same expression. For example, -we want ``*`` to *take precedence* over ``+`` , because of this, we define it like this: +we want ``*`` to *take precedence* over ``+`` we write: .. code-block:: idris From c4c79c9115fb2fcea80b7b968b8827564d61ab2a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:32:51 +0900 Subject: [PATCH 37/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index ecf9ea09c..1d2159817 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -55,7 +55,7 @@ we want ``*`` to *take precedence* over ``+`` we write: This way, the expression ``n + m * x`` is correctly interpreted as ``n + (m * x)``. -Fixities declarations are optional and only change how a file is parsed, but you can +Fixity declarations are optional and change how a file is parsed, but you can use any function defined with operator symbols with parenthesis around it: .. code-block:: idris From ef0ca76b4c29f9ccbe4986629996d55346b81d50 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:33:08 +0900 Subject: [PATCH 38/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index 1d2159817..e8b616702 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -65,7 +65,7 @@ use any function defined with operator symbols with parenthesis around it: (+) n 3 Because fixities are separated from the function definitions, a single operator -can have 0 or multiple fixity definitions. In the next section we explain how to +can have 0 or multiple fixity definitions. In the next section, we explain how to deal with multiple fixity definitions. Fixity & Precedence Namespacing From b22c7f82895ac4ad1bc628dbb23d4b416354a32c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:33:37 +0900 Subject: [PATCH 39/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index e8b616702..baf030eea 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -99,7 +99,7 @@ example: test = 1 - 3 - 10 This program will raise a warning on the last line of module ``C`` because -there are two conflicting fixities in scope, should we parse the expression +there are two conflicting fixities in scope. Should we parse the expression as ``(1 - 3) - 10`` or as ``1 - (3 - 10)``? In those cases, you can hide the extra fixity you do not wish to use by using ``%hide``: From 23b2ddd1794d9a6cbfc7280e14eacfb8b6e72aaf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:36:54 +0900 Subject: [PATCH 40/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index baf030eea..0ccafa44e 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -239,7 +239,7 @@ right-associative. Typebind Operators ================== -In dependently-typed programming, we have the ability define types which +In dependently-typed programming, we have the ability define constructors which first argument is a type and the second is a lambda using the first argument as it's type. A typical example of this is the dependent linear arrow: From e1f4fbba836295de59e6cef5463220eb8bb8a7f0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:37:49 +0900 Subject: [PATCH 41/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index 0ccafa44e..22bc2d233 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -240,8 +240,8 @@ Typebind Operators ================== In dependently-typed programming, we have the ability define constructors which -first argument is a type and the second is a lambda using the first argument -as it's type. A typical example of this is the dependent linear arrow: +first argument is a type and the second is a type indexed over the first argument. +A typical example of this is the dependent linear arrow: .. code-block:: idris From e7a4914f113d3049e3200186c7516513354fdcac Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:38:12 +0900 Subject: [PATCH 42/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index 22bc2d233..66cb0b72b 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -250,8 +250,8 @@ A typical example of this is the dependent linear arrow: (=@) x f = (1 v : x) -> f v -However, we cannot use as is because the second argument is -a lambda, and writing out any value using this operator will look a bit awkward: +However, when trying to use it in infix position, we have to use a lambda to populate the +second argument: .. code-block:: idris From 93de3e0e07fd24859d508e636958bd7351ecf185 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:39:04 +0900 Subject: [PATCH 43/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index 66cb0b72b..27cedc4f0 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -289,9 +289,10 @@ only ever be ``infixr`` with precedence 0. Autobind Operators ================== -Typebind operators allow to bind a *type* on the left side of an operator, but -sometimes, there is no dependency between the first argument, and the expression -on the right side of the operator. For those case, we use ``autobind``. +Typebind operators allow to bind a *type* on the left side of an operator, so that is can +be used as the index of the second argument. But sometimes, there is no dependency +between the first and second argument, yet we still want to use binding sytnax. For those +case, we use ``autobind``. An example of this is a DSL for a dependently-typed programming language where the constructor for ``Pi`` takes terms on the left and lambdas on the right: From 48618490d02a3ba213328e670877eb158fc8bf24 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:39:32 +0900 Subject: [PATCH 44/57] Update src/Idris/Doc/Keywords.idr --- src/Idris/Doc/Keywords.idr | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Idris/Doc/Keywords.idr b/src/Idris/Doc/Keywords.idr index 0f1d4a408..0c5ecf2b5 100644 --- a/src/Idris/Doc/Keywords.idr +++ b/src/Idris/Doc/Keywords.idr @@ -480,8 +480,8 @@ autobindDoc = """ operator sections. `autobind` operators are desugared as a lambda: - `(x := expr) =|> fn x` -> `(expr =@ (\x : ? =|> fn x))` - `(x : ty := expr) =|> fn x` -> `(expr =@ (\x : ty =|> fn x))` + `(x := expr) =|> fn x` -> `(expr =|> (\x : ? => fn x))` + `(x : ty := expr) =|> fn x` -> `(expr =|> (\x : ty => fn x))` """ typebindDoc : Doc IdrisDocAnn typebindDoc = """ From 58bcfb73c1d46ac58e8c838e1e931312b8673d80 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:39:44 +0900 Subject: [PATCH 45/57] Update tests/idris2/operators/operators001/Test.idr --- tests/idris2/operators/operators001/Test.idr | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/tests/idris2/operators/operators001/Test.idr b/tests/idris2/operators/operators001/Test.idr index e3d5b8f5a..c33218298 100644 --- a/tests/idris2/operators/operators001/Test.idr +++ b/tests/idris2/operators/operators001/Test.idr @@ -31,7 +31,7 @@ map3 : (x : a) =@ b -@ (y : List a) =@ List b map4 : (x : a) =@ (b -@ (y : List a) =@ List b) --- this could be possible if we allowed binding operatprs with binding operators +-- this could be possible if we allowed binding operators -- with higher precedences -- test : Test.map === Test.map2 -- failing From 8c69b3e749ee33a5d30053a8a5dff44c3628d972 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:39:55 +0900 Subject: [PATCH 46/57] Update src/Idris/Desugar.idr --- src/Idris/Desugar.idr | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 6cc08edd1..0f140c11e 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -807,7 +807,7 @@ mutual = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (PApp loc (PApp loc (PRef opFC op) l') r') - -- (x : ty) =@ f x ==>> (=@) exp (\x : ty => f x) + -- (x : ty) =@ f x ==>> (=@) ty (\x : ty => f x) desugarTree side ps (Infix loc opFC (op, Just (BindType pat lhs)) l r) = do l' <- desugarTree side ps l body <- desugarTree side ps r From 2859fcee08befef384ceeed82a5ee544982dfd22 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:40:12 +0900 Subject: [PATCH 47/57] Update docs/source/reference/operators.rst --- docs/source/reference/operators.rst | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index 27cedc4f0..d34635bcf 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -310,8 +310,8 @@ where the constructor for ``Pi`` takes terms on the left and lambdas on the righ We would like to use a custom operator to build values using ``VPi``, but its signature does not follow the pattern that ``typebind`` uses. Instead, we use -``autobind`` to tell the compiler that the type of the lambda is not given -by the first argument. For this we use ``:=`` instead of ``:``: +``autobind`` to tell the compiler that the type of the lambda must be inferred. +For this we use ``:=`` instead of ``:``: .. code-block:: idris From b48a2d11f0020d2fed6d61d1aab112234e883e55 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:46:07 +0900 Subject: [PATCH 48/57] Update src/Idris/Doc/Keywords.idr --- src/Idris/Doc/Keywords.idr | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Idris/Doc/Keywords.idr b/src/Idris/Doc/Keywords.idr index 0c5ecf2b5..6dd55cacf 100644 --- a/src/Idris/Doc/Keywords.idr +++ b/src/Idris/Doc/Keywords.idr @@ -498,7 +498,7 @@ typebindDoc = """ If you declare a new operator to be typebind you can use it the same way. - Start by defining `typebind infixr 1 =@`, and then you can use it + Start by defining `typebind infixr 0 =@`, and then you can use it like so: `(n : Nat) =@ f n` `typebind` only applies to `infixr` operators and cannot be used as From 1631326887adeb9c7bf08e437ded799525d167dd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:46:21 +0900 Subject: [PATCH 49/57] Update src/Idris/Error.idr --- src/Idris/Error.idr | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index a20be6fce..f26c6dc49 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -644,7 +644,7 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs candi spellingCandidates = case candidates of [] => [] [x] => ["Did you mean" <++> enclose "'" "'" (pretty0 x) <++> "?"] - xs => ["Did you mean one of:" <++> hcat (punctuate ", " + xs => ["Did you mean either of:" <++> hcat (punctuate ", " (map (enclose "'" "'" . pretty0) xs)) <++> "?"] From 0dfecf87df5cd6619cfd0e7f4261de61d695c7bd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:46:53 +0900 Subject: [PATCH 50/57] Update src/Idris/Doc/Keywords.idr --- src/Idris/Doc/Keywords.idr | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Idris/Doc/Keywords.idr b/src/Idris/Doc/Keywords.idr index 6dd55cacf..84c5113d2 100644 --- a/src/Idris/Doc/Keywords.idr +++ b/src/Idris/Doc/Keywords.idr @@ -505,7 +505,7 @@ typebindDoc = """ operator sections. `typebind` operators are desugared as a lambda: - `(x : ty) =@ fn x` -> `(ty =@ (\x : ty =@ fn x))` + `(x : ty) =@ fn x` -> `ty =@ (\x : ty =@ fn x)` """ rewriteeq : Doc IdrisDocAnn From c4eaacc493fc4a816212f139a476e741fcc580ba Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:49:27 +0900 Subject: [PATCH 51/57] Update src/Idris/Doc/Keywords.idr --- src/Idris/Doc/Keywords.idr | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Idris/Doc/Keywords.idr b/src/Idris/Doc/Keywords.idr index 84c5113d2..1ef73ff55 100644 --- a/src/Idris/Doc/Keywords.idr +++ b/src/Idris/Doc/Keywords.idr @@ -493,7 +493,7 @@ typebindDoc = """ use it on the right-hand-side. A typical example of a typebind operator is `(**)` the type constructor - for dependent pairs. It is used like this: `(x : Nat ** Vect x String)` + for dependent pairs. It is used like this: `(x : Nat) ** Vect x String` If you declare a new operator to be typebind you can use it the same way. From bc375909cfbfccaae6ae52d8b0e151aeed0f1ba3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:55:22 +0900 Subject: [PATCH 52/57] Update src/Idris/Desugar.idr --- src/Idris/Desugar.idr | 1 - 1 file changed, 1 deletion(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 0f140c11e..bbb41a2fd 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -42,7 +42,6 @@ import Data.Maybe import Data.List import Data.List.Views import Data.String -import Debug.Trace -- Convert high level Idris declarations (PDecl from Idris.Syntax) into -- TTImp, recording any high level syntax info on the way (e.g. infix From d51a5c91bbd8238e139ebdca22f1d43cc6f56aa7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 16:22:37 +0000 Subject: [PATCH 53/57] fix review comments --- src/Core/Binary.idr | 2 +- src/Core/TT.idr | 18 +++++++++--------- src/Idris/Desugar.idr | 17 ++++++++++------- src/Idris/Parser.idr | 6 +++--- src/Idris/Pretty.idr | 6 +++--- src/Idris/Syntax.idr | 6 +++--- src/Idris/Syntax/Traversals.idr | 4 ++-- tests/idris2/operators/operators007/expected | 2 +- 8 files changed, 32 insertions(+), 29 deletions(-) diff --git a/src/Core/Binary.idr b/src/Core/Binary.idr index 6f985cde4..813dd9a96 100644 --- a/src/Core/Binary.idr +++ b/src/Core/Binary.idr @@ -29,7 +29,7 @@ import public Libraries.Utils.Binary ||| version number if you're changing the version more than once in the same day. export ttcVersion : Int -ttcVersion = 2023_11_12_00 +ttcVersion = 2024_01_23_00 export checkTTCVersion : String -> Int -> Int -> Core () diff --git a/src/Core/TT.idr b/src/Core/TT.idr index 68bc6c6dd..cab8c5323 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -189,11 +189,11 @@ data OperatorLHSInfo : tm -> Type where -- Traditional operator wihtout binding, carries the lhs NoBinder : (lhs : tm) -> OperatorLHSInfo tm -- (nm : ty) =@ fn x - BindType : (name : WithBounds tm) -> (ty : tm) -> OperatorLHSInfo tm + BindType : (name : tm) -> (ty : tm) -> OperatorLHSInfo tm -- (nm := exp) =@ fn nm - BindExpr : (name : WithBounds tm) -> (expr : tm) -> OperatorLHSInfo tm + BindExpr : (name : tm) -> (expr : tm) -> OperatorLHSInfo tm -- (nm : ty := exp) =@ fn nm - BindExplicitType : (name : WithBounds tm) -> (type, expr : tm) -> OperatorLHSInfo tm + BindExplicitType : (name : tm) -> (type, expr : tm) -> OperatorLHSInfo tm export Show (OperatorLHSInfo tm) where @@ -207,9 +207,9 @@ Show (OperatorLHSInfo tm) where export Functor OperatorLHSInfo where map f (NoBinder lhs) = NoBinder $ f lhs - map f (BindType nm lhs) = BindType (map f nm) (f lhs) - map f (BindExpr nm lhs) = BindExpr (map f nm) (f lhs) - map f (BindExplicitType nm ty lhs) = BindExplicitType (map f nm) (f ty) (f lhs) + map f (BindType nm lhs) = BindType (f nm) (f lhs) + map f (BindExpr nm lhs) = BindExpr (f nm) (f lhs) + map f (BindExplicitType nm ty lhs) = BindExplicitType (f nm) (f ty) (f lhs) export (.getLhs) : OperatorLHSInfo tm -> tm @@ -221,9 +221,9 @@ export export (.getBoundPat) : OperatorLHSInfo tm -> Maybe tm (.getBoundPat) (NoBinder lhs) = Nothing -(.getBoundPat) (BindType name ty) = Just name.val -(.getBoundPat) (BindExpr name expr) = Just name.val -(.getBoundPat) (BindExplicitType name type expr) = Just name.val +(.getBoundPat) (BindType name ty) = Just name +(.getBoundPat) (BindExpr name expr) = Just name +(.getBoundPat) (BindExplicitType name type expr) = Just name export (.getBinder) : OperatorLHSInfo tm -> BindingModifier diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index bbb41a2fd..46d7088ff 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -171,10 +171,9 @@ checkConflictingFixities isPrefix exprFC opn checkConflictingBinding : Ref Ctxt Defs => Ref Syn SyntaxInfo => - Side -> FC -> Name -> (foundFixity : BacktickOrOperatorFixity) -> + FC -> Name -> (foundFixity : BacktickOrOperatorFixity) -> (usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core () -checkConflictingBinding LHS fc opName foundFixity use_site rhs = pure () -- don't check if we're on the LHS -checkConflictingBinding side fc opName foundFixity use_site rhs +checkConflictingBinding fc opName foundFixity use_site rhs = if isCompatible foundFixity use_site then pure () else throw $ OperatorBindingMismatch @@ -229,7 +228,11 @@ parameters (side : Side) PTerm -> Core (List (Tok ((OpStr, BacktickOrOperatorFixity), Maybe (OperatorLHSInfo PTerm)) PTerm)) toTokList (POp fc opFC l opn r) = do (precInfo, fixInfo) <- checkConflictingFixities False fc opn - checkConflictingBinding side opFC opn fixInfo l r + unless (side == LHS) -- do not check for conflicting fixity on the LHS + -- This is because we do not parse binders on the lhs + -- and so, if we check, we will find uses of regular + -- operator when binding is expected. + (checkConflictingBinding opFC opn fixInfo l r) rtoks <- toTokList r pure (Expr l.getLhs :: Op fc opFC ((opn, fixInfo), Just l) precInfo :: rtoks) toTokList (PPrefixOp fc opFC opn arg) @@ -811,20 +814,20 @@ mutual = do l' <- desugarTree side ps l body <- desugarTree side ps r pure $ PApp loc (PApp loc (PRef opFC op) l') - (PLam loc top Explicit pat.val l' body) + (PLam loc top Explicit pat l' body) -- (x := exp) =@ f x ==>> (=@) exp (\x : ? => f x) desugarTree side ps (Infix loc opFC (op, Just (BindExpr pat lhs)) l r) = do l' <- desugarTree side ps l body <- desugarTree side ps r pure $ PApp loc (PApp loc (PRef opFC op) l') - (PLam loc top Explicit pat.val (PInfer opFC) body) + (PLam loc top Explicit pat (PInfer opFC) body) -- (x : ty := exp) =@ f x ==>> (=@) exp (\x : ty => f x) desugarTree side ps (Infix loc opFC (op, Just (BindExplicitType pat ty expr)) l r) = do l' <- desugarTree side ps l body <- desugarTree side ps r pure $ PApp loc (PApp loc (PRef opFC op) l') - (PLam loc top Explicit pat.val ty body) + (PLam loc top Explicit pat ty body) desugarTree side ps (Infix loc opFC (op, Nothing) _ r) = throw $ InternalError "illegal fixity: Parsed as infix but no binding information" diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 1e1979111..e8910f52f 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -333,13 +333,13 @@ mutual ty <- typeExpr pdef fname indents decoratedSymbol fname ":=" exp <- expr pdef fname indents - pure (BindExplicitType boundName ty exp) + pure (BindExplicitType boundName.val ty exp) <|> do decoratedSymbol fname ":=" exp <- expr pdef fname indents - pure (BindExpr boundName exp) + pure (BindExpr boundName.val exp) <|> do decoratedSymbol fname ":" ty <- typeExpr pdef fname indents - pure (BindType boundName ty) + pure (BindType boundName.val ty) opBinder : OriginDesc -> IndentInfo -> Rule (OperatorLHSInfo PTerm) opBinder fname indents diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index 3280d50c9..34f0eb421 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -333,15 +333,15 @@ mutual prettyPrec d (PImplicit _) = "_" prettyPrec d (PInfer _) = annotate Hole $ "?" prettyPrec d (POp _ _ (BindType nm left) op right) = - group $ parens (prettyPrec d nm.val <++> ":" <++> pretty left) + group $ parens (prettyPrec d nm <++> ":" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (BindExpr nm left) op right) = - group $ parens (prettyPrec d nm.val <++> ":=" <++> pretty left) + group $ parens (prettyPrec d nm <++> ":=" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (BindExplicitType nm ty left) op right) = - group $ parens (prettyPrec d nm.val <++> ":" <++> pretty ty <++> ":=" <++> pretty left) + group $ parens (prettyPrec d nm <++> ":" <++> pretty ty <++> ":=" <++> pretty left) <++> prettyOp op <++> pretty right prettyPrec d (POp _ _ (NoBinder x) op y) = diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index 7226a7b28..56bd1b7c4 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -773,11 +773,11 @@ parameters {0 nm : Type} (toName : nm -> Name) showPTermPrec d (POp _ _ (NoBinder left) op right) = showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right showPTermPrec d (POp _ _ (BindType nm left) op right) - = "(" ++ showPTermPrec d nm.val ++ " : " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ showPTermPrec d nm ++ " : " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (POp _ _ (BindExpr nm left) op right) - = "(" ++ showPTermPrec d nm.val ++ " := " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ showPTermPrec d nm ++ " := " ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (POp _ _ (BindExplicitType nm ty left) op right) - = "(" ++ showPTermPrec d nm.val ++ " : " ++ showPTermPrec d ty ++ ":=" ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" + = "(" ++ showPTermPrec d nm ++ " : " ++ showPTermPrec d ty ++ ":=" ++ showPTermPrec d left ++ " " ++ showOpPrec d op ++ " " ++ showPTermPrec d right ++ ")" showPTermPrec d (PPrefixOp _ _ op x) = showOpPrec d op ++ showPTermPrec d x showPTermPrec d (PSectionL _ _ op x) = "(" ++ showOpPrec d op ++ " " ++ showPTermPrec d x ++ ")" showPTermPrec d (PSectionR _ _ x op) = "(" ++ showPTermPrec d x ++ " " ++ showOpPrec d op ++ ")" diff --git a/src/Idris/Syntax/Traversals.idr b/src/Idris/Syntax/Traversals.idr index b2aefb1e1..f44230973 100644 --- a/src/Idris/Syntax/Traversals.idr +++ b/src/Idris/Syntax/Traversals.idr @@ -107,7 +107,7 @@ mapPTermM f = goPTerm where >>= f goPTerm (POp fc opFC (BindType nm left) op right) = POp fc opFC - <$> (BindType nm <$> goPTerm left) + <$> (BindType <$> goPTerm nm <*> goPTerm left) <*> pure op <*> goPTerm right >>= f @@ -119,7 +119,7 @@ mapPTermM f = goPTerm where >>= f goPTerm (POp fc opFC (BindExplicitType nm ty left) op right) = POp fc opFC - <$> (BindExplicitType nm <$> goPTerm ty <*> goPTerm left) + <$> (BindExplicitType <$> goPTerm nm <*> goPTerm ty <*> goPTerm left) <*> pure op <*> goPTerm right >>= f diff --git a/tests/idris2/operators/operators007/expected b/tests/idris2/operators/operators007/expected index d006e178e..8644014d4 100644 --- a/tests/idris2/operators/operators007/expected +++ b/tests/idris2/operators/operators007/expected @@ -31,4 +31,4 @@ Possible solutions: - Write the expression using regular syntax: 'm >>= Just (x + y)'. - Change the fixity defined at Prelude.Ops:20:1--20:37 to 'export autobind infixl 1 >>='. - Hide or remove the fixity at Prelude.Ops:20:1--20:37 and import a module that exports a compatible fixity. - - Did you mean one of: '>=', '>>' ? + - Did you mean either of: '>=', '>>' ? From 6485a0023fba8cbe319d8c98aead468a5210e626 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 16:38:50 +0000 Subject: [PATCH 54/57] update confusing error message --- src/Libraries/Utils/Shunting.idr | 14 +++++----- tests/idris2/operators/operators004/Test.idr | 11 ++++++-- tests/idris2/operators/operators004/expected | 29 ++++++++++++++------ 3 files changed, 37 insertions(+), 17 deletions(-) diff --git a/src/Libraries/Utils/Shunting.idr b/src/Libraries/Utils/Shunting.idr index f6cc37904..9413daecf 100644 --- a/src/Libraries/Utils/Shunting.idr +++ b/src/Libraries/Utils/Shunting.idr @@ -90,20 +90,20 @@ isLAssoc : OpPrec -> Bool isLAssoc (AssocL _) = True isLAssoc _ = False --- Return whether the first operator should be applied before the second, --- assuming +-- Return whether the first operator should be applied before the second. +-- Interpolation to show the operator naked, show to print the operator with its location higher : Interpolation op => (showLoc : Show op) => FC -> op -> OpPrec -> op -> OpPrec -> Core Bool higher loc opx op opy (Prefix p) = pure False higher loc opx (NonAssoc x) opy oy = if x == getPrec oy - then throw (GenericMsgSol loc ("Operator \{opx} is non-associative") - ["Add brackets around \{opy}" - , "Change the fixity of \{show opy} to `infixl` or `infixr`"]) + then throw (GenericMsgSol loc ( "Operator \{opx} is non-associative") + [ "Add brackets around every use of \{opx}" + , "Change the fixity of \{show opx} to `infixl` or `infixr`"]) else pure (x > getPrec oy) higher loc opx ox opy (NonAssoc y) = if getPrec ox == y - then throw (GenericMsgSol loc ("Operator \{opy} is non-associative") - ["Add brackets around \{opy}" + then throw (GenericMsgSol loc ( "Operator \{opy} is non-associative") + [ "Add brackets around every use of \{opy}" , "Change the fixity of \{show opy} to `infixl` or `infixr`"]) else pure (getPrec ox > y) higher loc opl l opr r diff --git a/tests/idris2/operators/operators004/Test.idr b/tests/idris2/operators/operators004/Test.idr index 485981d8a..25093fc0e 100644 --- a/tests/idris2/operators/operators004/Test.idr +++ b/tests/idris2/operators/operators004/Test.idr @@ -1,9 +1,16 @@ infix 5 -:- +infix 5 :-: + (-:-) : a -> List a -> List a (-:-) = (::) -test : List Nat -test = 4 -:- 3 -:- [] +(:-:) : a -> List a -> List a +(:-:) = (::) +test : List Nat +test = 4 -:- 3 :-: [] + +test2 : List Nat +test2 = 4 :-: 3 -:- [] diff --git a/tests/idris2/operators/operators004/expected b/tests/idris2/operators/operators004/expected index 70d253093..f976037ea 100644 --- a/tests/idris2/operators/operators004/expected +++ b/tests/idris2/operators/operators004/expected @@ -1,14 +1,27 @@ 1/1: Building Test (Test.idr) Error: Operator -:- is non-associative -Test:8:8--8:22 - 4 | (-:-) : a -> List a -> List a - 5 | (-:-) = (::) - 6 | - 7 | test : List Nat - 8 | test = 4 -:- 3 -:- [] - ^^^^^^^^^^^^^^ +Test:13:8--13:22 + 09 | (:-:) : a -> List a -> List a + 10 | (:-:) = (::) + 11 | + 12 | test : List Nat + 13 | test = 4 -:- 3 :-: [] + ^^^^^^^^^^^^^^ Possible solutions: - - Add brackets around -:- + - Add brackets around every use of -:- - Change the fixity of -:- at Test:2:1--2:12 to `infixl` or `infixr` +Error: Operator :-: is non-associative + +Test:16:9--16:23 + 12 | test : List Nat + 13 | test = 4 -:- 3 :-: [] + 14 | + 15 | test2 : List Nat + 16 | test2 = 4 :-: 3 -:- [] + ^^^^^^^^^^^^^^ + +Possible solutions: + - Add brackets around every use of :-: + - Change the fixity of :-: at Test:4:1--4:12 to `infixl` or `infixr` From 58c361c57d1838c5d47fa3780212e5edcb49c85e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 23 Jan 2024 19:04:43 +0000 Subject: [PATCH 55/57] correctly carry the position of desugared equality --- src/Idris/Desugar.idr | 3 ++- tests/ideMode/ideMode005/expected | 6 ------ 2 files changed, 2 insertions(+), 7 deletions(-) diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 46d7088ff..a61fd910e 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -799,7 +799,8 @@ mutual Side -> List Name -> Tree (OpStr, Maybe $ OperatorLHSInfo PTerm) PTerm -> Core PTerm desugarTree side ps (Infix loc eqFC (UN $ Basic "=", _) l r) -- special case since '=' is special syntax - = pure $ PEq loc !(desugarTree side ps l) !(desugarTree side ps r) + = pure $ PEq eqFC !(desugarTree side ps l) !(desugarTree side ps r) + desugarTree side ps (Infix loc _ (UN $ Basic "$", _) l r) -- special case since '$' is special syntax = do l' <- desugarTree side ps l r' <- desugarTree side ps r diff --git a/tests/ideMode/ideMode005/expected b/tests/ideMode/ideMode005/expected index 299fc037f..e69de29bb 100644 --- a/tests/ideMode/ideMode005/expected +++ b/tests/ideMode/ideMode005/expected @@ -1,6 +0,0 @@ -24a25 -> 0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 40) (:end 2 49)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -49a50 -> 0000dd(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 7) (:end 17 47)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -40a41 -> 0000d3(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 8 27) (:end 8 44)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) From 8ea1a092e0755eda3abd936a5bd653ae0404e23c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 24 Feb 2024 12:20:30 +0000 Subject: [PATCH 56/57] add hiding test --- tests/idris2/operators/operators008/Test.idr | 13 +++++++++++++ tests/idris2/operators/operators008/Test2.idr | 4 ++++ tests/idris2/operators/operators008/expected | 2 ++ tests/idris2/operators/operators008/run | 3 +++ 4 files changed, 22 insertions(+) create mode 100644 tests/idris2/operators/operators008/Test.idr create mode 100644 tests/idris2/operators/operators008/Test2.idr create mode 100644 tests/idris2/operators/operators008/expected create mode 100755 tests/idris2/operators/operators008/run diff --git a/tests/idris2/operators/operators008/Test.idr b/tests/idris2/operators/operators008/Test.idr new file mode 100644 index 000000000..328dfb3c2 --- /dev/null +++ b/tests/idris2/operators/operators008/Test.idr @@ -0,0 +1,13 @@ +module Test + +import Test2 + +%hide Test2.infixr.(@>) + +infixl 9 @> + +(@>) : a -> b -> (a, b) +(@>) = MkPair + +test : 3 @> 2 @> 1 === (3 @> 2) @> 1 +test = Refl diff --git a/tests/idris2/operators/operators008/Test2.idr b/tests/idris2/operators/operators008/Test2.idr new file mode 100644 index 000000000..d85ce2bcf --- /dev/null +++ b/tests/idris2/operators/operators008/Test2.idr @@ -0,0 +1,4 @@ +module Test2 + +export autobind infixr 0 @> + diff --git a/tests/idris2/operators/operators008/expected b/tests/idris2/operators/operators008/expected new file mode 100644 index 000000000..84943dbde --- /dev/null +++ b/tests/idris2/operators/operators008/expected @@ -0,0 +1,2 @@ +1/2: Building Test2 (Test2.idr) +2/2: Building Test (Test.idr) diff --git a/tests/idris2/operators/operators008/run b/tests/idris2/operators/operators008/run new file mode 100755 index 000000000..d35387bb3 --- /dev/null +++ b/tests/idris2/operators/operators008/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Test.idr From 7c70762f2529a617760f6827b08b9ee6431a61c3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Sat, 24 Feb 2024 12:21:24 +0000 Subject: [PATCH 57/57] fix typos --- docs/source/reference/operators.rst | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/docs/source/reference/operators.rst b/docs/source/reference/operators.rst index d34635bcf..fd9484793 100644 --- a/docs/source/reference/operators.rst +++ b/docs/source/reference/operators.rst @@ -151,7 +151,7 @@ looks like main = do print (a &&& b) -- won't work print ((&&&) a b) -- ok -In what follows we have two examples of programs that benefit from +In what follows, we have two examples of programs that benefit from declaring a fixity ``private`` rather than ``export``. Private record fixity pattern @@ -181,7 +181,7 @@ as a binary infix operator impossible, since it now has 3 arguments. natRel : Nat -> Nat -> Type natRel x y = (<@>) lteRel x y -What we really want to write is ``natRel x y = <@> x y`` but +What we really want to write is ``natRel x y = (<@>) x y`` but ``(<@>)`` now has type ``SomeRelation a -> a -> a -> Type``. The solution is to define a private field with a private fixity @@ -240,7 +240,7 @@ Typebind Operators ================== In dependently-typed programming, we have the ability define constructors which -first argument is a type and the second is a type indexed over the first argument. +first argument is a type and the second is a type indexed over the first argument. A typical example of this is the dependent linear arrow: .. code-block:: idris @@ -290,9 +290,9 @@ Autobind Operators ================== Typebind operators allow to bind a *type* on the left side of an operator, so that is can -be used as the index of the second argument. But sometimes, there is no dependency -between the first and second argument, yet we still want to use binding sytnax. For those -case, we use ``autobind``. +be used as the index of the second argument. But sometimes, there is no dependency +between the first and second argument, yet we still want to use binding syntax. For those +cases, we use ``autobind``. An example of this is a DSL for a dependently-typed programming language where the constructor for ``Pi`` takes terms on the left and lambdas on the right: @@ -310,7 +310,7 @@ where the constructor for ``Pi`` takes terms on the left and lambdas on the righ We would like to use a custom operator to build values using ``VPi``, but its signature does not follow the pattern that ``typebind`` uses. Instead, we use -``autobind`` to tell the compiler that the type of the lambda must be inferred. +``autobind`` to tell the compiler that the type of the lambda must be inferred. For this we use ``:=`` instead of ``:``: .. code-block:: idris @@ -326,7 +326,7 @@ For this we use ``:=`` instead of ``:``: (sndTy := (_ := fstTy) =>> VStar) =>> (val1 := fstTy) =>> (val2 := sndTy `vapp` val1) =>> - VSgima fstTy sndTy + VSigma fstTy sndTy This new syntax is much closer to what the code is meant to look like for users accustomed to dependently-typed programming languages.