From 48e0a895c814cd942b64f931a35863991ad4767f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Tue, 17 Dec 2024 12:25:32 -0500 Subject: [PATCH 1/4] update parser --- src/Idris/Desugar.idr | 85 ++- src/Idris/Desugar/Mutual.idr | 9 +- src/Idris/Parser.idr | 684 +++++++++--------- src/Idris/Pretty.idr | 20 + src/Idris/Pretty/Annotations.idr | 4 + src/Idris/Resugar.idr | 20 +- src/Idris/Syntax.idr | 108 ++- src/Idris/Syntax/Traversals.idr | 121 ++-- src/Idris/Syntax/Views.idr | 3 +- .../Text/PrettyPrint/Prettyprinter/Doc.idr | 4 + src/Parser/Rule/Source.idr | 8 + src/TTImp/Elab/Quote.idr | 2 +- src/TTImp/ProcessDecls.idr | 2 +- src/TTImp/TTImp.idr | 4 +- src/TTImp/Utils.idr | 2 +- tests/idris2/error/perror034/Error.idr | 5 + tests/idris2/error/perror034/expected | 26 + tests/idris2/error/perror034/run | 3 + 18 files changed, 685 insertions(+), 425 deletions(-) create mode 100644 tests/idris2/error/perror034/Error.idr create mode 100644 tests/idris2/error/perror034/expected create mode 100644 tests/idris2/error/perror034/run diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 86ebd41bf0..8abcc99e2c 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -306,6 +306,38 @@ mutual case x == pur of -- implicitly add namespace to unqualified occurrences of `pure` in a qualified do-block False => pure $ IVar fc x True => pure $ IVar fc (maybe pur (`NS` pur) ns) + + -- Desugaring forall n1, n2, n3 . s into + -- {0 n1 : ?} -> {0 n2 : ?} -> {0 n3 : ?} -> s + desugarB side ps (Forall (MkFCVal fc (names, scope))) + = desugarForallNames ps (forget names) + where + desugarForallNames : (ctx : List Name) -> + (names : List (WithFC Name)) -> Core RawImp + desugarForallNames ctx [] = desugarB side ctx scope + desugarForallNames ctx (x :: xs) + = IPi x.fc erased Implicit (Just x.val) + <$> desugarB side ps (PImplicit x.fc) + <*> desugarForallNames (x.val :: ctx) xs + + -- Desugaring (n1, n2, n3 : t) -> s into + -- (n1 : t) -> (n2 : t) -> (n3 : t) -> s + desugarB side ps + (NewPi (MkFCVal fc + (MkPBinderScope (MkPBinder info (MkBasicMultiBinder rig names type)) scope))) + = desugarMultiBinder ps (forget names) + where + desugarMultiBinder : (ctx : List Name) -> List (WithFC Name) -> Core RawImp + desugarMultiBinder ctx [] + = desugarB side ctx scope + desugarMultiBinder ctx (name :: xs) + = let extendedCtx = name.val :: ps + in IPi fc rig + <$> mapDesugarPiInfo extendedCtx info + <*> (pure (Just name.val)) + <*> desugarB side ps type + <*> desugarMultiBinder extendedCtx xs + desugarB side ps (PPi fc rig p mn argTy retTy) = let ps' = maybe ps (:: ps) mn in pure $ IPi fc rig !(traverse (desugar side ps') p) @@ -1105,29 +1137,32 @@ mutual desugarDecl ps (MkFCVal fc $ PParameters params pds) = do params' <- getArgs params - pds' <- traverse (desugarDecl (ps ++ map fst params')) pds + let paramList = forget params' + pds' <- traverse (desugarDecl (ps ++ map fst paramList)) pds -- Look for implicitly bindable names in the parameters pnames <- ifThenElse (not !isUnboundImplicits) (pure []) $ map concat - $ for (map (Builtin.snd . Builtin.snd . Builtin.snd) params') - $ findUniqueBindableNames fc True (ps ++ map Builtin.fst params') [] + $ for (map (Builtin.snd . Builtin.snd . Builtin.snd) paramList) + $ findUniqueBindableNames fc True (ps ++ map Builtin.fst paramList) [] let paramsb = map (\(n, rig, info, tm) => (n, rig, info, doBind pnames tm)) params' pure [IParameters fc paramsb (concat pds')] where - getArgs : Either (List (Name, PTerm)) - (List (Name, RigCount, PiInfo (PTerm' Name), PTerm' Name)) -> - Core (List (ImpParameter' Name)) + getArgs : Either (List1 PlainBinder) + (List1 PBinder) -> + Core (List1 (ImpParameter' Name)) getArgs (Left params) - = traverse (\(n, ty) => do + = traverseList1 (\(MkWithName n ty) => do ty' <- desugar AnyExpr ps ty - pure (n, top, Explicit, ty')) params + pure (n.val, top, Explicit, ty')) params getArgs (Right params) - = traverse (\(n, rig, info, ntm) => do + = join <$> traverseList1 (\(MkPBinder info (MkBasicMultiBinder rig n ntm)) => do tm' <- desugar AnyExpr ps ntm i' <- traverse (desugar AnyExpr ps) info - pure (n, rig, i', tm')) params + let allbinders = map (\nn => (nn.val, rig, i', tm')) n + pure allbinders) params + desugarDecl ps (MkFCVal fc $ PUsing uimpls uds) = do syn <- get Syn let oldu = usingImpl syn @@ -1140,16 +1175,17 @@ mutual pure (concat uds') desugarDecl ps (MkFCVal fc $ PInterface vis cons_in tn doc params det conname body) = do addDocString tn doc - let paramNames = map fst params + let paramNames = concatMap (map val . forget . names) params let cons = concatMap expandConstraint cons_in cons' <- traverse (\ ntm => do tm' <- desugar AnyExpr (ps ++ paramNames) (snd ntm) pure (fst ntm, tm')) cons - params' <- traverse (\ (nm, (rig, tm)) => + params' <- concat <$> traverse (\ (MkBasicMultiBinder rig nm tm) => do tm' <- desugar AnyExpr ps tm - pure (nm, (rig, tm'))) + pure $ map (\n => (n, (rig, tm'))) (forget nm)) params + let _ = the (List (WithFC Name, RigCount, RawImp)) params' -- Look for bindable names in all the constraints and parameters let mnames = map dropNS (definedIn (map val body)) bnames <- ifThenElse (not !isUnboundImplicits) (pure []) @@ -1159,7 +1195,7 @@ mutual let paramsb = map (\ (nm, (rig, tm)) => let tm' = doBind bnames tm in - (nm, (rig, tm'))) + (nm.val, (rig, tm'))) params' let consb = map (\ (nm, tm) => (nm, doBind bnames tm)) cons' @@ -1226,24 +1262,29 @@ mutual desugarDecl ps (MkFCVal fc $ PRecord doc vis mbtot (MkPRecordLater tn params)) = desugarDecl ps (MkFCVal fc $ PData doc vis mbtot (MkPLater fc tn (mkRecType params))) where - mkRecType : List (Name, RigCount, PiInfo PTerm, PTerm) -> PTerm + mkRecType : List PBinder -> PTerm mkRecType [] = PType fc - mkRecType ((n, c, p, t) :: ts) = PPi fc c p (Just n) t (mkRecType ts) + mkRecType (MkPBinder p (MkBasicMultiBinder c (n ::: []) t) :: ts) + = PPi fc c p (Just n.val) t (mkRecType ts) + mkRecType (MkPBinder p (MkBasicMultiBinder c (n ::: x :: xs) t) :: ts) + = PPi fc c p (Just n.val) t (mkRecType (MkPBinder p (MkBasicMultiBinder c (x ::: xs) t) :: ts)) desugarDecl ps (MkFCVal fc $ PRecord doc vis mbtot (MkPRecord tn params opts conname_in fields)) = do addDocString tn doc - params' <- traverse (\ (n,c,p,tm) => + params' <- concat <$> traverse (\ (MkPBinder info (MkBasicMultiBinder rig names tm)) => do tm' <- desugar AnyExpr ps tm - p' <- mapDesugarPiInfo ps p - pure (n, c, p', tm')) + p' <- mapDesugarPiInfo ps info + let allBinders = map (\nn => (nn.val, rig, p', tm')) (forget names) + pure allBinders) params let _ = the (List (Name, RigCount, PiInfo RawImp, RawImp)) params' let fnames = concat $ map getfname fields + let paramNames = concatMap (map val . forget . names . bind) params let _ = the (List Name) fnames -- Look for bindable names in the parameters let bnames = if !isUnboundImplicits then concatMap (findBindableNames True - (ps ++ fnames ++ map fst params) []) + (ps ++ fnames ++ paramNames) []) (map (\(_,_,_,d) => d) params') else [] let _ = the (List (String, String)) bnames @@ -1251,8 +1292,8 @@ mutual let paramsb = map (\ (n, c, p, tm) => (n, c, p, doBind bnames tm)) params' let _ = the (List (Name, RigCount, PiInfo RawImp, RawImp)) paramsb let recName = nameRoot tn - fields' <- traverse (desugarField (ps ++ fnames ++ - map fst params) (mkNamespace recName)) + fields' <- traverse (desugarField (ps ++ fnames ++ paramNames + ) (mkNamespace recName)) fields let _ = the (List $ List IField) fields' let conname = maybe (mkConName tn) snd conname_in diff --git a/src/Idris/Desugar/Mutual.idr b/src/Idris/Desugar/Mutual.idr index 693d930044..bfca852448 100644 --- a/src/Idris/Desugar/Mutual.idr +++ b/src/Idris/Desugar/Mutual.idr @@ -1,6 +1,7 @@ module Idris.Desugar.Mutual import Idris.Syntax +import Data.List1 %default total @@ -26,9 +27,13 @@ getDecl AsType d@(MkFCVal _ $ PInterface _ _ _ _ _ _ _ _) = Just d getDecl AsType (MkFCVal fc $ PRecord doc vis mbtot (MkPRecord n ps _ _ _)) = Just (MkFCVal fc $ PData doc vis mbtot (MkPLater fc n (mkRecType ps))) where - mkRecType : List (Name, RigCount, PiInfo PTerm, PTerm) -> PTerm + mkRecType : List PBinder -> PTerm mkRecType [] = PType fc - mkRecType ((n, c, p, t) :: ts) = PPi fc c p (Just n) t (mkRecType ts) + mkRecType (MkPBinder p (MkBasicMultiBinder c (n ::: []) t) :: ts) + = PPi fc c p (Just n.val) t (mkRecType ts) + mkRecType (MkPBinder p (MkBasicMultiBinder c (n ::: x :: xs) t) :: ts) + = PPi fc c p (Just n.val) t + (assert_total $ mkRecType (MkPBinder p (MkBasicMultiBinder c (x ::: xs) t) :: ts)) getDecl AsType d@(MkFCVal _ $ PFixity _ ) = Just d getDecl AsType d@(MkFCVal _ $ PDirective _) = Just d getDecl AsType d = Nothing diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index fcd008a257..13b3a65218 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -75,6 +75,12 @@ parens fname p <*> p <* decoratedSymbol fname ")" +curly : {b : _} -> OriginDesc -> BRule b a -> Rule a +curly fname p + = pure id <* decoratedSymbol fname "{" + <*> p + <* decoratedSymbol fname "}" + decoratedDataTypeName : OriginDesc -> Rule Name decoratedDataTypeName fname = decorate fname Typ dataTypeName @@ -84,6 +90,9 @@ decoratedDataConstructorName fname = decorate fname Data dataConstructorName decoratedSimpleBinderName : OriginDesc -> Rule String decoratedSimpleBinderName fname = decorate fname Bound unqualifiedName +decoratedSimpleBinderUName : OriginDesc -> Rule Name +decoratedSimpleBinderUName fname = UN . Basic <$> decorate fname Bound unqualifiedName + decoratedSimpleNamedArg : OriginDesc -> Rule String decoratedSimpleNamedArg fname = decorate fname Bound unqualifiedName @@ -352,14 +361,14 @@ mutual autobindOp : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm autobindOp q fname indents - = do b <- fcBounds $ do - binder <- fcBounds $ parens fname (opBinder fname indents) - continue indents - op <- fcBounds iOperator - commit - e <- expr q fname indents - pure (binder, op, e) - pure (POp b.fc (fst b.val) (fst (snd b.val)) (snd (snd b.val))) + = do b <- fcBounds $ do + binder <- fcBounds $ parens fname (opBinder fname indents) + continue indents + op <- fcBounds iOperator + commit + e <- expr q fname indents + pure (binder, op, e) + pure (POp b.fc (fst b.val) (fst (snd b.val)) (snd (snd b.val))) opExprBase : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm opExprBase q fname indents @@ -659,13 +668,6 @@ mutual Nothing => pure top _ => fail "Invalid multiplicity (must be 0 or 1)" - pibindAll : OriginDesc -> PiInfo PTerm -> - List (RigCount, WithBounds (Maybe Name), PTerm) -> - PTerm -> PTerm - pibindAll fname p [] scope = scope - pibindAll fname p ((rig, n, ty) :: rest) scope - = PPi (boundToFC fname n) rig p (n.val) ty (pibindAll fname p rest scope) - bindList : OriginDesc -> IndentInfo -> Rule (List (RigCount, WithBounds PTerm, PTerm)) bindList fname indents @@ -677,104 +679,110 @@ mutual (decoratedSymbol fname ":" *> opExpr pdef fname indents) pure (rig, pat, ty)) + ||| A list of names bound to the same type + ||| BNF: + ||| pibindListName := qty name (, name)* ':' typeExpr pibindListName : OriginDesc -> IndentInfo -> - Rule (List (RigCount, WithBounds Name, PTerm)) + Rule BasicMultiBinder pibindListName fname indents = do rig <- multiplicity fname ns <- sepBy1 (decoratedSymbol fname ",") - (bounds $ UN <$> binderName) - let ns = forget ns - decorateBoundedNames fname Bound ns + (fcBounds $ UN <$> binderName) decoratedSymbol fname ":" ty <- typeExpr pdef fname indents atEnd indents - pure (map (\n => (rig, n, ty)) ns) + pure (MkBasicMultiBinder rig ns ty) where -- _ gets treated specially here, it means "I don't care about the name" binderName : Rule UserName - binderName = Basic <$> unqualifiedName + binderName = Basic <$> decoratedSimpleBinderName fname <|> symbol "_" $> Underscore - PiBindList : Type - PiBindList = List (RigCount, WithBounds (Maybe Name), PTerm) - - pibindList : OriginDesc -> IndentInfo -> - Rule PiBindList - pibindList fname indents - = do params <- pibindListName fname indents - pure $ map (\(rig, n, ty) => (rig, map Just n, ty)) params - + ||| The arrow used after an explicit binder + ||| BNF: + ||| bindSymbol := '->' | '=>' bindSymbol : OriginDesc -> Rule (PiInfo PTerm) bindSymbol fname = (decoratedSymbol fname "->" $> Explicit) <|> (decoratedSymbol fname "=>" $> AutoImplicit) + ||| An explicit pi-type + ||| BNF: + ||| explicitPi := '(' pibindListName ')' bindSymbol typeExpr explicitPi : OriginDesc -> IndentInfo -> Rule PTerm explicitPi fname indents - = do b <- bounds $ parens fname $ pibindList fname indents + = NewPi <$> fcBounds (do + b <- bounds $ parens fname $ pibindListName 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) + pure (MkPBinderScope (MkPBinder exp b.val) scope)) + ||| An auto-implicit pi-type + ||| BNF: + ||| autoImplicitPi := '{' 'auto' pibindListName '}' '->' typeExpr autoImplicitPi : OriginDesc -> IndentInfo -> Rule PTerm autoImplicitPi fname indents - = do b <- bounds $ do - decoratedSymbol fname "{" + = NewPi <$> fcBounds (do + b <- bounds $ curly fname $ do decoratedKeyword fname "auto" commit - binders <- pibindList fname indents - decoratedSymbol fname "}" + binders <- pibindListName fname indents pure binders mustWorkBecause b.bounds "Cannot return an auto implicit argument" $ decoratedSymbol fname "->" scope <- mustWork $ typeExpr pdef fname indents - pure (pibindAll fname AutoImplicit b.val scope) + pure (MkPBinderScope (MkPBinder AutoImplicit b.val) scope) + ) + ||| An default implicit pi-type + ||| BNF: + ||| defaultImplicitPi := '{' 'default' simpleExpr pibindListName '}' '->' typeExpr defaultImplicitPi : OriginDesc -> IndentInfo -> Rule PTerm defaultImplicitPi fname indents - = do b <- bounds $ do - decoratedSymbol fname "{" + = NewPi <$> fcBounds (do + b <- bounds $ curly fname $ do decoratedKeyword fname "default" commit t <- simpleExpr fname indents - binders <- pibindList fname indents - decoratedSymbol fname "}" - pure (t, binders) + binders <- pibindListName fname indents + pure (MkPBinder (DefImplicit t) binders) mustWorkBecause b.bounds "Cannot return a default implicit argument" $ decoratedSymbol fname "->" scope <- mustWork $ typeExpr pdef fname indents - pure $ let (t, binders) = b.val in - pibindAll fname (DefImplicit t) binders scope + pure (MkPBinderScope b.val scope) + ) + ||| Forall definition that automatically binds the names + ||| BNF: + ||| forall_ := 'forall' name (, name)* '.' typeExpr forall_ : OriginDesc -> IndentInfo -> Rule PTerm forall_ fname indents - = do b <- bounds $ do + = Forall <$> fcBounds (do + b <- bounds $ do decoratedKeyword fname "forall" commit ns <- sepBy1 (decoratedSymbol fname ",") - (bounds (decoratedSimpleBinderName fname)) - pure $ map (\n => ( erased {a=RigCount} - , map (Just . UN . Basic) n - , PImplicit (boundToFC fname n)) - ) (forget ns) + (fcBounds (decoratedSimpleBinderUName fname)) + pure ns b' <- bounds peek mustWorkBecause b'.bounds "Expected ',' or '.'" $ decoratedSymbol fname "." scope <- mustWork $ typeExpr pdef fname indents - pure (pibindAll fname Implicit b.val scope) + pure (b.val, scope)) + ||| implicit pi-type + ||| BNF: + ||| implicitPi := '{' pibindListName '}' '->' typeExpr implicitPi : OriginDesc -> IndentInfo -> Rule PTerm implicitPi fname indents - = do b <- bounds $ do - decoratedSymbol fname "{" - binders <- pibindList fname indents - decoratedSymbol fname "}" - pure binders + = NewPi <$> fcBounds (do + b <- bounds $ curly fname $ pibindListName fname indents mustWorkBecause b.bounds "Cannot return an implicit argument" $ decoratedSymbol fname "->" scope <- mustWork $ typeExpr pdef fname indents - pure (pibindAll fname Implicit b.val scope) + pure (MkPBinderScope (MkPBinder Implicit b.val) scope) + ) lam : OriginDesc -> IndentInfo -> Rule PTerm lam fname indents @@ -904,7 +912,7 @@ mutual )) <|> bounds (body False)) - pure (PUpdate (boundToFC fname b) b.val) + pure (PUpdate (boundToFC fname b) (forget b.val)) where oldSyntaxWarning : String oldSyntaxWarning = unlines @@ -913,13 +921,10 @@ mutual , #" and "{ f $= v } p" instead of "record { f $= v } p""# ] - body : Bool -> Rule (List PFieldUpdate) - body kw = do - decoratedSymbol fname "{" + body : Bool -> Rule (List1 PFieldUpdate) + body kw = curly fname $ do commit - fs <- sepBy1 (decoratedSymbol fname ",") (field kw fname indents) - decoratedSymbol fname "}" - pure $ forget fs + sepBy1 (decoratedSymbol fname ",") (field kw fname indents) field : Bool -> OriginDesc -> IndentInfo -> Rule PFieldUpdate field kw fname indents @@ -1195,6 +1200,21 @@ exportVisibility fname = (specified <$> visOption fname) <|> pure defaulted +plainBinder : (fname : OriginDesc) => (indents : IndentInfo) => Rule PlainBinder +plainBinder = do name <- fcBounds (decoratedSimpleBinderUName fname) + decoratedSymbol fname ":" + ty <- typeExpr pdef fname indents + pure $ MkWithName name ty + +basicMultiBinder : (fname : OriginDesc) => (indents : IndentInfo) => Rule BasicMultiBinder +basicMultiBinder + = do rig <- multiplicity fname + names <- sepBy1 (decoratedSymbol fname ",") + $ fcBounds (decoratedSimpleBinderUName fname) + decoratedSymbol fname ":" + ty <- typeExpr pdef fname indents + pure $ MkBasicMultiBinder rig names ty + tyDecls : Rule Name -> String -> OriginDesc -> IndentInfo -> Rule PTypeDecl tyDecls declName predoc fname indents = do bs <- bounds $ do @@ -1398,8 +1418,8 @@ dataVisOpt fname <|> do { tot <- totalityOpt fname ; vis <- visibility fname ; pure (vis, Just tot) } <|> pure (defaulted, Nothing) -dataDecl : OriginDesc -> IndentInfo -> Rule PDeclNoFC -dataDecl fname indents +dataDecl : (fname : OriginDesc) => (indents : IndentInfo) => Rule PDeclNoFC +dataDecl = do doc <- optDocumentation fname (vis,mbTot) <- dataVisOpt fname dat <- dataDeclBody fname indents @@ -1432,8 +1452,8 @@ logLevel fname pure (Just (mkLogLevel' topic (fromInteger lvl))) <|> fail "expected a log level" -directive : OriginDesc -> IndentInfo -> Rule Directive -directive fname indents +directive : (fname : OriginDesc) => (indents : IndentInfo) => Rule Directive +directive = do decoratedPragma fname "hide" n <- (fixityNS <|> (HideName <$> name)) atEnd indents @@ -1561,71 +1581,77 @@ namespaceHead fname = do decoratedKeyword fname "namespace" decorate fname Namespace $ mustWork namespaceId -namespaceDecl : OriginDesc -> IndentInfo -> Rule PDeclNoFC -namespaceDecl fname indents - = do doc <- optDocumentation fname - col <- column - ns <- namespaceHead fname - ds <- blockAfter col (topDecl fname) - pure (PNamespace ns (collectDefs ds)) - -transformDecl : OriginDesc -> IndentInfo -> Rule PDeclNoFC -transformDecl fname indents - = do decoratedPragma fname "transform" - n <- simpleStr - lhs <- expr plhs fname indents - decoratedSymbol fname "=" - rhs <- expr pnowith fname indents - pure (PTransform n lhs rhs) - -runElabDecl : OriginDesc -> IndentInfo -> Rule PDeclNoFC -runElabDecl fname indents - = do decoratedPragma fname "runElab" - tm <- expr pnowith fname indents - pure (PRunElabDecl tm) - -failDecls : OriginDesc -> IndentInfo -> Rule PDeclNoFC -failDecls fname indents - = do col <- column - decoratedKeyword fname "failing" - commit - msg <- optional (decorate fname Data simpleStr) - ds <- nonEmptyBlockAfter col (topDecl fname) - pure $ PFail msg (collectDefs (forget ds)) - -mutualDecls : OriginDesc -> IndentInfo -> Rule PDeclNoFC -mutualDecls fname indents - = do col <- column - decoratedKeyword fname "mutual" - commit - ds <- nonEmptyBlockAfter col (topDecl fname) - pure (PMutual (forget ds)) - -usingDecls : OriginDesc -> IndentInfo -> Rule PDeclNoFC -usingDecls fname indents - = do - col <- column - decoratedKeyword fname "using" - commit - decoratedSymbol fname "(" - us <- sepBy (decoratedSymbol fname ",") - (do n <- optional $ do - x <- unqualifiedName - decoratedSymbol fname ":" - pure (UN $ Basic x) - ty <- typeExpr pdef fname indents - pure (n, ty)) - decoratedSymbol fname ")" - ds <- nonEmptyBlockAfter col (topDecl fname) - pure (PUsing us (collectDefs (forget ds))) - -builtinDecl : OriginDesc -> IndentInfo -> Rule PDeclNoFC -builtinDecl fname indents - = do decoratedPragma fname "builtin" - commit - t <- builtinType - n <- name - pure $ PBuiltin t n +parameters {auto fname : OriginDesc} {auto indents : IndentInfo} + namespaceDecl : Rule PDeclNoFC + namespaceDecl + = do doc <- optDocumentation fname -- documentation is not recoded??? + col <- column + ns <- namespaceHead fname + ds <- blockAfter col (topDecl fname) + pure (PNamespace ns (collectDefs ds)) + + transformDecl : Rule PDeclNoFC + transformDecl + = do decoratedPragma fname "transform" + n <- simpleStr + lhs <- expr plhs fname indents + decoratedSymbol fname "=" + rhs <- expr pnowith fname indents + pure (PTransform n lhs rhs) + + runElabDecl : Rule PDeclNoFC + runElabDecl + = do + decoratedPragma fname "runElab" + tm <- expr pnowith fname indents + pure (PRunElabDecl tm) + + ||| failDecls := 'failing' simpleStr? nonEmptyBlock + failDecls : Rule PDeclNoFC + failDecls + = do + col <- column + decoratedKeyword fname "failing" + commit + msg <- optional (decorate fname Data simpleStr) + ds <- nonEmptyBlockAfter col (topDecl fname) + pure $ PFail msg (collectDefs $ forget ds) + + ||| mutualDecls := 'mutual' nonEmptyBlock + mutualDecls : Rule PDeclNoFC + mutualDecls + = do + col <- column + decoratedKeyword fname "mutual" + commit + ds <- nonEmptyBlockAfter col (topDecl fname) + pure (PMutual (forget ds)) + + usingDecls : Rule PDeclNoFC + usingDecls + = do col <- column + decoratedKeyword fname "using" + commit + decoratedSymbol fname "(" + us <- sepBy (decoratedSymbol fname ",") + (do n <- optional $ do + x <- unqualifiedName + decoratedSymbol fname ":" + pure (UN $ Basic x) + ty <- typeExpr pdef fname indents + pure (n, ty)) + decoratedSymbol fname ")" + ds <- nonEmptyBlockAfter col (topDecl fname) + pure (PUsing us (collectDefs (forget ds))) + + ||| builtinDecl := 'builtin' builtinType name + builtinDecl : Rule PDeclNoFC + builtinDecl + = do decoratedPragma fname "builtin" + commit + t <- builtinType + n <- name + pure $ PBuiltin t n visOpt : OriginDesc -> Rule (Either Visibility PFnOpt) visOpt fname @@ -1681,93 +1707,30 @@ constraints fname indents implBinds : OriginDesc -> IndentInfo -> (namedImpl : Bool) -> EmptyRule (List (FC, RigCount, Name, PiInfo PTerm, PTerm)) implBinds fname indents namedImpl = concatMap (map adjust) <$> go where - adjust : (RigCount, WithBounds Name, a) -> (FC, RigCount, Name, a) - adjust (r, wn, ty) = (virtualiseFC (boundToFC fname wn), r, wn.val, ty) + adjust : (RigCount, WithFC Name, a) -> (FC, RigCount, Name, a) + adjust (r, wn, ty) = (virtualiseFC wn.fc, r, wn.val, ty) isDefaultImplicit : PiInfo a -> Bool isDefaultImplicit (DefImplicit _) = True isDefaultImplicit _ = False - go : EmptyRule (List (List (RigCount, WithBounds Name, PiInfo PTerm, PTerm))) + go : EmptyRule (List (List (RigCount, WithFC Name, PiInfo PTerm, PTerm))) go = do decoratedSymbol fname "{" piInfo <- bounds $ option Implicit $ defImplicitField fname indents when (not namedImpl && isDefaultImplicit piInfo.val) $ fatalLoc piInfo.bounds "Default implicits are allowed only for named implementations" - ns <- map @{Compose} (\(rc, wb, n) => (rc, wb, piInfo.val, n)) $ pibindListName fname indents + ns <- map + (\case (MkBasicMultiBinder rig names type) => map (\nm => (rig, nm, piInfo.val, type)) (forget names)) + (pibindListName fname indents) + let ns = the (List (ZeroOneOmega, (WithFC Name, (PiInfo (PTerm' Name), PTerm' Name)))) ns commitSymbol fname "}" commitSymbol fname "->" more <- go pure (ns :: more) <|> pure [] -ifaceParam : OriginDesc -> IndentInfo -> Rule (List Name, (RigCount, PTerm)) -ifaceParam fname indents - = do decoratedSymbol fname "(" - rig <- multiplicity fname - ns <- sepBy1 (decoratedSymbol fname ",") (decorate fname Bound name) - decoratedSymbol fname ":" - tm <- typeExpr pdef fname indents - decoratedSymbol fname ")" - pure (forget ns, (rig, tm)) - <|> do n <- bounds (decorate fname Bound name) - pure ([n.val], (erased, PInfer (boundToFC fname n))) - -ifaceDecl : OriginDesc -> IndentInfo -> Rule PDeclNoFC -ifaceDecl fname indents - = do doc <- optDocumentation fname - vis <- visibility fname - col <- column - decoratedKeyword fname "interface" - commit - cons <- constraints fname indents - n <- decorate fname Typ name - paramss <- many (ifaceParam fname indents) - let params = concatMap (\ (ns, rt) => map (\ n => (n, rt)) ns) paramss - det <- option [] $ decoratedSymbol fname "|" *> sepBy (decoratedSymbol fname ",") (decorate fname Bound name) - decoratedKeyword fname "where" - dc <- optional (recordConstructor fname) - body <- blockAfter col (topDecl fname) - pure (PInterface vis cons n doc params det dc (collectDefs body)) - -implDecl : OriginDesc -> IndentInfo -> Rule PDeclNoFC -implDecl fname indents - = do doc <- optDocumentation fname - visOpts <- many (visOpt fname) - vis <- getVisibility Nothing visOpts - let opts = mapMaybe getRight visOpts - col <- column - option () (decoratedKeyword fname "implementation") - iname <- optional $ decoratedSymbol fname "[" - *> decorate fname Function name - <* decoratedSymbol fname "]" - impls <- implBinds fname indents (isJust iname) - cons <- constraints fname indents - n <- decorate fname Typ name - params <- many (continue indents *> simpleExpr fname indents) - nusing <- option [] $ decoratedKeyword fname "using" - *> forget <$> some (decorate fname Function name) - body <- optional $ decoratedKeyword fname "where" *> blockAfter col (topDecl fname) - atEnd indents - pure - (PImplementation vis opts Single impls cons n params iname nusing - (map collectDefs body)) - -localClaim : OriginDesc -> IndentInfo -> Rule (WithFC PClaimData) -localClaim fname indents - = do bs <- bounds (do - doc <- optDocumentation fname - visOpts <- many (visOpt fname) - vis <- getVisibility Nothing visOpts - let opts = mapMaybe getRight visOpts - rig <- multiplicity fname - cls <- tyDecls (decorate fname Function name) - doc fname indents - pure $ MkPClaim rig vis opts cls - ) - pure bs.withFC - -fieldDecl : OriginDesc -> IndentInfo -> Rule PField -fieldDecl fname indents +fieldDecl : (fname : OriginDesc) => IndentInfo -> Rule PField +fieldDecl indents = do doc <- optDocumentation fname decoratedSymbol fname "{" commit @@ -1794,116 +1757,173 @@ fieldDecl fname indents pure (MkRecordField doc rig p (forget ns) ty)) pure b.withFC -typedArg : OriginDesc -> IndentInfo -> Rule (List (Name, RigCount, PiInfo PTerm, PTerm)) -typedArg fname indents - = do params <- parens fname $ pibindListName fname indents - pure $ map (\(c, n, tm) => (n.val, c, Explicit, tm)) params - <|> do decoratedSymbol fname "{" - commit - info <- - (pure AutoImplicit <* decoratedKeyword fname "auto" - <|> (decoratedKeyword fname "default" *> DefImplicit <$> simpleExpr fname indents) - <|> pure Implicit) - params <- pibindListName fname indents - decoratedSymbol fname "}" - pure $ map (\(c, n, tm) => (n.val, c, info, tm)) params - -recordParam : OriginDesc -> IndentInfo -> Rule (List (Name, RigCount, PiInfo PTerm, PTerm)) -recordParam fname indents - = typedArg fname indents - <|> do n <- bounds (UN . Basic <$> decoratedSimpleBinderName fname) - pure [(n.val, top, Explicit, PInfer (boundToFC fname n))] - --- A record without a where is a forward declaration -recordBody : OriginDesc -> IndentInfo -> - String -> WithDefault Visibility Private -> - Maybe TotalReq -> - Int -> - Name -> - List (Name, RigCount, PiInfo PTerm, PTerm) -> - EmptyRule PDeclNoFC -recordBody fname indents doc vis mbtot col n params - = do atEndIndent indents - pure (PRecord doc vis mbtot (MkPRecordLater n params)) - <|> do mustWork $ decoratedKeyword fname "where" - opts <- dataOpts fname - dcflds <- blockWithOptHeaderAfter col - (\ idt => recordConstructor fname <* atEnd idt) - (fieldDecl fname) - pure (PRecord doc vis mbtot - (MkPRecord n params opts (fst dcflds) (snd dcflds))) - -recordDecl : OriginDesc -> IndentInfo -> Rule PDeclNoFC -recordDecl fname indents - = do doc <- optDocumentation fname - (vis,mbtot) <- dataVisOpt fname - col <- column - decoratedKeyword fname "record" - n <- mustWork (decoratedDataTypeName fname) - paramss <- many (continue indents >> recordParam fname indents) - let params = concat paramss - recordBody fname indents doc vis mbtot col n params - -paramDecls : OriginDesc -> IndentInfo -> Rule PDeclNoFC -paramDecls fname indents = do - startCol <- column - _ <- decoratedKeyword fname "parameters" - _ <- commit - params <- Right <$> newParamDecls fname indents - <|> Left <$> withWarning "DEPRECATED: old parameter syntax https://github.com/idris-lang/Idris2/issues/3447" (oldParamDecls fname indents) - commit - declarations <- bounds $ nonEmptyBlockAfter startCol (topDecl fname) - pure (PParameters params - (collectDefs (forget declarations.val))) +parameters {auto fname : OriginDesc} {auto indents : IndentInfo} + + ifaceParam : Rule BasicMultiBinder + ifaceParam + = parens fname basicMultiBinder + <|> do n <- fcBounds (decorate fname Bound name) + pure (MkBasicMultiBinder erased (singleton n) (PInfer n.fc)) + + ifaceDecl : Rule PDeclNoFC + ifaceDecl + = do doc <- optDocumentation fname + vis <- visibility fname + col <- column + decoratedKeyword fname "interface" + commit + cons <- constraints fname indents + n <- decorate fname Typ name + params <- many ifaceParam + det <- option [] $ decoratedSymbol fname "|" *> sepBy (decoratedSymbol fname ",") (decorate fname Bound name) + decoratedKeyword fname "where" + dc <- optional (recordConstructor fname) + body <- blockAfter col (topDecl fname) + pure (PInterface + vis cons n doc params det dc (collectDefs body)) + + implDecl : Rule PDeclNoFC + implDecl + = do doc <- optDocumentation fname + visOpts <- many (visOpt fname) + vis <- getVisibility Nothing visOpts + let opts = mapMaybe getRight visOpts + col <- column + option () (decoratedKeyword fname "implementation") + iname <- optional $ decoratedSymbol fname "[" + *> decorate fname Function name + <* decoratedSymbol fname "]" + impls <- implBinds fname indents (isJust iname) + cons <- constraints fname indents + n <- decorate fname Typ name + params <- many (continue indents *> simpleExpr fname indents) + nusing <- option [] $ decoratedKeyword fname "using" + *> forget <$> some (decorate fname Function name) + body <- optional $ decoratedKeyword fname "where" *> blockAfter col (topDecl fname) + atEnd indents + pure $ + PImplementation vis opts Single impls cons n params iname nusing + (map collectDefs body) + + localClaim : Rule PClaimData + localClaim + = do doc <- optDocumentation fname + visOpts <- many (visOpt fname) + vis <- getVisibility Nothing visOpts + let opts = mapMaybe getRight visOpts + rig <- multiplicity fname + cls <- tyDecls (decorate fname Function name) + doc fname indents + pure $ MkPClaim rig vis opts cls + + + -- A Single binder with multiple names + typedArg : Rule PBinder + typedArg + = do params <- parens fname $ pibindListName fname indents + pure $ MkPBinder Explicit params + <|> do decoratedSymbol fname "{" + commit + info <- + (pure AutoImplicit <* decoratedKeyword fname "auto" + <|> (decoratedKeyword fname "default" *> DefImplicit <$> simpleExpr fname indents) + <|> pure Implicit) + params <- pibindListName fname indents + decoratedSymbol fname "}" + pure $ MkPBinder info params + + recordParam : Rule (PBinder) + recordParam + = typedArg + <|> do n <- fcBounds (decoratedSimpleBinderUName fname) + pure (MkFullBinder Explicit top n $ PInfer n.fc) + + -- A record without a where is a forward declaration + recordBody : String -> WithDefault Visibility Private -> + Maybe TotalReq -> + Int -> + Name -> + List PBinder -> + EmptyRule PDeclNoFC + recordBody doc vis mbtot col n params + = do atEndIndent indents + pure (PRecord doc vis mbtot (MkPRecordLater n params)) + <|> do mustWork $ decoratedKeyword fname "where" + opts <- dataOpts fname + dcflds <- blockWithOptHeaderAfter col + (\ idt => recordConstructor fname <* atEnd idt) + fieldDecl + pure (PRecord doc vis mbtot + (MkPRecord n params opts (fst dcflds) (snd dcflds))) + + recordDecl : Rule PDeclNoFC + recordDecl + = do doc <- optDocumentation fname + (vis,mbtot) <- dataVisOpt fname + col <- column + decoratedKeyword fname "record" + n <- mustWork (decoratedDataTypeName fname) + paramss <- many (continue indents >> recordParam) + recordBody doc vis mbtot col n paramss + + ||| Parameter blocks + ||| BNF: + ||| paramDecls := 'parameters' (oldParamDecls | newParamDecls) indentBlockDefs + paramDecls : Rule PDeclNoFC + paramDecls = do + startCol <- column + b1 <- decoratedKeyword fname "parameters" + commit + args <- Right <$> newParamDecls + <|> Left <$> withWarning "DEPRECATED: old parameter syntax https://github.com/idris-lang/Idris2/issues/3447" oldParamDecls + commit + declarations <- nonEmptyBlockAfter startCol (topDecl fname) + pure (PParameters args + (collectDefs (forget declarations))) - where - oldParamDecls : OriginDesc -> IndentInfo -> Rule (List (Name, PTerm)) - oldParamDecls fname indents - = do decoratedSymbol fname "(" - ps <- sepBy (decoratedSymbol fname ",") - (do x <- UN . Basic <$> decoratedSimpleBinderName fname - decoratedSymbol fname ":" - ty <- typeExpr pdef fname indents - pure (x, ty)) - decoratedSymbol fname ")" - pure ps - - newParamDecls : OriginDesc -> IndentInfo -> Rule (List (Name, RigCount, PiInfo PTerm, PTerm)) - newParamDecls fname indents - = map concat (some $ typedArg fname indents) - - --- topLevelClaim is for claims appearing at the top-level of the file --- localClaim is for claims appearing in nested positions, like `let` or `record` in the future -topLevelClaim : OriginDesc -> IndentInfo -> Rule PDecl -topLevelClaim o i = mapFC PClaim <$> localClaim o i - -definition : OriginDesc -> IndentInfo -> Rule PDecl -definition fname indents - = do nd <- fcBounds (clause 0 Nothing fname indents) - pure (mapFC (PDef . singleton) nd) - -operatorBindingKeyword : OriginDesc -> EmptyRule BindingModifier -operatorBindingKeyword fname - = (decoratedKeyword fname "autobind" >> pure Autobind) - <|> (decoratedKeyword fname "typebind" >> pure Typebind) - <|> pure NotBinding - -fixDecl : OriginDesc -> IndentInfo -> Rule PDecl -fixDecl fname indents - = do vis <- exportVisibility fname - binding <- operatorBindingKeyword fname - b <- fcBounds (do fixity <- decorate fname Keyword $ fix - commit - prec <- decorate fname Keyword $ intLit - ops <- sepBy1 (decoratedSymbol fname ",") iOperator - pure (MkPFixityData vis binding fixity (fromInteger prec) ops) + where + oldParamDecls : Rule (List1 PlainBinder) + oldParamDecls + = parens fname $ sepBy1 (decoratedSymbol fname ",") plainBinder + + newParamDecls : Rule (List1 PBinder) + newParamDecls = some typedArg + + + -- topLevelClaim is for claims appearing at the top-level of the file + -- localClaim is for claims appearing in nested positions, like `let` or `record` in the future + topLevelClaim : Rule PDeclNoFC + topLevelClaim = PClaim <$> localClaim + + definition : Rule PDeclNoFC + definition + = do nd <- clause 0 Nothing fname indents + pure (PDef (singleton nd)) + + operatorBindingKeyword : EmptyRule BindingModifier + operatorBindingKeyword + = (decoratedKeyword fname "autobind" >> pure Autobind) + <|> (decoratedKeyword fname "typebind" >> pure Typebind) + <|> pure NotBinding + + fixDecl : Rule PDecl + fixDecl + = do vis <- exportVisibility fname + binding <- operatorBindingKeyword + b <- fcBounds (do fixity <- decorate fname Keyword $ fix + commit + prec <- decorate fname Keyword $ intLit + ops <- sepBy1 (decoratedSymbol fname ",") iOperator + pure (MkPFixityData vis binding fixity (fromInteger prec) ops) ) - pure (mapFC PFixity b) + pure (mapFC PFixity b) -directiveDecl : OriginDesc -> IndentInfo -> Rule PDecl -directiveDecl fname indents - = fcBounds $ PDirective <$> directive fname indents +cgDirectiveDecl : Rule PDeclNoFC +cgDirectiveDecl + = (>>=) {c1 = True, c2 = False} cgDirective $ \dir => + let (cg1, cg2) = span isAlphaNum dir + in the (EmptyRule PDeclNoFC) $ pure $ + PDirective (CGAction cg1 (stripBraces (trim cg2))) -- Declared at the top -- topDecl : OriginDesc -> IndentInfo -> Rule (List PDecl) @@ -1912,29 +1932,23 @@ topDecl fname indents -- i.e. the claim "String : Type" is a parse error, but the underlying reason may not be clear to new users. = do id <- anyReservedIdent the (Rule PDecl) $ fatalLoc id.bounds "Cannot begin a declaration with a reserved identifier" - <|> fcBounds (dataDecl fname indents) - <|> topLevelClaim fname indents - <|> directiveDecl fname indents - <|> fcBounds (implDecl fname indents) - <|> definition fname indents - <|> fixDecl fname indents - <|> fcBounds (ifaceDecl fname indents) - <|> fcBounds (recordDecl fname indents) - <|> fcBounds (namespaceDecl fname indents) - <|> fcBounds (failDecls fname indents) - <|> fcBounds (mutualDecls fname indents) - <|> fcBounds (paramDecls fname indents) - <|> fcBounds (usingDecls fname indents) - <|> fcBounds (builtinDecl fname indents) - <|> fcBounds (runElabDecl fname indents) - <|> fcBounds (transformDecl fname indents) - <|> do dstr <- bounds (terminal "Expected CG directive" - (\case - CGDirective d => Just d - _ => Nothing)) - pure (let cgrest = span isAlphaNum dstr.val in - MkFCVal (boundToFC fname dstr) $ PDirective $ - (CGAction (fst cgrest) (stripBraces (trim (snd cgrest))))) + <|> fcBounds dataDecl + <|> fcBounds topLevelClaim + <|> fcBounds (PDirective <$> directive) + <|> fcBounds implDecl + <|> fcBounds definition + <|> fixDecl + <|> fcBounds ifaceDecl + <|> fcBounds recordDecl + <|> fcBounds namespaceDecl + <|> fcBounds failDecls + <|> fcBounds mutualDecls + <|> fcBounds paramDecls + <|> fcBounds usingDecls + <|> fcBounds builtinDecl + <|> fcBounds runElabDecl + <|> fcBounds transformDecl + <|> fcBounds cgDirectiveDecl -- If the user tried to begin a declaration with any other keyword, then show a more informative error. <|> do kw <- bounds anyKeyword the (Rule PDecl) $ fatalLoc kw.bounds "Keyword '\{kw.val}' is not a valid start to a declaration" diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index 2248884ae9..41d8be91f3 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -109,6 +109,9 @@ export code : Doc IdrisAnn -> Doc IdrisAnn code = annotate Code +Pretty i a => Pretty i (WithFC a) where + pretty x = pretty x.val + mutual prettyAlt : PClause' KindedName -> Doc IdrisSyntax @@ -174,9 +177,26 @@ mutual then annotateM (kindAnn op) $ pretty0 nm else Chara '`' <+> annotateM (kindAnn op) (pretty0 nm) <+> Chara '`' + Pretty IdrisSyntax (BasicMultiBinder' KindedName) where + pretty (MkBasicMultiBinder rig names type) + = prettyRig rig <++> commaSep (forget $ map (prettyBinder . val) names) + <++> colon <++> pretty type + export Pretty IdrisSyntax IPTerm where prettyPrec d (PRef _ nm) = annotateM (kindAnn nm) $ cast $ prettyOp False nm.rawName + prettyPrec d (NewPi (MkFCVal fc (MkPBinderScope (MkPBinder Implicit bind) scope))) = + lcurly <+> pretty bind <+> rcurly <++> arrow <++> prettyPrec d scope + prettyPrec d (NewPi (MkFCVal fc (MkPBinderScope (MkPBinder Explicit bind) scope))) = + lparen <+> pretty bind <+> rparen <++> arrow <++> prettyPrec d scope + prettyPrec d (NewPi (MkFCVal fc (MkPBinderScope (MkPBinder AutoImplicit bind) scope))) = + lcurly <+> auto_ <++> pretty bind <+> rcurly <++> arrow <++> prettyPrec d scope + prettyPrec d (NewPi (MkFCVal fc (MkPBinderScope (MkPBinder (DefImplicit x) bind) scope))) = + lcurly <+> default_ <++> prettyPrec appPrec x + <++> pretty bind <+> rcurly <++> arrow <++> prettyPrec d scope + prettyPrec d (Forall (MkFCVal fc (names, scope))) = + parenthesise (d > startPrec) $ group $ + forall_ <++> commaSep (map (prettyBinder . val) (forget names)) <++> dot <++> pretty scope prettyPrec d (PPi _ rig Explicit Nothing arg ret) = parenthesise (d > startPrec) $ group $ branchVal diff --git a/src/Idris/Pretty/Annotations.idr b/src/Idris/Pretty/Annotations.idr index 208d077bb4..fdddf01b18 100644 --- a/src/Idris/Pretty/Annotations.idr +++ b/src/Idris/Pretty/Annotations.idr @@ -84,6 +84,10 @@ export impossible_ : Doc IdrisSyntax impossible_ = keyword "impossible" +export +forall_ : Doc IdrisSyntax +forall_ = keyword "forall" + export auto_ : Doc IdrisSyntax auto_ = keyword "auto" diff --git a/src/Idris/Resugar.idr b/src/Idris/Resugar.idr index 1bc5afdf34..17fd416537 100644 --- a/src/Idris/Resugar.idr +++ b/src/Idris/Resugar.idr @@ -47,7 +47,8 @@ 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 (MkFCVal opFC $ NoBinder (unbracketApp x)) (MkFCVal opFC $ pop kn) (unbracketApp y) + let asOp = POp fc (MkFCVal opFC + $ NoBinder (unbracketApp x)) (MkFCVal opFC (pop kn)) (unbracketApp y) if not (null (lookupName rootName (infixes syn))) then pure asOp else case dropNS raw of @@ -532,18 +533,25 @@ mutual toPDecl (IData fc vis mbtot d) = pure (Just (MkFCVal fc $ PData "" vis mbtot !(toPData d))) toPDecl (IDef fc n cs) - = pure (Just (MkFCVal fc $ PDef $ !(traverse toPClause cs))) + = pure (Just (MkFCVal fc $ PDef !(traverse toPClause cs))) toPDecl (IParameters fc ps ds) = do ds' <- traverse toPDecl ds args <- - traverse (\(n, rig, info, tpe) => + traverseList1 (\(n, rig, info, tpe) => do info' <- traverse (toPTerm startPrec) info type' <- toPTerm startPrec tpe - pure (n, rig, info', type')) ps - pure (Just $ MkFCVal fc (PParameters (Right args) (catMaybes ds'))) + pure (MkFullBinder info' rig (NoFC n) type')) ps + pure (Just (MkFCVal fc (PParameters (Right args) (catMaybes ds')))) toPDecl (IRecord fc _ vis mbtot r) = do (n, ps, opts, con, fs) <- toPRecord r - pure (Just $ MkFCVal fc (PRecord "" vis mbtot (MkPRecord n ps opts con fs))) + pure (Just (MkFCVal fc $ PRecord "" vis mbtot (MkPRecord n (map toBinder ps) opts con fs))) + where + toBinder : (Name, ZeroOneOmega, PiInfo (PTerm' KindedName), PTerm' KindedName) -> PBinder' KindedName + toBinder (n, rig, info, ty) + = MkFullBinder info rig (NoFC n) ty + -- ^^^^ + -- we should know this location + toPDecl (IFail fc msg ds) = do ds' <- traverse toPDecl ds pure (Just (MkFCVal fc $ PFail msg (catMaybes ds'))) diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index dae4bfac66..f7385e2131 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -111,6 +111,9 @@ mutual -- Direct (more or less) translations to RawImp PRef : FC -> nm -> PTerm' nm + NewPi : WithFC (PBinderScope' nm) -> PTerm' nm + Forall : WithFC (List1 (WithFC Name), PTerm' nm) -> PTerm' nm + PPi : FC -> RigCount -> PiInfo (PTerm' nm) -> Maybe Name -> (argTy : PTerm' nm) -> (retTy : PTerm' nm) -> PTerm' nm PLam : FC -> RigCount -> PiInfo (PTerm' nm) -> (pat : PTerm' nm) -> @@ -188,6 +191,8 @@ mutual export getPTermLoc : PTerm' nm -> FC getPTermLoc (PRef fc _) = fc + getPTermLoc (NewPi x) = x.fc + getPTermLoc (Forall x) = x.fc getPTermLoc (PPi fc _ _ _ _ _) = fc getPTermLoc (PLam fc _ _ _ _ _) = fc getPTermLoc (PLet fc _ _ _ _ _ _) = fc @@ -272,6 +277,82 @@ mutual StrLiteral : FC -> String -> PStr' nm StrInterp : FC -> PTerm' nm -> PStr' nm + public export + PlainMultiBinder : Type + PlainMultiBinder = PlainMultiBinder' Name + + ||| A plain binder without information about its use + ||| plainBinder := name ':' term + public export + PlainMultiBinder' : (nm : Type) -> Type + PlainMultiBinder' nm = List1 (WithName (PTerm' nm)) + + public export + PlainBinder : Type + PlainBinder = PlainBinder' Name + + ||| A plain binder without information about its use + ||| plainBinder := name ':' term + public export + PlainBinder' : (nm : Type) -> Type + PlainBinder' nm = WithName (PTerm' nm) + + public export + BasicMultiBinder : Type + BasicMultiBinder = BasicMultiBinder' Name + + ||| A binder with quantity information attached + ||| basicBinder := qty plainBinder + public export + record BasicMultiBinder' (nm : Type) where + constructor MkBasicMultiBinder + rig : RigCount + names : List1 (WithFC Name) + type : PTerm' nm + + public export + BasicBinder : Type + BasicBinder = BasicBinder' Name + + ||| A binder with quantity information attached + ||| basicBinder := qty plainBinder + public export + record BasicBinder' (nm : Type) where + constructor MkBasicBinder + rig : RigCount + name : WithFC Name + type : PTerm' nm + + public export + PBinder : Type + PBinder = PBinder' Name + + ||| A binder with information about how it is binding + ||| pbinder := '{' basicBinder '}' + ||| | '{' auto basicBinder '}' + ||| | '{' default term basicBinder '}' + ||| | '(' basicBinder ')' + public export + record PBinder' (nm : Type) where + constructor MkPBinder + info : PiInfo (PTerm' nm) + bind : BasicMultiBinder' nm + + public export + PBinderScope : Type + PBinderScope = PBinderScope' Name + + public export + record PBinderScope' (nm : Type) where + constructor MkPBinderScope + binder : PBinder' nm + scope : PTerm' nm + + public export + MkFullBinder : PiInfo (PTerm' nm) -> RigCount -> WithFC Name -> PTerm' nm -> PBinder' nm + MkFullBinder info rig x y = MkPBinder info (MkBasicMultiBinder rig (singleton x) y) + + export getLoc : PDo' nm -> FC getLoc (DoExp fc _) = fc @@ -335,13 +416,13 @@ mutual public export data PRecordDecl' : Type -> Type where MkPRecord : (tyname : Name) -> - (params : List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) -> + (params : List (PBinder' nm)) -> (opts : List DataOpt) -> (conName : Maybe (String, Name)) -> (decls : List (PField' nm)) -> PRecordDecl' nm MkPRecordLater : (tyname : Name) -> - (params : List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) -> + (params : List (PBinder' nm)) -> PRecordDecl' nm export @@ -498,8 +579,8 @@ mutual PDef : List (PClause' nm) -> PDeclNoFC' nm PData : (doc : String) -> WithDefault Visibility Private -> Maybe TotalReq -> PDataDecl' nm -> PDeclNoFC' nm - PParameters : Either (List (Name, PTerm' nm)) - (List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) -> + PParameters : Either (List1 (PlainBinder' nm)) + (List1 (PBinder' nm)) -> List (PDecl' nm) -> PDeclNoFC' nm PUsing : List (Maybe Name, PTerm' nm) -> List (PDecl' nm) -> PDeclNoFC' nm @@ -507,7 +588,7 @@ mutual (constraints : List (Maybe Name, PTerm' nm)) -> Name -> (doc : String) -> - (params : List (Name, RigCount, PTerm' nm)) -> + (params : List (BasicMultiBinder' nm)) -> (det : List Name) -> (conName : Maybe (String, Name)) -> List (PDecl' nm) -> @@ -752,6 +833,8 @@ parameters {0 nm : Type} (toName : nm -> Name) showUpdate : PFieldUpdate' nm -> String showPTermPrec : Prec -> PTerm' nm -> String showOpPrec : Prec -> OpStr' nm -> String + showPBinder : Prec -> PBinder' nm -> String + showBasicMultiBinder : BasicMultiBinder' nm -> String showPTerm : PTerm' nm -> String showPTerm = showPTermPrec Open @@ -783,7 +866,22 @@ parameters {0 nm : Type} (toName : nm -> Name) showUpdate (PSetField p v) = showSep "." p ++ " = " ++ showPTerm v showUpdate (PSetFieldApp p v) = showSep "." p ++ " $= " ++ showPTerm v + showBasicMultiBinder (MkBasicMultiBinder rig names type) + = "\{showCount rig} \{showNames}: \{showPTerm type}" + where + showNames : String + showNames = concat $ intersperse ", " $ map (show . val) (forget names) + + showPBinder d (MkPBinder Implicit bind) = "{\{showBasicMultiBinder bind}}" + showPBinder d (MkPBinder Explicit bind) = "(\{showBasicMultiBinder bind})" + showPBinder d (MkPBinder AutoImplicit bind) = "{auto \{showBasicMultiBinder bind}}" + showPBinder d (MkPBinder (DefImplicit x) bind) = "{default \{showPTerm x} \{ showBasicMultiBinder bind}}" + showPTermPrec d (PRef _ n) = showPrec d (toName n) + showPTermPrec d (Forall (MkFCVal _ (names, scope))) + = "forall " ++ concat (intersperse ", " (map (show . val) (forget names))) ++ " . " ++ showPTermPrec d scope + showPTermPrec d (NewPi (MkFCVal _ (MkPBinderScope binder scope))) + = showPBinder d binder ++ " -> " ++ showPTermPrec d scope showPTermPrec d (PPi _ rig Explicit Nothing arg ret) = showPTermPrec d arg ++ " -> " ++ showPTermPrec d ret showPTermPrec d (PPi _ rig Explicit (Just n) arg ret) diff --git a/src/Idris/Syntax/Traversals.idr b/src/Idris/Syntax/Traversals.idr index 50d5c40aea..6cd96abcb5 100644 --- a/src/Idris/Syntax/Traversals.idr +++ b/src/Idris/Syntax/Traversals.idr @@ -17,6 +17,10 @@ mapPTermM f = goPTerm where goPTerm : PTerm' nm -> Core (PTerm' nm) goPTerm t@(PRef _ _) = f t + goPTerm (NewPi x) = + NewPi <$> traverseFC goPBinderScope x + goPTerm (Forall x) = + Forall <$> traverseFC (\(a, b) => MkPair a <$> goPTerm b) x goPTerm (PPi fc x info z argTy retTy) = PPi fc x <$> goPiInfo info <*> pure z @@ -42,7 +46,7 @@ mapPTermM f = goPTerm where <*> goPClauses xs >>= f goPTerm (PLocal fc xs scope) = - PLocal fc <$> goPDecls xs + PLocal fc <$> traverse (traverseFC goPDecl) xs <*> goPTerm scope >>= f goPTerm (PUpdate fc xs) = @@ -81,7 +85,7 @@ mapPTermM f = goPTerm where >>= f goPTerm t@(PQuoteName _ _) = f t goPTerm (PQuoteDecl fc x) = - PQuoteDecl fc <$> traverse (traverseFC goPDecl) x + PQuoteDecl fc <$> goPDecls x >>= f goPTerm (PUnquote fc x) = PUnquote fc <$> goPTerm x @@ -99,27 +103,9 @@ mapPTermM f = goPTerm where >>= f goPTerm t@(PImplicit _) = f t goPTerm t@(PInfer _) = f t - goPTerm (POp fc (MkFCVal opFC $ NoBinder left) op right) = + goPTerm (POp fc bind op right) = POp fc - <$> (MkFCVal opFC <$> NoBinder <$> goPTerm left) - <*> pure op - <*> goPTerm right - >>= f - goPTerm (POp fc (MkFCVal opFC $ BindType nm left) op right) = - POp fc - <$> (MkFCVal opFC <$> (BindType <$> goPTerm nm <*> goPTerm left)) - <*> pure op - <*> goPTerm right - >>= f - goPTerm (POp fc (MkFCVal opFC $ BindExpr nm left) op right) = - POp fc - <$> (MkFCVal opFC <$> BindExpr nm <$> goPTerm left) - <*> pure op - <*> goPTerm right - >>= f - goPTerm (POp fc (MkFCVal opFC $ BindExplicitType nm ty left) op right) = - POp fc - <$> (MkFCVal opFC <$> (BindExplicitType <$> goPTerm nm <*> goPTerm ty <*> goPTerm left)) + <$> traverseFC goOpBinder bind <*> pure op <*> goPTerm right >>= f @@ -206,6 +192,13 @@ mapPTermM f = goPTerm where PWithUnambigNames fc ns <$> goPTerm rhs >>= f + goOpBinder : OperatorLHSInfo (PTerm' nm) -> Core (OperatorLHSInfo (PTerm' nm)) + goOpBinder (NoBinder lhs) = NoBinder <$> goPTerm lhs + goOpBinder (BindType name ty) = BindType <$> goPTerm name <*> goPTerm ty + goOpBinder (BindExpr name expr) = BindExpr <$> goPTerm name <*> goPTerm expr + goOpBinder (BindExplicitType name type expr) + = BindExplicitType <$> goPTerm name <*> goPTerm type <*> goPTerm expr + goPFieldUpdate : PFieldUpdate' nm -> Core (PFieldUpdate' nm) goPFieldUpdate (PSetField p t) = PSetField p <$> goPTerm t goPFieldUpdate (PSetFieldApp p t) = PSetFieldApp p <$> goPTerm t @@ -258,18 +251,31 @@ mapPTermM f = goPTerm where MkPClaim c v <$> goPFnOpts opts <*> traverseFC goPTypeDecl tdecl - goPDecl' : PDecl' nm -> Core (PDecl' nm) + goPlainBinder : PlainBinder' nm -> Core (PlainBinder' nm) + goPlainBinder = traverseWName f + + goBasicMultiBinder : BasicMultiBinder' nm -> Core (BasicMultiBinder' nm) + goBasicMultiBinder (MkBasicMultiBinder rig names type) + = MkBasicMultiBinder rig names <$> goPTerm type + + goPBinder : PBinder' nm -> Core (PBinder' nm) + goPBinder (MkPBinder info bind) + = MkPBinder <$> goPiInfo info <*> goBasicMultiBinder bind + + goPBinderScope : PBinderScope' nm -> Core (PBinderScope' nm) + goPBinderScope (MkPBinderScope binder scope) + = MkPBinderScope <$> goPBinder binder <*> goPTerm scope + + -- goParamArgs : goPDecl : PDeclNoFC' nm -> Core (PDeclNoFC' nm) goPDecl (PClaim claim) = PClaim <$> goPClaim claim - goPDecl (PDef cls) = PDef <$> traverse goPClause cls + goPDecl (PDef cls) = PDef <$> goPClauses cls goPDecl (PData doc v mbt d) = PData doc v mbt <$> goPDataDecl d - goPDecl (PParameters (Left nts) ps) = - PParameters <$> Left <$> traverse (traversePair goPTerm) nts - <*> goPDecls ps - goPDecl (PParameters (Right nts) ps) = - PParameters <$> (Right <$> (go4TupledPTerms nts)) + goPDecl (PParameters nts ps) = + PParameters <$> either (\x => Left <$> traverseList1 goPlainBinder x) + (\x => Right <$> traverseList1 goPBinder x) nts -- go4TupledPTerms nts <*> goPDecls ps goPDecl (PUsing mnts ps) = PUsing <$> goPairedPTerms mnts @@ -278,7 +284,7 @@ mapPTermM f = goPTerm where PInterface v <$> goPairedPTerms mnts <*> pure n <*> pure doc - <*> go3TupledPTerms nrts + <*> traverse goBasicMultiBinder nrts <*> pure ns <*> pure mn <*> goPDecls ps @@ -291,12 +297,12 @@ mapPTermM f = goPTerm where <*> pure ns <*> goMPDecls mps goPDecl (PRecord doc v tot (MkPRecord n nts opts mn fs)) = - pure $ PRecord doc v tot !(MkPRecord n <$> go4TupledPTerms nts + pure $ PRecord doc v tot !(MkPRecord n <$> traverse goPBinder nts <*> pure opts <*> pure mn <*> goPFields fs) goPDecl (PRecord doc v tot (MkPRecordLater n nts)) = - pure $ PRecord doc v tot (MkPRecordLater n !(go4TupledPTerms nts)) + pure $ PRecord doc v tot (MkPRecordLater n !(traverse goPBinder nts)) goPDecl (PFail msg ps) = PFail msg <$> goPDecls ps goPDecl (PMutual ps) = PMutual <$> goPDecls ps goPDecl (PFixity p) = pure (PFixity p) @@ -420,9 +426,12 @@ mapPTerm f = goPTerm where mutual - goPTerm : PTerm' nm -> PTerm' nm goPTerm t@(PRef _ _) = f t + goPTerm (Forall binderScope) + = Forall (mapFC (mapSnd f) binderScope) + goPTerm (NewPi binderScope) + = NewPi (mapFC goPBinderScope binderScope) goPTerm (PPi fc x info z argTy retTy) = f $ PPi fc x (goPiInfo info) z (goPTerm argTy) (goPTerm retTy) goPTerm (PLam fc x info z argTy scope) @@ -550,39 +559,51 @@ mapPTerm f = goPTerm where goPClaim : PClaimData' nm -> PClaimData' nm goPClaim (MkPClaim c v opts tdecl) = MkPClaim c v (goPFnOpt <$> opts) (mapFC goPTypeDecl tdecl) - goPDeclFC : PDecl' nm -> PDecl' nm + goPlainBinder : PlainBinder' nm -> PlainBinder' nm + goPlainBinder = mapWName goPTerm + + goPBinderScope : PBinderScope' nm -> PBinderScope' nm + goPBinderScope (MkPBinderScope binder scope) + = MkPBinderScope (goPBinder binder) (goPTerm scope) goPDecl : PDeclNoFC' nm -> PDeclNoFC' nm goPDecl (PClaim claim) = PClaim $ goPClaim claim - goPDecl (PDef cls) = PDef $ map goPClause cls + goPDecl (PDef cls) = PDef $ (map goPClause) cls goPDecl (PData doc v mbt d) = PData doc v mbt $ goPDataDecl d - goPDecl (PParameters (Left nts) ps) - = PParameters (Left $ map (map goPTerm) nts) (goPDeclFC <$> ps) - goPDecl (PParameters (Right nts) ps) - = PParameters (Right $ go4TupledPTerms nts) (goPDeclFC <$> ps) + goPDecl (PParameters nts ps) + = PParameters (bimap (map goPlainBinder) (map goPBinder) nts) (mapFC goPDecl <$> ps) goPDecl (PUsing mnts ps) - = PUsing (goPairedPTerms mnts) (goPDeclFC <$> ps) + = PUsing (goPairedPTerms mnts) (mapFC goPDecl <$> ps) goPDecl (PInterface v mnts n doc nrts ns mn ps) - = PInterface v (goPairedPTerms mnts) n doc (go3TupledPTerms nrts) ns mn (goPDeclFC <$> ps) + = PInterface v (goPairedPTerms mnts) n doc (goBasicMultiBinder <$> nrts) ns mn (mapFC goPDecl <$> ps) goPDecl (PImplementation v opts p is cs n ts mn ns mps) = PImplementation v opts p (goImplicits is) (goPairedPTerms cs) - n (goPTerm <$> ts) mn ns (map (goPDeclFC <$>) mps) + n (goPTerm <$> ts) mn ns (map (mapFC goPDecl <$>) mps) goPDecl (PRecord doc v tot (MkPRecord n nts opts mn fs)) = PRecord doc v tot - (MkPRecord n (go4TupledPTerms nts) opts mn (map (mapFC goRecordField) fs)) + (MkPRecord n (map goPBinder nts) opts mn (map (mapFC goRecordField) fs)) goPDecl (PRecord doc v tot (MkPRecordLater n nts)) - = PRecord doc v tot (MkPRecordLater n (go4TupledPTerms nts)) - goPDecl (PFail msg ps) = PFail msg $ goPDeclFC <$> ps - goPDecl (PMutual ps) = PMutual $ goPDeclFC <$> ps + = PRecord doc v tot (MkPRecordLater n (goPBinder <$> nts )) + goPDecl (PFail msg ps) = PFail msg $ mapFC goPDecl <$> ps + goPDecl (PMutual ps) = PMutual $ map (mapFC goPDecl) ps goPDecl (PFixity p) = PFixity p - goPDecl (PNamespace strs ps) = PNamespace strs $ goPDeclFC <$> ps + goPDecl (PNamespace strs ps) = PNamespace strs $ mapFC goPDecl <$> ps goPDecl (PTransform n a b) = PTransform n (goPTerm a) (goPTerm b) goPDecl (PRunElabDecl a) = PRunElabDecl $ goPTerm a goPDecl (PDirective d) = PDirective d goPDecl p@(PBuiltin _ _) = p - goPDeclFC = mapFC goPDecl + goPBinder : PBinder' nm -> PBinder' nm + goPBinder (MkPBinder info bind) = MkPBinder (goPiInfo info) (goBasicMultiBinder bind) + + goBasicMultiBinder : BasicMultiBinder' nm -> BasicMultiBinder' nm + goBasicMultiBinder (MkBasicMultiBinder rig names type) + = MkBasicMultiBinder rig names (goPTerm type) + + goBasicBinder : BasicBinder' nm -> BasicBinder' nm + goBasicBinder (MkBasicBinder rig name type) + = MkBasicBinder rig name (goPTerm type) goPTypeDecl : PTypeDeclData' nm -> PTypeDeclData' nm goPTypeDecl (MkPTy n d t) = MkPTy n d $ goPTerm t @@ -639,6 +660,8 @@ export substFC : FC -> PTerm' nm -> PTerm' nm substFC fc = mapPTerm $ \case PRef _ x => PRef fc x + NewPi x => NewPi (setFC fc x) + Forall x => Forall (setFC fc x) PPi _ x y z argTy retTy => PPi fc x y z argTy retTy PLam _ x y pat argTy scope => PLam fc x y pat argTy scope PLet _ x pat nTy nVal scope alts => PLet fc x pat nTy nVal scope alts @@ -665,7 +688,7 @@ substFC fc = mapPTerm $ \case PDotted _ x => PDotted fc x PImplicit _ => PImplicit fc PInfer _ => PInfer fc - POp _ ab nm r => POp fc ab nm r + POp _ ab nm r => POp fc (setFC fc ab) nm r PPrefixOp _ x y => PPrefixOp fc (setFC fc x) y PSectionL _ x y => PSectionL fc (setFC fc x) y PSectionR _ x y => PSectionR fc x (setFC fc y) diff --git a/src/Idris/Syntax/Views.idr b/src/Idris/Syntax/Views.idr index 53b3c3e6ed..a704adbe4d 100644 --- a/src/Idris/Syntax/Views.idr +++ b/src/Idris/Syntax/Views.idr @@ -27,7 +27,8 @@ 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 leftSide op r) = (PRef op.fc op.val.toName,) . (Explicit fc leftSide.val.getLhs ::) . (Explicit fc r ::) + go (POp fc leftSide op r) = + (PRef op.fc op.val.toName,) . (Explicit fc leftSide.val.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/src/Libraries/Text/PrettyPrint/Prettyprinter/Doc.idr b/src/Libraries/Text/PrettyPrint/Prettyprinter/Doc.idr index 61e39077bd..ac7cd7bbb0 100644 --- a/src/Libraries/Text/PrettyPrint/Prettyprinter/Doc.idr +++ b/src/Libraries/Text/PrettyPrint/Prettyprinter/Doc.idr @@ -418,6 +418,10 @@ export FromString (Doc ann) where fromString = pretty0 +export +commaSep : List (Doc ann) -> Doc ann +commaSep = concatWith (\x, y => x <+> pretty0 "," <++> y) + ||| Variant of `encloseSep` with braces and comma as separator. export list : List (Doc ann) -> Doc ann diff --git a/src/Parser/Rule/Source.idr b/src/Parser/Rule/Source.idr index 4a04ad5c6d..e3fe541427 100644 --- a/src/Parser/Rule/Source.idr +++ b/src/Parser/Rule/Source.idr @@ -261,6 +261,14 @@ pragma n = Pragma s => guard (s == n) _ => Nothing +export +cgDirective : Rule String +cgDirective = + terminal "Expected CG directive" $ + \case + CGDirective d => Just d + _ => Nothing + export builtinType : Rule BuiltinType builtinType = diff --git a/src/TTImp/Elab/Quote.idr b/src/TTImp/Elab/Quote.idr index 7588a740f1..ca67e82472 100644 --- a/src/TTImp/Elab/Quote.idr +++ b/src/TTImp/Elab/Quote.idr @@ -156,7 +156,7 @@ mutual = pure $ IDef fc v !(traverse getUnquoteClause d) getUnquoteDecl (IParameters fc ps ds) = pure $ IParameters fc - !(traverse unqTuple ps) + !(traverseList1 unqTuple ps) !(traverse getUnquoteDecl ds) where unqTuple : (Name, RigCount, PiInfo RawImp, RawImp) -> Core (Name, RigCount, PiInfo RawImp, RawImp) diff --git a/src/TTImp/ProcessDecls.idr b/src/TTImp/ProcessDecls.idr index 9caad56208..169db6a2d6 100644 --- a/src/TTImp/ProcessDecls.idr +++ b/src/TTImp/ProcessDecls.idr @@ -125,7 +125,7 @@ process eopts nest env (IData fc vis mbtot ddef) process eopts nest env (IDef fc fname def) = processDef eopts nest env fc fname def process eopts nest env (IParameters fc ps decls) - = processParams nest env fc ps decls + = processParams nest env fc (forget ps) decls process eopts nest env (IRecord fc ns vis mbtot rec) = processRecord eopts nest env ns vis mbtot rec process eopts nest env (IFail fc msg decls) diff --git a/src/TTImp/TTImp.idr b/src/TTImp/TTImp.idr index 056263a6f7..0a4443bf3e 100644 --- a/src/TTImp/TTImp.idr +++ b/src/TTImp/TTImp.idr @@ -10,7 +10,7 @@ import Core.TT import Core.Value import Data.List -import Data.List1 +import public Data.List1 import Data.Maybe import Libraries.Data.SortedSet @@ -476,7 +476,7 @@ mutual Maybe TotalReq -> ImpData' nm -> ImpDecl' nm IDef : FC -> Name -> List (ImpClause' nm) -> ImpDecl' nm IParameters : FC -> - List (ImpParameter' nm) -> + List1 (ImpParameter' nm) -> List (ImpDecl' nm) -> ImpDecl' nm IRecord : FC -> Maybe String -> -- nested namespace diff --git a/src/TTImp/Utils.idr b/src/TTImp/Utils.idr index e9e636f6bd..95c5bd593c 100644 --- a/src/TTImp/Utils.idr +++ b/src/TTImp/Utils.idr @@ -34,7 +34,7 @@ rawImpFromDecl decl = case decl of => maybe id (::) tycon $ map type datacons IData fc1 y _ (MkImpLater fc2 n tycon) => [tycon] IDef fc1 y ys => getFromClause !ys - IParameters fc1 ys zs => rawImpFromDecl !zs ++ map getParamTy ys + IParameters fc1 ys zs => rawImpFromDecl !zs ++ map getParamTy (forget ys) IRecord fc1 y z _ (MkImpRecord fc n params opts conName fields) => do (a, b) <- map (snd . snd) params getFromPiInfo a ++ [b] ++ getFromIField !fields diff --git a/tests/idris2/error/perror034/Error.idr b/tests/idris2/error/perror034/Error.idr new file mode 100644 index 0000000000..89633a7aca --- /dev/null +++ b/tests/idris2/error/perror034/Error.idr @@ -0,0 +1,5 @@ +unnamed : () +unnamed = (Nat) -> Nat + +named : () +named = (arg : Nat) -> Nat diff --git a/tests/idris2/error/perror034/expected b/tests/idris2/error/perror034/expected new file mode 100644 index 0000000000..7642d5fd70 --- /dev/null +++ b/tests/idris2/error/perror034/expected @@ -0,0 +1,26 @@ +1/1: Building Error (Error.idr) +Error: While processing right hand side of unnamed. When unifying: + Type +and: + () +Mismatch between: Type and (). + +Error:2:11--2:23 + 1 | unnamed : () + 2 | unnamed = (Nat) -> Nat + ^^^^^^^^^^^^ + +Error: While processing right hand side of named. When unifying: + Type +and: + () +Mismatch between: Type and (). + +Error:5:9--5:27 + 1 | unnamed : () + 2 | unnamed = (Nat) -> Nat + 3 | + 4 | named : () + 5 | named = (arg : Nat) -> Nat + ^^^^^^^^^^^^^^^^^^ + diff --git a/tests/idris2/error/perror034/run b/tests/idris2/error/perror034/run new file mode 100644 index 0000000000..b1c49389fc --- /dev/null +++ b/tests/idris2/error/perror034/run @@ -0,0 +1,3 @@ +. ../../../testutils.sh + +check Error.idr From 22fa54cca68093b6fb7139e1b23d0e56edf5e9f1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Fri, 20 Dec 2024 00:34:02 -0500 Subject: [PATCH 2/4] split up ideMode tests --- tests/ideMode/ideMode005/README.md | 17 ----- tests/ideMode/ideMode005/SimpleData.idr | 8 --- tests/ideMode/ideMode005/expected | 0 tests/ideMode/ideMode005/expected3 | 64 ------------------- tests/ideMode/ideMode005/input3 | 1 - tests/ideMode/ideMode005/regenerate | 24 ------- tests/ideMode/ideMode005/run | 60 ----------------- .../{ideMode005 => ideMode007}/Syntax.idr | 0 .../{ideMode005 => ideMode007}/dummy.ipkg | 0 .../expected1 => ideMode007/expected} | 2 +- .../{ideMode005/input1 => ideMode007/input} | 0 tests/ideMode/ideMode007/run | 3 + .../{ideMode005 => ideMode008}/Interface.idr | 0 tests/ideMode/ideMode008/dummy.ipkg | 1 + .../expected2 => ideMode008/expected} | 0 .../{ideMode005/input2 => ideMode008/input} | 0 tests/ideMode/ideMode008/run | 3 + .../{ideMode005 => ideMode009}/Ambiguity.idr | 0 .../expected4 => ideMode009/expected} | 0 .../{ideMode005/input4 => ideMode009/input} | 0 tests/ideMode/ideMode009/run | 3 + .../{ideMode005 => ideMode010}/Rainbow.idr | 0 tests/ideMode/ideMode010/dummy.ipkg | 1 + .../expected5 => ideMode010/expected} | 0 .../{ideMode005/input5 => ideMode010/input} | 0 tests/ideMode/ideMode010/run | 3 + .../Implementation.idr | 0 tests/ideMode/ideMode011/dummy.ipkg | 1 + .../expected6 => ideMode011/expected} | 0 .../{ideMode005/input6 => ideMode011/input} | 0 tests/ideMode/ideMode011/run | 3 + .../{ideMode005 => ideMode012}/Case.idr | 0 tests/ideMode/ideMode012/dummy.ipkg | 1 + .../expected7 => ideMode012/expected} | 0 .../{ideMode005/input7 => ideMode012/input} | 0 tests/ideMode/ideMode012/run | 3 + .../RecordUpdate.idr | 0 tests/ideMode/ideMode013/dummy.ipkg | 1 + .../expected8 => ideMode013/expected} | 0 .../{ideMode005/input8 => ideMode013/input} | 0 tests/ideMode/ideMode013/run | 3 + .../{ideMode005 => ideMode014}/With.idr | 0 tests/ideMode/ideMode014/dummy.ipkg | 1 + .../expected9 => ideMode014/expected} | 0 .../{ideMode005/input9 => ideMode014/input} | 0 tests/ideMode/ideMode014/run | 3 + .../{ideMode005 => ideMode015}/Ranges.idr | 0 tests/ideMode/ideMode015/dummy.ipkg | 1 + .../expectedA => ideMode015/expected} | 0 .../{ideMode005/inputA => ideMode015/input} | 0 tests/ideMode/ideMode015/run | 3 + .../StringLiterals.idr | 0 tests/ideMode/ideMode016/dummy.ipkg | 1 + .../expectedB => ideMode016/expected} | 0 .../{ideMode005/inputB => ideMode016/input} | 0 tests/ideMode/ideMode016/run | 3 + .../{ideMode005 => ideMode017}/WithApp.idr | 0 tests/ideMode/ideMode017/dummy.ipkg | 1 + .../expectedC => ideMode017/expected} | 0 .../{ideMode005/inputC => ideMode017/input} | 0 tests/ideMode/ideMode017/run | 3 + .../{ideMode005 => ideMode018}/Rewrite.idr | 0 tests/ideMode/ideMode018/dummy.ipkg | 1 + .../expectedD => ideMode018/expected} | 0 .../{ideMode005/inputD => ideMode018/input} | 0 tests/ideMode/ideMode018/run | 3 + .../RecordProjections.idr | 0 tests/ideMode/ideMode019/dummy.ipkg | 1 + .../expectedE => ideMode019/expected} | 0 .../{ideMode005/inputE => ideMode019/input} | 0 tests/ideMode/ideMode019/run | 3 + .../{ideMode005 => ideMode020}/Tuples.idr | 0 tests/ideMode/ideMode020/dummy.ipkg | 1 + .../expectedF => ideMode020/expected} | 0 .../{ideMode005/inputF => ideMode020/input} | 0 tests/ideMode/ideMode020/run | 3 + .../{ideMode005 => ideMode021}/LetBinders.idr | 0 tests/ideMode/ideMode021/dummy.ipkg | 1 + .../expectedG => ideMode021/expected} | 0 .../{ideMode005/inputG => ideMode021/input} | 0 tests/ideMode/ideMode021/run | 3 + .../SnocRainbow.idr | 0 tests/ideMode/ideMode022/dummy.ipkg | 1 + .../expectedH => ideMode022/expected} | 0 .../{ideMode005/inputH => ideMode022/input} | 0 tests/ideMode/ideMode022/run | 3 + .../{ideMode005 => ideMode023}/Fail.idr | 0 tests/ideMode/ideMode023/dummy.ipkg | 1 + .../expectedI => ideMode023/expected} | 0 .../{ideMode005/inputI => ideMode023/input} | 0 tests/ideMode/ideMode023/run | 3 + .../{ideMode005 => ideMode024}/Holes.idr | 0 tests/ideMode/ideMode024/dummy.ipkg | 1 + .../expectedJ => ideMode024/expected} | 0 .../{ideMode005/inputJ => ideMode024/input} | 0 tests/ideMode/ideMode024/run | 3 + 96 files changed, 71 insertions(+), 175 deletions(-) delete mode 100644 tests/ideMode/ideMode005/README.md delete mode 100644 tests/ideMode/ideMode005/SimpleData.idr delete mode 100644 tests/ideMode/ideMode005/expected delete mode 100644 tests/ideMode/ideMode005/expected3 delete mode 100644 tests/ideMode/ideMode005/input3 delete mode 100755 tests/ideMode/ideMode005/regenerate delete mode 100755 tests/ideMode/ideMode005/run rename tests/ideMode/{ideMode005 => ideMode007}/Syntax.idr (100%) rename tests/ideMode/{ideMode005 => ideMode007}/dummy.ipkg (100%) rename tests/ideMode/{ideMode005/expected1 => ideMode007/expected} (99%) rename tests/ideMode/{ideMode005/input1 => ideMode007/input} (100%) create mode 100755 tests/ideMode/ideMode007/run rename tests/ideMode/{ideMode005 => ideMode008}/Interface.idr (100%) create mode 100644 tests/ideMode/ideMode008/dummy.ipkg rename tests/ideMode/{ideMode005/expected2 => ideMode008/expected} (100%) rename tests/ideMode/{ideMode005/input2 => ideMode008/input} (100%) create mode 100755 tests/ideMode/ideMode008/run rename tests/ideMode/{ideMode005 => ideMode009}/Ambiguity.idr (100%) rename tests/ideMode/{ideMode005/expected4 => ideMode009/expected} (100%) rename tests/ideMode/{ideMode005/input4 => ideMode009/input} (100%) create mode 100755 tests/ideMode/ideMode009/run rename tests/ideMode/{ideMode005 => ideMode010}/Rainbow.idr (100%) create mode 100644 tests/ideMode/ideMode010/dummy.ipkg rename tests/ideMode/{ideMode005/expected5 => ideMode010/expected} (100%) rename tests/ideMode/{ideMode005/input5 => ideMode010/input} (100%) create mode 100755 tests/ideMode/ideMode010/run rename tests/ideMode/{ideMode005 => ideMode011}/Implementation.idr (100%) create mode 100644 tests/ideMode/ideMode011/dummy.ipkg rename tests/ideMode/{ideMode005/expected6 => ideMode011/expected} (100%) rename tests/ideMode/{ideMode005/input6 => ideMode011/input} (100%) create mode 100755 tests/ideMode/ideMode011/run rename tests/ideMode/{ideMode005 => ideMode012}/Case.idr (100%) create mode 100644 tests/ideMode/ideMode012/dummy.ipkg rename tests/ideMode/{ideMode005/expected7 => ideMode012/expected} (100%) rename tests/ideMode/{ideMode005/input7 => ideMode012/input} (100%) create mode 100755 tests/ideMode/ideMode012/run rename tests/ideMode/{ideMode005 => ideMode013}/RecordUpdate.idr (100%) create mode 100644 tests/ideMode/ideMode013/dummy.ipkg rename tests/ideMode/{ideMode005/expected8 => ideMode013/expected} (100%) rename tests/ideMode/{ideMode005/input8 => ideMode013/input} (100%) create mode 100755 tests/ideMode/ideMode013/run rename tests/ideMode/{ideMode005 => ideMode014}/With.idr (100%) create mode 100644 tests/ideMode/ideMode014/dummy.ipkg rename tests/ideMode/{ideMode005/expected9 => ideMode014/expected} (100%) rename tests/ideMode/{ideMode005/input9 => ideMode014/input} (100%) create mode 100755 tests/ideMode/ideMode014/run rename tests/ideMode/{ideMode005 => ideMode015}/Ranges.idr (100%) create mode 100644 tests/ideMode/ideMode015/dummy.ipkg rename tests/ideMode/{ideMode005/expectedA => ideMode015/expected} (100%) rename tests/ideMode/{ideMode005/inputA => ideMode015/input} (100%) create mode 100755 tests/ideMode/ideMode015/run rename tests/ideMode/{ideMode005 => ideMode016}/StringLiterals.idr (100%) create mode 100644 tests/ideMode/ideMode016/dummy.ipkg rename tests/ideMode/{ideMode005/expectedB => ideMode016/expected} (100%) rename tests/ideMode/{ideMode005/inputB => ideMode016/input} (100%) create mode 100755 tests/ideMode/ideMode016/run rename tests/ideMode/{ideMode005 => ideMode017}/WithApp.idr (100%) create mode 100644 tests/ideMode/ideMode017/dummy.ipkg rename tests/ideMode/{ideMode005/expectedC => ideMode017/expected} (100%) rename tests/ideMode/{ideMode005/inputC => ideMode017/input} (100%) create mode 100755 tests/ideMode/ideMode017/run rename tests/ideMode/{ideMode005 => ideMode018}/Rewrite.idr (100%) create mode 100644 tests/ideMode/ideMode018/dummy.ipkg rename tests/ideMode/{ideMode005/expectedD => ideMode018/expected} (100%) rename tests/ideMode/{ideMode005/inputD => ideMode018/input} (100%) create mode 100755 tests/ideMode/ideMode018/run rename tests/ideMode/{ideMode005 => ideMode019}/RecordProjections.idr (100%) create mode 100644 tests/ideMode/ideMode019/dummy.ipkg rename tests/ideMode/{ideMode005/expectedE => ideMode019/expected} (100%) rename tests/ideMode/{ideMode005/inputE => ideMode019/input} (100%) create mode 100755 tests/ideMode/ideMode019/run rename tests/ideMode/{ideMode005 => ideMode020}/Tuples.idr (100%) create mode 100644 tests/ideMode/ideMode020/dummy.ipkg rename tests/ideMode/{ideMode005/expectedF => ideMode020/expected} (100%) rename tests/ideMode/{ideMode005/inputF => ideMode020/input} (100%) create mode 100755 tests/ideMode/ideMode020/run rename tests/ideMode/{ideMode005 => ideMode021}/LetBinders.idr (100%) create mode 100644 tests/ideMode/ideMode021/dummy.ipkg rename tests/ideMode/{ideMode005/expectedG => ideMode021/expected} (100%) rename tests/ideMode/{ideMode005/inputG => ideMode021/input} (100%) create mode 100755 tests/ideMode/ideMode021/run rename tests/ideMode/{ideMode005 => ideMode022}/SnocRainbow.idr (100%) create mode 100644 tests/ideMode/ideMode022/dummy.ipkg rename tests/ideMode/{ideMode005/expectedH => ideMode022/expected} (100%) rename tests/ideMode/{ideMode005/inputH => ideMode022/input} (100%) create mode 100755 tests/ideMode/ideMode022/run rename tests/ideMode/{ideMode005 => ideMode023}/Fail.idr (100%) create mode 100644 tests/ideMode/ideMode023/dummy.ipkg rename tests/ideMode/{ideMode005/expectedI => ideMode023/expected} (100%) rename tests/ideMode/{ideMode005/inputI => ideMode023/input} (100%) create mode 100755 tests/ideMode/ideMode023/run rename tests/ideMode/{ideMode005 => ideMode024}/Holes.idr (100%) create mode 100644 tests/ideMode/ideMode024/dummy.ipkg rename tests/ideMode/{ideMode005/expectedJ => ideMode024/expected} (100%) rename tests/ideMode/{ideMode005/inputJ => ideMode024/input} (100%) create mode 100755 tests/ideMode/ideMode024/run diff --git a/tests/ideMode/ideMode005/README.md b/tests/ideMode/ideMode005/README.md deleted file mode 100644 index 640fab8cd6..0000000000 --- a/tests/ideMode/ideMode005/README.md +++ /dev/null @@ -1,17 +0,0 @@ -README -====== - -Structure of this test ----------------------- - -This test case is a bit special because of the big output generated by the -various test cases. Instead of having a single pair of `expected` and `output` -files, we use one pair per test file. - -The diffs for each pair is then appended to the main `output` file. So the -`expected` file is empty and should stay empty. - -How to fix this test? ---------------------- - -Run `./regenerate PATH/TO/IDRIS2` diff --git a/tests/ideMode/ideMode005/SimpleData.idr b/tests/ideMode/ideMode005/SimpleData.idr deleted file mode 100644 index f5fffe297f..0000000000 --- a/tests/ideMode/ideMode005/SimpleData.idr +++ /dev/null @@ -1,8 +0,0 @@ -module SimpleData - -data Fwd a = FNil | (<|) a (Fwd a) -data Bwd a = BNil | (|>) (Bwd a) a - -data Tree n l - = Leaf l - | Node (Tree n l) n (Tree n l) diff --git a/tests/ideMode/ideMode005/expected b/tests/ideMode/ideMode005/expected deleted file mode 100644 index e69de29bb2..0000000000 diff --git a/tests/ideMode/ideMode005/expected3 b/tests/ideMode/ideMode005/expected3 deleted file mode 100644 index 5f61a558cc..0000000000 --- a/tests/ideMode/ideMode005/expected3 +++ /dev/null @@ -1,64 +0,0 @@ -000018(:protocol-version 2 1) -00003e(:write-string "1/1: Building SimpleData (SimpleData.idr)" 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 0 7) (:end 0 17)) ((:decor :module)))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 0) (:end 2 4)) ((:decor :keyword)))))) 1) -0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 5) (:end 2 8)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 5) (:end 2 8)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 5) (:end 2 8)) ((:decor :type)))))) 1) -0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 9) (:end 2 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 9) (:end 2 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 9) (:end 2 10)) ((:decor :bound)))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 11) (:end 2 12)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 13) (:end 2 17)) ((:decor :data)))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 18) (:end 2 19)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 20) (:end 2 24)) ((:decor :data)))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 25) (:end 2 26)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 27) (:end 2 28)) ((:decor :keyword)))))) 1) -0000d7(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 28) (:end 2 31)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 32) (:end 2 33)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 33) (:end 2 34)) ((:decor :keyword)))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 0) (:end 3 4)) ((:decor :keyword)))))) 1) -0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 5) (:end 3 8)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 5) (:end 3 8)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 5) (:end 3 8)) ((:decor :type)))))) 1) -0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 9) (:end 3 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 9) (:end 3 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 9) (:end 3 10)) ((:decor :bound)))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 11) (:end 3 12)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 13) (:end 3 17)) ((:decor :data)))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 18) (:end 3 19)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 20) (:end 3 24)) ((:decor :data)))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 25) (:end 3 26)) ((:decor :keyword)))))) 1) -0000d7(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 26) (:end 3 29)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 30) (:end 3 31)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 31) (:end 3 32)) ((:decor :keyword)))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 33) (:end 3 34)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 0) (:end 5 4)) ((:decor :keyword)))))) 1) -0000d6(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 5) (:end 5 9)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000d6(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 5) (:end 5 9)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 5) (:end 5 9)) ((:decor :type)))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 10) (:end 5 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 10) (:end 5 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 10) (:end 5 11)) ((:decor :bound)))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 12) (:end 5 13)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 12) (:end 5 13)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 12) (:end 5 13)) ((:decor :bound)))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 2) (:end 6 3)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 4) (:end 6 8)) ((:decor :data)))))) 1) -0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 9) (:end 6 10)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 2) (:end 7 3)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 4) (:end 7 8)) ((:decor :data)))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 9) (:end 7 10)) ((:decor :keyword)))))) 1) -0000d8(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 10) (:end 7 14)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 15) (:end 7 16)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 17) (:end 7 18)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 18) (:end 7 19)) ((:decor :keyword)))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 20) (:end 7 21)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 22) (:end 7 23)) ((:decor :keyword)))))) 1) -0000d8(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 23) (:end 7 27)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 28) (:end 7 29)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 30) (:end 7 31)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 31) (:end 7 32)) ((:decor :keyword)))))) 1) -000015(:return (:ok ()) 1) -Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/input3 b/tests/ideMode/ideMode005/input3 deleted file mode 100644 index 2fd7d3b40c..0000000000 --- a/tests/ideMode/ideMode005/input3 +++ /dev/null @@ -1 +0,0 @@ -000021((:load-file "SimpleData.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/regenerate b/tests/ideMode/ideMode005/regenerate deleted file mode 100755 index a1dac444cf..0000000000 --- a/tests/ideMode/ideMode005/regenerate +++ /dev/null @@ -1,24 +0,0 @@ -. ../../testutils.sh - -idris2 --ide-mode < input1 > expected1 -idris2 --ide-mode < input2 > expected2 -idris2 --ide-mode < input3 > expected3 -idris2 --ide-mode < input4 > expected4 -idris2 --ide-mode < input5 > expected5 -idris2 --ide-mode < input6 > expected6 -idris2 --ide-mode < input7 > expected7 -idris2 --ide-mode < input8 > expected8 -idris2 --ide-mode < input9 > expected9 -idris2 --ide-mode < inputA > expectedA -idris2 --ide-mode < inputB > expectedB -idris2 --ide-mode < inputC > expectedC -idris2 --ide-mode < inputD > expectedD -idris2 --ide-mode < inputE > expectedE -idris2 --ide-mode < inputF > expectedF -idris2 --ide-mode < inputG > expectedG -idris2 --ide-mode < inputH > expectedH -idris2 --ide-mode < inputI > expectedI -idris2 --ide-mode < inputJ > expectedJ - -rm -f expected output* -touch expected diff --git a/tests/ideMode/ideMode005/run b/tests/ideMode/ideMode005/run deleted file mode 100755 index 3cb5147f28..0000000000 --- a/tests/ideMode/ideMode005/run +++ /dev/null @@ -1,60 +0,0 @@ -. ../../testutils.sh - -rm -f output* - -idris2 --ide-mode < input1 > output1 -diff expected1 output1 >> output - -idris2 --ide-mode < input2 > output2 -diff expected2 output2 >> output - -idris2 --ide-mode < input3 > output3 -diff expected3 output3 >> output - -idris2 --ide-mode < input4 > output4 -diff expected4 output4 >> output - -idris2 --ide-mode < input5 > output5 -diff expected5 output5 >> output - -idris2 --ide-mode < input6 > output6 -diff expected6 output6 >> output - -idris2 --ide-mode < input7 > output7 -diff expected7 output7 >> output - -idris2 --ide-mode < input8 > output8 -diff expected8 output8 >> output - -idris2 --ide-mode < input9 > output9 -diff expected9 output9 >> output - -idris2 --ide-mode < inputA > outputA -diff expectedA outputA >> output - -idris2 --ide-mode < inputB > outputB -diff expectedB outputB >> output - -idris2 --ide-mode < inputC > outputC -diff expectedC outputC >> output - -idris2 --ide-mode < inputD > outputD -diff expectedD outputD >> output - -idris2 --ide-mode < inputE > outputE -diff expectedE outputE >> output - -idris2 --ide-mode < inputF > outputF -diff expectedF outputF >> output - -idris2 --ide-mode < inputG > outputG -diff expectedG outputG >> output - -idris2 --ide-mode < inputH > outputH -diff expectedH outputH >> output - -idris2 --ide-mode < inputI > outputI -diff expectedI outputI >> output - -idris2 --ide-mode < inputJ > outputJ -diff expectedJ outputJ >> output diff --git a/tests/ideMode/ideMode005/Syntax.idr b/tests/ideMode/ideMode007/Syntax.idr similarity index 100% rename from tests/ideMode/ideMode005/Syntax.idr rename to tests/ideMode/ideMode007/Syntax.idr diff --git a/tests/ideMode/ideMode005/dummy.ipkg b/tests/ideMode/ideMode007/dummy.ipkg similarity index 100% rename from tests/ideMode/ideMode005/dummy.ipkg rename to tests/ideMode/ideMode007/dummy.ipkg diff --git a/tests/ideMode/ideMode005/expected1 b/tests/ideMode/ideMode007/expected similarity index 99% rename from tests/ideMode/ideMode005/expected1 rename to tests/ideMode/ideMode007/expected index 1aea867122..91dbbe6913 100644 --- a/tests/ideMode/ideMode005/expected1 +++ b/tests/ideMode/ideMode007/expected @@ -306,4 +306,4 @@ 000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 30) (:end 46 31)) ((:decor :keyword)))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 8) (:end 47 15)) ((:name "doBlock") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000015(:return (:ok ()) 1) -Alas the file is done, aborting +Alas the file is done, aborting \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input1 b/tests/ideMode/ideMode007/input similarity index 100% rename from tests/ideMode/ideMode005/input1 rename to tests/ideMode/ideMode007/input diff --git a/tests/ideMode/ideMode007/run b/tests/ideMode/ideMode007/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode007/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Interface.idr b/tests/ideMode/ideMode008/Interface.idr similarity index 100% rename from tests/ideMode/ideMode005/Interface.idr rename to tests/ideMode/ideMode008/Interface.idr diff --git a/tests/ideMode/ideMode008/dummy.ipkg b/tests/ideMode/ideMode008/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode008/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expected2 b/tests/ideMode/ideMode008/expected similarity index 100% rename from tests/ideMode/ideMode005/expected2 rename to tests/ideMode/ideMode008/expected diff --git a/tests/ideMode/ideMode005/input2 b/tests/ideMode/ideMode008/input similarity index 100% rename from tests/ideMode/ideMode005/input2 rename to tests/ideMode/ideMode008/input diff --git a/tests/ideMode/ideMode008/run b/tests/ideMode/ideMode008/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode008/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Ambiguity.idr b/tests/ideMode/ideMode009/Ambiguity.idr similarity index 100% rename from tests/ideMode/ideMode005/Ambiguity.idr rename to tests/ideMode/ideMode009/Ambiguity.idr diff --git a/tests/ideMode/ideMode005/expected4 b/tests/ideMode/ideMode009/expected similarity index 100% rename from tests/ideMode/ideMode005/expected4 rename to tests/ideMode/ideMode009/expected diff --git a/tests/ideMode/ideMode005/input4 b/tests/ideMode/ideMode009/input similarity index 100% rename from tests/ideMode/ideMode005/input4 rename to tests/ideMode/ideMode009/input diff --git a/tests/ideMode/ideMode009/run b/tests/ideMode/ideMode009/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode009/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Rainbow.idr b/tests/ideMode/ideMode010/Rainbow.idr similarity index 100% rename from tests/ideMode/ideMode005/Rainbow.idr rename to tests/ideMode/ideMode010/Rainbow.idr diff --git a/tests/ideMode/ideMode010/dummy.ipkg b/tests/ideMode/ideMode010/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode010/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expected5 b/tests/ideMode/ideMode010/expected similarity index 100% rename from tests/ideMode/ideMode005/expected5 rename to tests/ideMode/ideMode010/expected diff --git a/tests/ideMode/ideMode005/input5 b/tests/ideMode/ideMode010/input similarity index 100% rename from tests/ideMode/ideMode005/input5 rename to tests/ideMode/ideMode010/input diff --git a/tests/ideMode/ideMode010/run b/tests/ideMode/ideMode010/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode010/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Implementation.idr b/tests/ideMode/ideMode011/Implementation.idr similarity index 100% rename from tests/ideMode/ideMode005/Implementation.idr rename to tests/ideMode/ideMode011/Implementation.idr diff --git a/tests/ideMode/ideMode011/dummy.ipkg b/tests/ideMode/ideMode011/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode011/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expected6 b/tests/ideMode/ideMode011/expected similarity index 100% rename from tests/ideMode/ideMode005/expected6 rename to tests/ideMode/ideMode011/expected diff --git a/tests/ideMode/ideMode005/input6 b/tests/ideMode/ideMode011/input similarity index 100% rename from tests/ideMode/ideMode005/input6 rename to tests/ideMode/ideMode011/input diff --git a/tests/ideMode/ideMode011/run b/tests/ideMode/ideMode011/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode011/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Case.idr b/tests/ideMode/ideMode012/Case.idr similarity index 100% rename from tests/ideMode/ideMode005/Case.idr rename to tests/ideMode/ideMode012/Case.idr diff --git a/tests/ideMode/ideMode012/dummy.ipkg b/tests/ideMode/ideMode012/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode012/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expected7 b/tests/ideMode/ideMode012/expected similarity index 100% rename from tests/ideMode/ideMode005/expected7 rename to tests/ideMode/ideMode012/expected diff --git a/tests/ideMode/ideMode005/input7 b/tests/ideMode/ideMode012/input similarity index 100% rename from tests/ideMode/ideMode005/input7 rename to tests/ideMode/ideMode012/input diff --git a/tests/ideMode/ideMode012/run b/tests/ideMode/ideMode012/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode012/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/RecordUpdate.idr b/tests/ideMode/ideMode013/RecordUpdate.idr similarity index 100% rename from tests/ideMode/ideMode005/RecordUpdate.idr rename to tests/ideMode/ideMode013/RecordUpdate.idr diff --git a/tests/ideMode/ideMode013/dummy.ipkg b/tests/ideMode/ideMode013/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode013/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expected8 b/tests/ideMode/ideMode013/expected similarity index 100% rename from tests/ideMode/ideMode005/expected8 rename to tests/ideMode/ideMode013/expected diff --git a/tests/ideMode/ideMode005/input8 b/tests/ideMode/ideMode013/input similarity index 100% rename from tests/ideMode/ideMode005/input8 rename to tests/ideMode/ideMode013/input diff --git a/tests/ideMode/ideMode013/run b/tests/ideMode/ideMode013/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode013/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/With.idr b/tests/ideMode/ideMode014/With.idr similarity index 100% rename from tests/ideMode/ideMode005/With.idr rename to tests/ideMode/ideMode014/With.idr diff --git a/tests/ideMode/ideMode014/dummy.ipkg b/tests/ideMode/ideMode014/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode014/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expected9 b/tests/ideMode/ideMode014/expected similarity index 100% rename from tests/ideMode/ideMode005/expected9 rename to tests/ideMode/ideMode014/expected diff --git a/tests/ideMode/ideMode005/input9 b/tests/ideMode/ideMode014/input similarity index 100% rename from tests/ideMode/ideMode005/input9 rename to tests/ideMode/ideMode014/input diff --git a/tests/ideMode/ideMode014/run b/tests/ideMode/ideMode014/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode014/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Ranges.idr b/tests/ideMode/ideMode015/Ranges.idr similarity index 100% rename from tests/ideMode/ideMode005/Ranges.idr rename to tests/ideMode/ideMode015/Ranges.idr diff --git a/tests/ideMode/ideMode015/dummy.ipkg b/tests/ideMode/ideMode015/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode015/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedA b/tests/ideMode/ideMode015/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedA rename to tests/ideMode/ideMode015/expected diff --git a/tests/ideMode/ideMode005/inputA b/tests/ideMode/ideMode015/input similarity index 100% rename from tests/ideMode/ideMode005/inputA rename to tests/ideMode/ideMode015/input diff --git a/tests/ideMode/ideMode015/run b/tests/ideMode/ideMode015/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode015/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/StringLiterals.idr b/tests/ideMode/ideMode016/StringLiterals.idr similarity index 100% rename from tests/ideMode/ideMode005/StringLiterals.idr rename to tests/ideMode/ideMode016/StringLiterals.idr diff --git a/tests/ideMode/ideMode016/dummy.ipkg b/tests/ideMode/ideMode016/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode016/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedB b/tests/ideMode/ideMode016/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedB rename to tests/ideMode/ideMode016/expected diff --git a/tests/ideMode/ideMode005/inputB b/tests/ideMode/ideMode016/input similarity index 100% rename from tests/ideMode/ideMode005/inputB rename to tests/ideMode/ideMode016/input diff --git a/tests/ideMode/ideMode016/run b/tests/ideMode/ideMode016/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode016/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/WithApp.idr b/tests/ideMode/ideMode017/WithApp.idr similarity index 100% rename from tests/ideMode/ideMode005/WithApp.idr rename to tests/ideMode/ideMode017/WithApp.idr diff --git a/tests/ideMode/ideMode017/dummy.ipkg b/tests/ideMode/ideMode017/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode017/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedC b/tests/ideMode/ideMode017/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedC rename to tests/ideMode/ideMode017/expected diff --git a/tests/ideMode/ideMode005/inputC b/tests/ideMode/ideMode017/input similarity index 100% rename from tests/ideMode/ideMode005/inputC rename to tests/ideMode/ideMode017/input diff --git a/tests/ideMode/ideMode017/run b/tests/ideMode/ideMode017/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode017/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Rewrite.idr b/tests/ideMode/ideMode018/Rewrite.idr similarity index 100% rename from tests/ideMode/ideMode005/Rewrite.idr rename to tests/ideMode/ideMode018/Rewrite.idr diff --git a/tests/ideMode/ideMode018/dummy.ipkg b/tests/ideMode/ideMode018/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode018/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedD b/tests/ideMode/ideMode018/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedD rename to tests/ideMode/ideMode018/expected diff --git a/tests/ideMode/ideMode005/inputD b/tests/ideMode/ideMode018/input similarity index 100% rename from tests/ideMode/ideMode005/inputD rename to tests/ideMode/ideMode018/input diff --git a/tests/ideMode/ideMode018/run b/tests/ideMode/ideMode018/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode018/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/RecordProjections.idr b/tests/ideMode/ideMode019/RecordProjections.idr similarity index 100% rename from tests/ideMode/ideMode005/RecordProjections.idr rename to tests/ideMode/ideMode019/RecordProjections.idr diff --git a/tests/ideMode/ideMode019/dummy.ipkg b/tests/ideMode/ideMode019/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode019/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedE b/tests/ideMode/ideMode019/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedE rename to tests/ideMode/ideMode019/expected diff --git a/tests/ideMode/ideMode005/inputE b/tests/ideMode/ideMode019/input similarity index 100% rename from tests/ideMode/ideMode005/inputE rename to tests/ideMode/ideMode019/input diff --git a/tests/ideMode/ideMode019/run b/tests/ideMode/ideMode019/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode019/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Tuples.idr b/tests/ideMode/ideMode020/Tuples.idr similarity index 100% rename from tests/ideMode/ideMode005/Tuples.idr rename to tests/ideMode/ideMode020/Tuples.idr diff --git a/tests/ideMode/ideMode020/dummy.ipkg b/tests/ideMode/ideMode020/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode020/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedF b/tests/ideMode/ideMode020/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedF rename to tests/ideMode/ideMode020/expected diff --git a/tests/ideMode/ideMode005/inputF b/tests/ideMode/ideMode020/input similarity index 100% rename from tests/ideMode/ideMode005/inputF rename to tests/ideMode/ideMode020/input diff --git a/tests/ideMode/ideMode020/run b/tests/ideMode/ideMode020/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode020/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/LetBinders.idr b/tests/ideMode/ideMode021/LetBinders.idr similarity index 100% rename from tests/ideMode/ideMode005/LetBinders.idr rename to tests/ideMode/ideMode021/LetBinders.idr diff --git a/tests/ideMode/ideMode021/dummy.ipkg b/tests/ideMode/ideMode021/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode021/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedG b/tests/ideMode/ideMode021/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedG rename to tests/ideMode/ideMode021/expected diff --git a/tests/ideMode/ideMode005/inputG b/tests/ideMode/ideMode021/input similarity index 100% rename from tests/ideMode/ideMode005/inputG rename to tests/ideMode/ideMode021/input diff --git a/tests/ideMode/ideMode021/run b/tests/ideMode/ideMode021/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode021/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/SnocRainbow.idr b/tests/ideMode/ideMode022/SnocRainbow.idr similarity index 100% rename from tests/ideMode/ideMode005/SnocRainbow.idr rename to tests/ideMode/ideMode022/SnocRainbow.idr diff --git a/tests/ideMode/ideMode022/dummy.ipkg b/tests/ideMode/ideMode022/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode022/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedH b/tests/ideMode/ideMode022/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedH rename to tests/ideMode/ideMode022/expected diff --git a/tests/ideMode/ideMode005/inputH b/tests/ideMode/ideMode022/input similarity index 100% rename from tests/ideMode/ideMode005/inputH rename to tests/ideMode/ideMode022/input diff --git a/tests/ideMode/ideMode022/run b/tests/ideMode/ideMode022/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode022/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Fail.idr b/tests/ideMode/ideMode023/Fail.idr similarity index 100% rename from tests/ideMode/ideMode005/Fail.idr rename to tests/ideMode/ideMode023/Fail.idr diff --git a/tests/ideMode/ideMode023/dummy.ipkg b/tests/ideMode/ideMode023/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode023/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedI b/tests/ideMode/ideMode023/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedI rename to tests/ideMode/ideMode023/expected diff --git a/tests/ideMode/ideMode005/inputI b/tests/ideMode/ideMode023/input similarity index 100% rename from tests/ideMode/ideMode005/inputI rename to tests/ideMode/ideMode023/input diff --git a/tests/ideMode/ideMode023/run b/tests/ideMode/ideMode023/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode023/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input diff --git a/tests/ideMode/ideMode005/Holes.idr b/tests/ideMode/ideMode024/Holes.idr similarity index 100% rename from tests/ideMode/ideMode005/Holes.idr rename to tests/ideMode/ideMode024/Holes.idr diff --git a/tests/ideMode/ideMode024/dummy.ipkg b/tests/ideMode/ideMode024/dummy.ipkg new file mode 100644 index 0000000000..c31d6563c7 --- /dev/null +++ b/tests/ideMode/ideMode024/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expectedJ b/tests/ideMode/ideMode024/expected similarity index 100% rename from tests/ideMode/ideMode005/expectedJ rename to tests/ideMode/ideMode024/expected diff --git a/tests/ideMode/ideMode005/inputJ b/tests/ideMode/ideMode024/input similarity index 100% rename from tests/ideMode/ideMode005/inputJ rename to tests/ideMode/ideMode024/input diff --git a/tests/ideMode/ideMode024/run b/tests/ideMode/ideMode024/run new file mode 100755 index 0000000000..3d64bf912d --- /dev/null +++ b/tests/ideMode/ideMode024/run @@ -0,0 +1,3 @@ +. ../../testutils.sh + +idris2 --ide-mode < input From 778249b31f3777f779ba48c4b1b209aece10d67b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Fri, 20 Dec 2024 01:52:49 -0500 Subject: [PATCH 3/4] Carry names across semantic decorations * Attempt to remove stateful semantic decoration * Remove redundant semantic coloring functions --- src/Idris/Parser.idr | 47 +++++++++++-------------------- src/Parser/Rule/Source.idr | 4 ++- tests/ideMode/ideMode001/expected | 8 +++--- tests/ideMode/ideMode003/expected | 8 +++--- tests/ideMode/ideMode006/expected | 14 ++++----- tests/ideMode/ideMode007/expected | 24 ++++++++-------- tests/ideMode/ideMode008/expected | 14 ++++----- tests/ideMode/ideMode010/expected | 22 +++++++-------- tests/ideMode/ideMode011/expected | 18 ++++++------ tests/ideMode/ideMode012/expected | 6 ++-- tests/ideMode/ideMode013/expected | 32 ++++++++++----------- tests/ideMode/ideMode014/expected | 22 +++++++-------- tests/ideMode/ideMode015/expected | 4 +-- tests/ideMode/ideMode016/expected | 10 +++---- tests/ideMode/ideMode017/expected | 10 +++---- tests/ideMode/ideMode018/expected | 4 +-- tests/ideMode/ideMode019/expected | 14 ++++----- tests/ideMode/ideMode020/expected | 6 ++-- tests/ideMode/ideMode021/expected | 14 ++++----- tests/ideMode/ideMode022/expected | 22 +++++++-------- tests/ideMode/ideMode023/expected | 10 +++---- tests/ideMode/ideMode024/expected | 6 ++-- 22 files changed, 153 insertions(+), 166 deletions(-) diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 13b3a65218..5e5c73d37e 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -28,29 +28,13 @@ import Idris.Parser.Let fcBounds : OriginDesc => Rule a -> Rule (WithFC a) fcBounds a = (.withFC) <$> bounds a -decorate : OriginDesc -> Decoration -> Rule a -> Rule a +decorate : {a : Type} -> OriginDesc -> Decoration -> Rule a -> Rule a decorate fname decor rule = do res <- bounds rule actD (decorationFromBounded fname decor res) pure res.val -boundedNameDecoration : OriginDesc -> Decoration -> WithBounds Name -> ASemanticDecoration -boundedNameDecoration fname decor bstr = ((fname, start bstr, end bstr) - , decor - , Just bstr.val) - -decorateBoundedNames : OriginDesc -> Decoration -> List (WithBounds Name) -> EmptyRule () -decorateBoundedNames fname decor bns - = act $ MkState (cast (map (boundedNameDecoration fname decor) bns)) [] - -decorateBoundedName : OriginDesc -> Decoration -> WithBounds Name -> EmptyRule () -decorateBoundedName fname decor bn = actD (boundedNameDecoration fname decor bn) - -decorateKeywords : OriginDesc -> List (WithBounds a) -> EmptyRule () -decorateKeywords fname xs - = act $ MkState (cast (map (decorationFromBounded fname Keyword) xs)) [] - -dependentDecorate : OriginDesc -> Rule a -> (a -> Decoration) -> Rule a +dependentDecorate : {a : Type} -> OriginDesc -> Rule a -> (a -> Decoration) -> Rule a dependentDecorate fname rule decor = do res <- bounds rule actD (decorationFromBounded fname (decor res.val) res) @@ -59,6 +43,10 @@ dependentDecorate fname rule decor = do decoratedKeyword : OriginDesc -> String -> Rule () decoratedKeyword fname kwd = decorate fname Keyword (keyword kwd) +decorateKeywords : {a : Type} -> OriginDesc -> List (WithBounds a) -> EmptyRule () +decorateKeywords fname xs + = act $ MkState (cast (map (decorationFromBounded fname Keyword) xs)) [] + decoratedPragma : OriginDesc -> String -> Rule () decoratedPragma fname prg = decorate fname Keyword (pragma prg) @@ -87,11 +75,8 @@ decoratedDataTypeName fname = decorate fname Typ dataTypeName decoratedDataConstructorName : OriginDesc -> Rule Name decoratedDataConstructorName fname = decorate fname Data dataConstructorName -decoratedSimpleBinderName : OriginDesc -> Rule String -decoratedSimpleBinderName fname = decorate fname Bound unqualifiedName - decoratedSimpleBinderUName : OriginDesc -> Rule Name -decoratedSimpleBinderUName fname = UN . Basic <$> decorate fname Bound unqualifiedName +decoratedSimpleBinderUName fname = decorate fname Bound (UN . Basic <$> unqualifiedName) decoratedSimpleNamedArg : OriginDesc -> Rule String decoratedSimpleNamedArg fname @@ -403,7 +388,7 @@ mutual dpairType : OriginDesc -> WithBounds t -> IndentInfo -> Rule PTerm dpairType fname start indents - = do loc <- bounds (do x <- UN . Basic <$> decoratedSimpleBinderName fname + = do loc <- bounds (do x <- decoratedSimpleBinderUName fname decoratedSymbol fname ":" ty <- typeExpr pdef fname indents pure (x, ty)) @@ -598,7 +583,7 @@ mutual simplerExpr : OriginDesc -> IndentInfo -> Rule PTerm simplerExpr fname indents - = do b <- bounds (do x <- bounds (UN . Basic <$> decoratedSimpleBinderName fname) + = do b <- bounds (do x <- bounds (decoratedSimpleBinderUName fname) decoratedSymbol fname "@" commit expr <- simpleExpr fname indents @@ -687,16 +672,16 @@ mutual pibindListName fname indents = do rig <- multiplicity fname ns <- sepBy1 (decoratedSymbol fname ",") - (fcBounds $ UN <$> binderName) + (fcBounds binderName) decoratedSymbol fname ":" ty <- typeExpr pdef fname indents atEnd indents pure (MkBasicMultiBinder rig ns ty) where -- _ gets treated specially here, it means "I don't care about the name" - binderName : Rule UserName - binderName = Basic <$> decoratedSimpleBinderName fname - <|> symbol "_" $> Underscore + binderName : Rule Name + binderName = decoratedSimpleBinderUName fname + <|> decorate fname Bound (UN <$> symbol "_" $> Underscore) ||| The arrow used after an explicit binder ||| BNF: @@ -1241,7 +1226,7 @@ withProblem fname col indents wval <- bracketedExpr fname start indents prf <- optional $ do decoratedKeyword fname "proof" - pure (!(multiplicity fname), UN $ Basic !(decoratedSimpleBinderName fname)) + pure (!(multiplicity fname), !(decoratedSimpleBinderUName fname)) pure (MkPWithProblem rig wval prf) mutual @@ -1549,9 +1534,9 @@ directive <|> do decoratedPragma fname "name" n <- name ns <- sepBy1 (decoratedSymbol fname ",") - (decoratedSimpleBinderName fname) + (decoratedSimpleBinderUName fname) atEnd indents - pure (Names n (forget ns)) + pure (Names n (forget (map nameRoot ns))) <|> do decoratedPragma fname "start" e <- expr pdef fname indents atEnd indents diff --git a/src/Parser/Rule/Source.idr b/src/Parser/Rule/Source.idr index e3fe541427..5e34b41e24 100644 --- a/src/Parser/Rule/Source.idr +++ b/src/Parser/Rule/Source.idr @@ -107,7 +107,9 @@ documentation' = terminal "Expected documentation comment" $ _ => Nothing export -decorationFromBounded : OriginDesc -> Decoration -> WithBounds a -> ASemanticDecoration +decorationFromBounded : {a : Type} -> OriginDesc -> Decoration -> WithBounds a -> ASemanticDecoration +decorationFromBounded {a = Name} fname decor bnds + = ((fname, start bnds, end bnds), decor, Just bnds.val) decorationFromBounded fname decor bnds = ((fname, start bnds, end bnds), decor, Nothing) diff --git a/tests/ideMode/ideMode001/expected b/tests/ideMode/ideMode001/expected index 1c041e9265..9c71489afd 100644 --- a/tests/ideMode/ideMode001/expected +++ b/tests/ideMode/ideMode001/expected @@ -1,7 +1,7 @@ 000018(:protocol-version 2 1) 000038(:write-string "1/1: Building LocType (LocType.idr)" 1) 000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 0) (:end 0 4)) ((:decor :keyword)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 5) (:end 0 9)) ((:decor :type)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 5) (:end 0 9)) ((:name "Vect") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 10) (:end 0 11)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 12) (:end 0 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 16) (:end 0 18)) ((:decor :keyword)))))) 1) @@ -9,12 +9,12 @@ 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 24) (:end 0 26)) ((:decor :keyword)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 27) (:end 0 31)) ((:decor :type)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 32) (:end 0 37)) ((:decor :keyword)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 5) (:end 1 8)) ((:decor :data)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 5) (:end 1 8)) ((:name "Nil") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 9) (:end 1 10)) ((:decor :keyword)))))) 1) 0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 15)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 16) (:end 1 17)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 18) (:end 1 19)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 5) (:end 2 9)) ((:decor :data)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 5) (:end 2 9)) ((:name "::") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 11)) ((:decor :keyword)))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 12) (:end 2 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 14) (:end 2 16)) ((:decor :keyword)))))) 1) @@ -28,7 +28,7 @@ 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 37) (:end 2 38)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 38) (:end 2 39)) ((:decor :keyword)))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 40) (:end 2 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 0) (:end 4 6)) ((:decor :function)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 0) (:end 4 6)) ((:name "append") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1) 0000ce(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 9) (:end 4 13)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 14) (:end 4 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode003/expected b/tests/ideMode/ideMode003/expected index 08909bb39d..6e26c518ca 100644 --- a/tests/ideMode/ideMode003/expected +++ b/tests/ideMode/ideMode003/expected @@ -1,7 +1,7 @@ 000018(:protocol-version 2 1) 000038(:write-string "1/1: Building LocType (LocType.idr)" 1) 000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 0) (:end 0 4)) ((:decor :keyword)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 5) (:end 0 9)) ((:decor :type)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 5) (:end 0 9)) ((:name "Vect") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 10) (:end 0 11)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 12) (:end 0 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 16) (:end 0 18)) ((:decor :keyword)))))) 1) @@ -9,12 +9,12 @@ 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 24) (:end 0 26)) ((:decor :keyword)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 27) (:end 0 31)) ((:decor :type)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 32) (:end 0 37)) ((:decor :keyword)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 5) (:end 1 8)) ((:decor :data)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 5) (:end 1 8)) ((:name "Nil") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 9) (:end 1 10)) ((:decor :keyword)))))) 1) 0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 15)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 16) (:end 1 17)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 18) (:end 1 19)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 5) (:end 2 9)) ((:decor :data)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 5) (:end 2 9)) ((:name "::") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 11)) ((:decor :keyword)))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 12) (:end 2 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 14) (:end 2 16)) ((:decor :keyword)))))) 1) @@ -28,7 +28,7 @@ 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 37) (:end 2 38)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 38) (:end 2 39)) ((:decor :keyword)))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 40) (:end 2 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 0) (:end 4 6)) ((:decor :function)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 0) (:end 4 6)) ((:name "append") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1) 0000ce(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 9) (:end 4 13)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 14) (:end 4 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode006/expected b/tests/ideMode/ideMode006/expected index 4b4f064c90..b9e7ddf413 100644 --- a/tests/ideMode/ideMode006/expected +++ b/tests/ideMode/ideMode006/expected @@ -5,7 +5,7 @@ 000071(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 2 0) (:end 2 8)) ((:decor :keyword)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 2 9) (:end 2 14)) ((:decor :keyword)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 4 0) (:end 4 6)) ((:decor :keyword)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 4 7) (:end 4 13)) ((:decor :type)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 4 7) (:end 4 13)) ((:name "Exists") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 4 14) (:end 4 15)) ((:decor :keyword)))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 4 15) (:end 4 16)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 4 17) (:end 4 18)) ((:decor :keyword)))))) 1) @@ -15,17 +15,17 @@ 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 4 28) (:end 4 29)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 4 30) (:end 4 35)) ((:decor :keyword)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 5 2) (:end 5 13)) ((:decor :keyword)))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 5 14) (:end 5 22)) ((:decor :data)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 5 14) (:end 5 22)) ((:name "MkExists") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 6 2) (:end 6 3)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 6 4) (:end 6 5)) ((:decor :function)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 6 4) (:end 6 5)) ((:name "x") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 6 6) (:end 6 7)) ((:decor :keyword)))))) 1) 0000c6(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 6 8) (:end 6 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 7 2) (:end 7 5)) ((:decor :function)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 7 2) (:end 7 5)) ((:name "prf") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 7 6) (:end 7 7)) ((:decor :keyword)))))) 1) 0000c6(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 7 8) (:end 7 9)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 7 10) (:end 7 11)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 9 0) (:end 9 7)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 10 0) (:end 10 3)) ((:decor :function)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 10 0) (:end 10 3)) ((:name "lpo") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 10 4) (:end 10 5)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 10 6) (:end 10 7)) ((:decor :keyword)))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 10 7) (:end 10 8)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -58,7 +58,7 @@ 000075(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 10 78) (:end 10 79)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 10 79) (:end 10 80)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 12 0) (:end 12 7)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 13 0) (:end 13 3)) ((:decor :function)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 13 0) (:end 13 3)) ((:name "ext") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 13 4) (:end 13 5)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 13 6) (:end 13 7)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 13 7) (:end 13 8)) ((:decor :keyword)))))) 1) @@ -90,7 +90,7 @@ 0000ca(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 13 55) (:end 13 56)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 13 57) (:end 13 60)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000ca(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 13 61) (:end 13 62)) ((:name "g") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 15 0) (:end 15 7)) ((:decor :function)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 15 0) (:end 15 7)) ((:name "example") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 15 8) (:end 15 9)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 15 10) (:end 15 11)) ((:decor :keyword)))))) 1) 0000ca(:output (:ok (:highlight-source ((((:filename "Unsafe.idr") (:start 15 11) (:end 15 12)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode007/expected b/tests/ideMode/ideMode007/expected index 91dbbe6913..591edad928 100644 --- a/tests/ideMode/ideMode007/expected +++ b/tests/ideMode/ideMode007/expected @@ -4,7 +4,7 @@ 000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 0 7) (:end 0 13)) ((:decor :module)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 2 0) (:end 2 8)) ((:decor :keyword)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 2 9) (:end 2 14)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 0) (:end 4 9)) ((:decor :function)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 0) (:end 4 9)) ((:name "showMaybe") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 10) (:end 4 11)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 12) (:end 4 13)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 13) (:end 4 14)) ((:decor :keyword)))))) 1) @@ -25,7 +25,7 @@ 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 61) (:end 4 63)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 64) (:end 4 70)) ((:decor :type)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 0) (:end 5 9)) ((:name "showMaybe") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 10) (:end 5 11)) ((:decor :bound)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 10) (:end 5 11)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 11) (:end 5 12)) ((:decor :keyword)))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 12) (:end 5 14)) ((:name "ma") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 15) (:end 5 16)) ((:decor :keyword)))))) 1) @@ -48,7 +48,7 @@ 0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 22) (:end 7 24)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 25) (:end 7 29)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 30) (:end 7 31)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 0) (:end 9 4)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 0) (:end 9 4)) ((:name "nats") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 5) (:end 9 6)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 7) (:end 9 11)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 12) (:end 9 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -106,16 +106,16 @@ 0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 38) (:end 15 40)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 41) (:end 15 43)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 0) (:end 18 6)) ((:decor :keyword)))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 7) (:end 18 11)) ((:decor :type)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 7) (:end 18 11)) ((:name "ANat") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 12) (:end 18 17)) ((:decor :keyword)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 2) (:end 19 13)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 14) (:end 19 20)) ((:decor :data)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 2) (:end 20 6)) ((:decor :function)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 14) (:end 19 20)) ((:name "MkANat") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 2) (:end 20 6)) ((:name "aNat") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 7) (:end 20 8)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 9) (:end 20 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 9) (:end 20 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 9) (:end 20 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 0) (:end 22 7)) ((:decor :function)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 0) (:end 22 7)) ((:name "doBlock") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 8) (:end 22 9)) ((:decor :keyword)))))) 1) 0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 10) (:end 22 15)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d2(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 16) (:end 22 20)) ((:name "ANat") (:namespace "Syntax") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -190,7 +190,7 @@ 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 22) (:end 31 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 22) (:end 31 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 25) (:end 31 26)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 2) (:end 33 6)) ((:decor :function)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 2) (:end 33 6)) ((:name "add3") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 7) (:end 33 8)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 9) (:end 33 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 2) (:end 34 6)) ((:name "add3") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -217,7 +217,7 @@ 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 32) (:end 36 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 32) (:end 36 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 35) (:end 36 36)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 2) (:end 38 6)) ((:decor :function)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 2) (:end 38 6)) ((:name "add4") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 7) (:end 38 8)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 9) (:end 38 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 2) (:end 39 6)) ((:name "add4") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -229,11 +229,11 @@ 0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 17) (:end 39 18)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 19) (:end 39 20)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 21) (:end 39 22)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 0) (:end 41 7)) ((:decor :function)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 0) (:end 41 7)) ((:name "anonLam") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 8) (:end 41 9)) ((:decor :keyword)))))) 1) 0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 10) (:end 41 15)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 16) (:end 41 17)) ((:decor :type)))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 17) (:end 41 18)) ((:decor :bound)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 17) (:end 41 18)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 19) (:end 41 20)) ((:decor :keyword)))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 21) (:end 41 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 21) (:end 41 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -306,4 +306,4 @@ 000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 30) (:end 46 31)) ((:decor :keyword)))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 8) (:end 47 15)) ((:name "doBlock") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000015(:return (:ok ()) 1) -Alas the file is done, aborting \ No newline at end of file +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode008/expected b/tests/ideMode/ideMode008/expected index 043e9d100b..ffad34efda 100644 --- a/tests/ideMode/ideMode008/expected +++ b/tests/ideMode/ideMode008/expected @@ -8,29 +8,29 @@ 0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 15) (:end 2 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 15) (:end 2 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 17) (:end 2 19)) ((:decor :keyword)))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 20) (:end 2 26)) ((:decor :type)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 20) (:end 2 26)) ((:name "Pretty") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 27) (:end 2 28)) ((:decor :keyword)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 28) (:end 2 29)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 30) (:end 2 31)) ((:decor :bound)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 30) (:end 2 31)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 32) (:end 2 33)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 34) (:end 2 38)) ((:decor :type)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 38) (:end 2 39)) ((:decor :keyword)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 40) (:end 2 45)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 2) (:end 3 13)) ((:decor :keyword)))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 14) (:end 3 22)) ((:decor :data)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 14) (:end 3 22)) ((:name "MkPretty") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 2) (:end 5 3)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 4) (:end 5 7)) ((:name "Doc") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 4) (:end 5 7)) ((:decor :function)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 4) (:end 5 7)) ((:name "Doc") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 8) (:end 5 9)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 10) (:end 5 14)) ((:decor :type)))))) 1) 0000d9(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 2) (:end 6 7)) ((:name "toDoc") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 2) (:end 6 7)) ((:decor :function)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 2) (:end 6 7)) ((:name "toDoc") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 8) (:end 6 9)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 10) (:end 6 16)) ((:decor :type)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 17) (:end 6 19)) ((:decor :keyword)))))) 1) 0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 20) (:end 6 23)) ((:name "Doc") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000da(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 2) (:end 8 8)) ((:name "pretty") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 2) (:end 8 8)) ((:decor :function)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 2) (:end 8 8)) ((:name "pretty") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 9) (:end 8 10)) ((:decor :keyword)))))) 1) 0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 11) (:end 8 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 11) (:end 8 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -46,7 +46,7 @@ 0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 25) (:end 9 26)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 26) (:end 9 27)) ((:decor :keyword)))))) 1) 0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 2) (:end 11 9)) ((:name "prettys") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 2) (:end 11 9)) ((:decor :function)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 2) (:end 11 9)) ((:name "prettys") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 10) (:end 11 11)) ((:decor :keyword)))))) 1) 0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 12) (:end 11 16)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 12) (:end 11 16)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode010/expected b/tests/ideMode/ideMode010/expected index 15217351bc..d8e94f4d34 100644 --- a/tests/ideMode/ideMode010/expected +++ b/tests/ideMode/ideMode010/expected @@ -5,7 +5,7 @@ 000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 1 2) (:end 1 8)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 1 9) (:end 1 15)) ((:decor :keyword)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 2) (:end 2 6)) ((:decor :keyword)))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 7) (:end 2 12)) ((:decor :type)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 7) (:end 2 12)) ((:name "List0") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 13) (:end 2 14)) ((:decor :keyword)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 15) (:end 2 19)) ((:decor :type)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 5 0) (:end 5 9)) ((:decor :keyword)))))) 1) @@ -13,14 +13,14 @@ 000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 6 2) (:end 6 8)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 6 9) (:end 6 15)) ((:decor :keyword)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 2) (:end 7 6)) ((:decor :keyword)))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 7) (:end 7 12)) ((:decor :type)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 7) (:end 7 12)) ((:name "List1") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 13) (:end 7 14)) ((:decor :keyword)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 15) (:end 7 19)) ((:decor :type)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 20) (:end 7 25)) ((:decor :keyword)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 4) (:end 8 7)) ((:decor :data)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 4) (:end 8 7)) ((:name "Nil") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 8) (:end 8 9)) ((:decor :keyword)))))) 1) 0000d3(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 10) (:end 8 15)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 4) (:end 9 9)) ((:decor :data)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 4) (:end 9 9)) ((:name "Cons1") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 12) (:end 9 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 16) (:end 9 18)) ((:decor :keyword)))))) 1) @@ -29,7 +29,7 @@ 0000d3(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 28) (:end 9 33)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 11 2) (:end 11 8)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 11 9) (:end 11 15)) ((:decor :keyword)))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 2) (:end 12 6)) ((:decor :function)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 2) (:end 12 6)) ((:name "::") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 7) (:end 12 8)) ((:decor :keyword)))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 9) (:end 12 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 13) (:end 12 15)) ((:decor :keyword)))))) 1) @@ -44,7 +44,7 @@ 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 16 2) (:end 16 8)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 16 9) (:end 16 15)) ((:decor :keyword)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 2) (:end 17 6)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 7) (:end 17 11)) ((:decor :type)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 7) (:end 17 11)) ((:name "::") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 12) (:end 17 13)) ((:decor :keyword)))))) 1) 0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 14) (:end 17 17)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 18) (:end 17 20)) ((:decor :keyword)))))) 1) @@ -57,28 +57,28 @@ 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 20 2) (:end 20 8)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 20 9) (:end 20 15)) ((:decor :keyword)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 2) (:end 21 6)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 7) (:end 21 12)) ((:decor :type)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 7) (:end 21 12)) ((:name "List0") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 13) (:end 21 14)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 15) (:end 21 19)) ((:decor :type)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 20) (:end 21 25)) ((:decor :keyword)))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 4) (:end 22 7)) ((:decor :data)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 4) (:end 22 7)) ((:name "Nil") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 8) (:end 22 9)) ((:decor :keyword)))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 10) (:end 22 15)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 4) (:end 23 8)) ((:decor :data)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 4) (:end 23 8)) ((:name "::") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 9) (:end 23 10)) ((:decor :keyword)))))) 1) 0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 11) (:end 23 14)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 15) (:end 23 17)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 17) (:end 23 21)) ((:decor :type)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 22) (:end 23 24)) ((:decor :keyword)))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 25) (:end 23 30)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 25 0) (:end 25 1)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 25 0) (:end 25 1)) ((:name "m") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 25 2) (:end 25 3)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 25 4) (:end 25 7)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 0) (:end 26 1)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 2) (:end 26 3)) ((:decor :keyword)))))) 1) 0000de(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 4) (:end 26 14)) ((:name "believe_me") (:namespace "Builtin") (:decor :postulate) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 15) (:end 26 23)) ((:decor :data)))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 0) (:end 28 7)) ((:decor :function)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 0) (:end 28 7)) ((:name "Rainbow") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 8) (:end 28 9)) ((:decor :keyword)))))) 1) 0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 10) (:end 28 13)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 14) (:end 28 16)) ((:decor :keyword)))))) 1) diff --git a/tests/ideMode/ideMode011/expected b/tests/ideMode/ideMode011/expected index 6dc244115c..5587d9b0e0 100644 --- a/tests/ideMode/ideMode011/expected +++ b/tests/ideMode/ideMode011/expected @@ -8,27 +8,27 @@ 0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:decor :type)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:name "OneTwoThree") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:decor :bound)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 2) (:end 5 3)) ((:decor :keyword)))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 4) (:end 5 7)) ((:decor :data)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 4) (:end 5 7)) ((:name "One") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000ce(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 8) (:end 5 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 2) (:end 6 3)) ((:decor :keyword)))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 4) (:end 6 7)) ((:decor :data)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 4) (:end 6 7)) ((:name "Two") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000ce(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 8) (:end 6 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 10) (:end 6 11)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 2) (:end 7 3)) ((:decor :keyword)))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 4) (:end 7 9)) ((:decor :data)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 4) (:end 7 9)) ((:name "Three") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 10) (:end 7 11)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 12) (:end 7 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 14) (:end 7 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 0) (:end 9 1)) ((:decor :keyword)))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 1) (:end 9 4)) ((:decor :function)))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 1) (:end 9 4)) ((:name "OTT") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 4) (:end 9 5)) ((:decor :keyword)))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 6) (:end 9 13)) ((:decor :type)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 6) (:end 9 13)) ((:name "Functor") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e7(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 14) (:end 9 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e7(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 14) (:end 9 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e7(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 14) (:end 9 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -84,7 +84,7 @@ 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 5) (:end 15 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 5) (:end 15 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 7) (:end 15 9)) ((:decor :keyword)))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 10) (:end 15 14)) ((:decor :type)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 10) (:end 15 14)) ((:name "Show") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 15) (:end 15 16)) ((:decor :keyword)))))) 1) 0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 16) (:end 15 27)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 16) (:end 15 27)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -149,7 +149,7 @@ 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 3) (:end 20 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 3) (:end 20 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 5) (:end 20 7)) ((:decor :keyword)))))) 1) -000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 8) (:end 20 10)) ((:decor :type)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 8) (:end 20 10)) ((:name "Eq") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 11) (:end 20 12)) ((:decor :keyword)))))) 1) 0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 12) (:end 20 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 12) (:end 20 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode012/expected b/tests/ideMode/ideMode012/expected index 8ed4e76884..3041537c88 100644 --- a/tests/ideMode/ideMode012/expected +++ b/tests/ideMode/ideMode012/expected @@ -2,7 +2,7 @@ 000032(:write-string "1/1: Building Case (Case.idr)" 1) 00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1) 00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 0 7) (:end 0 11)) ((:decor :module)))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 0) (:end 2 6)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 0) (:end 2 6)) ((:name "listId") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 7) (:end 2 8)) ((:decor :keyword)))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 9) (:end 2 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 14) (:end 2 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -28,7 +28,7 @@ 0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 17) (:end 5 19)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d2(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 20) (:end 5 26)) ((:name "listId") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 27) (:end 5 29)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 0) (:end 7 7)) ((:decor :function)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 0) (:end 7 7)) ((:name "listRev") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 8) (:end 7 9)) ((:decor :keyword)))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 10) (:end 7 14)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 15) (:end 7 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -54,7 +54,7 @@ 0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 29) (:end 10 30)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 30) (:end 10 31)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 31) (:end 10 32)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 0) (:end 12 11)) ((:decor :function)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 0) (:end 12 11)) ((:name "listFilter2") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 12) (:end 12 13)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 14) (:end 12 15)) ((:decor :keyword)))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 15) (:end 12 16)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode013/expected b/tests/ideMode/ideMode013/expected index a6e61bb737..85ec0900bf 100644 --- a/tests/ideMode/ideMode013/expected +++ b/tests/ideMode/ideMode013/expected @@ -3,17 +3,17 @@ 000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 0 7) (:end 0 19)) ((:decor :module)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 2 0) (:end 2 6)) ((:decor :keyword)))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 2 7) (:end 2 17)) ((:decor :type)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 2 7) (:end 2 17)) ((:name "Attributes") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 2 18) (:end 2 23)) ((:decor :keyword)))))) 1) -000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 2) (:end 3 6)) ((:decor :function)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 2) (:end 3 6)) ((:name "font") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 7) (:end 3 8)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 9) (:end 3 15)) ((:decor :type)))))) 1) -000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 2) (:end 4 6)) ((:decor :function)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 2) (:end 4 6)) ((:name "size") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1) 0000db(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 9) (:end 4 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000db(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 9) (:end 4 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000db(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 9) (:end 4 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 0) (:end 6 7)) ((:decor :function)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 0) (:end 6 7)) ((:name "bigMono") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 8) (:end 6 9)) ((:decor :keyword)))))) 1) 0000e2(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 10) (:end 6 20)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 21) (:end 6 23)) ((:decor :keyword)))))) 1) @@ -21,18 +21,18 @@ 0000e1(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 0) (:end 7 7)) ((:name "bigMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 8) (:end 7 9)) ((:decor :keyword)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 10) (:end 7 11)) ((:decor :keyword)))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 12) (:end 7 16)) ((:decor :function)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 12) (:end 7 16)) ((:name "font") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 17) (:end 7 19)) ((:decor :keyword)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 20) (:end 7 21)) ((:decor :keyword)))))) 1) 0000e6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 21) (:end 7 23)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 24) (:end 7 30)) ((:decor :data)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 30) (:end 7 31)) ((:decor :keyword)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 10) (:end 8 11)) ((:decor :keyword)))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 12) (:end 8 16)) ((:decor :function)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 12) (:end 8 16)) ((:name "size") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 17) (:end 8 19)) ((:decor :keyword)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 20) (:end 8 22)) ((:decor :data)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 0) (:end 11 9)) ((:decor :function)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 0) (:end 11 9)) ((:name "smallMono") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 10) (:end 11 11)) ((:decor :keyword)))))) 1) 0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 12) (:end 11 22)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 23) (:end 11 25)) ((:decor :keyword)))))) 1) @@ -40,7 +40,7 @@ 0000e5(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 0) (:end 12 9)) ((:name "smallMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 10) (:end 12 11)) ((:decor :keyword)))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 12) (:end 12 13)) ((:decor :keyword)))))) 1) -00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 14) (:end 12 18)) ((:decor :function)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 14) (:end 12 18)) ((:name "size") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 19) (:end 12 21)) ((:decor :keyword)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 22) (:end 12 23)) ((:decor :data)))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 24) (:end 12 25)) ((:decor :keyword)))))) 1) @@ -48,17 +48,17 @@ 0000e5(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 28) (:end 12 35)) ((:name "bigMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 14 0) (:end 14 23)) ((:decor :comment)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 15 0) (:end 15 6)) ((:decor :keyword)))))) 1) -000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 15 7) (:end 15 11)) ((:decor :type)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 15 7) (:end 15 11)) ((:name "Text") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 15 12) (:end 15 17)) ((:decor :keyword)))))) 1) -00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 2) (:end 16 12)) ((:decor :function)))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 2) (:end 16 12)) ((:name "attributes") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 13) (:end 16 14)) ((:decor :keyword)))))) 1) 0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 15) (:end 16 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 15) (:end 16 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 15) (:end 16 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 2) (:end 17 9)) ((:decor :function)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 2) (:end 17 9)) ((:name "content") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 10) (:end 17 11)) ((:decor :keyword)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 12) (:end 17 18)) ((:decor :type)))))) 1) -00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 0) (:end 19 10)) ((:decor :function)))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 0) (:end 19 10)) ((:name "setArial10") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 11) (:end 19 12)) ((:decor :keyword)))))) 1) 0000de(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 13) (:end 19 17)) ((:name "Text") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 18) (:end 19 20)) ((:decor :keyword)))))) 1) @@ -66,14 +66,14 @@ 0000e7(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 0) (:end 20 10)) ((:name "setArial10") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 11) (:end 20 12)) ((:decor :keyword)))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 13) (:end 20 14)) ((:decor :keyword)))))) 1) -00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 15) (:end 20 25)) ((:decor :function)))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 15) (:end 20 25)) ((:name "attributes") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 25) (:end 20 27)) ((:decor :keyword)))))) 1) -00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 27) (:end 20 31)) ((:decor :function)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 27) (:end 20 31)) ((:name "font") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 32) (:end 20 34)) ((:decor :keyword)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 35) (:end 20 42)) ((:decor :data)))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 13) (:end 21 14)) ((:decor :keyword)))))) 1) -00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 15) (:end 21 25)) ((:decor :function)))))) 1) -00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 25) (:end 21 30)) ((:decor :function)))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 15) (:end 21 25)) ((:name "attributes") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 25) (:end 21 30)) ((:name "size") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 31) (:end 21 33)) ((:decor :keyword)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 34) (:end 21 36)) ((:decor :data)))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 22 13) (:end 22 14)) ((:decor :keyword)))))) 1) diff --git a/tests/ideMode/ideMode014/expected b/tests/ideMode/ideMode014/expected index 285e04eb3c..5a0cfaba85 100644 --- a/tests/ideMode/ideMode014/expected +++ b/tests/ideMode/ideMode014/expected @@ -2,7 +2,7 @@ 000032(:write-string "1/1: Building With (With.idr)" 1) 00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1) 00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 0 7) (:end 0 11)) ((:decor :module)))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 0) (:end 2 1)) ((:decor :function)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 0) (:end 2 1)) ((:name "f") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 2) (:end 2 3)) ((:decor :keyword)))))) 1) 00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 4) (:end 2 5)) ((:decor :keyword)))))) 1) 0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 5) (:end 2 6)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -11,12 +11,12 @@ 000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 12) (:end 2 13)) ((:decor :keyword)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 14) (:end 2 16)) ((:decor :keyword)))))) 1) 00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 17) (:end 2 18)) ((:decor :type)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 18) (:end 2 19)) ((:decor :bound)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 18) (:end 2 19)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 20) (:end 2 21)) ((:decor :keyword)))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 22) (:end 2 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 22) (:end 2 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 26) (:end 2 28)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 29) (:end 2 30)) ((:decor :bound)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 29) (:end 2 30)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 31) (:end 2 32)) ((:decor :keyword)))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 33) (:end 2 36)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 33) (:end 2 36)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -43,7 +43,7 @@ 0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 14) (:end 3 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 17) (:end 3 22)) ((:decor :keyword)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 23) (:end 3 25)) ((:decor :bound)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 23) (:end 3 25)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 2) (:end 4 3)) ((:name "f") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 4) (:end 4 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 4) (:end 4 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -72,7 +72,7 @@ 0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 26) (:end 5 28)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 29) (:end 5 32)) ((:name "sym") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 33) (:end 5 35)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 0) (:end 7 1)) ((:decor :function)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 0) (:end 7 1)) ((:name "g") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 2) (:end 7 3)) ((:decor :keyword)))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 4) (:end 7 8)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c5(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 9) (:end 7 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -110,7 +110,7 @@ 0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 16) (:end 10 20)) ((:name "asas") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 21) (:end 10 22)) ((:decor :keyword)))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 23) (:end 10 24)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 0) (:end 12 6)) ((:decor :function)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 0) (:end 12 6)) ((:name "nested") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 7) (:end 12 8)) ((:decor :keyword)))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 9) (:end 12 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 13) (:end 12 15)) ((:decor :keyword)))))) 1) @@ -258,13 +258,13 @@ 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 17) (:end 21 18)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 19) (:end 21 20)) ((:decor :data)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 0) (:end 23 4)) ((:decor :keyword)))))) 1) -00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 5) (:end 23 9)) ((:decor :type)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 5) (:end 23 9)) ((:name "ANat") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 10) (:end 23 11)) ((:decor :keyword)))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 12) (:end 23 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 16) (:end 23 18)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 19) (:end 23 23)) ((:decor :type)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 24) (:end 23 29)) ((:decor :keyword)))))) 1) -00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 2) (:end 24 8)) ((:decor :data)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 2) (:end 24 8)) ((:name "MkANat") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 9) (:end 24 10)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 11) (:end 24 12)) ((:decor :keyword)))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 12) (:end 24 13)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -274,7 +274,7 @@ 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 21) (:end 24 23)) ((:decor :keyword)))))) 1) 0000ce(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 24) (:end 24 28)) ((:name "ANat") (:namespace "With") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 29) (:end 24 30)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 0) (:end 26 8)) ((:decor :function)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 0) (:end 26 8)) ((:name "someNats") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 9) (:end 26 10)) ((:decor :keyword)))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 11) (:end 26 14)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 15) (:end 26 17)) ((:decor :keyword)))))) 1) @@ -293,7 +293,7 @@ 0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 11) (:end 28 12)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 11) (:end 28 12)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 13) (:end 28 14)) ((:decor :keyword)))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 15) (:end 28 16)) ((:decor :bound)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 15) (:end 28 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 16) (:end 28 17)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 17) (:end 28 18)) ((:decor :keyword)))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 18) (:end 28 24)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -311,7 +311,7 @@ 0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 13) (:end 29 14)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 13) (:end 29 14)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 15) (:end 29 16)) ((:decor :keyword)))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 17) (:end 29 18)) ((:decor :bound)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 17) (:end 29 18)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 18) (:end 29 19)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 19) (:end 29 20)) ((:decor :keyword)))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 20) (:end 29 26)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode015/expected b/tests/ideMode/ideMode015/expected index 9eb01ff246..958e6055e1 100644 --- a/tests/ideMode/ideMode015/expected +++ b/tests/ideMode/ideMode015/expected @@ -1,6 +1,6 @@ 000018(:protocol-version 2 1) 000036(:write-string "1/1: Building Ranges (Ranges.idr)" 1) -000072(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 0) (:end 0 5)) ((:decor :function)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 0) (:end 0 5)) ((:name "hours") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 6) (:end 0 7)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 8) (:end 0 12)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 13) (:end 0 16)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -11,7 +11,7 @@ 000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 10) (:end 1 12)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 12) (:end 1 14)) ((:decor :data)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 14) (:end 1 15)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 0) (:end 3 4)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 0) (:end 3 4)) ((:name "nats") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 5) (:end 3 6)) ((:decor :keyword)))))) 1) 0000df(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 7) (:end 3 13)) ((:name "Stream") (:namespace "Prelude.Types.Stream") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 14) (:end 3 17)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode016/expected b/tests/ideMode/ideMode016/expected index dca191e69c..b4bfbc76ef 100644 --- a/tests/ideMode/ideMode016/expected +++ b/tests/ideMode/ideMode016/expected @@ -2,7 +2,7 @@ 000046(:write-string "1/1: Building StringLiterals (StringLiterals.idr)" 1) 000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 0 7) (:end 0 21)) ((:decor :module)))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 0) (:end 2 5)) ((:decor :function)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 0) (:end 2 5)) ((:name "hello") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 6) (:end 2 7)) ((:decor :keyword)))))) 1) 0000de(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 8) (:end 2 18)) ((:name "FromString") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 19) (:end 2 20)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -11,7 +11,7 @@ 0000e3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 0) (:end 3 5)) ((:name "hello") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 6) (:end 3 7)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 8) (:end 3 15)) ((:decor :data)))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 0) (:end 5 9)) ((:decor :function)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 0) (:end 5 9)) ((:name "helloName") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 10) (:end 5 11)) ((:decor :keyword)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 12) (:end 5 18)) ((:decor :type)))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 19) (:end 5 21)) ((:decor :keyword)))))) 1) @@ -27,7 +27,7 @@ 000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 31) (:end 6 37)) ((:decor :type)))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 37) (:end 6 38)) ((:decor :keyword)))))) 1) 0000d3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 42) (:end 6 46)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 0) (:end 8 11)) ((:decor :function)))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 0) (:end 8 11)) ((:name "welcomeName") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 12) (:end 8 13)) ((:decor :keyword)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 14) (:end 8 20)) ((:decor :type)))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 21) (:end 8 23)) ((:decor :keyword)))))) 1) @@ -38,13 +38,13 @@ 000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 19) (:end 12 5)) ((:decor :data)))))) 1) 0000ea(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 10 4) (:end 10 13)) ((:name "helloName") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 10 14) (:end 10 18)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 14 0) (:end 14 11)) ((:decor :function)))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 14 0) (:end 14 11)) ((:name "scareQuotes") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 14 12) (:end 14 13)) ((:decor :keyword)))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 14 14) (:end 14 20)) ((:decor :type)))))) 1) 0000ec(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 0) (:end 15 11)) ((:name "scareQuotes") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 12) (:end 15 13)) ((:decor :keyword)))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 14) (:end 15 25)) ((:decor :data)))))) 1) -00007c(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 0) (:end 17 4)) ((:decor :function)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 0) (:end 17 4)) ((:name "test") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 5) (:end 17 6)) ((:decor :keyword)))))) 1) 0000ec(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 7) (:end 17 33)) ((:name "scareQuotes") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 34) (:end 17 35)) ((:decor :keyword)))))) 1) diff --git a/tests/ideMode/ideMode017/expected b/tests/ideMode/ideMode017/expected index 46454d2de9..a857df1178 100644 --- a/tests/ideMode/ideMode017/expected +++ b/tests/ideMode/ideMode017/expected @@ -1,17 +1,17 @@ 000018(:protocol-version 2 1) 000038(:write-string "1/1: Building WithApp (WithApp.idr)" 1) 000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 0) (:end 0 4)) ((:decor :keyword)))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 5) (:end 0 10)) ((:decor :type)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 5) (:end 0 10)) ((:name "Natty") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 11) (:end 0 12)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 13) (:end 0 16)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 17) (:end 0 19)) ((:decor :keyword)))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 20) (:end 0 24)) ((:decor :type)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 25) (:end 0 30)) ((:decor :keyword)))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 2) (:end 1 3)) ((:decor :data)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 2) (:end 1 3)) ((:name "Z") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 4) (:end 1 5)) ((:decor :keyword)))))) 1) 0000cf(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 6) (:end 1 11)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 12) (:end 1 13)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00006f(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 2) (:end 2 3)) ((:decor :data)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 2) (:end 2 3)) ((:name "S") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 4) (:end 2 5)) ((:decor :keyword)))))) 1) 0000cf(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 6) (:end 2 11)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 12) (:end 2 13)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -21,7 +21,7 @@ 0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 24) (:end 2 25)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 26) (:end 2 27)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 27) (:end 2 28)) ((:decor :keyword)))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 0) (:end 4 4)) ((:decor :function)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 0) (:end 4 4)) ((:name "view") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 5) (:end 4 6)) ((:decor :keyword)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1) 0000c7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 8) (:end 4 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -46,7 +46,7 @@ 0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 16) (:end 6 20)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 21) (:end 6 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 22) (:end 6 23)) ((:decor :keyword)))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 0) (:end 8 2)) ((:decor :function)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 0) (:end 8 2)) ((:name "id") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 3) (:end 8 4)) ((:decor :keyword)))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 5) (:end 8 8)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 9) (:end 8 11)) ((:decor :keyword)))))) 1) diff --git a/tests/ideMode/ideMode018/expected b/tests/ideMode/ideMode018/expected index 7a297532c5..9288590571 100644 --- a/tests/ideMode/ideMode018/expected +++ b/tests/ideMode/ideMode018/expected @@ -1,6 +1,6 @@ 000018(:protocol-version 2 1) 000038(:write-string "1/1: Building Rewrite (Rewrite.idr)" 1) -000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 0) (:end 0 9)) ((:decor :function)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 0) (:end 0 9)) ((:name "transport") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 10) (:end 0 11)) ((:decor :keyword)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 12) (:end 0 13)) ((:decor :keyword)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 13) (:end 0 14)) ((:decor :keyword)))))) 1) @@ -23,7 +23,7 @@ 0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 29) (:end 1 31)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 32) (:end 1 34)) ((:decor :keyword)))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 35) (:end 1 36)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 0) (:end 3 6)) ((:decor :function)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 0) (:end 3 6)) ((:name "nested") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 7) (:end 3 8)) ((:decor :keyword)))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 9) (:end 3 10)) ((:decor :keyword)))))) 1) 000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 10) (:end 3 11)) ((:decor :keyword)))))) 1) diff --git a/tests/ideMode/ideMode019/expected b/tests/ideMode/ideMode019/expected index e6f7dc9d8c..c3f6c2314d 100644 --- a/tests/ideMode/ideMode019/expected +++ b/tests/ideMode/ideMode019/expected @@ -1,7 +1,7 @@ 000018(:protocol-version 2 1) 00004c(:write-string "1/1: Building RecordProjections (RecordProjections.idr)" 1) 00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1) -00007a(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 7) (:end 0 11)) ((:decor :type)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 7) (:end 0 11)) ((:name "Pack") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 12) (:end 0 13)) ((:decor :keyword)))))) 1) 0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 13) (:end 0 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 15) (:end 0 16)) ((:decor :keyword)))))) 1) @@ -9,16 +9,16 @@ 00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 21) (:end 0 22)) ((:decor :keyword)))))) 1) 00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 23) (:end 0 28)) ((:decor :keyword)))))) 1) 00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 2) (:end 1 13)) ((:decor :keyword)))))) 1) -00007b(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 14) (:end 1 20)) ((:decor :data)))))) 1) -00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 2 2) (:end 2 8)) ((:decor :function)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 14) (:end 1 20)) ((:name "MkPack") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 2 2) (:end 2 8)) ((:name "nested") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 2 9) (:end 2 10)) ((:decor :keyword)))))) 1) 0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 2 11) (:end 2 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 2) (:end 3 5)) ((:decor :function)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 2) (:end 3 5)) ((:name "nat") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 6) (:end 3 7)) ((:decor :keyword)))))) 1) 0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 8) (:end 3 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 8) (:end 3 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 8) (:end 3 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 0) (:end 5 4)) ((:decor :function)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 0) (:end 5 4)) ((:name "proj") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 5) (:end 5 6)) ((:decor :keyword)))))) 1) 0000e2(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 7) (:end 5 11)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 12) (:end 5 13)) ((:decor :keyword)))))) 1) @@ -53,7 +53,7 @@ 0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 41) (:end 7 43)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000dd(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 44) (:end 7 48)) ((:name "proj") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 49) (:end 7 51)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 0) (:end 9 8)) ((:decor :function)))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 0) (:end 9 8)) ((:name "ununpack") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 9) (:end 9 10)) ((:decor :keyword)))))) 1) 0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 11) (:end 9 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 16) (:end 9 17)) ((:decor :keyword)))))) 1) @@ -73,7 +73,7 @@ 0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 16) (:end 10 23)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 23) (:end 10 30)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 30) (:end 10 31)) ((:decor :keyword)))))) 1) -00007f(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 0) (:end 12 8)) ((:decor :function)))))) 1) +0000dd(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 0) (:end 12 8)) ((:name "deepNats") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007f(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 9) (:end 12 10)) ((:decor :keyword)))))) 1) 0000e5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 11) (:end 12 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 16) (:end 12 17)) ((:decor :keyword)))))) 1) diff --git a/tests/ideMode/ideMode020/expected b/tests/ideMode/ideMode020/expected index 439439373e..5ad92608ea 100644 --- a/tests/ideMode/ideMode020/expected +++ b/tests/ideMode/ideMode020/expected @@ -1,6 +1,6 @@ 000018(:protocol-version 2 1) 000036(:write-string "1/1: Building Tuples (Tuples.idr)" 1) -000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 0) (:end 0 4)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 0) (:end 0 4)) ((:name "init") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 5) (:end 0 6)) ((:decor :keyword)))))) 1) 0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 7) (:end 0 8)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 9) (:end 0 11)) ((:decor :keyword)))))) 1) @@ -40,7 +40,7 @@ 0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 47) (:end 1 48)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 48) (:end 1 49)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 49) (:end 1 50)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 0) (:end 3 5)) ((:decor :function)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 0) (:end 3 5)) ((:name "unzip") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 6) (:end 3 7)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 8) (:end 3 12)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 13) (:end 3 14)) ((:decor :keyword)))))) 1) @@ -94,7 +94,7 @@ 0000d6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 16) (:end 7 18)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 19) (:end 7 21)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 21) (:end 7 22)) ((:decor :keyword)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 0) (:end 9 6)) ((:decor :function)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 0) (:end 9 6)) ((:name "unzip4") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 7) (:end 9 8)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 9) (:end 9 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 14) (:end 9 15)) ((:decor :keyword)))))) 1) diff --git a/tests/ideMode/ideMode021/expected b/tests/ideMode/ideMode021/expected index f7bbd2ddeb..b0125a2296 100644 --- a/tests/ideMode/ideMode021/expected +++ b/tests/ideMode/ideMode021/expected @@ -6,7 +6,7 @@ 000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 8) (:end 2 13)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 14) (:end 2 15)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 0) (:end 3 6)) ((:decor :keyword)))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 7) (:end 3 12)) ((:decor :type)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 7) (:end 3 12)) ((:name "List1") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 13) (:end 3 14)) ((:decor :keyword)))))) 1) 0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 14) (:end 3 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 16) (:end 3 17)) ((:decor :keyword)))))) 1) @@ -14,17 +14,17 @@ 000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 22) (:end 3 23)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 24) (:end 3 29)) ((:decor :keyword)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 2) (:end 4 13)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 14) (:end 4 19)) ((:decor :data)))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 5 2) (:end 5 6)) ((:decor :function)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 14) (:end 4 19)) ((:name ":::") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 5 2) (:end 5 6)) ((:name "head") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 5 7) (:end 5 8)) ((:decor :keyword)))))) 1) 0000cb(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 5 9) (:end 5 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 2) (:end 6 6)) ((:decor :function)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 2) (:end 6 6)) ((:name "tail") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 7) (:end 6 8)) ((:decor :keyword)))))) 1) 0000db(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 9) (:end 6 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000db(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 9) (:end 6 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000db(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 9) (:end 6 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 14) (:end 6 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 0) (:end 8 4)) ((:decor :function)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 0) (:end 8 4)) ((:name "swap") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 5) (:end 8 6)) ((:decor :keyword)))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 7) (:end 8 12)) ((:name "List1") (:namespace "LetBinders") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 13) (:end 8 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -66,7 +66,7 @@ 0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 5) (:end 14 6)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 7) (:end 14 10)) ((:name ":::") (:namespace "LetBinders") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d1(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 11) (:end 14 15)) ((:name "rest") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000078(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 0) (:end 16 8)) ((:decor :function)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 0) (:end 16 8)) ((:name "identity") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 9) (:end 16 10)) ((:decor :keyword)))))) 1) 0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 11) (:end 16 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 16) (:end 16 17)) ((:decor :keyword)))))) 1) @@ -83,7 +83,7 @@ 0000e0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 0) (:end 17 8)) ((:name "identity") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 9) (:end 17 10)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 18 2) (:end 18 5)) ((:decor :keyword)))))) 1) -000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 4) (:end 20 13)) ((:decor :function)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 4) (:end 20 13)) ((:name "replicate") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 14) (:end 20 15)) ((:decor :keyword)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 16) (:end 20 17)) ((:decor :keyword)))))) 1) 0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 17) (:end 20 18)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode022/expected b/tests/ideMode/ideMode022/expected index b0a66c6a83..e2b178d7ff 100644 --- a/tests/ideMode/ideMode022/expected +++ b/tests/ideMode/ideMode022/expected @@ -5,7 +5,7 @@ 000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 1 2) (:end 1 8)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 1 9) (:end 1 15)) ((:decor :keyword)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 2) (:end 2 6)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 7) (:end 2 12)) ((:decor :type)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 7) (:end 2 12)) ((:name "List0") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 13) (:end 2 14)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 15) (:end 2 19)) ((:decor :type)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 5 0) (:end 5 9)) ((:decor :keyword)))))) 1) @@ -13,14 +13,14 @@ 000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 6 2) (:end 6 8)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 6 9) (:end 6 15)) ((:decor :keyword)))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 2) (:end 7 6)) ((:decor :keyword)))))) 1) -000074(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 7) (:end 7 12)) ((:decor :type)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 7) (:end 7 12)) ((:name "List1") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 13) (:end 7 14)) ((:decor :keyword)))))) 1) 000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 15) (:end 7 19)) ((:decor :type)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 20) (:end 7 25)) ((:decor :keyword)))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 4) (:end 8 7)) ((:decor :data)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 4) (:end 8 7)) ((:name "Lin") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 8) (:end 8 9)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 10) (:end 8 15)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 4) (:end 9 9)) ((:decor :data)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 4) (:end 9 9)) ((:name "Cons1") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 12) (:end 9 17)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 18) (:end 9 20)) ((:decor :keyword)))))) 1) @@ -29,7 +29,7 @@ 0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 28) (:end 9 33)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 11 2) (:end 11 8)) ((:decor :keyword)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 11 9) (:end 11 15)) ((:decor :keyword)))))) 1) -000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 2) (:end 12 6)) ((:decor :function)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 2) (:end 12 6)) ((:name ":<") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 7) (:end 12 8)) ((:decor :keyword)))))) 1) 0000d8(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 9) (:end 12 14)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 15) (:end 12 17)) ((:decor :keyword)))))) 1) @@ -44,7 +44,7 @@ 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 16 2) (:end 16 8)) ((:decor :keyword)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 16 9) (:end 16 15)) ((:decor :keyword)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 2) (:end 17 6)) ((:decor :keyword)))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 7) (:end 17 11)) ((:decor :type)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 7) (:end 17 11)) ((:name ":<") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 12) (:end 17 13)) ((:decor :keyword)))))) 1) 0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 14) (:end 17 19)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 20) (:end 17 22)) ((:decor :keyword)))))) 1) @@ -57,28 +57,28 @@ 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 20 2) (:end 20 8)) ((:decor :keyword)))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 20 9) (:end 20 15)) ((:decor :keyword)))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 2) (:end 21 6)) ((:decor :keyword)))))) 1) -000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 7) (:end 21 12)) ((:decor :type)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 7) (:end 21 12)) ((:name "List0") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 13) (:end 21 14)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 15) (:end 21 19)) ((:decor :type)))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 20) (:end 21 25)) ((:decor :keyword)))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 4) (:end 22 7)) ((:decor :data)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 4) (:end 22 7)) ((:name "Lin") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 8) (:end 22 9)) ((:decor :keyword)))))) 1) 0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 10) (:end 22 15)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 4) (:end 23 8)) ((:decor :data)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 4) (:end 23 8)) ((:name ":<") (:namespace "") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 9) (:end 23 10)) ((:decor :keyword)))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 11) (:end 23 15)) ((:decor :type)))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 16) (:end 23 18)) ((:decor :keyword)))))) 1) 0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 19) (:end 23 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 23) (:end 23 25)) ((:decor :keyword)))))) 1) 0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 26) (:end 23 31)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 25 0) (:end 25 1)) ((:decor :function)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 25 0) (:end 25 1)) ((:name "m") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 25 2) (:end 25 3)) ((:decor :keyword)))))) 1) 0000db(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 25 4) (:end 25 7)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 0) (:end 26 1)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 2) (:end 26 3)) ((:decor :keyword)))))) 1) 0000e2(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 4) (:end 26 14)) ((:name "believe_me") (:namespace "Builtin") (:decor :postulate) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 15) (:end 26 23)) ((:decor :data)))))) 1) -000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 0) (:end 28 7)) ((:decor :function)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 0) (:end 28 7)) ((:name "Rainbow") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 8) (:end 28 9)) ((:decor :keyword)))))) 1) 0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 10) (:end 28 13)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 14) (:end 28 16)) ((:decor :keyword)))))) 1) diff --git a/tests/ideMode/ideMode023/expected b/tests/ideMode/ideMode023/expected index 0f8b9c8aa3..6a774d52f9 100644 --- a/tests/ideMode/ideMode023/expected +++ b/tests/ideMode/ideMode023/expected @@ -4,7 +4,7 @@ 000070(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 0 9) (:end 0 14)) ((:decor :keyword)))))) 1) 00006f(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 2 0) (:end 2 7)) ((:decor :keyword)))))) 1) 00006d(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 2 8) (:end 2 36)) ((:decor :data)))))) 1) -000070(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 4 2) (:end 4 8)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 4 2) (:end 4 8)) ((:name "revAcc") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 4 9) (:end 4 10)) ((:decor :keyword)))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 4 11) (:end 4 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c6(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 4 16) (:end 4 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -36,7 +36,7 @@ 000071(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 6 44) (:end 6 45)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 8 2) (:end 8 34)) ((:decor :comment)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 9 2) (:end 9 48)) ((:decor :comment)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 10 2) (:end 10 5)) ((:decor :function)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 10 2) (:end 10 5)) ((:name "rev") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 10 6) (:end 10 7)) ((:decor :keyword)))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 10 8) (:end 10 12)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c8(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 10 13) (:end 10 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -48,7 +48,7 @@ 0000d3(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 11 8) (:end 11 14)) ((:name "revAcc") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d7(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 11 15) (:end 11 17)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 13 2) (:end 13 31)) ((:decor :comment)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 14 2) (:end 14 6)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 14 2) (:end 14 6)) ((:name "oops") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 14 7) (:end 14 8)) ((:decor :keyword)))))) 1) 0000d0(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 14 9) (:end 14 12)) ((:name "rev") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 14 13) (:end 14 14)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -69,7 +69,7 @@ 00006f(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 17 8) (:end 17 29)) ((:decor :data)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 19 2) (:end 19 78)) ((:decor :comment)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 20 2) (:end 20 69)) ((:decor :comment)))))) 1) -000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 21 2) (:end 21 6)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 21 2) (:end 21 6)) ((:name "oops") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000071(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 21 7) (:end 21 8)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 21 14) (:end 21 15)) ((:decor :data)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 21 16) (:end 21 17)) ((:decor :data)))))) 1) @@ -84,7 +84,7 @@ 000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 28 4) (:end 28 44)) ((:decor :comment)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 29 4) (:end 29 55)) ((:decor :comment)))))) 1) 000072(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 30 4) (:end 30 55)) ((:decor :comment)))))) 1) -000073(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 32 4) (:end 32 11)) ((:decor :function)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 32 4) (:end 32 11)) ((:name "success") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000073(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 32 12) (:end 32 13)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 32 14) (:end 32 17)) ((:decor :data)))))) 1) 0000d4(:output (:ok (:highlight-source ((((:filename "Fail.idr") (:start 32 18) (:end 32 21)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) diff --git a/tests/ideMode/ideMode024/expected b/tests/ideMode/ideMode024/expected index e23c1e7a2a..b9a0456182 100644 --- a/tests/ideMode/ideMode024/expected +++ b/tests/ideMode/ideMode024/expected @@ -2,7 +2,7 @@ 000034(:write-string "1/1: Building Holes (Holes.idr)" 1) 000070(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 0 7) (:end 0 12)) ((:decor :module)))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 2 0) (:end 2 4)) ((:decor :function)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 2 0) (:end 2 4)) ((:name "nats") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 2 5) (:end 2 6)) ((:decor :keyword)))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 2 7) (:end 2 11)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 2 12) (:end 2 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -15,7 +15,7 @@ 00006f(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 3 17) (:end 3 18)) ((:decor :data)))))) 1) 0000d5(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 3 19) (:end 3 21)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000d6(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 3 22) (:end 3 24)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 5 0) (:end 5 4)) ((:decor :function)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 5 0) (:end 5 4)) ((:name "goal") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 5 5) (:end 5 6)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 5 7) (:end 5 8)) ((:decor :keyword)))))) 1) 0000c5(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 5 8) (:end 5 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) @@ -28,7 +28,7 @@ 0000d0(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 6 0) (:end 6 4)) ((:name "goal") (:namespace "Holes") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 0000c5(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 6 5) (:end 6 6)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 6 7) (:end 6 8)) ((:decor :keyword)))))) 1) -000071(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 8 0) (:end 8 5)) ((:decor :function)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 8 0) (:end 8 5)) ((:name "goal2") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 8 6) (:end 8 7)) ((:decor :keyword)))))) 1) 000070(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 8 8) (:end 8 9)) ((:decor :keyword)))))) 1) 0000c7(:output (:ok (:highlight-source ((((:filename "Holes.idr") (:start 8 9) (:end 8 11)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) From 104f40b89e64fdc5ee9318910310d8c009de8b79 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Andr=C3=A9=20Videla?= Date: Fri, 20 Dec 2024 20:48:33 -0500 Subject: [PATCH 4/4] update doc comments --- src/Idris/Parser.idr | 21 ++++++++++++++------- 1 file changed, 14 insertions(+), 7 deletions(-) diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 5e5c73d37e..0a106d3a49 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -1185,12 +1185,18 @@ exportVisibility fname = (specified <$> visOption fname) <|> pure defaulted +||| A binder with only one name and one type +||| BNF: +||| plainBinder := name ':' typeExpr plainBinder : (fname : OriginDesc) => (indents : IndentInfo) => Rule PlainBinder plainBinder = do name <- fcBounds (decoratedSimpleBinderUName fname) decoratedSymbol fname ":" ty <- typeExpr pdef fname indents pure $ MkWithName name ty +||| A binder with multiple names and one type +||| BNF: +||| basicMultiBinder := name (, name)* ':' typeExpr basicMultiBinder : (fname : OriginDesc) => (indents : IndentInfo) => Rule BasicMultiBinder basicMultiBinder = do rig <- multiplicity fname @@ -1817,7 +1823,10 @@ parameters {auto fname : OriginDesc} {auto indents : IndentInfo} decoratedSymbol fname "}" pure $ MkPBinder info params - recordParam : Rule (PBinder) + ||| Record parameter, can be either a typed binder or a name + ||| BNF: + ||| recordParam := typedArg | name + recordParam : Rule PBinder recordParam = typedArg <|> do n <- fcBounds (decoratedSimpleBinderUName fname) @@ -1875,11 +1884,6 @@ parameters {auto fname : OriginDesc} {auto indents : IndentInfo} newParamDecls = some typedArg - -- topLevelClaim is for claims appearing at the top-level of the file - -- localClaim is for claims appearing in nested positions, like `let` or `record` in the future - topLevelClaim : Rule PDeclNoFC - topLevelClaim = PClaim <$> localClaim - definition : Rule PDeclNoFC definition = do nd <- clause 0 Nothing fname indents @@ -1903,6 +1907,9 @@ parameters {auto fname : OriginDesc} {auto indents : IndentInfo} ) pure (mapFC PFixity b) +-- The compiler cannot infer the values for c1 and c2 so I had to write it +-- this way. +-- - Andre cgDirectiveDecl : Rule PDeclNoFC cgDirectiveDecl = (>>=) {c1 = True, c2 = False} cgDirective $ \dir => @@ -1918,7 +1925,7 @@ topDecl fname indents = do id <- anyReservedIdent the (Rule PDecl) $ fatalLoc id.bounds "Cannot begin a declaration with a reserved identifier" <|> fcBounds dataDecl - <|> fcBounds topLevelClaim + <|> fcBounds (PClaim <$> localClaim) <|> fcBounds (PDirective <$> directive) <|> fcBounds implDecl <|> fcBounds definition