From e58bcfc7efbd789e9621622e87cd44195309acb4 Mon Sep 17 00:00:00 2001 From: Ohad Kammar Date: Mon, 10 May 2021 09:05:43 +0100 Subject: [PATCH] Semantic highlighting (#1335) Co-authored-by: Guillaume ALLAIS --- .gitignore | 15 +- docs/source/implementation/ide-protocol.rst | 12 +- idris2api.ipkg | 1 - libs/base/Language/Reflection/TT.idr | 4 +- src/Core/Core.idr | 4 + src/Core/FC.idr | 32 +- src/Core/Metadata.idr | 131 +++- src/Core/Name.idr | 32 +- src/Core/Reflect.idr | 5 + src/Core/TT.idr | 32 + src/Core/TTC.idr | 5 + src/Idris/Desugar.idr | 222 +++--- src/Idris/DocString.idr | 4 +- src/Idris/Driver.idr | 2 +- src/Idris/Elab/Implementation.idr | 88 +-- src/Idris/Elab/Interface.idr | 116 ++-- src/Idris/Error.idr | 25 +- src/Idris/IDEMode/CaseSplit.idr | 2 +- src/Idris/IDEMode/Commands.idr | 5 +- src/Idris/IDEMode/Parser.idr | 9 +- src/Idris/IDEMode/REPL.idr | 13 +- src/Idris/IDEMode/SyntaxHighlight.idr | 146 ++-- src/Idris/ModTree.idr | 4 +- src/Idris/Package.idr | 6 +- src/Idris/Parser.idr | 725 +++++++++++--------- src/Idris/Parser/Let.idr | 2 +- src/Idris/Pretty.idr | 14 +- src/Idris/ProcessIdr.idr | 9 +- src/Idris/REPL.idr | 11 +- src/Idris/REPL/Common.idr | 11 +- src/Idris/Resugar.idr | 34 +- src/Idris/Syntax.idr | 81 ++- src/Libraries/Data/PosMap.idr | 19 +- src/Libraries/Text/Parser.idr | 125 ++-- src/Libraries/Text/Parser/Core.idr | 246 +++---- src/Libraries/Utils/Path.idr | 23 +- src/Libraries/Utils/Shunting.idr | 48 +- src/Parser/Rule/Common.idr | 26 - src/Parser/Rule/Package.idr | 7 +- src/Parser/Rule/Source.idr | 85 ++- src/Parser/Source.idr | 13 +- src/TTImp/Elab/App.idr | 30 +- src/TTImp/Elab/Case.idr | 2 +- src/TTImp/Elab/Check.idr | 7 +- src/TTImp/Elab/ImplicitBind.idr | 6 + src/TTImp/Elab/Record.idr | 15 +- src/TTImp/Elab/Rewrite.idr | 36 +- src/TTImp/Interactive/CaseSplit.idr | 2 +- src/TTImp/Parser.idr | 13 +- src/TTImp/ProcessDecls.idr | 2 +- src/TTImp/ProcessDef.idr | 47 +- src/TTImp/ProcessRecord.idr | 35 +- src/TTImp/ProcessType.idr | 2 + src/TTImp/TTImp.idr | 22 +- src/TTImp/WithClause.idr | 92 ++- src/Yaffle/Main.idr | 2 +- src/Yaffle/REPL.idr | 2 +- tests/Main.idr | 2 +- tests/ideMode/ideMode001/expected | 85 ++- tests/ideMode/ideMode003/expected | 85 ++- tests/ideMode/ideMode005/Ambiguity.idr | 10 + tests/ideMode/ideMode005/Case.idr | 22 + tests/ideMode/ideMode005/Implementation.idr | 25 + tests/ideMode/ideMode005/Interface.idr | 14 + tests/ideMode/ideMode005/README.md | 17 + tests/ideMode/ideMode005/Rainbow.idr | 30 + tests/ideMode/ideMode005/Ranges.idr | 5 + tests/ideMode/ideMode005/RecordUpdate.idr | 23 + tests/ideMode/ideMode005/Rewrite.idr | 9 + tests/ideMode/ideMode005/SimpleData.idr | 8 + tests/ideMode/ideMode005/StringLiterals.idr | 19 + tests/ideMode/ideMode005/Syntax.idr | 52 ++ tests/ideMode/ideMode005/With.idr | 30 + tests/ideMode/ideMode005/WithApp.idr | 12 + tests/ideMode/ideMode005/dummy.ipkg | 1 + tests/ideMode/ideMode005/expected | 0 tests/ideMode/ideMode005/expected1 | 302 ++++++++ tests/ideMode/ideMode005/expected2 | 77 +++ tests/ideMode/ideMode005/expected3 | 63 ++ tests/ideMode/ideMode005/expected4 | 1 + tests/ideMode/ideMode005/expected5 | 89 +++ tests/ideMode/ideMode005/expected6 | 215 ++++++ tests/ideMode/ideMode005/expected7 | 106 +++ tests/ideMode/ideMode005/expected8 | 80 +++ tests/ideMode/ideMode005/expected9 | 327 +++++++++ tests/ideMode/ideMode005/expectedA | 27 + tests/ideMode/ideMode005/expectedB | 55 ++ tests/ideMode/ideMode005/expectedC | 81 +++ tests/ideMode/ideMode005/expectedD | 70 ++ tests/ideMode/ideMode005/input1 | 1 + tests/ideMode/ideMode005/input2 | 1 + tests/ideMode/ideMode005/input3 | 1 + tests/ideMode/ideMode005/input4 | 1 + tests/ideMode/ideMode005/input5 | 1 + tests/ideMode/ideMode005/input6 | 1 + tests/ideMode/ideMode005/input7 | 1 + tests/ideMode/ideMode005/input8 | 1 + tests/ideMode/ideMode005/input9 | 1 + tests/ideMode/ideMode005/inputA | 1 + tests/ideMode/ideMode005/inputB | 1 + tests/ideMode/ideMode005/inputC | 1 + tests/ideMode/ideMode005/inputD | 1 + tests/ideMode/ideMode005/regenerate | 19 + tests/ideMode/ideMode005/run | 42 ++ tests/idris2/reflection001/expected | 8 +- tests/idris2/reg007/expected | 4 +- tests/idris2/with003/expected | 4 +- tests/typedd-book/chapter10/expected | 4 +- 108 files changed, 3618 insertions(+), 1127 deletions(-) delete mode 100644 src/Parser/Rule/Common.idr create mode 100644 tests/ideMode/ideMode005/Ambiguity.idr create mode 100644 tests/ideMode/ideMode005/Case.idr create mode 100644 tests/ideMode/ideMode005/Implementation.idr create mode 100644 tests/ideMode/ideMode005/Interface.idr create mode 100644 tests/ideMode/ideMode005/README.md create mode 100644 tests/ideMode/ideMode005/Rainbow.idr create mode 100644 tests/ideMode/ideMode005/Ranges.idr create mode 100644 tests/ideMode/ideMode005/RecordUpdate.idr create mode 100644 tests/ideMode/ideMode005/Rewrite.idr create mode 100644 tests/ideMode/ideMode005/SimpleData.idr create mode 100644 tests/ideMode/ideMode005/StringLiterals.idr create mode 100644 tests/ideMode/ideMode005/Syntax.idr create mode 100644 tests/ideMode/ideMode005/With.idr create mode 100644 tests/ideMode/ideMode005/WithApp.idr create mode 100644 tests/ideMode/ideMode005/dummy.ipkg create mode 100644 tests/ideMode/ideMode005/expected create mode 100644 tests/ideMode/ideMode005/expected1 create mode 100644 tests/ideMode/ideMode005/expected2 create mode 100644 tests/ideMode/ideMode005/expected3 create mode 100644 tests/ideMode/ideMode005/expected4 create mode 100644 tests/ideMode/ideMode005/expected5 create mode 100644 tests/ideMode/ideMode005/expected6 create mode 100644 tests/ideMode/ideMode005/expected7 create mode 100644 tests/ideMode/ideMode005/expected8 create mode 100644 tests/ideMode/ideMode005/expected9 create mode 100644 tests/ideMode/ideMode005/expectedA create mode 100644 tests/ideMode/ideMode005/expectedB create mode 100644 tests/ideMode/ideMode005/expectedC create mode 100644 tests/ideMode/ideMode005/expectedD create mode 100644 tests/ideMode/ideMode005/input1 create mode 100644 tests/ideMode/ideMode005/input2 create mode 100644 tests/ideMode/ideMode005/input3 create mode 100644 tests/ideMode/ideMode005/input4 create mode 100644 tests/ideMode/ideMode005/input5 create mode 100644 tests/ideMode/ideMode005/input6 create mode 100644 tests/ideMode/ideMode005/input7 create mode 100644 tests/ideMode/ideMode005/input8 create mode 100644 tests/ideMode/ideMode005/input9 create mode 100644 tests/ideMode/ideMode005/inputA create mode 100644 tests/ideMode/ideMode005/inputB create mode 100644 tests/ideMode/ideMode005/inputC create mode 100644 tests/ideMode/ideMode005/inputD create mode 100755 tests/ideMode/ideMode005/regenerate create mode 100755 tests/ideMode/ideMode005/run diff --git a/.gitignore b/.gitignore index 27ce73be9..72cee0897 100644 --- a/.gitignore +++ b/.gitignore @@ -17,7 +17,7 @@ idris2docs_venv /libs/**/build /tests/**/build -/tests/**/output +/tests/**/output* /tests/**/*.so /tests/**/*.dylib /tests/**/*.dll @@ -39,9 +39,16 @@ idris2docs_venv /result # Editor/IDE Related -*~ # Vim swap file -.\#* # Emacs swap file -.vscode/* # VS Code +# WARNING: +# do not put comments on the same line as a regex +# git seems to ignore the pattern in this case + +# Vim swap file +*~ +# Emacs swap file +.\#* +# VS Code +.vscode/* # macOS .DS_Store diff --git a/docs/source/implementation/ide-protocol.rst b/docs/source/implementation/ide-protocol.rst index d2fe14831..88555b499 100644 --- a/docs/source/implementation/ide-protocol.rst +++ b/docs/source/implementation/ide-protocol.rst @@ -203,7 +203,17 @@ The following keys are available: ``implicit`` provides a Boolean value that is True if the region is the name of an implicit argument ``decor`` - describes the category of a token, which can be ``type``, ``function``, ``data``, ``keyword``, or ``bound``. + describes the category of a token, which can be: + + ``type``: type constructors + + ``function``: defined functions + + ``data``: data constructors + + ``bound``: bound variables, or + + ``keyword`` ``source-loc`` states that the region refers to a source code location. Its body is a collection of key-value pairs, with the following possibilities: diff --git a/idris2api.ipkg b/idris2api.ipkg index 2aef2470b..70894a155 100644 --- a/idris2api.ipkg +++ b/idris2api.ipkg @@ -169,7 +169,6 @@ modules = Parser.Lexer.Package, Parser.Lexer.Source, - Parser.Rule.Common, Parser.Rule.Package, Parser.Rule.Source, diff --git a/libs/base/Language/Reflection/TT.idr b/libs/base/Language/Reflection/TT.idr index 5c3faa411..1a01e780d 100644 --- a/libs/base/Language/Reflection/TT.idr +++ b/libs/base/Language/Reflection/TT.idr @@ -14,8 +14,8 @@ FilePos = (Int, Int) -- the second 'FilePos' indicates the start of the next term. public export data FC : Type where - MkFC : String -> FilePos -> FilePos -> FC - EmptyFC : FC + MkFC : String -> FilePos -> FilePos -> FC + EmptyFC : FC public export emptyFC : FC diff --git a/src/Core/Core.idr b/src/Core/Core.idr index 46247455e..5dbc83178 100644 --- a/src/Core/Core.idr +++ b/src/Core/Core.idr @@ -463,6 +463,10 @@ export %inline (<$>) : (a -> b) -> Core a -> Core b (<$>) f (MkCore a) = MkCore (map (map f) a) +export %inline +(<$) : b -> Core a -> Core b +(<$) = (<$>) . const + export %inline ignore : Core a -> Core () ignore = map (\ _ => ()) diff --git a/src/Core/FC.idr b/src/Core/FC.idr index 215fe4170..71401b45f 100644 --- a/src/Core/FC.idr +++ b/src/Core/FC.idr @@ -25,7 +25,10 @@ FileName = String ||| file or by the compiler. That makes it useful to have the notion of ||| `EmptyFC` as part of the type. public export -data FC = MkFC FileName FilePos FilePos +data FC = MkFC FileName FilePos FilePos + | ||| Virtual FCs are FC attached to desugared/generated code. They can help with marking + ||| errors, but we shouldn't attach semantic highlighting metadata to them. + MkVirtualFC FileName FilePos FilePos | EmptyFC ||| A version of a file context that cannot be empty @@ -45,8 +48,22 @@ justFC (fname, start, end) = MkFC fname start end export isNonEmptyFC : FC -> Maybe NonEmptyFC isNonEmptyFC (MkFC fn start end) = Just (fn, start, end) +isNonEmptyFC (MkVirtualFC fn start end) = Just (fn, start, end) isNonEmptyFC EmptyFC = Nothing +||| A view checking whether an arbitrary FC originates from a source location +export +isConcreteFC : FC -> Maybe NonEmptyFC +isConcreteFC (MkFC fn start end) = Just (fn, start, end) +isConcreteFC _ = Nothing + +||| Turn an FC into a virtual one +export +virtualiseFC : FC -> FC +virtualiseFC (MkFC fn start end) = MkVirtualFC fn start end +virtualiseFC fc = fc + + export defaultFC : NonEmptyFC defaultFC = ("", (0, 0), (0, 0)) @@ -140,6 +157,7 @@ mergeFC _ _ = Nothing export Eq FC where (==) (MkFC n s e) (MkFC n' s' e') = n == n' && s == s' && e == e' + (==) (MkVirtualFC n s e) (MkVirtualFC n' s' e') = n == n' && s == s' && e == e' (==) EmptyFC EmptyFC = True (==) _ _ = False @@ -149,6 +167,12 @@ Show FC where show (MkFC file startPos endPos) = file ++ ":" ++ showPos startPos ++ "--" ++ showPos endPos + show (MkVirtualFC file startPos endPos) = file ++ ":" ++ + showPos startPos ++ "--" ++ + showPos endPos + +prettyPos : FilePos -> Doc ann +prettyPos (l, c) = pretty (l + 1) <+> colon <+> pretty (c + 1) export Pretty FC where @@ -156,6 +180,6 @@ Pretty FC where pretty (MkFC file startPos endPos) = pretty file <+> colon <+> prettyPos startPos <+> pretty "--" <+> prettyPos endPos - where - prettyPos : FilePos -> Doc ann - prettyPos (l, c) = pretty (l + 1) <+> colon <+> pretty (c + 1) + pretty (MkVirtualFC file startPos endPos) = pretty file <+> colon + <+> prettyPos startPos <+> pretty "--" + <+> prettyPos endPos diff --git a/src/Core/Metadata.idr b/src/Core/Metadata.idr index fb768c4c3..0a274aa52 100644 --- a/src/Core/Metadata.idr +++ b/src/Core/Metadata.idr @@ -17,7 +17,60 @@ import Libraries.Utils.Binary %default covering --- Additional data we keep about the context to support interactive editing +public export +data Decoration : Type where + Typ : Decoration + Function : Decoration + Data : Decoration + Keyword : Decoration + Bound : Decoration + +export +nameTypeDecoration : NameType -> Decoration +nameTypeDecoration Bound = Bound +nameTypeDecoration Func = Function +nameTypeDecoration (DataCon _ _) = Data +nameTypeDecoration (TyCon _ _ ) = Typ + +public export +ASemanticDecoration : Type +ASemanticDecoration = (NonEmptyFC, Decoration, Maybe Name) + +public export +SemanticDecorations : Type +SemanticDecorations = List ASemanticDecoration + +public export +Eq Decoration where + Typ == Typ = True + Function == Function = True + Data == Data = True + Keyword == Keyword = True + Bound == Bound = True + _ == _ = False + +public export +Show Decoration where + show Typ = "type" + show Function = "function" + show Data = "data" + show Keyword = "keyword" + show Bound = "bound" + +TTC Decoration where + toBuf b Typ = tag 0 + toBuf b Function = tag 1 + toBuf b Data = tag 2 + toBuf b Keyword = tag 3 + toBuf b Bound = tag 4 + fromBuf b + = case !getTag of + 0 => pure Typ + 1 => pure Function + 2 => pure Data + 3 => pure Keyword + 4 => pure Bound + _ => corrupt "Decoration" public export record Metadata where @@ -43,26 +96,41 @@ record Metadata where currentLHS : Maybe ClosedTerm holeLHS : List (Name, ClosedTerm) nameLocMap : PosMap (NonEmptyFC, Name) + sourcefile : String + + -- Semantic Highlighting + -- Posmap of known semantic decorations + semanticHighlighting : PosMap ASemanticDecoration + -- Posmap of aliases (in `with` clauses the LHS disapear during + -- elaboration after making sure that they match their parents' + semanticAliases : PosMap (NonEmptyFC, NonEmptyFC) Show Metadata where - show (MkMetadata apps names tydecls currentLHS holeLHS nameLocMap) + show (MkMetadata apps names tydecls currentLHS holeLHS nameLocMap + fname semanticHighlighting semanticAliases) = "Metadata:\n" ++ " lhsApps: " ++ show apps ++ "\n" ++ " names: " ++ show names ++ "\n" ++ " type declarations: " ++ show tydecls ++ "\n" ++ " current LHS: " ++ show currentLHS ++ "\n" ++ " holes: " ++ show holeLHS ++ "\n" ++ - " nameLocMap: " ++ show nameLocMap + " nameLocMap: " ++ show nameLocMap ++ "\n" ++ + " sourcefile: " ++ fname ++ + " semanticHighlighting: " ++ show semanticHighlighting ++ + " semanticAliases: " ++ show semanticAliases export -initMetadata : Metadata -initMetadata = MkMetadata +initMetadata : String -> Metadata +initMetadata fname = MkMetadata { lhsApps = [] , names = [] , tydecls = [] , currentLHS = Nothing , holeLHS = [] , nameLocMap = empty + , sourcefile = fname + , semanticHighlighting = empty + , semanticAliases = empty } -- A label for metadata in the global state @@ -76,6 +144,9 @@ TTC Metadata where toBuf b (tydecls m) toBuf b (holeLHS m) toBuf b (nameLocMap m) + toBuf b (sourcefile m) + toBuf b (semanticHighlighting m) + toBuf b (semanticAliases m) fromBuf b = do apps <- fromBuf b @@ -83,7 +154,10 @@ TTC Metadata where tys <- fromBuf b hlhs <- fromBuf b dlocs <- fromBuf b - pure (MkMetadata apps ns tys Nothing hlhs dlocs) + fname <- fromBuf b + semhl <- fromBuf b + semal <- fromBuf b + pure (MkMetadata apps ns tys Nothing hlhs dlocs fname semhl semal) export addLHS : {vars : _} -> @@ -213,6 +287,31 @@ findHoleLHS hn = do meta <- get MD pure (lookupBy (\x, y => dropNS x == dropNS y) hn (holeLHS meta)) +export +addSemanticAlias : {auto m : Ref MD Metadata} -> + NonEmptyFC -> NonEmptyFC -> Core () +addSemanticAlias from to + = do meta <- get MD + put MD $ { semanticAliases $= insert (from, to) } meta + +export +addSemanticDecorations : {auto m : Ref MD Metadata} -> + {auto c : Ref Ctxt Defs} -> + SemanticDecorations -> Core () +addSemanticDecorations decors + = do meta <- get MD + let posmap = meta.semanticHighlighting + let (newDecors,droppedDecors) = span + ((meta.sourcefile ==) + . (\((fn, _), _) => fn)) + decors + unless (isNil droppedDecors) + $ log "ide-mode.highlight" 19 $ "ignored adding decorations to " + ++ meta.sourcefile ++ ": " ++ show droppedDecors + put MD $ record {semanticHighlighting + = (fromList newDecors) `union` posmap} meta + + -- Normalise all the types of the names, since they might have had holes -- when added and the holes won't necessarily get saved normaliseTypes : {auto m : Ref MD Metadata} -> @@ -249,13 +348,14 @@ TTC TTMFile where pure (MkTTMFile ver md) HasNames Metadata where - full gam (MkMetadata lhs ns tys clhs hlhs dlocs) - = pure $ MkMetadata !(traverse fullLHS lhs) - !(traverse fullTy ns) - !(traverse fullTy tys) - Nothing - !(traverse fullHLHS hlhs) - (fromList !(traverse fullDecls (toList dlocs))) + full gam md + = pure $ record { lhsApps = !(traverse fullLHS $ md.lhsApps) + , names = !(traverse fullTy $ md.names) + , tydecls = !(traverse fullTy $ md.tydecls) + , currentLHS = Nothing + , holeLHS = !(traverse fullHLHS $ md.holeLHS) + , nameLocMap = fromList !(traverse fullDecls (toList $ md.nameLocMap)) + } md where fullLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm)) fullLHS (fc, (i, tm)) = pure (fc, (i, !(full gam tm))) @@ -269,13 +369,16 @@ HasNames Metadata where fullDecls : (NonEmptyFC, Name) -> Core (NonEmptyFC, Name) fullDecls (fc, n) = pure (fc, !(full gam n)) - resolved gam (MkMetadata lhs ns tys clhs hlhs dlocs) + resolved gam (MkMetadata lhs ns tys clhs hlhs dlocs fname semhl semal) = pure $ MkMetadata !(traverse resolvedLHS lhs) !(traverse resolvedTy ns) !(traverse resolvedTy tys) Nothing !(traverse resolvedHLHS hlhs) (fromList !(traverse resolvedDecls (toList dlocs))) + fname + semhl + semal where resolvedLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm)) resolvedLHS (fc, (i, tm)) = pure (fc, (i, !(resolved gam tm))) diff --git a/src/Core/Name.idr b/src/Core/Name.idr index f0f447528..ec0fb3908 100644 --- a/src/Core/Name.idr +++ b/src/Core/Name.idr @@ -2,6 +2,7 @@ module Core.Name import Data.List import Data.Strings +import Data.Maybe import Decidable.Equality import Libraries.Text.PrettyPrint.Prettyprinter import Libraries.Text.PrettyPrint.Prettyprinter.Util @@ -75,6 +76,21 @@ isUserName (NS _ n) = isUserName n isUserName (DN _ n) = isUserName n isUserName _ = True +||| True iff name can be traced back to a source name. +||| Used to determine whether it needs semantic highlighting. +export +isSourceName : Name -> Bool +isSourceName (NS _ n) = isSourceName n +isSourceName (UN _) = True +isSourceName (MN _ _) = False +isSourceName (PV n _) = isSourceName n +isSourceName (DN _ n) = isSourceName n +isSourceName (RF _) = True +isSourceName (Nested _ n) = isSourceName n +isSourceName (CaseBlock _ _) = False +isSourceName (WithBlock _ _) = False +isSourceName (Resolved _) = False + export isUN : Name -> Maybe String isUN (UN str) = Just str @@ -93,6 +109,19 @@ nameRoot (CaseBlock n _) = "$" ++ show n nameRoot (WithBlock n _) = "$" ++ show n nameRoot (Resolved i) = "$" ++ show i +export +displayName : Name -> (Maybe Namespace, String) +displayName (NS ns n) = mapFst (pure . maybe ns (ns <.>)) $ displayName n +displayName (UN n) = (Nothing, n) +displayName (MN n _) = (Nothing, n) +displayName (PV n _) = displayName n +displayName (DN n _) = (Nothing, n) +displayName (RF n) = (Nothing, n) +displayName (Nested _ n) = displayName n +displayName (CaseBlock outer _) = (Nothing, "case block in " ++ show outer) +displayName (WithBlock outer _) = (Nothing, "with block in " ++ show outer) +displayName (Resolved i) = (Nothing, "$resolved" ++ show i) + --- Drop a namespace from a name export dropNS : Name -> Name @@ -129,7 +158,8 @@ Pretty Name where pretty (PV n d) = braces (pretty 'P' <+> colon <+> pretty n <+> colon <+> pretty d) pretty (DN str _) = pretty str pretty (RF n) = "." <+> pretty n - pretty (Nested (outer, idx) inner) = pretty outer <+> colon <+> pretty idx <+> colon <+> pretty inner + pretty (Nested (outer, idx) inner) + = pretty outer <+> colon <+> pretty idx <+> colon <+> pretty inner pretty (CaseBlock outer _) = reflow "case block in" <++> pretty outer pretty (WithBlock outer _) = reflow "with block in" <++> pretty outer pretty (Resolved x) = pretty "$resolved" <+> pretty x diff --git a/src/Core/Reflect.idr b/src/Core/Reflect.idr index d18b03f2f..5b35c7201 100644 --- a/src/Core/Reflect.idr +++ b/src/Core/Reflect.idr @@ -594,6 +594,11 @@ Reflect FC where start' <- reflect fc defs lhs env start end' <- reflect fc defs lhs env end appCon fc defs (reflectiontt "MkFC") [fn', start', end'] + reflect fc defs lhs env (MkVirtualFC fn start end) + = do fn' <- reflect fc defs lhs env fn + start' <- reflect fc defs lhs env start + end' <- reflect fc defs lhs env end + appCon fc defs (reflectiontt "MkFC") [fn', start', end'] reflect fc defs lhs env EmptyFC = getCon fc defs (reflectiontt "EmptyFC") {- diff --git a/src/Core/TT.idr b/src/Core/TT.idr index 73909cd8b..dcd07c71d 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -82,6 +82,38 @@ isConstantType (UN n) = case n of _ => Nothing isConstantType _ = Nothing +export +isPrimType : Constant -> Bool +isPrimType (I x) = False +isPrimType (I8 x) = False +isPrimType (I16 x) = False +isPrimType (I32 x) = False +isPrimType (I64 x) = False +isPrimType (BI x) = False +isPrimType (B8 x) = False +isPrimType (B16 x) = False +isPrimType (B32 x) = False +isPrimType (B64 x) = False +isPrimType (Str x) = False +isPrimType (Ch x) = False +isPrimType (Db x) = False +isPrimType WorldVal = False + +isPrimType Int8Type = True +isPrimType Int16Type = True +isPrimType Int32Type = True +isPrimType Int64Type = True +isPrimType IntType = True +isPrimType IntegerType = True +isPrimType Bits8Type = True +isPrimType Bits16Type = True +isPrimType Bits32Type = True +isPrimType Bits64Type = True +isPrimType StringType = True +isPrimType CharType = True +isPrimType DoubleType = True +isPrimType WorldType = True + export constantEq : (x, y : Constant) -> Maybe (x = y) constantEq (I x) (I y) = case decEq x y of diff --git a/src/Core/TTC.idr b/src/Core/TTC.idr index 075d894db..bd3834c48 100644 --- a/src/Core/TTC.idr +++ b/src/Core/TTC.idr @@ -22,6 +22,8 @@ export TTC FC where toBuf b (MkFC file startPos endPos) = do tag 0; toBuf b file; toBuf b startPos; toBuf b endPos + toBuf b (MkVirtualFC file startPos endPos) + = do tag 2; toBuf b file; toBuf b startPos; toBuf b endPos toBuf b EmptyFC = tag 1 fromBuf b @@ -30,6 +32,9 @@ TTC FC where s <- fromBuf b; e <- fromBuf b pure (MkFC f s e) 1 => pure EmptyFC + 2 => do f <- fromBuf b; + s <- fromBuf b; e <- fromBuf b + pure (MkVirtualFC f s e) _ => corrupt "FC" export TTC Namespace where diff --git a/src/Idris/Desugar.idr b/src/Idris/Desugar.idr index 6b3935ca0..175e95c6d 100644 --- a/src/Idris/Desugar.idr +++ b/src/Idris/Desugar.idr @@ -87,7 +87,7 @@ mkPrec Prefix p = Prefix p toTokList : {auto s : Ref Syn SyntaxInfo} -> PTerm -> Core (List (Tok OpStr PTerm)) -toTokList (POp fc opn l r) +toTokList (POp fc opFC opn l r) = do syn <- get Syn let op = nameRoot opn case lookup op (infixes syn) of @@ -95,16 +95,16 @@ toTokList (POp fc opn l r) if any isOpChar (fastUnpack op) then throw (GenericMsg fc $ "Unknown operator '" ++ op ++ "'") else do rtoks <- toTokList r - pure (Expr l :: Op fc opn backtickPrec :: rtoks) + pure (Expr l :: Op fc opFC opn backtickPrec :: rtoks) Just (Prefix, _) => throw (GenericMsg fc $ "'" ++ op ++ "' is a prefix operator") Just (fix, prec) => do rtoks <- toTokList r - pure (Expr l :: Op fc opn (mkPrec fix prec) :: rtoks) + pure (Expr l :: Op fc opFC opn (mkPrec fix prec) :: rtoks) where backtickPrec : OpPrec backtickPrec = NonAssoc 1 -toTokList (PPrefixOp fc opn arg) +toTokList (PPrefixOp fc opFC opn arg) = do syn <- get Syn let op = nameRoot opn case lookup op (prefixes syn) of @@ -112,7 +112,7 @@ toTokList (PPrefixOp fc opn arg) throw (GenericMsg fc $ "'" ++ op ++ "' is not a prefix operator") Just prec => do rtoks <- toTokList arg - pure (Op fc opn (Prefix prec) :: rtoks) + pure (Op fc opFC opn (Prefix prec) :: rtoks) toTokList t = pure [Expr t] record BangData where @@ -123,21 +123,39 @@ record BangData where initBangs : BangData initBangs = MkBangData 0 [] +addNS : Maybe Namespace -> Name -> Name +addNS (Just ns) n@(NS _ _) = n +addNS (Just ns) n = NS ns n +addNS _ n = n + +bindFun : FC -> Maybe Namespace -> RawImp -> RawImp -> RawImp +bindFun fc ns ma f = + let fc = virtualiseFC fc in + IApp fc (IApp fc (IVar fc (addNS ns $ UN ">>=")) ma) f + +seqFun : FC -> Maybe Namespace -> RawImp -> RawImp -> RawImp +seqFun fc ns ma mb = + let fc = virtualiseFC fc in + IApp fc (IApp fc (IVar fc (addNS ns (UN ">>"))) ma) mb + bindBangs : List (Name, FC, RawImp) -> RawImp -> RawImp bindBangs [] tm = tm bindBangs ((n, fc, btm) :: bs) tm - = bindBangs bs $ IApp fc (IApp fc (IVar fc (UN ">>=")) btm) - (ILam EmptyFC top Explicit (Just n) - (Implicit fc False) tm) + = bindBangs bs + $ bindFun fc Nothing btm + $ ILam EmptyFC top Explicit (Just n) (Implicit fc False) tm idiomise : FC -> RawImp -> RawImp idiomise fc (IAlternative afc u alts) = IAlternative afc (mapAltType (idiomise afc) u) (idiomise afc <$> alts) idiomise fc (IApp afc f a) - = IApp fc (IApp fc (IVar fc (UN "<*>")) - (idiomise afc f)) - a -idiomise fc fn = IApp fc (IVar fc (UN "pure")) fn + = let fc = virtualiseFC fc in + IApp fc (IApp fc (IVar fc (UN "<*>")) + (idiomise afc f)) + a +idiomise fc fn + = let fc = virtualiseFC fc in + IApp fc (IVar fc (UN "pure")) fn pairname : Name pairname = NS builtinNS (UN "Pair") @@ -166,9 +184,11 @@ mutual pure $ IPi fc rig !(traverse (desugar side ps') p) mn !(desugarB side ps argTy) !(desugarB side ps' retTy) - desugarB side ps (PLam fc rig p pat@(PRef _ n@(UN nm)) argTy scope) + desugarB side ps (PLam fc rig p pat@(PRef prefFC n@(UN nm)) argTy scope) = if lowerFirst nm || nm == "_" - then pure $ ILam fc rig !(traverse (desugar side ps) p) + then do whenJust (isConcreteFC prefFC) \nfc + => addSemanticDecorations [(nfc, Bound, Just n)] + pure $ ILam fc rig !(traverse (desugar side ps) p) (Just n) !(desugarB side ps argTy) !(desugar side (n :: ps) scope) else pure $ ILam EmptyFC rig !(traverse (desugar side ps) p) @@ -189,15 +209,17 @@ mutual ICase fc (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False) [snd !(desugarClause ps True (MkPatClause fc pat scope []))] desugarB side ps (PLet fc rig (PRef prefFC n) nTy nVal scope []) - = pure $ ILet fc prefFC rig n !(desugarB side ps nTy) !(desugarB side ps nVal) - !(desugar side (n :: ps) scope) + = do whenJust (isConcreteFC prefFC) \nfc => + addSemanticDecorations [(nfc, Bound, Just n)] + pure $ ILet fc prefFC rig n !(desugarB side ps nTy) !(desugarB side ps nVal) + !(desugar side (n :: ps) scope) desugarB side ps (PLet fc rig pat nTy nVal scope alts) = pure $ ICase fc !(desugarB side ps nVal) !(desugarB side ps nTy) !(traverse (map snd . desugarClause ps True) (MkPatClause fc pat scope [] :: alts)) desugarB side ps (PCase fc x xs) = pure $ ICase fc !(desugarB side ps x) - (Implicit fc False) + (Implicit (virtualiseFC fc) False) !(traverse (map snd . desugarClause ps True) xs) desugarB side ps (PLocal fc xs scope) = let ps' = definedIn xs ++ ps in @@ -207,8 +229,10 @@ mutual = pure $ IUpdate pfc !(traverse (desugarUpdate side ps) fs) !(desugarB side ps rec) desugarB side ps (PUpdate fc fs) - = desugarB side ps (PLam fc top Explicit (PRef fc (MN "rec" 0)) (PImplicit fc) - (PApp fc (PUpdate fc fs) (PRef fc (MN "rec" 0)))) + = desugarB side ps + $ let vfc = virtualiseFC fc in + PLam vfc top Explicit (PRef vfc (MN "rec" 0)) (PImplicit vfc) + $ PApp vfc (PUpdate fc fs) (PRef vfc (MN "rec" 0)) desugarB side ps (PApp fc x y) = pure $ IApp fc !(desugarB side ps x) !(desugarB side ps y) desugarB side ps (PAutoApp fc x y) @@ -230,24 +254,24 @@ mutual [apply (IVar fc (UN "===")) [l', r'], apply (IVar fc (UN "~=~")) [l', r']] desugarB side ps (PBracketed fc e) = desugarB side ps e - desugarB side ps (POp fc op l r) - = do ts <- toTokList (POp fc op l r) + desugarB side ps (POp fc opFC op l r) + = do ts <- toTokList (POp fc opFC op l r) desugarTree side ps !(parseOps ts) - desugarB side ps (PPrefixOp fc op arg) - = do ts <- toTokList (PPrefixOp fc op arg) + desugarB side ps (PPrefixOp fc opFC op arg) + = do ts <- toTokList (PPrefixOp fc opFC op arg) desugarTree side ps !(parseOps ts) - desugarB side ps (PSectionL fc op arg) + desugarB side ps (PSectionL fc opFC op arg) = do syn <- get Syn -- It might actually be a prefix argument rather than a section -- so check that first, otherwise desugar as a lambda case lookup (nameRoot op) (prefixes syn) of Nothing => desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) - (POp fc op (PRef fc (MN "arg" 0)) arg)) - Just prec => desugarB side ps (PPrefixOp fc op arg) - desugarB side ps (PSectionR fc arg op) + (POp fc opFC op (PRef fc (MN "arg" 0)) arg)) + Just prec => desugarB side ps (PPrefixOp fc opFC op arg) + desugarB side ps (PSectionR fc opFC arg op) = desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) - (POp fc op arg (PRef fc (MN "arg" 0)))) + (POp fc opFC op arg (PRef fc (MN "arg" 0)))) desugarB side ps (PSearch fc depth) = pure $ ISearch fc depth desugarB side ps (PPrimVal fc (BI x)) = case !fromIntegerName of @@ -255,20 +279,23 @@ mutual pure $ IAlternative fc (UniqueDefault (IPrimVal fc (BI x))) [IPrimVal fc (BI x), IPrimVal fc (I (fromInteger x))] - Just fi => pure $ IApp fc (IVar fc fi) - (IPrimVal fc (BI x)) + Just fi => + let vfc = virtualiseFC fc in + pure $ IApp vfc (IVar vfc fi) (IPrimVal fc (BI x)) desugarB side ps (PPrimVal fc (Ch x)) = case !fromCharName of Nothing => pure $ IPrimVal fc (Ch x) - Just f => pure $ IApp fc (IVar fc f) - (IPrimVal fc (Ch x)) + Just f => + let vfc = virtualiseFC fc in + pure $ IApp vfc (IVar vfc f) (IPrimVal fc (Ch x)) desugarB side ps (PPrimVal fc (Db x)) = case !fromDoubleName of Nothing => pure $ IPrimVal fc (Db x) - Just f => pure $ IApp fc (IVar fc f) - (IPrimVal fc (Db x)) + Just f => + let vfc = virtualiseFC fc in + pure $ IApp vfc (IVar vfc f) (IPrimVal fc (Db x)) desugarB side ps (PPrimVal fc x) = pure $ IPrimVal fc x desugarB side ps (PQuote fc tm) = pure $ IQuote fc !(desugarB side ps tm) @@ -313,40 +340,41 @@ mutual let val = idiomise fc itm logRaw "desugar.idiom" 10 "Desugared to" val pure val - desugarB side ps (PList fc args) - = expandList side ps fc args + desugarB side ps (PList fc nilFC args) + = expandList side ps nilFC args desugarB side ps (PPair fc l r) = do l' <- desugarB side ps l r' <- desugarB side ps r let pval = apply (IVar fc mkpairname) [l', r'] pure $ IAlternative fc (UniqueDefault pval) [apply (IVar fc pairname) [l', r'], pval] - desugarB side ps (PDPair fc (PRef nfc (UN n)) (PImplicit _) r) + desugarB side ps (PDPair fc opFC (PRef nfc (UN n)) (PImplicit _) r) = do r' <- desugarB side ps r - let pval = apply (IVar fc mkdpairname) [IVar nfc (UN n), r'] + let pval = apply (IVar opFC mkdpairname) [IVar nfc (UN n), r'] pure $ IAlternative fc (UniqueDefault pval) [apply (IVar fc dpairname) [Implicit nfc False, ILam nfc top Explicit (Just (UN n)) (Implicit nfc False) r'], pval] - desugarB side ps (PDPair fc (PRef nfc (UN n)) ty r) + desugarB side ps (PDPair fc opFC (PRef nfc (UN n)) ty r) = do ty' <- desugarB side ps ty r' <- desugarB side ps r - pure $ apply (IVar fc dpairname) + pure $ apply (IVar opFC dpairname) [ty', ILam nfc top Explicit (Just (UN n)) ty' r'] - desugarB side ps (PDPair fc l (PImplicit _) r) + desugarB side ps (PDPair fc opFC l (PImplicit _) r) = do l' <- desugarB side ps l r' <- desugarB side ps r - pure $ apply (IVar fc mkdpairname) [l', r'] - desugarB side ps (PDPair fc l ty r) + pure $ apply (IVar opFC mkdpairname) [l', r'] + desugarB side ps (PDPair fc opFC l ty r) = throw (GenericMsg fc "Invalid dependent pair type") desugarB side ps (PUnit fc) = pure $ IAlternative fc (UniqueDefault (IVar fc (UN "MkUnit"))) [IVar fc (UN "Unit"), IVar fc (UN "MkUnit")] desugarB side ps (PIfThenElse fc x t e) - = pure $ ICase fc !(desugarB side ps x) (IVar fc (UN "Bool")) + = let fc = virtualiseFC fc in + pure $ ICase fc !(desugarB side ps x) (IVar fc (UN "Bool")) [PatClause fc (IVar fc (UN "True")) !(desugar side ps t), PatClause fc (IVar fc (UN "False")) !(desugar side ps e)] desugarB side ps (PComprehension fc ret conds) @@ -361,22 +389,15 @@ mutual desugarB side ps (PRewrite fc rule tm) = pure $ IRewrite fc !(desugarB side ps rule) !(desugarB side ps tm) desugarB side ps (PRange fc start next end) - = case next of - Nothing => - desugarB side ps (PApp fc - (PApp fc (PRef fc (UN "rangeFromTo")) - start) end) - Just n => - desugarB side ps (PApp fc - (PApp fc - (PApp fc (PRef fc (UN "rangeFromThenTo")) - start) n) end) + = let fc = virtualiseFC fc in + desugarB side ps $ case next of + Nothing => papply fc (PRef fc (UN "rangeFromTo")) [start,end] + Just n => papply fc (PRef fc (UN "rangeFromThenTo")) [start, n, end] desugarB side ps (PRangeStream fc start next) - = case next of - Nothing => - desugarB side ps (PApp fc (PRef fc (UN "rangeFrom")) start) - Just n => - desugarB side ps (PApp fc (PApp fc (PRef fc (UN "rangeFromThen")) start) n) + = let fc = virtualiseFC fc in + desugarB side ps $ case next of + Nothing => papply fc (PRef fc (UN "rangeFrom")) [start] + Just n => papply fc (PRef fc (UN "rangeFromThen")) [start, n] desugarB side ps (PUnifyLog fc lvl tm) = pure $ IUnifyLog fc lvl !(desugarB side ps tm) desugarB side ps (PPostfixApp fc rec projs) @@ -404,23 +425,21 @@ mutual {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> {auto m : Ref MD Metadata} -> - Side -> List Name -> FC -> List PTerm -> Core RawImp - expandList side ps fc [] = pure (IVar fc (UN "Nil")) - expandList side ps fc (x :: xs) - = pure $ apply (IVar fc (UN "::")) - [!(desugarB side ps x), !(expandList side ps fc xs)] - - addNS : Maybe Namespace -> Name -> Name - addNS (Just ns) n@(NS _ _) = n - addNS (Just ns) n = NS ns n - addNS _ n = n + Side -> List Name -> + (nilFC : FC) -> List (FC, PTerm) -> Core RawImp + expandList side ps nilFC [] = pure (IVar nilFC (UN "Nil")) + expandList side ps nilFC ((consFC, x) :: xs) + = pure $ apply (IVar consFC (UN "::")) + [!(desugarB side ps x), !(expandList side ps nilFC xs)] addFromString : {auto c : Ref Ctxt Defs} -> FC -> RawImp -> Core RawImp addFromString fc tm = pure $ case !fromStringName of Nothing => tm - Just f => IApp fc (IVar fc f) tm + Just f => + let fc = virtualiseFC fc in + IApp fc (IVar fc f) tm expandString : {auto s : Ref Syn SyntaxInfo} -> {auto b : Ref Bang BangData} -> @@ -428,9 +447,11 @@ mutual {auto m : Ref MD Metadata} -> {auto u : Ref UST UState} -> Side -> List Name -> FC -> List PStr -> Core RawImp - expandString side ps fc xs = pure $ case !(traverse toRawImp (filter notEmpty $ mergeStrLit xs)) of - [] => IPrimVal fc (Str "") - xs@(_::_) => foldr1 concatStr xs + expandString side ps fc xs + = do xs <- traverse toRawImp (filter notEmpty $ mergeStrLit xs) + pure $ case xs of + [] => IPrimVal fc (Str "") + (_ :: _) => foldr1 concatStr xs where toRawImp : PStr -> Core RawImp toRawImp (StrLiteral fc str) = pure $ IPrimVal fc (Str str) @@ -449,7 +470,10 @@ mutual notEmpty (StrInterp _ _) = True concatStr : RawImp -> RawImp -> RawImp - concatStr a b = IApp (getFC a) (IApp (getFC b) (IVar (getFC b) (UN "++")) a) b + concatStr a b = + let aFC = virtualiseFC (getFC a) + bFC = virtualiseFC (getFC b) + in IApp aFC (IApp bFC (IVar bFC (UN "++")) a) b trimMultiline : FC -> Nat -> List (List PStr) -> Core (List PStr) trimMultiline fc indent lines @@ -469,8 +493,7 @@ mutual then throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters" else pure initLines trimLast _ (initLines `snoc` xs) | Snoc xs initLines _ - = let fc = fromMaybe fc $ findBy (\case StrInterp fc _ => Just fc; - StrLiteral _ _ => Nothing) xs in + = let fc = fromMaybe fc $ findBy isStrInterp xs in throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters" dropLastNL : List PStr -> List PStr @@ -508,13 +531,14 @@ mutual = do tm' <- desugar side ps tm rest' <- expandDo side ps topfc ns rest gam <- get Ctxt - pure $ IApp fc (IApp fc (IVar fc (addNS ns (UN ">>"))) tm') rest' + pure $ seqFun fc ns tm' rest' expandDo side ps topfc ns (DoBind fc nameFC n tm :: rest) = do tm' <- desugar side ps tm rest' <- expandDo side ps topfc ns rest - pure $ IApp fc (IApp fc (IVar fc (addNS ns (UN ">>="))) tm') - (ILam nameFC top Explicit (Just n) - (Implicit fc False) rest') + whenJust (isConcreteFC nameFC) \nfc => addSemanticDecorations [(nfc, Bound, Just n)] + pure $ bindFun fc ns tm' + $ ILam nameFC top Explicit (Just n) + (Implicit (virtualiseFC fc) False) rest' expandDo side ps topfc ns (DoBindPat fc pat exp alts :: rest) = do pat' <- desugar LHS ps pat (newps, bpat) <- bindNames False pat' @@ -522,19 +546,22 @@ mutual alts' <- traverse (map snd . desugarClause ps True) alts let ps' = newps ++ ps rest' <- expandDo side ps' topfc ns rest - pure $ IApp fc (IApp fc (IVar fc (addNS ns (UN ">>="))) exp') - (ILam EmptyFC top Explicit (Just (MN "_" 0)) + let fcOriginal = fc + let fc = virtualiseFC fc + pure $ bindFun fc ns exp' + $ ILam EmptyFC top Explicit (Just (MN "_" 0)) (Implicit fc False) (ICase fc (IVar EmptyFC (MN "_" 0)) (Implicit fc False) - (PatClause fc bpat rest' - :: alts'))) + (PatClause fcOriginal bpat rest' + :: alts')) expandDo side ps topfc ns (DoLet fc lhsFC n rig ty tm :: rest) = do b <- newRef Bang initBangs tm' <- desugarB side ps tm ty' <- desugar side ps ty rest' <- expandDo side ps topfc ns rest - let bind = ILet fc lhsFC rig n ty' tm' rest' + whenJust (isConcreteFC lhsFC) \nfc => addSemanticDecorations [(nfc, Bound, Just n)] + let bind = ILet fc (virtualiseFC lhsFC) rig n ty' tm' rest' bd <- get Bang pure $ bindBangs (bangNames bd) bind expandDo side ps topfc ns (DoLetPat fc pat ty tm alts :: rest) @@ -547,6 +574,7 @@ mutual let ps' = newps ++ ps rest' <- expandDo side ps' topfc ns rest bd <- get Bang + let fc = virtualiseFC fc pure $ bindBangs (bangNames bd) $ ICase fc tm' ty' (PatClause fc bpat rest' @@ -566,20 +594,20 @@ mutual {auto u : Ref UST UState} -> {auto m : Ref MD Metadata} -> Side -> List Name -> Tree OpStr PTerm -> Core RawImp - desugarTree side ps (Infix loc (UN "=") l r) -- special case since '=' is special syntax + desugarTree side ps (Infix loc eqFC (UN "=") l r) -- special case since '=' is special syntax = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IAlternative loc FirstSuccess - [apply (IVar loc (UN "===")) [l', r'], - apply (IVar loc (UN "~=~")) [l', r']]) - desugarTree side ps (Infix loc (UN "$") l r) -- special case since '$' is special syntax + [apply (IVar eqFC (UN "===")) [l', r'], + apply (IVar eqFC (UN "~=~")) [l', r']]) + desugarTree side ps (Infix loc _ (UN "$") l r) -- special case since '$' is special syntax = do l' <- desugarTree side ps l r' <- desugarTree side ps r pure (IApp loc l' r') - desugarTree side ps (Infix loc op l r) + desugarTree side ps (Infix loc opFC op l r) = do l' <- desugarTree side ps l r' <- desugarTree side ps r - pure (IApp loc (IApp loc (IVar loc op) l') r') + pure (IApp loc (IApp loc (IVar opFC op) l') r') -- negation is a special case, since we can't have an operator with -- two meanings otherwise @@ -587,7 +615,7 @@ mutual -- Note: In case of negated signed integer literals, we apply the -- negation directly. Otherwise, the literal might be -- truncated to 0 before being passed on to `negate`. - desugarTree side ps (Pre loc (UN "-") $ Leaf $ PPrimVal fc c) + desugarTree side ps (Pre loc opFC (UN "-") $ Leaf $ PPrimVal fc c) = let newFC = fromMaybe EmptyFC (mergeFC loc fc) continue = desugarTree side ps . Leaf . PPrimVal newFC in case c of @@ -601,15 +629,15 @@ mutual -- not a signed integer literal. proceed by desugaring -- and applying to `negate`. _ => do arg' <- desugarTree side ps (Leaf $ PPrimVal fc c) - pure (IApp loc (IVar loc (UN "negate")) arg') + pure (IApp loc (IVar opFC (UN "negate")) arg') - desugarTree side ps (Pre loc (UN "-") arg) + desugarTree side ps (Pre loc opFC (UN "-") arg) = do arg' <- desugarTree side ps arg - pure (IApp loc (IVar loc (UN "negate")) arg') + pure (IApp loc (IVar opFC (UN "negate")) arg') - desugarTree side ps (Pre loc op arg) + desugarTree side ps (Pre loc opFC op arg) = do arg' <- desugarTree side ps arg - pure (IApp loc (IVar loc op) arg') + pure (IApp loc (IVar opFC op) arg') desugarTree side ps (Leaf t) = desugarB side ps t desugarType : {auto s : Ref Syn SyntaxInfo} -> diff --git a/src/Idris/DocString.idr b/src/Idris/DocString.idr index abbeb7924..4d38a9412 100644 --- a/src/Idris/DocString.idr +++ b/src/Idris/DocString.idr @@ -286,9 +286,9 @@ getDocsForPTerm (PRef fc name) = pure $ [!(render styleAnn !(getDocsForName fc n getDocsForPTerm (PPrimVal _ constant) = getDocsForPrimitive constant getDocsForPTerm (PType _) = pure ["Type : Type\n\tThe type of all types is Type. The type of Type is Type."] getDocsForPTerm (PString _ _) = pure ["String Literal\n\tDesugars to a fromString call"] -getDocsForPTerm (PList _ _) = pure ["List Literal\n\tDesugars to (::) and Nil"] +getDocsForPTerm (PList _ _ _) = pure ["List Literal\n\tDesugars to (::) and Nil"] getDocsForPTerm (PPair _ _ _) = pure ["Pair Literal\n\tDesugars to MkPair or Pair"] -getDocsForPTerm (PDPair _ _ _ _) = pure ["Dependant Pair Literal\n\tDesugars to MkDPair or DPair"] +getDocsForPTerm (PDPair _ _ _ _ _) = pure ["Dependant Pair Literal\n\tDesugars to MkDPair or DPair"] getDocsForPTerm (PUnit _) = pure ["Unit Literal\n\tDesugars to MkUnit or Unit"] getDocsForPTerm pterm = pure ["Docs not implemented for " ++ show pterm ++ " yet"] diff --git a/src/Idris/Driver.idr b/src/Idris/Driver.idr index d7e260902..790c551c8 100644 --- a/src/Idris/Driver.idr +++ b/src/Idris/Driver.idr @@ -178,7 +178,7 @@ stMain cgs opts when (checkVerbose opts) $ -- override Quiet if implicitly set setOutput (REPL False) u <- newRef UST initUState - m <- newRef MD initMetadata + m <- newRef MD (initMetadata $ fromMaybe "(interactive)" fname) updateREPLOpts session <- getSession when (not $ nobanner session) $ do diff --git a/src/Idris/Elab/Implementation.idr b/src/Idris/Elab/Implementation.idr index 41d3eaafb..2920846df 100644 --- a/src/Idris/Elab/Implementation.idr +++ b/src/Idris/Elab/Implementation.idr @@ -124,7 +124,7 @@ elabImplementation : {vars : _} -> Maybe (List ImpDecl) -> Core () -- TODO: Refactor all these steps into separate functions -elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named impName_in nusing mbody +elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named impName_in nusing mbody = do -- let impName_in = maybe (mkImplName fc iname ps) id impln -- If we're in a nested block, update the name let impName_nest = case lookup impName_in (names nest) of @@ -138,15 +138,15 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im inames <- lookupCtxtName iname (gamma defs) let [cndata] = concatMap (\n => lookupName n (ifaces syn)) (map fst inames) - | [] => undefinedName fc iname - | ns => throw (AmbiguousName fc (map fst ns)) + | [] => undefinedName vfc iname + | ns => throw (AmbiguousName vfc (map fst ns)) let cn : Name = fst cndata let cdata : IFaceInfo = snd cndata Just ity <- lookupTyExact cn (gamma defs) - | Nothing => undefinedName fc cn + | Nothing => undefinedName vfc cn Just conty <- lookupTyExact (iconstructor cdata) (gamma defs) - | Nothing => undefinedName fc (iconstructor cdata) + | Nothing => undefinedName vfc (iconstructor cdata) let impsp = nub (concatMap findIBinds ps ++ concatMap findIBinds (map snd cons)) @@ -171,14 +171,14 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im then [Inline] else [Inline, Hint True] - let initTy = bindImpls fc is $ bindConstraints fc AutoImplicit cons - (apply (IVar fc iname) ps) + let initTy = bindImpls vfc is $ bindConstraints vfc AutoImplicit cons + (apply (IVar vfc iname) ps) let paramBinds = if !isUnboundImplicits then findBindableNames True vars [] initTy else [] let impTy = doBind paramBinds initTy - let impTyDecl = IClaim fc top vis opts (MkImpTy EmptyFC EmptyFC impName impTy) + let impTyDecl = IClaim vfc top vis opts (MkImpTy EmptyFC EmptyFC impName impTy) log "elab.implementation" 5 $ "Implementation type: " ++ show impTy @@ -186,17 +186,17 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im -- If the body is empty, we're done for now (just declaring that -- the implementation exists and define it later) - when (defPass pass) $ maybe (pure ()) - (\body_in => do + when (defPass pass) $ + whenJust mbody $ \body_in => do defs <- get Ctxt Just impTyc <- lookupTyExact impName (gamma defs) | Nothing => throw (InternalError ("Can't happen, can't find type of " ++ show impName)) methImps <- getMethImps [] impTyc log "elab.implementation" 3 $ "Bind implicits to each method: " ++ show methImps - -- 1.5. Lookup default definitions and add them to to body + -- 1.5. Lookup default definitions and add them to the body let (body, missing) - = addDefaults fc impName + = addDefaults vfc impName (zip (params cdata) ps) (map (dropNS . name) (methods cdata)) (defaults cdata) body_in @@ -208,7 +208,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im defs <- get Ctxt let hs = openHints defs -- snapshot open hint state log "elab.implementation" 10 $ "Open hints: " ++ (show (impName :: nusing)) - traverse_ (\n => do n' <- checkUnambig fc n + traverse_ (\n => do n' <- checkUnambig vfc n addOpenHint n') nusing -- 2. Elaborate top level function types for this interface @@ -223,16 +223,16 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im let mtops = map (Builtin.fst . snd) fns let con = iconstructor cdata let ilhs = impsApply (IVar EmptyFC impName) - (map (\x => (x, IBindVar fc (show x))) + (map (\x => (x, IBindVar vfc (show x))) (map fst methImps)) -- RHS is the constructor applied to a search for the necessary -- parent constraints, then the method implementations defs <- get Ctxt let fldTys = getFieldArgs !(normaliseHoles defs [] conty) log "elab.implementation" 5 $ "Field types " ++ show fldTys - let irhs = apply (autoImpsApply (IVar fc con) $ map (const (ISearch fc 500)) (parents cdata)) + let irhs = apply (autoImpsApply (IVar vfc con) $ map (const (ISearch vfc 500)) (parents cdata)) (map (mkMethField methImps fldTys) fns) - let impFn = IDef fc impName [PatClause fc ilhs irhs] + let impFn = IDef vfc impName [PatClause vfc ilhs irhs] log "elab.implementation" 5 $ "Implementation record: " ++ show impFn -- If it's a named implementation, add it as a global hint while @@ -241,7 +241,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im -- Make sure we don't use this name to solve parent constraints -- when elaborating the record, or we'll end up in a cycle! - setFlag fc impName BlockedHint + setFlag vfc impName BlockedHint -- Update nested names so we elaborate the body in the right -- environment @@ -249,11 +249,11 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im let nest' = record { names $= (names' ++) } nest traverse_ (processDecl [] nest' env) [impFn] - unsetFlag fc impName BlockedHint + unsetFlag vfc impName BlockedHint - setFlag fc impName TCInline + setFlag vfc impName TCInline -- it's the methods we're interested in, not the implementation - setFlag fc impName (SetTotal PartialOK) + setFlag vfc impName (SetTotal PartialOK) -- 4. (TODO: Order method bodies to be in declaration order, in -- case of dependencies) @@ -270,19 +270,22 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im -- inline flag has done its job, and outside the interface -- it can hurt, so unset it now - unsetFlag fc impName TCInline + unsetFlag vfc impName TCInline -- Reset the open hints (remove the named implementation) setOpenHints hs - pure ()) mbody - where + + where + vfc : FC + vfc = virtualiseFC ifc + applyEnv : Name -> Core (Name, (Maybe Name, List (Var vars), FC -> NameType -> Term vars)) applyEnv n = do n' <- resolveName n pure (Resolved n', (Nothing, reverse (allVars env), - \fn, nt => applyToFull fc - (Ref fc nt (Resolved n')) env)) + \fn, nt => applyToFull vfc + (Ref vfc nt (Resolved n')) env)) -- For the method fields in the record, get the arguments we need to abstract -- over @@ -299,7 +302,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im impsApply : RawImp -> List (Name, RawImp) -> RawImp impsApply fn [] = fn impsApply fn ((n, arg) :: ns) - = impsApply (INamedApp fc fn n arg) ns + = impsApply (INamedApp vfc fn n arg) ns autoImpsApply : RawImp -> List RawImp -> RawImp autoImpsApply f [] = f @@ -308,7 +311,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im mkLam : List (Name, RigCount, PiInfo RawImp) -> RawImp -> RawImp mkLam [] tm = tm mkLam ((x, c, p) :: xs) tm - = ILam EmptyFC c p (Just x) (Implicit fc False) (mkLam xs tm) + = ILam EmptyFC c p (Just x) (Implicit vfc False) (mkLam xs tm) applyTo : RawImp -> List (Name, RigCount, PiInfo RawImp) -> RawImp applyTo tm [] = tm @@ -335,7 +338,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im mkLam argns (impsApply (applyTo (IVar EmptyFC n) argns) - (map (\n => (n, IVar fc (UN (show n)))) imps)) + (map (\n => (n, IVar vfc (UN (show n)))) imps)) where applyUpdate : (Name, RigCount, PiInfo RawImp) -> (Name, RigCount, PiInfo RawImp) @@ -354,12 +357,12 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im applyCon : Name -> Name -> Core (Name, RawImp) applyCon impl n = do mn <- inCurrentNS (methName n) - pure (dropNS n, IVar fc mn) + pure (dropNS n, IVar vfc mn) bindImps : List (Name, RigCount, RawImp) -> RawImp -> RawImp bindImps [] ty = ty bindImps ((n, c, t) :: ts) ty - = IPi fc c Implicit (Just n) t (bindImps ts ty) + = IPi vfc c Implicit (Just n) t (bindImps ts ty) -- Return method name, specialised method name, implicit name updates, -- and method type. Also return how the method name should be updated @@ -390,7 +393,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im -- substitute in the explicit parameters. let mty_iparams = substBindVars vars - (map (\n => (n, Implicit fc False)) imppnames) + (map (\n => (n, Implicit vfc False)) imppnames) mty_in let mty_params = substNames vars (zip pnames ps) mty_iparams @@ -400,7 +403,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im show mty_params let mbase = bindImps methImps $ - bindConstraints fc AutoImplicit cons $ + bindConstraints vfc AutoImplicit cons $ mty_params let ibound = findImplicits mbase @@ -417,8 +420,8 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im log "elab.implementation" 10 $ "Used names " ++ show ibound let ibinds = map fst methImps let methupds' = if isNil ibinds then [] - else [(n, impsApply (IVar fc n) - (map (\x => (x, IBindVar fc (show x))) ibinds))] + else [(n, impsApply (IVar vfc n) + (map (\x => (x, IBindVar vfc (show x))) ibinds))] pure ((meth.name, n, upds, meth.count, meth.totalReq, mty), methupds') @@ -437,7 +440,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im mkTopMethDecl : (Name, Name, List (String, String), RigCount, Maybe TotalReq, RawImp) -> ImpDecl mkTopMethDecl (mn, n, upds, c, treq, mty) = let opts = maybe opts_in (\t => Totality t :: opts_in) treq in - IClaim fc c vis opts (MkImpTy EmptyFC EmptyFC n mty) + IClaim vfc c vis opts (MkImpTy EmptyFC EmptyFC n mty) -- Given the method type (result of topMethType) return the mapping from -- top level method name to current implementation's method name @@ -489,9 +492,10 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im updateBody ns (IDef fc n cs) = do cs' <- traverse (updateClause ns) cs n' <- findMethName ns fc n + log "ide-mode.highlight" 1 $ show (n, n', fc) pure (IDef fc n' cs') - updateBody ns _ - = throw (GenericMsg fc + updateBody ns e + = throw (GenericMsg (getFC e) "Implementation body can only contain definitions") addTransform : Name -> List (Name, Name) -> @@ -501,16 +505,16 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im = do log "elab.implementation" 3 $ "Adding transform for " ++ show meth.name ++ " : " ++ show meth.type ++ "\n\tfor " ++ show iname ++ " in " ++ show ns - let lhs = INamedApp fc (IVar fc meth.name) - (UN "__con") - (IVar fc iname) + let lhs = INamedApp vfc (IVar vfc meth.name) + (UN "__con") + (IVar vfc iname) let Just mname = lookup (dropNS meth.name) ns | Nothing => pure () - let rhs = IVar fc mname + let rhs = IVar vfc mname log "elab.implementation" 5 $ show lhs ++ " ==> " ++ show rhs handleUnify (processDecl [] nest env - (ITransform fc (UN (show meth.name ++ " " ++ show iname)) lhs rhs)) + (ITransform vfc (UN (show meth.name ++ " " ++ show iname)) lhs rhs)) (\err => log "elab.implementation" 5 $ "Can't add transform " ++ show lhs ++ " ==> " ++ show rhs ++ diff --git a/src/Idris/Elab/Interface.idr b/src/Idris/Elab/Interface.idr index 16a2ca45b..c50aeca2e 100644 --- a/src/Idris/Elab/Interface.idr +++ b/src/Idris/Elab/Interface.idr @@ -4,6 +4,7 @@ import Core.Binary import Core.Context import Core.Context.Log import Core.Core +import Core.Name import Core.Env import Core.Metadata import Core.TT @@ -111,20 +112,24 @@ mkIfaceData : {vars : _} -> List (Maybe Name, RigCount, RawImp) -> Name -> Name -> List (Name, (RigCount, RawImp)) -> List Name -> List (Name, RigCount, RawImp) -> Core ImpDecl -mkIfaceData {vars} fc vis env constraints n conName ps dets meths +mkIfaceData {vars} ifc vis env constraints n conName ps dets meths = let opts = if isNil dets then [NoHints, UniqueSearch] else [NoHints, UniqueSearch, SearchBy dets] pNames = map fst ps - retty = apply (IVar fc n) (map (IVar EmptyFC) pNames) + retty = apply (IVar vfc n) (map (IVar EmptyFC) pNames) conty = mkTy Implicit (map jname ps) $ mkTy AutoImplicit (map bhere constraints) (mkTy Explicit (map bname meths) retty) con = MkImpTy EmptyFC EmptyFC conName !(bindTypeNames [] (pNames ++ map fst meths ++ vars) conty) in - pure $ IData fc vis (MkImpData fc n + pure $ IData vfc vis (MkImpData vfc n !(bindTypeNames [] (pNames ++ map fst meths ++ vars) - (mkDataTy fc ps)) + (mkDataTy vfc ps)) opts [con]) where + + vfc : FC + vfc = virtualiseFC ifc + jname : (Name, (RigCount, RawImp)) -> (Maybe Name, RigCount, RawImp) jname (n, rig, t) = (Just n, rig, t) @@ -138,7 +143,7 @@ mkIfaceData {vars} fc vis env constraints n conName ps dets meths List (Maybe Name, RigCount, RawImp) -> RawImp -> RawImp mkTy imp [] ret = ret mkTy imp ((n, c, argty) :: args) ret - = IPi fc c imp n argty (mkTy imp args ret) + = IPi vfc c imp n argty (mkTy imp args ret) -- Get the implicit arguments for a method declaration or constraint hint -- to allow us to build the data declaration @@ -181,31 +186,31 @@ getMethToplevel : {vars : _} -> Core (List ImpDecl) getMethToplevel {vars} env vis iname cname constraints allmeths params sig = do let paramNames = map fst params - let ity = apply (IVar fc iname) (map (IVar EmptyFC) paramNames) + let ity = apply (IVar vfc iname) (map (IVar EmptyFC) paramNames) -- Make the constraint application explicit for any method names -- which appear in other method types let ty_constr = substNames vars (map applyCon allmeths) sig.type - ty_imp <- bindTypeNames [] vars (bindPs params $ bindIFace fc ity ty_constr) + ty_imp <- bindTypeNames [] vars (bindPs params $ bindIFace vfc ity ty_constr) cn <- inCurrentNS sig.name - let tydecl = IClaim fc sig.count vis (if sig.isData then [Inline, Invertible] + let tydecl = IClaim vfc sig.count vis (if sig.isData then [Inline, Invertible] else [Inline]) - (MkImpTy fc sig.nameLoc cn ty_imp) - let conapp = apply (IVar fc cname) - (map (IBindVar EmptyFC) (map bindName allmeths)) + (MkImpTy vfc sig.nameLoc cn ty_imp) + let conapp = apply (IVar vfc cname) + (map (IBindVar EmptyFC) (map bindName allmeths)) let argns = getExplicitArgs 0 sig.type -- eta expand the RHS so that we put implicits in the right place - let fnclause = PatClause fc (INamedApp fc (IVar fc cn) + let fnclause = PatClause vfc (INamedApp vfc (IVar sig.location cn) (UN "__con") conapp) (mkLam argns (apply (IVar EmptyFC (methName sig.name)) (map (IVar EmptyFC) argns))) - let fndef = IDef fc cn [fnclause] + let fndef = IDef vfc cn [fnclause] pure [tydecl, fndef] where - fc : FC - fc = sig.location + vfc : FC + vfc = virtualiseFC sig.location -- Bind the type parameters given explicitly - there might be information -- in there that we can't infer after all @@ -216,7 +221,7 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig applyCon : Name -> (Name, RawImp) applyCon n = let name = UN "__con" in - (n, INamedApp fc (IVar fc n) name (IVar fc name)) + (n, INamedApp vfc (IVar vfc n) name (IVar vfc name)) getExplicitArgs : Int -> RawImp -> List Name getExplicitArgs i (IPi _ _ Explicit n _ sc) @@ -227,7 +232,7 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig mkLam : List Name -> RawImp -> RawImp mkLam [] tm = tm mkLam (x :: xs) tm - = ILam EmptyFC top Explicit (Just x) (Implicit fc False) (mkLam xs tm) + = ILam EmptyFC top Explicit (Just x) (Implicit vfc False) (mkLam xs tm) bindName : Name -> String bindName (UN n) = "__bind_" ++ n @@ -337,10 +342,10 @@ elabInterface : {vars : _} -> (conName : Maybe Name) -> List ImpDecl -> Core () -elabInterface {vars} fc vis env nest constraints iname params dets mcon body +elabInterface {vars} ifc vis env nest constraints iname params dets mcon body = do fullIName <- getFullName iname ns_iname <- inCurrentNS fullIName - let conName_in = maybe (mkCon fc fullIName) id mcon + let conName_in = maybe (mkCon vfc fullIName) id mcon -- Machine generated names need to be qualified when looking them up conName <- inCurrentNS conName_in let meth_sigs = mapMaybe getSig body @@ -357,13 +362,16 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body pure (record { name = n } mt)) meth_decls defs <- get Ctxt Just ty <- lookupTyExact ns_iname (gamma defs) - | Nothing => undefinedName fc iname + | Nothing => undefinedName ifc iname let implParams = getImplParams ty updateIfaceSyn ns_iname conName implParams paramNames (map snd constraints) ns_meths ds where + vfc : FC + vfc = virtualiseFC ifc + paramNames : List Name paramNames = map fst params @@ -387,7 +395,7 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body consts <- traverse (getMethDecl env nest params meth_names . (top,)) constraints log "elab.interface" 5 $ "Constraints: " ++ show consts - dt <- mkIfaceData fc vis env consts iname conName params + dt <- mkIfaceData vfc vis env consts iname conName params dets meths log "elab.interface" 10 $ "Methods: " ++ show meths log "elab.interface" 5 $ "Making interface data type " ++ show dt @@ -406,9 +414,9 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body log "elab.interface" 5 $ "Top level methods: " ++ show fns traverse_ (processDecl [] nest env) fns traverse_ (\n => do mn <- inCurrentNS n - setFlag fc mn Inline - setFlag fc mn TCInline - setFlag fc mn Overloadable) meth_names + setFlag vfc mn Inline + setFlag vfc mn TCInline + setFlag vfc mn Overloadable) meth_names -- Check that a default definition is correct. We just discard it here once -- we know it's okay, since we'll need to re-elaborate it for each @@ -416,8 +424,9 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body elabDefault : List Declaration -> (FC, List FnOpt, Name, List ImpClause) -> Core (Name, List ImpClause) - elabDefault tydecls (fc, opts, n, cs) + elabDefault tydecls (dfc, opts, n, cs) = do -- orig <- branch + let dn_in = MN ("Default implementation of " ++ show n) 0 dn <- inCurrentNS dn_in @@ -425,37 +434,40 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body the (Core (RigCount, RawImp)) $ case findBy (\ d => d <$ guard (n == d.name)) tydecls of Just d => pure (d.count, d.type) - Nothing => throw (GenericMsg fc ("No method named " ++ show n ++ " in interface " ++ show iname)) + Nothing => throw (GenericMsg dfc ("No method named " ++ show n ++ " in interface " ++ show iname)) - let ity = apply (IVar fc iname) (map (IVar fc) paramNames) + let ity = apply (IVar vdfc iname) (map (IVar vdfc) paramNames) -- Substitute the method names with their top level function -- name, so they don't get implicitly bound in the name methNameMap <- traverse (\d => do let n = d.name cn <- inCurrentNS n - pure (n, applyParams (IVar fc cn) paramNames)) + pure (n, applyParams (IVar vdfc cn) paramNames)) tydecls - let dty = bindPs params -- bind parameters - $ bindIFace fc ity -- bind interface (?!) + let dty = bindPs params -- bind parameters + $ bindIFace vdfc ity -- bind interface (?!) $ substNames vars methNameMap dty dty_imp <- bindTypeNames [] (map name tydecls ++ vars) dty log "elab.interface.default" 5 $ "Default method " ++ show dn ++ " : " ++ show dty_imp - let dtydecl = IClaim fc rig vis [] (MkImpTy EmptyFC EmptyFC dn dty_imp) + let dtydecl = IClaim vdfc rig vis [] (MkImpTy EmptyFC EmptyFC dn dty_imp) processDecl [] nest env dtydecl - let cs' = map (changeName dn) cs + cs' <- traverse (changeName dn) cs log "elab.interface.default" 5 $ "Default method body " ++ show cs' - processDecl [] nest env (IDef fc dn cs') + processDecl [] nest env (IDef vdfc dn cs') -- Reset the original context, we don't need to keep the definition -- Actually we do for the metadata and name map! -- put Ctxt orig pure (n, cs) where + vdfc : FC + vdfc = virtualiseFC dfc + -- Bind the type parameters given explicitly - there might be information -- in there that we can't infer after all bindPs : List (Name, (RigCount, RawImp)) -> RawImp -> RawImp @@ -466,38 +478,48 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body applyParams : RawImp -> List Name -> RawImp applyParams tm [] = tm applyParams tm (UN n :: ns) - = applyParams (INamedApp fc tm (UN n) (IBindVar fc n)) ns + = applyParams (INamedApp vdfc tm (UN n) (IBindVar vdfc n)) ns applyParams tm (_ :: ns) = applyParams tm ns - changeNameTerm : Name -> RawImp -> RawImp + changeNameTerm : Name -> RawImp -> Core RawImp changeNameTerm dn (IVar fc n') - = if n == n' then IVar fc dn else IVar fc n' + = do if n /= n' then pure (IVar fc n') else do + log "ide-mode.highlight" 7 $ + "elabDefault is trying to add Function: " ++ show n ++ " (" ++ show fc ++")" + whenJust (isConcreteFC fc) \nfc => do + log "ide-mode.highlight" 7 $ "elabDefault is adding Function: " ++ show n + addSemanticDecorations [(nfc, Function, Just n)] + pure (IVar fc dn) changeNameTerm dn (IApp fc f arg) - = IApp fc (changeNameTerm dn f) arg + = IApp fc <$> changeNameTerm dn f <*> pure arg changeNameTerm dn (IAutoApp fc f arg) - = IAutoApp fc (changeNameTerm dn f) arg + = IAutoApp fc <$> changeNameTerm dn f <*> pure arg changeNameTerm dn (INamedApp fc f x arg) - = INamedApp fc (changeNameTerm dn f) x arg - changeNameTerm dn tm = tm + = INamedApp fc <$> changeNameTerm dn f <*> pure x <*> pure arg + changeNameTerm dn tm = pure tm - changeName : Name -> ImpClause -> ImpClause + changeName : Name -> ImpClause -> Core ImpClause changeName dn (PatClause fc lhs rhs) - = PatClause fc (changeNameTerm dn lhs) rhs + = PatClause fc <$> changeNameTerm dn lhs <*> pure rhs changeName dn (WithClause fc lhs wval prf flags cs) - = WithClause fc (changeNameTerm dn lhs) wval - prf flags (map (changeName dn) cs) + = WithClause fc + <$> changeNameTerm dn lhs + <*> pure wval + <*> pure prf + <*> pure flags + <*> traverse (changeName dn) cs changeName dn (ImpossibleClause fc lhs) - = ImpossibleClause fc (changeNameTerm dn lhs) + = ImpossibleClause fc <$> changeNameTerm dn lhs elabConstraintHints : (conName : Name) -> List Name -> Core () elabConstraintHints conName meth_names = do let nconstraints = nameCons 0 constraints - chints <- traverse (getConstraintHint fc env vis iname conName + chints <- traverse (getConstraintHint vfc env vis iname conName (map fst nconstraints) meth_names paramNames) nconstraints log "elab.interface" 5 $ "Constraint hints from " ++ show constraints ++ ": " ++ show chints traverse_ (processDecl [] nest env) (concatMap snd chints) traverse_ (\n => do mn <- inCurrentNS n - setFlag fc mn TCInline) (map fst chints) + setFlag vfc mn TCInline) (map fst chints) diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index 09d0eabf6..fef3d6490 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -66,8 +66,9 @@ pshowNoNorm env tm ploc : {auto o : Ref ROpts REPLOpts} -> FC -> Core (Doc IdrisAnn) -ploc EmptyFC = pure emptyDoc -ploc fc@(MkFC fn s e) = do +ploc fc = do + let Just (fn, s, e) = isNonEmptyFC fc + | Nothing => pure emptyDoc let (sr, sc) = mapHom (fromInteger . cast) s let (er, ec) = mapHom (fromInteger . cast) e let nsize = length $ show (er + 1) @@ -91,10 +92,12 @@ ploc fc@(MkFC fn s e) = do -- Assumes the two FCs are sorted ploc2 : {auto o : Ref ROpts REPLOpts} -> FC -> FC -> Core (Doc IdrisAnn) -ploc2 fc EmptyFC = ploc fc -ploc2 EmptyFC fc = ploc fc -ploc2 (MkFC fn1 s1 e1) (MkFC fn2 s2 e2) = - do let (sr1, sc1) = mapHom (fromInteger . cast) s1 +ploc2 fc1 fc2 = + do let Just (fn1, s1, e1) = isNonEmptyFC fc1 + | Nothing => ploc fc2 + let Just (fn2, s2, e2) = isNonEmptyFC fc2 + | Nothing => ploc fc1 + let (sr1, sc1) = mapHom (fromInteger . cast) s1 let (sr2, sc2) = mapHom (fromInteger . cast) s2 let (er1, ec1) = mapHom (fromInteger . cast) e1 let (er2, ec2) = mapHom (fromInteger . cast) e2 @@ -179,10 +182,12 @@ perror (PatternVariableUnifies fc env n tm) prettyVar (PV n _) = prettyVar n prettyVar n = pretty n order : FC -> FC -> (FC, FC) - order EmptyFC fc2 = (EmptyFC, fc2) - order fc1 EmptyFC = (fc1, EmptyFC) - order fc1@(MkFC _ sr1 sc1) fc2@(MkFC _ sr2 sc2) = - if sr1 < sr2 then (fc1, fc2) else if sr1 == sr2 && sc1 < sc2 then (fc1, fc2) else (fc2, fc1) + order fc1 fc2 = + let Just (_, sr1, sc1) = isNonEmptyFC fc1 + | Nothing => (EmptyFC, fc2) + Just (_, sr2, sc2) = isNonEmptyFC fc2 + | Nothing => (fc1, EmptyFC) + in if sr1 < sr2 then (fc1, fc2) else if sr1 == sr2 && sc1 < sc2 then (fc1, fc2) else (fc2, fc1) perror (CyclicMeta fc env n tm) = pure $ errorDesc (reflow "Cycle detected in solution of metavariable" <++> meta (pretty !(prettyName n)) <++> equals <++> code !(pshow env tm)) <+> line <+> !(ploc fc) diff --git a/src/Idris/IDEMode/CaseSplit.idr b/src/Idris/IDEMode/CaseSplit.idr index 802e58545..75af5fdd2 100644 --- a/src/Idris/IDEMode/CaseSplit.idr +++ b/src/Idris/IDEMode/CaseSplit.idr @@ -48,7 +48,7 @@ toStrUpdate (UN n, term) where bracket : PTerm -> PTerm bracket tm@(PRef _ _) = tm - bracket tm@(PList _ _) = tm + bracket tm@(PList _ _ _) = tm bracket tm@(PPair _ _ _) = tm bracket tm@(PUnit _) = tm bracket tm@(PComprehension _ _ _) = tm diff --git a/src/Idris/IDEMode/Commands.idr b/src/Idris/IDEMode/Commands.idr index d75923875..ccf85c208 100644 --- a/src/Idris/IDEMode/Commands.idr +++ b/src/Idris/IDEMode/Commands.idr @@ -1,6 +1,8 @@ module Idris.IDEMode.Commands import Core.Core +import Core.Context +import Core.Context.Log import Core.Name import public Idris.REPL.Opts import Libraries.Utils.Hex @@ -281,9 +283,10 @@ sendStr f st = map (const ()) (fPutStr f st) export -send : SExpable a => File -> a -> Core () +send : {auto c : Ref Ctxt Defs} -> SExpable a => File -> a -> Core () send f resp = do let r = show (toSExp resp) ++ "\n" + log "ide-mode.send" 20 r coreLift $ sendStr f $ leftPad '0' 6 (asHex (cast (length r))) coreLift $ sendStr f r coreLift $ fflush f diff --git a/src/Idris/IDEMode/Parser.idr b/src/Idris/IDEMode/Parser.idr index c6928105c..349332678 100644 --- a/src/Idris/IDEMode/Parser.idr +++ b/src/Idris/IDEMode/Parser.idr @@ -5,6 +5,9 @@ module Idris.IDEMode.Parser import Idris.IDEMode.Commands import Core.Core +import Core.Name +import Core.Metadata +import Core.FC import Data.Maybe import Data.List @@ -76,11 +79,11 @@ sexp pure (SExpList xs) ideParser : {e : _} -> - (fname : String) -> String -> Grammar Token e ty -> Either Error ty + (fname : String) -> String -> Grammar SemanticDecorations Token e ty -> Either Error ty ideParser fname str p = do toks <- mapError (fromLexError fname) $ idelex str - parsed <- mapError (fromParsingError fname) $ parse p toks - Right (fst parsed) + (decor, (parsed, _)) <- mapError (fromParsingError fname) $ parseWith p toks + Right parsed export diff --git a/src/Idris/IDEMode/REPL.idr b/src/Idris/IDEMode/REPL.idr index 170af5ad2..4ce0c76e9 100644 --- a/src/Idris/IDEMode/REPL.idr +++ b/src/Idris/IDEMode/REPL.idr @@ -133,7 +133,8 @@ getInput f pure (pack inp) ||| Do nothing and tell the user to wait for us to implmement this (or join the effort!) -todoCmd : {auto o : Ref ROpts REPLOpts} -> +todoCmd : {auto c : Ref Ctxt Defs} -> + {auto o : Ref ROpts REPLOpts} -> String -> Core () todoCmd cmdName = iputStrLn $ reflow $ cmdName ++ ": command not yet implemented. Hopefully soon!" @@ -260,24 +261,24 @@ processCatch cmd msg <- perror err pure $ REPL $ REPLError msg) -idePutStrLn : File -> Integer -> String -> Core () +idePutStrLn : {auto c : Ref Ctxt Defs} -> File -> Integer -> String -> Core () idePutStrLn outf i msg = send outf (SExpList [SymbolAtom "write-string", toSExp msg, toSExp i]) -returnFromIDE : File -> Integer -> SExp -> Core () +returnFromIDE : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core () returnFromIDE outf i msg = do send outf (SExpList [SymbolAtom "return", msg, toSExp i]) -printIDEResult : File -> Integer -> SExp -> Core () +printIDEResult : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core () printIDEResult outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "ok", toSExp msg]) -printIDEResultWithHighlight : File -> Integer -> SExp -> Core () +printIDEResultWithHighlight : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core () printIDEResultWithHighlight outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "ok", toSExp msg -- TODO return syntax highlighted result , SExpList []]) -printIDEError : Ref ROpts REPLOpts => File -> Integer -> Doc IdrisAnn -> Core () +printIDEError : Ref ROpts REPLOpts => {auto c : Ref Ctxt Defs} -> File -> Integer -> Doc IdrisAnn -> Core () printIDEError outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "error", toSExp !(renderWithoutColor msg) ]) SExpable REPLEval where diff --git a/src/Idris/IDEMode/SyntaxHighlight.idr b/src/Idris/IDEMode/SyntaxHighlight.idr index 7873762f4..3a61ac660 100644 --- a/src/Idris/IDEMode/SyntaxHighlight.idr +++ b/src/Idris/IDEMode/SyntaxHighlight.idr @@ -1,35 +1,34 @@ module Idris.IDEMode.SyntaxHighlight import Core.Context +import Core.Context.Log import Core.InitPrimitives import Core.Metadata import Core.TT import Idris.REPL +import Idris.Syntax +import Idris.DocString import Idris.IDEMode.Commands import Data.List +import Data.Maybe + +import Libraries.Data.PosMap %default covering -data Decoration : Type where - Typ : Decoration - Function : Decoration - Data : Decoration - Keyword : Decoration - Bound : Decoration - SExpable Decoration where - toSExp Typ = SymbolAtom "type" - toSExp Function = SymbolAtom "function" - toSExp Data = SymbolAtom "data" - toSExp Keyword = SymbolAtom "keyword" - toSExp Bound = SymbolAtom "bound" + toSExp Typ = SExpList [ SymbolAtom "decor", SymbolAtom "type"] + toSExp Function = SExpList [ SymbolAtom "decor", SymbolAtom "function"] + toSExp Data = SExpList [ SymbolAtom "decor", SymbolAtom "data"] + toSExp Keyword = SExpList [ SymbolAtom "decor", SymbolAtom "keyword"] + toSExp Bound = SExpList [ SymbolAtom "decor", SymbolAtom "bound"] record Highlight where constructor MkHighlight location : NonEmptyFC - name : Name + name : String isImplicit : Bool key : String decor : Decoration @@ -38,19 +37,26 @@ record Highlight where ns : String SExpable FC where - toSExp (MkFC fname (startLine, startCol) (endLine, endCol)) - = SExpList [ SExpList [ SymbolAtom "filename", StringAtom fname ] - , SExpList [ SymbolAtom "start", IntegerAtom (cast startLine + 1), IntegerAtom (cast startCol + 1) ] - , SExpList [ SymbolAtom "end", IntegerAtom (cast endLine + 1), IntegerAtom (cast endCol + 1) ] + toSExp fc = case isNonEmptyFC fc of + Just (fname , (startLine, startCol), (endLine, endCol)) => + SExpList [ SExpList [ SymbolAtom "filename", StringAtom fname ] + , SExpList [ SymbolAtom "start" + , IntegerAtom (cast startLine + 1) + , IntegerAtom (cast startCol + 1) + ] + , SExpList [ SymbolAtom "end" + , IntegerAtom (cast endLine + 1) + , IntegerAtom (cast endCol) + ] ] - toSExp EmptyFC = SExpList [] + Nothing => SExpList [] SExpable Highlight where toSExp (MkHighlight loc nam impl k dec doc t ns) = SExpList [ toSExp $ justFC loc - , SExpList [ SExpList [ SymbolAtom "name", StringAtom (show nam) ] + , SExpList [ SExpList [ SymbolAtom "name", StringAtom nam ] , SExpList [ SymbolAtom "namespace", StringAtom ns ] - , SExpList [ SymbolAtom "decor", toSExp dec ] + , toSExp dec , SExpList [ SymbolAtom "implicit", toSExp impl ] , SExpList [ SymbolAtom "key", StringAtom k ] , SExpList [ SymbolAtom "doc-overview", StringAtom doc ] @@ -58,12 +64,13 @@ SExpable Highlight where ] ] -inFile : String -> (NonEmptyFC, (Name, Nat, ClosedTerm)) -> Bool +inFile : (s : String) -> (NonEmptyFC, a) -> Bool inFile fname ((file, _, _), _) = file == fname ||| Output some data using current dialog index export -printOutput : {auto o : Ref ROpts REPLOpts} -> +printOutput : {auto c : Ref Ctxt Defs} -> + {auto o : Ref ROpts REPLOpts} -> SExp -> Core () printOutput msg = do opts <- get ROpts @@ -74,7 +81,8 @@ printOutput msg msg, toSExp i]) -outputHighlight : {auto opts : Ref ROpts REPLOpts} -> +outputHighlight : {auto c : Ref Ctxt Defs} -> + {auto opts : Ref ROpts REPLOpts} -> Highlight -> Core () outputHighlight h = printOutput $ SExpList [ SymbolAtom "ok" @@ -86,37 +94,46 @@ outputHighlight h = hlt : List Highlight hlt = [h] -outputNameSyntax : {auto opts : Ref ROpts REPLOpts} -> - (NonEmptyFC, (Name, Nat, ClosedTerm)) -> Core () -outputNameSyntax (fc, (name, _, term)) = - let dec = case term of - (Local fc x idx y) => Just Bound +lwOutputHighlight : + {auto c : Ref Ctxt Defs} -> + {auto opts : Ref ROpts REPLOpts} -> + (NonEmptyFC, Decoration) -> Core () +lwOutputHighlight (nfc,decor) = + printOutput $ SExpList [ SymbolAtom "ok" + , SExpList [ SymbolAtom "highlight-source" + , toSExp $ the (List _) [ + SExpList [ toSExp $ justFC nfc + , SExpList [ toSExp decor] + ]]]] - -- See definition of NameType in Core.TT for possible values of Ref's nametype field - -- data NameType : Type where - -- Bound : NameType - -- Func : NameType - -- DataCon : (tag : Int) -> (arity : Nat) -> NameType - -- TyCon : (tag : Int) -> (arity : Nat) -> NameType - (Ref fc Bound name) => Just Bound - (Ref fc Func name) => Just Function - (Ref fc (DataCon tag arity) name) => Just Data - (Ref fc (TyCon tag arity) name) => Just Typ - (Meta fc x y xs) => Just Bound - (Bind fc x b scope) => Just Bound - (App fc fn arg) => Just Bound - (As fc x as pat) => Just Bound - (TDelayed fc x y) => Nothing - (TDelay fc x ty arg) => Nothing - (TForce fc x y) => Nothing - (PrimVal fc c) => Just Typ - (Erased fc imp) => Just Bound - (TType fc) => Just Typ - hilite = Prelude.map (\ d => MkHighlight fc name False "" d "" (show term) "") dec - in maybe (pure ()) outputHighlight hilite + + +outputNameSyntax : {auto c : Ref Ctxt Defs} -> + {auto s : Ref Syn SyntaxInfo} -> + {auto opts : Ref ROpts REPLOpts} -> + (NonEmptyFC, Decoration, Name) -> Core () +outputNameSyntax (nfc, decor, nm) = do + defs <- get Ctxt + log "ide-mode.highlight" 20 $ "highlighting at " ++ show nfc + ++ ": " ++ show nm + ++ "\nAs: " ++ show decor + let fc = justFC nfc + let (mns, name) = displayName nm + outputHighlight $ MkHighlight + { location = nfc + , name + , isImplicit = False + , key = "" + , decor + , docOverview = "" --!(getDocsForName fc nm) + , typ = "" -- TODO: extract type maybe "" show !(lookupTyExact nm (gamma defs)) + , ns = maybe "" show mns + } export -outputSyntaxHighlighting : {auto m : Ref MD Metadata} -> +outputSyntaxHighlighting : {auto c : Ref Ctxt Defs} -> + {auto m : Ref MD Metadata} -> + {auto s : Ref Syn SyntaxInfo} -> {auto opts : Ref ROpts REPLOpts} -> String -> REPLResult -> @@ -124,9 +141,28 @@ outputSyntaxHighlighting : {auto m : Ref MD Metadata} -> outputSyntaxHighlighting fname loadResult = do opts <- get ROpts when (opts.synHighlightOn) $ do - allNames <- the (Core ?) $ filter (inFile fname) . names <$> get MD - -- decls <- filter (inFile fname) . tydecls <$> get MD - _ <- traverse outputNameSyntax allNames -- ++ decls) - pure () + meta <- get MD + let allNames = filter (inFile fname) $ toList meta.nameLocMap + --decls <- filter (inFile fname) . tydecls <$> get MD + --_ <- traverse outputNameSyntax allNames -- ++ decls) + + let semHigh = meta.semanticHighlighting + log "ide-mode.highlight" 19 $ + "Semantic metadata is: " ++ show semHigh + + let aliases + : List ASemanticDecoration + = flip foldMap meta.semanticAliases $ \ (from, to) => + let decors = uncurry exactRange (snd to) semHigh in + map (\ ((fnm, loc), rest) => ((fnm, snd from), rest)) decors + log "ide-mode.highlight.alias" 19 $ + "Semantic metadata from aliases is: " ++ show aliases + + traverse_ {b = Unit} + (\(nfc, decor, mn) => + case mn of + Nothing => lwOutputHighlight (nfc, decor) + Just n => outputNameSyntax (nfc, decor, n)) + (aliases ++ toList semHigh) pure loadResult diff --git a/src/Idris/ModTree.idr b/src/Idris/ModTree.idr index 55ece0467..54738ac28 100644 --- a/src/Idris/ModTree.idr +++ b/src/Idris/ModTree.idr @@ -178,7 +178,7 @@ buildMod loc num len mod Nothing => True Just t => any (\x => x > t) (srcTime :: map snd depTimes) u <- newRef UST initUState - m <- newRef MD initMetadata + m <- newRef MD (initMetadata src) put Syn initSyntax if needsBuilding @@ -220,7 +220,7 @@ buildDeps fname case ok of [] => do -- On success, reload the main ttc in a clean context clearCtxt; addPrimitives - put MD initMetadata + put MD (initMetadata fname) mainttc <- getTTCFileName fname "ttc" log "import" 10 $ "Reloading " ++ show mainttc ++ " from " ++ fname readAsMain mainttc diff --git a/src/Idris/Package.idr b/src/Idris/Package.idr index bfa17d988..af5444992 100644 --- a/src/Idris/Package.idr +++ b/src/Idris/Package.idr @@ -159,7 +159,7 @@ field fname pure [LT (MkPkgVersion (fromInteger <$> vs)) True, GT (MkPkgVersion (fromInteger <$> vs)) True] - mkBound : List Bound -> PkgVersionBounds -> PackageEmptyRule PkgVersionBounds + mkBound : List Bound -> PkgVersionBounds -> EmptyRule PkgVersionBounds mkBound (LT b i :: bs) pkgbs = maybe (mkBound bs (record { upperBound = Just b, upperInclusive = i } pkgbs)) @@ -285,7 +285,7 @@ compileMain : {auto c : Ref Ctxt Defs} -> {auto o : Ref ROpts REPLOpts} -> Name -> String -> String -> Core () compileMain mainn mmod exec - = do m <- newRef MD initMetadata + = do m <- newRef MD (initMetadata mmod) u <- newRef UST initUState ignore $ loadMainFile mmod ignore $ compileExp (PRef replFC mainn) exec @@ -559,7 +559,7 @@ runRepl : {auto c : Ref Ctxt Defs} -> Core () runRepl fname = do u <- newRef UST initUState - m <- newRef MD initMetadata + m <- newRef MD (initMetadata $ fromMaybe "(interactive)" fname) the (Core ()) $ case fname of Nothing => pure () diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index 43831a2ca..1812ef2e1 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -1,6 +1,7 @@ module Idris.Parser import Core.Options +import Core.Metadata import Idris.Syntax import public Parser.Source import Parser.Lexer.Source @@ -19,6 +20,52 @@ import Idris.Parser.Let %default covering +decorate : FileName -> Decoration -> Rule a -> Rule a +decorate fname decor rule = do + res <- bounds rule + act [((fname, (start res, end res)), decor, Nothing)] + pure res.val + + +decorationFromBounded : FileName -> Decoration -> WithBounds a -> ASemanticDecoration +decorationFromBounded fname decor bnds + = ((fname, start bnds, end bnds), decor, Nothing) + +boundedNameDecoration : FileName -> Decoration -> WithBounds Name -> ASemanticDecoration +boundedNameDecoration fname decor bstr = ((fname, start bstr, end bstr) + , decor + , Just bstr.val) + +decorateBoundedNames : FileName -> Decoration -> List (WithBounds Name) -> EmptyRule () +decorateBoundedNames fname decor bns = act $ map (boundedNameDecoration fname decor) bns + +decorateBoundedName : FileName -> Decoration -> WithBounds Name -> EmptyRule () +decorateBoundedName fname decor bn = act [boundedNameDecoration fname decor bn] + +decorateKeywords : FileName -> List (WithBounds a) -> EmptyRule () +decorateKeywords fname = act . map (decorationFromBounded fname Keyword) + +dependentDecorate : FileName -> Rule a -> (a -> Decoration) -> Rule a +dependentDecorate fname rule decor = do + res <- bounds rule + act [((fname, (start res, end res)), decor res.val, Nothing)] + pure res.val + +decoratedKeyword : FileName -> String -> Rule () +decoratedKeyword fname kwd = decorate fname Keyword (keyword kwd) + +decoratedSymbol : FileName -> String -> Rule () +decoratedSymbol fname smb = decorate fname Keyword (symbol smb) + +decoratedDataTypeName : FileName -> Rule Name +decoratedDataTypeName fname = decorate fname Typ dataTypeName + +decoratedDataConstructorName : FileName -> Rule Name +decoratedDataConstructorName fname = decorate fname Data dataConstructorName + +decoratedSimpleBinderName : FileName -> Rule String +decoratedSimpleBinderName fname = decorate fname Bound unqualifiedName + -- Forward declare since they're used in the parser topDecl : FileName -> IndentInfo -> Rule (List PDecl) collectDefs : List PDecl -> List PDecl @@ -58,36 +105,39 @@ plhs = MkParseOpts False False atom : FileName -> Rule PTerm atom fname - = do x <- bounds $ exactIdent "Type" + = do x <- bounds $ decorate fname Typ $ exactIdent "Type" pure (PType (boundToFC fname x)) <|> do x <- bounds $ name pure (PRef (boundToFC fname x) x.val) - <|> do x <- bounds $ constant + <|> do x <- bounds $ dependentDecorate fname constant \c => + if isPrimType c + then Typ + else Data pure (PPrimVal (boundToFC fname x) x.val) - <|> do x <- bounds $ symbol "_" + <|> do x <- bounds $ decoratedSymbol fname "_" pure (PImplicit (boundToFC fname x)) <|> do x <- bounds $ symbol "?" pure (PInfer (boundToFC fname x)) <|> do x <- bounds $ holeName pure (PHole (boundToFC fname x) False x.val) - <|> do x <- bounds $ pragma "MkWorld" + <|> do x <- bounds $ decorate fname Data $ pragma "MkWorld" pure (PPrimVal (boundToFC fname x) WorldVal) - <|> do x <- bounds $ pragma "World" + <|> do x <- bounds $ decorate fname Typ $ pragma "World" pure (PPrimVal (boundToFC fname x) WorldType) <|> do x <- bounds $ pragma "search" pure (PSearch (boundToFC fname x) 50) whereBlock : FileName -> Int -> Rule (List PDecl) whereBlock fname col - = do keyword "where" + = do decoratedKeyword fname "where" ds <- blockAfter col (topDecl fname) pure (collectDefs (concat ds)) -- Expect a keyword, but if we get anything else it's a fatal error -commitKeyword : IndentInfo -> String -> Rule () -commitKeyword indents req +commitKeyword : FileName -> IndentInfo -> String -> Rule () +commitKeyword fname indents req = do mustContinue indents (Just req) - keyword req + decoratedKeyword fname req mustContinue indents Nothing commitSymbol : String -> Rule () @@ -95,6 +145,11 @@ commitSymbol req = symbol req <|> fatalError ("Expected '" ++ req ++ "'") +continueWithDecorated : FileName -> IndentInfo -> String -> Rule () +continueWithDecorated fname indents req + = mustContinue indents (Just req) *> decoratedSymbol fname req + + continueWith : IndentInfo -> String -> Rule () continueWith indents req = mustContinue indents (Just req) *> symbol req @@ -127,9 +182,9 @@ mutual <|> do b <- bounds (MkPair <$> simpleExpr fname indents <*> many (argExpr q fname indents)) (f, args) <- pure b.val pure (applyExpImp (start b) (end b) f (concat args)) - <|> do b <- bounds (MkPair <$> iOperator <*> expr pdef fname indents) + <|> do b <- bounds (MkPair <$> bounds iOperator <*> expr pdef fname indents) (op, arg) <- pure b.val - pure (PPrefixOp (boundToFC fname b) op arg) + pure (PPrefixOp (boundToFC fname b) (boundToFC fname op) op.val arg) <|> fail "Expected 'case', 'if', 'do', application or operator expression" where applyExpImp : FilePos -> FilePos -> PTerm -> @@ -150,14 +205,14 @@ mutual argExpr q fname indents = do continue indents arg <- simpleExpr fname indents - the (SourceEmptyRule _) $ case arg of + the (EmptyRule _) $ case arg of PHole loc _ n => pure [UnnamedExpArg (PHole loc True n)] t => pure [UnnamedExpArg t] <|> do continue indents braceArgs fname indents <|> if withOK q then do continue indents - symbol "|" + decoratedSymbol fname "|" arg <- expr (record {withOK = False} q) fname indents pure [WithArg arg] else fail "| not allowed here" @@ -167,34 +222,34 @@ mutual braceArgs : FileName -> IndentInfo -> Rule (List ArgType) braceArgs fname indents - = do start <- bounds (symbol "{") - list <- sepBy (symbol ",") - $ do x <- bounds unqualifiedName + = do start <- bounds (decoratedSymbol fname "{") + list <- sepBy (decoratedSymbol fname ",") + $ do x <- bounds (decoratedSimpleBinderName fname) let fc = boundToFC fname x option (NamedArg (UN x.val) $ PRef fc (UN x.val)) - $ do tm <- symbol "=" *> expr pdef fname indents + $ do tm <- decoratedSymbol fname "=" *> expr pdef fname indents pure (NamedArg (UN x.val) tm) matchAny <- option [] (if isCons list then - do symbol "," - x <- bounds (symbol "_") + do decoratedSymbol fname "," + x <- bounds (decoratedSymbol fname "_") pure [underscore (boundToFC fname x)] else fail "non-empty list required") - end <- bounds (symbol "}") + end <- bounds (decoratedSymbol fname "}") matchAny <- do let fc = boundToFC fname (mergeBounds start end) pure $ if isNil list then [underscore fc] else matchAny pure $ matchAny ++ list - <|> do symbol "@{" + <|> do decoratedSymbol fname "@{" commit tm <- expr pdef fname indents - symbol "}" + decoratedSymbol fname "}" pure [UnnamedAutoArg tm] with_ : FileName -> IndentInfo -> Rule PTerm with_ fname indents - = do b <- bounds (do keyword "with" + = do b <- bounds (do decoratedKeyword fname "with" commit ns <- singleName <|> nameList end <- location @@ -210,38 +265,44 @@ mutual nameList : Rule (List Name) nameList = do - symbol "[" + decoratedSymbol fname "[" commit - ns <- sepBy1 (symbol ",") name - symbol "]" + ns <- sepBy1 (decoratedSymbol fname ",") name + decoratedSymbol fname "]" pure (forget ns) opExpr : ParseOpts -> FileName -> IndentInfo -> Rule PTerm opExpr q fname indents = do l <- bounds (appExpr q fname indents) (if eqOK q - then do r <- bounds (continue indents *> symbol "=" *> opExpr q fname indents) - pure (POp (boundToFC fname (mergeBounds l r)) (UN "=") l.val r.val) + then do r <- bounds (continue indents *> decoratedSymbol fname "=" *> opExpr q fname indents) + pure $ + let fc = boundToFC fname (mergeBounds l r) + opFC = virtualiseFC fc -- already been highlighted: we don't care + in POp fc opFC (UN "=") l.val r.val else fail "= not allowed") <|> - (do b <- bounds [| MkPair (continue indents *> iOperator) (opExpr q fname indents) |] + (do b <- bounds [| MkPair (continue indents *> bounds iOperator) (opExpr q fname indents) |] (op, r) <- pure b.val - pure (POp (boundToFC fname (mergeBounds l b)) op l.val r)) + let fc = boundToFC fname (mergeBounds l b) + let opFC = boundToFC fname op + pure (POp fc opFC op.val l.val r)) <|> pure l.val dpairType : FileName -> WithBounds t -> IndentInfo -> Rule PTerm dpairType fname start indents - = do loc <- bounds (do x <- unqualifiedName - symbol ":" + = do loc <- bounds (do x <- decoratedSimpleBinderName fname + decoratedSymbol fname ":" ty <- expr pdef fname indents pure (x, ty)) (x, ty) <- pure loc.val - (do symbol "**" - rest <- bounds (nestedDpair fname loc indents <|> expr pdef fname indents) - pure (PDPair (boundToFC fname (mergeBounds start rest)) - (PRef (boundToFC fname loc) (UN x)) - ty - rest.val)) + op <- bounds (symbol "**") + rest <- bounds (nestedDpair fname loc indents <|> expr pdef fname indents) + pure (PDPair (boundToFC fname (mergeBounds start rest)) + (boundToFC fname op) + (PRef (boundToFC fname loc) (UN x)) + ty + rest.val) nestedDpair : FileName -> WithBounds t -> IndentInfo -> Rule PTerm nestedDpair fname start indents @@ -250,6 +311,7 @@ mutual loc <- bounds (symbol "**") rest <- bounds (nestedDpair fname loc indents <|> expr pdef fname indents) pure (PDPair (boundToFC fname (mergeBounds start rest)) + (boundToFC fname loc) l (PImplicit (boundToFC fname (mergeBounds start rest))) rest.val) @@ -259,32 +321,41 @@ mutual -- left section. This may also be a prefix operator, but we'll sort -- that out when desugaring: if the operator is infix, treat it as a -- section otherwise treat it as prefix - = do b <- bounds (do op <- iOperator + = do b <- bounds (do op <- bounds iOperator e <- expr pdef fname indents - continueWith indents ")" + continueWithDecorated fname indents ")" pure (op, e)) (op, e) <- pure b.val - pure (PSectionL (boundToFC fname (mergeBounds s b)) op e) + act [(toNonEmptyFC $ boundToFC fname s, Keyword, Nothing)] + let fc = boundToFC fname (mergeBounds s b) + let opFC = boundToFC fname op + pure (PSectionL fc opFC op.val e) <|> do -- (.y.z) -- section of projection (chain) b <- bounds $ forget <$> some postfixProj - symbol ")" + decoratedSymbol fname ")" + act [(toNonEmptyFC $ boundToFC fname s, Keyword, Nothing)] pure $ PPostfixAppPartial (boundToFC fname b) b.val -- unit type/value <|> do b <- bounds (continueWith indents ")") pure (PUnit (boundToFC fname (mergeBounds s b))) -- dependent pairs with type annotation (so, the type form) - <|> do dpairType fname s indents <* symbol ")" + <|> do dpairType fname s indents <* (decorate fname Typ $ symbol ")") + <* act [(toNonEmptyFC $ boundToFC fname s, Typ, Nothing)] <|> do e <- bounds (expr pdef fname indents) -- dependent pairs with no type annotation (do loc <- bounds (symbol "**") rest <- bounds ((nestedDpair fname loc indents <|> expr pdef fname indents) <* symbol ")") pure (PDPair (boundToFC fname (mergeBounds s rest)) + (boundToFC fname loc) e.val (PImplicit (boundToFC fname (mergeBounds s rest))) rest.val)) <|> -- right sections - ((do op <- bounds (iOperator <* symbol ")") - pure (PSectionR (boundToFC fname (mergeBounds s op)) e.val op.val) + ((do op <- bounds (bounds iOperator <* decoratedSymbol fname ")") + act [(toNonEmptyFC $ boundToFC fname s, Keyword, Nothing)] + let fc = boundToFC fname (mergeBounds s op) + let opFC = boundToFC fname op.val + pure (PSectionR fc opFC e.val op.val.val) <|> -- all the other bracketed expressions tuple fname s indents e.val)) @@ -294,40 +365,50 @@ mutual ts <- bounds (nonEmptyTuple fname s indents var) pure (PLam fc top Explicit var (PInfer fc) ts.val) - getInitRange : List PTerm -> SourceEmptyRule (PTerm, Maybe PTerm) - getInitRange [x] = pure (x, Nothing) - getInitRange [x, y] = pure (x, Just y) + getInitRange : List (WithBounds PTerm) -> EmptyRule (PTerm, Maybe PTerm) + getInitRange [x] = pure (x.val, Nothing) + getInitRange [x,y] = pure (x.val, Just y.val) getInitRange _ = fatalError "Invalid list range syntax" - listRange : FileName -> WithBounds t -> IndentInfo -> List PTerm -> Rule PTerm + listRange : FileName -> WithBounds t -> IndentInfo -> List (WithBounds PTerm) -> Rule PTerm listRange fname s indents xs - = do b <- bounds (symbol "]") + = do b <- bounds (decoratedSymbol fname "]") let fc = boundToFC fname (mergeBounds s b) rstate <- getInitRange xs + decorateKeywords fname xs pure (PRangeStream fc (fst rstate) (snd rstate)) - <|> do y <- bounds (expr pdef fname indents <* symbol "]") + <|> do y <- bounds (expr pdef fname indents <* decoratedSymbol fname "]") let fc = boundToFC fname (mergeBounds s y) rstate <- getInitRange xs + decorateKeywords fname xs pure (PRange fc (fst rstate) (snd rstate) y.val) listExpr : FileName -> WithBounds t -> IndentInfo -> Rule PTerm listExpr fname s indents = do b <- bounds (do ret <- expr pnowith fname indents - symbol "|" - conds <- sepBy1 (symbol ",") (doAct fname indents) - symbol "]" + decoratedSymbol fname "|" + conds <- sepBy1 (decoratedSymbol fname ",") (doAct fname indents) + decoratedSymbol fname "]" pure (ret, conds)) (ret, conds) <- pure b.val pure (PComprehension (boundToFC fname (mergeBounds s b)) ret (concat conds)) - <|> do xs <- sepBy (symbol ",") (expr pdef fname indents) - (do symbol ".." + <|> do xs <- option [] $ do + hd <- expr pdef fname indents + tl <- many $ do b <- bounds (symbol ",") + x <- expr pdef fname indents + pure (x <$ b) + pure ((hd <$ s) :: tl) + (do decoratedSymbol fname ".." listRange fname s indents xs) <|> (do b <- bounds (symbol "]") - pure (PList (boundToFC fname (mergeBounds s b)) xs)) + pure $ + let fc = boundToFC fname (mergeBounds s b) + nilFC = if null xs then fc else boundToFC fname b + in PList fc nilFC (map (\ t => (boundToFC fname t, t.val)) xs)) nonEmptyTuple : FileName -> WithBounds t -> IndentInfo -> PTerm -> Rule PTerm nonEmptyTuple fname s indents e - = do rest <- bounds (forget <$> some (bounds (symbol "," *> optional (bounds (expr pdef fname indents)))) + = do rest <- bounds (forget <$> some (bounds (decoratedSymbol fname "," *> optional (bounds (expr pdef fname indents)))) <* continueWith indents ")") pure $ buildOutput rest (mergePairs 0 rest rest.val) where @@ -361,7 +442,8 @@ mutual tuple : FileName -> WithBounds t -> IndentInfo -> PTerm -> Rule PTerm tuple fname s indents e = nonEmptyTuple fname s indents e - <|> do end <- bounds (continueWith indents ")") + <|> do end <- bounds (continueWithDecorated fname indents ")") + act [(toNonEmptyFC $ boundToFC fname s, Keyword, Nothing)] pure (PBracketed (boundToFC fname (mergeBounds s end)) e) postfixProjection : FileName -> IndentInfo -> Rule PTerm @@ -385,8 +467,8 @@ mutual simplerExpr : FileName -> IndentInfo -> Rule PTerm simplerExpr fname indents - = do b <- bounds (do x <- bounds unqualifiedName - symbol "@" + = do b <- bounds (do x <- bounds (decoratedSimpleBinderName fname) + decoratedSymbol fname "@" commit expr <- simpleExpr fname indents pure (x, expr)) @@ -398,23 +480,23 @@ mutual <|> record_ fname indents <|> singlelineStr pdef fname indents <|> multilineStr pdef fname indents - <|> do b <- bounds (symbol ".(" *> commit *> expr pdef fname indents <* symbol ")") + <|> do b <- bounds (decoratedSymbol fname ".(" *> commit *> expr pdef fname indents <* decoratedSymbol fname ")") pure (PDotted (boundToFC fname b) b.val) - <|> do b <- bounds (symbol "`(" *> expr pdef fname indents <* symbol ")") + <|> do b <- bounds (decoratedSymbol fname "`(" *> expr pdef fname indents <* decoratedSymbol fname ")") pure (PQuote (boundToFC fname b) b.val) - <|> do b <- bounds (symbol "`{{" *> name <* symbol "}}") + <|> do b <- bounds (decoratedSymbol fname "`{{" *> name <* decoratedSymbol fname "}}") pure (PQuoteName (boundToFC fname b) b.val) - <|> do b <- bounds (symbol "`[" *> nonEmptyBlock (topDecl fname) <* symbol "]") + <|> do b <- bounds (decoratedSymbol fname "`[" *> nonEmptyBlock (topDecl fname) <* decoratedSymbol fname "]") pure (PQuoteDecl (boundToFC fname b) (collectDefs (concat b.val))) - <|> do b <- bounds (symbol "~" *> simpleExpr fname indents) + <|> do b <- bounds (decoratedSymbol fname "~" *> simpleExpr fname indents) pure (PUnquote (boundToFC fname b) b.val) <|> do start <- bounds (symbol "(") bracketedExpr fname start indents <|> do start <- bounds (symbol "[") listExpr fname start indents - <|> do b <- bounds (symbol "!" *> simpleExpr fname indents) - pure (PBang (boundToFC fname b) b.val) - <|> do b <- bounds (symbol "[|" *> expr pdef fname indents <* symbol "|]") + <|> do b <- bounds (decoratedSymbol fname "!" *> simpleExpr fname indents) + pure (PBang (virtualiseFC $ boundToFC fname b) b.val) + <|> do b <- bounds (decoratedSymbol fname "[|" *> expr pdef fname indents <* decoratedSymbol fname "|]") pure (PIdiom (boundToFC fname b) b.val) <|> do b <- bounds (pragma "runElab" *> expr pdef fname indents) pure (PRunElab (boundToFC fname b) b.val) @@ -426,9 +508,9 @@ mutual (lvl, e) <- pure b.val pure (PUnifyLog (boundToFC fname b) lvl e) - multiplicity : SourceEmptyRule RigCount - multiplicity - = case !(optional $ intLit) of + multiplicity : FileName -> EmptyRule RigCount + multiplicity fname + = case !(optional $ decorate fname Keyword intLit) of (Just 0) => pure erased (Just 1) => pure linear Nothing => pure top @@ -444,27 +526,29 @@ mutual bindList : FileName -> IndentInfo -> Rule (List (RigCount, WithBounds PTerm, PTerm)) bindList fname indents - = forget <$> sepBy1 (symbol ",") - (do rig <- multiplicity + = forget <$> sepBy1 (decoratedSymbol fname ",") + (do rig <- multiplicity fname pat <- bounds (simpleExpr fname indents) ty <- option (PInfer (boundToFC fname pat)) - (symbol ":" *> opExpr pdef fname indents) + (decoratedSymbol fname ":" *> opExpr pdef fname indents) pure (rig, pat, ty)) pibindListName : FileName -> IndentInfo -> Rule (List (RigCount, WithBounds Name, PTerm)) pibindListName fname indents - = do rig <- multiplicity - ns <- sepBy1 (symbol ",") (bounds binderName) - symbol ":" + = do rig <- multiplicity fname + ns <- sepBy1 (decoratedSymbol fname ",") (bounds binderName) + let ns = forget $ map (map UN) ns + decorateBoundedNames fname Bound ns + decoratedSymbol fname ":" ty <- expr pdef fname indents atEnd indents - pure (map (\n => (rig, map UN n, ty)) (forget ns)) - <|> forget <$> sepBy1 (symbol ",") - (do rig <- multiplicity - n <- bounds binderName - symbol ":" + pure (map (\n => (rig, n, ty)) ns) + <|> forget <$> sepBy1 (decoratedSymbol fname ",") + (do rig <- multiplicity fname + n <- bounds (decorate fname Bound binderName) + decoratedSymbol fname ":" ty <- expr pdef fname indents pure (rig, map UN n, ty)) where @@ -478,70 +562,71 @@ mutual = do params <- pibindListName fname indents pure $ map (\(rig, n, ty) => (rig, map Just n, ty)) params - bindSymbol : Rule (PiInfo PTerm) - bindSymbol - = (symbol "->" $> Explicit) - <|> (symbol "=>" $> AutoImplicit) + bindSymbol : FileName -> Rule (PiInfo PTerm) + bindSymbol fname + = (decoratedSymbol fname "->" $> Explicit) + <|> (decoratedSymbol fname "=>" $> AutoImplicit) explicitPi : FileName -> IndentInfo -> Rule PTerm explicitPi fname indents - = do symbol "(" + = do decoratedSymbol fname "(" binders <- pibindList fname indents - symbol ")" - exp <- bindSymbol + decoratedSymbol fname ")" + exp <- bindSymbol fname scope <- mustWork $ typeExpr pdef fname indents pure (pibindAll fname exp binders scope) autoImplicitPi : FileName -> IndentInfo -> Rule PTerm autoImplicitPi fname indents - = do symbol "{" - keyword "auto" + = do decoratedSymbol fname "{" + decoratedKeyword fname "auto" commit binders <- pibindList fname indents - symbol "}" - symbol "->" + decoratedSymbol fname "}" + decoratedSymbol fname "->" scope <- mustWork $ typeExpr pdef fname indents pure (pibindAll fname AutoImplicit binders scope) defaultImplicitPi : FileName -> IndentInfo -> Rule PTerm defaultImplicitPi fname indents - = do symbol "{" - keyword "default" + = do decoratedSymbol fname "{" + decoratedKeyword fname "default" commit t <- simpleExpr fname indents binders <- pibindList fname indents - symbol "}" - symbol "->" + decoratedSymbol fname "}" + decoratedSymbol fname "->" scope <- mustWork $ typeExpr pdef fname indents pure (pibindAll fname (DefImplicit t) binders scope) forall_ : FileName -> IndentInfo -> Rule PTerm forall_ fname indents - = do keyword "forall" + = do decoratedKeyword fname "forall" commit - ns <- sepBy1 (symbol ",") (bounds unqualifiedName) + ns <- sepBy1 (decoratedSymbol fname ",") + (bounds (decoratedSimpleBinderName fname)) let binders = map (\n => ( erased {a=RigCount} , map (Just . UN) n , PImplicit (boundToFC fname n)) ) (forget ns) - symbol "." + decoratedSymbol fname "." scope <- mustWork $ typeExpr pdef fname indents pure (pibindAll fname Implicit binders scope) implicitPi : FileName -> IndentInfo -> Rule PTerm implicitPi fname indents - = do symbol "{" + = do decoratedSymbol fname "{" binders <- pibindList fname indents - symbol "}" - symbol "->" + decoratedSymbol fname "}" + decoratedSymbol fname "->" scope <- mustWork $ typeExpr pdef fname indents pure (pibindAll fname Implicit binders scope) lam : FileName -> IndentInfo -> Rule PTerm lam fname indents - = do symbol "\\" + = do decoratedSymbol fname "\\" commit - switch <- optional (bounds $ keyword "case") + switch <- optional (bounds $ decoratedKeyword fname "case") case switch of Nothing => continueLam Just r => continueLamCase r @@ -557,7 +642,7 @@ mutual continueLam : Rule PTerm continueLam = do binders <- bindList fname indents - symbol "=>" + decoratedSymbol fname "=>" mustContinue indents Nothing scope <- expr pdef fname indents pure (bindAll binders scope) @@ -567,20 +652,20 @@ mutual b <- bounds (forget <$> nonEmptyBlock (caseAlt fname)) pure (let fc = boundToFC fname b - fcCase = boundToFC fname endCase + fcCase = virtualiseFC $ boundToFC fname endCase n = MN "lcase" 0 in PLam fcCase top Explicit (PRef fcCase n) (PInfer fcCase) $ - PCase fc (PRef fcCase n) b.val) + PCase (virtualiseFC fc) (PRef fcCase n) b.val) letBlock : FileName -> IndentInfo -> Rule (WithBounds (Either LetBinder LetDecl)) letBlock fname indents = bounds (letBinder <||> letDecl) where letBinder : Rule LetBinder - letBinder = do s <- bounds (MkPair <$> multiplicity <*> expr plhs fname indents) + letBinder = do s <- bounds (MkPair <$> multiplicity fname <*> expr plhs fname indents) (rig, pat) <- pure s.val ty <- option (PImplicit (boundToFC fname s)) - (symbol ":" *> typeExpr (pnoeq pdef) fname indents) - (symbol "=" <|> symbol ":=") + (decoratedSymbol fname ":" *> typeExpr (pnoeq pdef) fname indents) + (decoratedSymbol fname "=" <|> decoratedSymbol fname ":=") val <- expr pnowith fname indents alts <- block (patAlt fname) pure (MkLetBinder rig pat ty val alts) @@ -590,22 +675,23 @@ mutual let_ : FileName -> IndentInfo -> Rule PTerm let_ fname indents - = do keyword "let" + = do decoratedKeyword fname "let" commit res <- nonEmptyBlock (letBlock fname) - commitKeyword indents "in" + commitKeyword fname indents "in" scope <- typeExpr pdef fname indents pure (mkLets fname res scope) case_ : FileName -> IndentInfo -> Rule PTerm case_ fname indents - = do b <- bounds (do keyword "case" + = do b <- bounds (do decoratedKeyword fname "case" scr <- expr pdef fname indents - mustWork (commitKeyword indents "of") + mustWork (commitKeyword fname indents "of") alts <- block (caseAlt fname) pure (scr, alts)) (scr, alts) <- pure b.val - pure (PCase (boundToFC fname b) scr alts) + pure (PCase (virtualiseFC $ boundToFC fname b) scr alts) + caseAlt : FileName -> IndentInfo -> Rule PClause caseAlt fname indents @@ -614,21 +700,21 @@ mutual caseRHS : FileName -> WithBounds t -> IndentInfo -> PTerm -> Rule PClause caseRHS fname start indents lhs - = do rhs <- bounds (symbol "=>" *> mustContinue indents Nothing *> expr pdef fname indents) + = do rhs <- bounds (decoratedSymbol fname "=>" *> mustContinue indents Nothing *> expr pdef fname indents) atEnd indents let fc = boundToFC fname (mergeBounds start rhs) pure (MkPatClause fc lhs rhs.val []) - <|> do end <- bounds (keyword "impossible") + <|> do end <- bounds (decoratedKeyword fname "impossible") atEnd indents pure (MkImpossible (boundToFC fname (mergeBounds start end)) lhs) if_ : FileName -> IndentInfo -> Rule PTerm if_ fname indents - = do b <- bounds (do keyword "if" + = do b <- bounds (do decoratedKeyword fname "if" x <- expr pdef fname indents - commitKeyword indents "then" + commitKeyword fname indents "then" t <- expr pdef fname indents - commitKeyword indents "else" + commitKeyword fname indents "else" e <- expr pdef fname indents pure (x, t, e)) atEnd indents @@ -637,20 +723,22 @@ mutual record_ : FileName -> IndentInfo -> Rule PTerm record_ fname indents - = do b <- bounds (do kw <- option False (keyword "record" $> True) -- TODO deprecated - symbol "{" + = do b <- bounds (do kw <- option False + (decoratedKeyword fname "record" + $> True) -- TODO deprecated + decoratedSymbol fname "{" commit - fs <- sepBy1 (symbol ",") (field kw fname indents) - symbol "}" + fs <- sepBy1 (decoratedSymbol fname ",") (field kw fname indents) + decoratedSymbol fname "}" pure $ forget fs) pure (PUpdate (boundToFC fname b) b.val) field : Bool -> FileName -> IndentInfo -> Rule PFieldUpdate field kw fname indents - = do path <- map fieldName <$> [| name :: many recFieldCompat |] - upd <- (ifThenElse kw (symbol "=") (symbol ":=") $> PSetField) + = do path <- map fieldName <$> [| decorate fname Function name :: many recFieldCompat |] + upd <- (ifThenElse kw (decoratedSymbol fname "=") (decoratedSymbol fname ":=") $> PSetField) <|> - (symbol "$=" $> PSetFieldApp) + (decoratedSymbol fname "$=" $> PSetFieldApp) val <- opExpr plhs fname indents pure (upd path val) where @@ -662,13 +750,15 @@ mutual -- this allows the dotted syntax .field -- but also the arrowed syntax ->field for compatibility with Idris 1 recFieldCompat : Rule Name - recFieldCompat = postfixProj <|> (symbol "->" *> name) + recFieldCompat = decorate fname Function postfixProj + <|> (decoratedSymbol fname "->" + *> decorate fname Function name) rewrite_ : FileName -> IndentInfo -> Rule PTerm rewrite_ fname indents - = do b <- bounds (do keyword "rewrite" + = do b <- bounds (do decoratedKeyword fname "rewrite" rule <- expr pdef fname indents - commitKeyword indents "in" + commitKeyword fname indents "in" tm <- expr pdef fname indents pure (rule, tm)) (rule, tm) <- pure b.val @@ -678,17 +768,19 @@ mutual doBlock fname indents = do b <- bounds $ keyword "do" *> block (doAct fname) commit - pure (PDoBlock (boundToFC fname b) Nothing (concat b.val)) + pure (PDoBlock (virtualiseFC $ boundToFC fname b) Nothing (concat b.val)) <|> do nsdo <- bounds namespacedIdent - the (SourceEmptyRule PTerm) $ case nsdo.val of + -- TODO: need to attach metadata correctly here + the (EmptyRule PTerm) $ case nsdo.val of (ns, "do") => do commit actions <- Core.bounds (block (doAct fname)) - let fc = boundToFC fname (mergeBounds nsdo actions) + let fc = virtualiseFC $ + boundToFC fname (mergeBounds nsdo actions) pure (PDoBlock fc ns (concat actions.val)) _ => fail "Not a namespaced 'do'" - validPatternVar : Name -> SourceEmptyRule () + validPatternVar : Name -> EmptyRule () validPatternVar (UN n) = if lowerFirst n then pure () else fail "Not a pattern variable" @@ -700,42 +792,46 @@ mutual -- If the name doesn't begin with a lower case letter, we should -- treat this as a pattern, so fail validPatternVar n.val - symbol "<-" + decoratedSymbol fname "<-" val <- expr pdef fname indents pure (n, val)) atEnd indents let (n, val) = b.val pure [DoBind (boundToFC fname b) (boundToFC fname n) n.val val] - <|> do keyword "let" + <|> do decoratedKeyword fname "let" commit res <- nonEmptyBlock (letBlock fname) atEnd indents pure (mkDoLets fname res) - <|> do b <- bounds (keyword "rewrite" *> expr pdef fname indents) + <|> do b <- bounds (decoratedKeyword fname "rewrite" *> expr pdef fname indents) atEnd indents pure [DoRewrite (boundToFC fname b) b.val] <|> do e <- bounds (expr plhs fname indents) - (atEnd indents $> [DoExp (boundToFC fname e) e.val]) - <|> (do b <- bounds $ symbol "<-" *> [| (expr pnowith fname indents, block (patAlt fname)) |] + (atEnd indents $> [DoExp (virtualiseFC $ boundToFC fname e) e.val]) + <|> (do b <- bounds $ decoratedSymbol fname "<-" *> [| (expr pnowith fname indents, block (patAlt fname)) |] atEnd indents let (v, alts) = b.val - let fc = boundToFC fname (mergeBounds e b) + let fc = virtualiseFC $ boundToFC fname (mergeBounds e b) pure [DoBindPat fc e.val v alts]) patAlt : FileName -> IndentInfo -> Rule PClause patAlt fname indents - = do symbol "|" + = do decoratedSymbol fname "|" caseAlt fname indents lazy : FileName -> IndentInfo -> Rule PTerm lazy fname indents - = do tm <- bounds (exactIdent "Lazy" *> simpleExpr fname indents) + = do tm <- bounds (decorate fname Typ (exactIdent "Lazy") + *> simpleExpr fname indents) pure (PDelayed (boundToFC fname tm) LLazy tm.val) - <|> do tm <- bounds (exactIdent "Inf" *> simpleExpr fname indents) + <|> do tm <- bounds (decorate fname Typ (exactIdent "Inf") + *> simpleExpr fname indents) pure (PDelayed (boundToFC fname tm) LInf tm.val) - <|> do tm <- bounds (exactIdent "Delay" *> simpleExpr fname indents) + <|> do tm <- bounds (decorate fname Data (exactIdent "Delay") + *> simpleExpr fname indents) pure (PDelay (boundToFC fname tm) tm.val) - <|> do tm <- bounds (exactIdent "Force" *> simpleExpr fname indents) + <|> do tm <- bounds (decorate fname Data (exactIdent "Force") + *> simpleExpr fname indents) pure (PForce (boundToFC fname tm) tm.val) binder : FileName -> IndentInfo -> Rule PTerm @@ -752,7 +848,7 @@ mutual typeExpr q fname indents = do arg <- bounds (opExpr q fname indents) (do continue indents - rest <- some [| (bindSymbol, bounds $ opExpr pdef fname indents) |] + rest <- some [| (bindSymbol fname, bounds $ opExpr pdef fname indents) |] pure (mkPi arg (forget rest))) <|> pure arg.val where @@ -772,7 +868,8 @@ mutual export singlelineStr : ParseOpts -> FileName -> IndentInfo -> Rule PTerm singlelineStr q fname idents - = do b <- bounds $ do begin <- bounds strBegin + = decorate fname Data $ + do b <- bounds $ do begin <- bounds strBegin commit xs <- many $ bounds $ (interpBlock q fname idents) <||> strLitLines pstrs <- case traverse toPStr xs of @@ -791,7 +888,8 @@ mutual export multilineStr : ParseOpts -> FileName -> IndentInfo -> Rule PTerm multilineStr q fname idents - = do b <- bounds $ do multilineBegin + = decorate fname Data $ + do b <- bounds $ do multilineBegin commit xs <- many $ bounds $ (interpBlock q fname idents) <||> strLitLines endloc <- location @@ -810,21 +908,21 @@ mutual ((acc `snoc` (line `snoc` (StrLiteral (boundToFC fname x) str))) ++ ((\str => [StrLiteral (boundToFC fname x) str]) <$> (init strs))) -visOption : Rule Visibility -visOption - = (keyword "public" *> keyword "export" $> Public) - <|> (keyword "export" $> Export) - <|> (keyword "private" $> Private) +visOption : FileName -> Rule Visibility +visOption fname + = (decoratedKeyword fname "public" *> decoratedKeyword fname "export" $> Public) + <|> (decoratedKeyword fname "export" $> Export) + <|> (decoratedKeyword fname "private" $> Private) -visibility : SourceEmptyRule Visibility -visibility - = visOption +visibility : FileName -> EmptyRule Visibility +visibility fname + = visOption fname <|> pure Private tyDecls : Rule Name -> String -> FileName -> IndentInfo -> Rule (List1 PTypeDecl) tyDecls declName predoc fname indents - = do bs <- do docns <- sepBy1 (symbol ",") [| (option "" documentation, bounds declName) |] - symbol ":" + = do bs <- do docns <- sepBy1 (decoratedSymbol fname ",") [| (option "" documentation, bounds declName) |] + decoratedSymbol fname ":" mustWork $ do ty <- expr pdef fname indents pure $ map (\(doc, n) => (doc, n.val, boundToFC fname n, ty)) docns @@ -832,7 +930,7 @@ tyDecls declName predoc fname indents pure $ map (\(doc, n, nFC, ty) => (MkPTy nFC nFC n (predoc ++ doc) ty)) bs -withFlags : SourceEmptyRule (List WithFlag) +withFlags : EmptyRule (List WithFlag) withFlags = pragma "syntactic" *> (Syntactic ::) <$> withFlags <|> pure [] @@ -842,23 +940,24 @@ mutual FileName -> WithBounds t -> Int -> IndentInfo -> (lhs : PTerm) -> Rule PClause parseRHS withArgs fname start col indents lhs - = do b <- bounds $ symbol "=" *> mustWork [| (expr pdef fname indents, option [] $ whereBlock fname col) |] + = do b <- bounds $ decoratedSymbol fname "=" *> mustWork [| (expr pdef fname indents, option [] $ whereBlock fname col) |] atEnd indents (rhs, ws) <- pure b.val let fc = boundToFC fname (mergeBounds start b) pure (MkPatClause fc lhs rhs ws) - <|> do b <- bounds (do keyword "with" + <|> do b <- bounds (do decoratedKeyword fname "with" commit - flags <- bounds (withFlags) - symbol "(" - wval <- bracketedExpr fname flags indents - prf <- optional (keyword "proof" *> name) + flags <- withFlags + start <- bounds (decoratedSymbol fname "(") + wval <- bracketedExpr fname start indents + prf <- optional (decoratedKeyword fname "proof" + *> UN <$> decoratedSimpleBinderName fname) ws <- mustWork $ nonEmptyBlockAfter col (clause (S withArgs) fname) pure (prf, flags, wval, forget ws)) (prf, flags, wval, ws) <- pure b.val let fc = boundToFC fname (mergeBounds start b) - pure (MkWithClause fc lhs wval prf flags.val ws) - <|> do end <- bounds (keyword "impossible") + pure (MkWithClause fc lhs wval prf flags ws) + <|> do end <- bounds (decoratedKeyword fname "impossible") atEnd indents pure (MkImpossible (boundToFC fname (mergeBounds start end)) lhs) @@ -883,14 +982,16 @@ mutual parseWithArg : Rule (FC, PTerm) parseWithArg - = do symbol "|" + = do decoratedSymbol fname "|" tm <- bounds (expr plhs fname indents) pure (boundToFC fname tm, tm.val) -mkTyConType : FC -> List Name -> PTerm -mkTyConType fc [] = PType fc -mkTyConType fc (x :: xs) - = PPi fc top Explicit Nothing (PType fc) (mkTyConType fc xs) +mkTyConType : FileName -> FC -> List (WithBounds Name) -> PTerm +mkTyConType fname fc [] = PType (virtualiseFC fc) +mkTyConType fname fc (x :: xs) + = let bfc = boundToFC fname x in + PPi bfc top Explicit Nothing (PType (virtualiseFC fc)) + $ mkTyConType fname fc xs mkDataConType : FC -> PTerm -> List ArgType -> Maybe PTerm mkDataConType fc ret [] = Just ret @@ -904,7 +1005,7 @@ mkDataConType _ _ _ -- with and named applications not allowed in simple ADTs simpleCon : FileName -> PTerm -> IndentInfo -> Rule PTypeDecl simpleCon fname ret indents = do b <- bounds (do cdoc <- option "" documentation - cname <- bounds dataConstructorName + cname <- bounds $ decoratedDataConstructorName fname params <- many (argExpr plhs fname indents) pure (cdoc, cname.val, boundToFC fname cname, params)) atEnd indents @@ -913,17 +1014,20 @@ simpleCon fname ret indents fromMaybe (fatalError "Named arguments not allowed in ADT constructors") (pure . MkPTy cfc cnameFC cname cdoc <$> mkDataConType cfc ret (concat params)) -simpleData : FileName -> WithBounds t -> Name -> IndentInfo -> Rule PDataDecl -simpleData fname start n indents - = do b <- bounds (do params <- many name - tyend <- bounds (symbol "=") +simpleData : FileName -> WithBounds t -> + WithBounds Name -> IndentInfo -> Rule PDataDecl +simpleData fname start tyName indents + = do b <- bounds (do params <- many (bounds $ decorate fname Bound name) + tyend <- bounds (decoratedSymbol fname "=") let tyfc = boundToFC fname (mergeBounds start tyend) - let conRetTy = papply tyfc (PRef tyfc n) (map (PRef tyfc) params) - cons <- sepBy1 (symbol "|") (simpleCon fname conRetTy indents) + let tyCon = PRef (boundToFC fname tyName) tyName.val + let toPRef = \ t => PRef (boundToFC fname t) t.val + let conRetTy = papply tyfc tyCon (map toPRef params) + cons <- sepBy1 (decoratedSymbol fname "|") (simpleCon fname conRetTy indents) pure (params, tyfc, forget cons)) (params, tyfc, cons) <- pure b.val - pure (MkPData (boundToFC fname (mergeBounds start b)) n - (mkTyConType tyfc params) [] cons) + pure (MkPData (boundToFC fname (mergeBounds start b)) tyName.val + (mkTyConType fname tyfc params) [] cons) dataOpt : Rule DataOpt dataOpt @@ -934,29 +1038,30 @@ dataOpt <|> (exactIdent "noNewtype" $> NoNewtype) dataBody : FileName -> Int -> WithBounds t -> Name -> IndentInfo -> PTerm -> - SourceEmptyRule PDataDecl + EmptyRule PDataDecl dataBody fname mincol start n indents ty = do atEndIndent indents pure (MkPLater (boundToFC fname start) n ty) - <|> do b <- bounds (do keyword "where" - opts <- option [] $ symbol "[" *> forget <$> sepBy1 (symbol ",") dataOpt <* symbol "]" - cs <- blockAfter mincol (tyDecls (mustWork dataConstructorName) "" fname) + <|> do b <- bounds (do decoratedKeyword fname "where" + opts <- option [] $ decoratedSymbol fname "[" *> forget <$> sepBy1 (decoratedSymbol fname ",") dataOpt <* decoratedSymbol fname "]" + cs <- blockAfter mincol (tyDecls (mustWork $ decoratedDataConstructorName fname) "" fname) pure (opts, concatMap forget cs)) (opts, cs) <- pure b.val pure (MkPData (boundToFC fname (mergeBounds start b)) n ty opts cs) -gadtData : FileName -> Int -> WithBounds t -> Name -> IndentInfo -> Rule PDataDecl -gadtData fname mincol start n indents - = do symbol ":" +gadtData : FileName -> Int -> WithBounds t -> + WithBounds Name -> IndentInfo -> Rule PDataDecl +gadtData fname mincol start tyName indents + = do decoratedSymbol fname ":" commit ty <- expr pdef fname indents - dataBody fname mincol start n indents ty + dataBody fname mincol start tyName.val indents ty dataDeclBody : FileName -> IndentInfo -> Rule PDataDecl dataDeclBody fname indents = do b <- bounds (do col <- column - keyword "data" - n <- mustWork dataTypeName + decoratedKeyword fname "data" + n <- mustWork (bounds $ decoratedDataTypeName fname) pure (col, n)) (col, n) <- pure b.val simpleData fname b n indents <|> gadtData fname col b n indents @@ -964,7 +1069,7 @@ dataDeclBody fname indents dataDecl : FileName -> IndentInfo -> Rule PDecl dataDecl fname indents = do b <- bounds (do doc <- option "" documentation - vis <- visibility + vis <- visibility fname dat <- dataDeclBody fname indents pure (doc, vis, dat)) (doc, vis, dat) <- pure b.val @@ -987,11 +1092,11 @@ extension = (exactIdent "ElabReflection" $> ElabReflection) <|> (exactIdent "Borrowing" $> Borrowing) -totalityOpt : Rule TotalReq -totalityOpt - = (keyword "partial" $> PartialOK) - <|> (keyword "total" $> Total) - <|> (keyword "covering" $> CoveringOnly) +totalityOpt : FileName -> Rule TotalReq +totalityOpt fname + = (decoratedKeyword fname "partial" $> PartialOK) + <|> (decoratedKeyword fname "total" $> Total) + <|> (decoratedKeyword fname "covering" $> CoveringOnly) logLevel : Rule (Maybe LogLevel) logLevel @@ -1003,7 +1108,7 @@ logLevel directive : FileName -> IndentInfo -> Rule Directive directive fname indents - = do pragma "hide" + = do decorate fname Keyword $ pragma "hide" n <- name atEnd indents pure (Hide n) @@ -1011,80 +1116,81 @@ directive fname indents -- n <- name -- atEnd indents -- pure (Hide True n) - <|> do pragma "logging" + <|> do decorate fname Keyword $ pragma "logging" lvl <- logLevel atEnd indents pure (Logging lvl) - <|> do pragma "auto_lazy" + <|> do decorate fname Keyword $ pragma "auto_lazy" b <- onoff atEnd indents pure (LazyOn b) - <|> do pragma "unbound_implicits" + <|> do decorate fname Keyword $ pragma "unbound_implicits" b <- onoff atEnd indents pure (UnboundImplicits b) - <|> do pragma "prefix_record_projections" + <|> do decorate fname Keyword $ pragma "prefix_record_projections" b <- onoff atEnd indents pure (PrefixRecordProjections b) - <|> do pragma "ambiguity_depth" - lvl <- intLit + <|> do decorate fname Keyword $ pragma "ambiguity_depth" + lvl <- decorate fname Keyword $ intLit atEnd indents pure (AmbigDepth (fromInteger lvl)) - <|> do pragma "auto_implicit_depth" - dpt <- intLit + <|> do decorate fname Keyword $ pragma "auto_implicit_depth" + dpt <- decorate fname Keyword $ intLit atEnd indents pure (AutoImplicitDepth (fromInteger dpt)) - <|> do pragma "nf_metavar_threshold" - dpt <- intLit + <|> do decorate fname Keyword $ pragma "nf_metavar_threshold" + dpt <- decorate fname Keyword $ intLit atEnd indents pure (NFMetavarThreshold (fromInteger dpt)) - <|> do pragma "pair" + <|> do decorate fname Keyword $ pragma "pair" ty <- name f <- name s <- name atEnd indents pure (PairNames ty f s) - <|> do pragma "rewrite" + <|> do decorate fname Keyword $ pragma "rewrite" eq <- name rw <- name atEnd indents pure (RewriteName eq rw) - <|> do pragma "integerLit" + <|> do decorate fname Keyword $ pragma "integerLit" n <- name atEnd indents pure (PrimInteger n) - <|> do pragma "stringLit" + <|> do decorate fname Keyword $ pragma "stringLit" n <- name atEnd indents pure (PrimString n) - <|> do pragma "charLit" + <|> do decorate fname Keyword $ pragma "charLit" n <- name atEnd indents pure (PrimChar n) - <|> do pragma "doubleLit" + <|> do decorate fname Keyword $ pragma "doubleLit" n <- name atEnd indents pure (PrimDouble n) - <|> do pragma "name" + <|> do decorate fname Keyword $ pragma "name" n <- name - ns <- sepBy1 (symbol ",") unqualifiedName + ns <- sepBy1 (decoratedSymbol fname ",") + (decoratedSimpleBinderName fname) atEnd indents pure (Names n (forget ns)) - <|> do pragma "start" + <|> do decorate fname Keyword $ pragma "start" e <- expr pdef fname indents atEnd indents pure (StartExpr e) - <|> do pragma "allow_overloads" + <|> do decorate fname Keyword $ pragma "allow_overloads" n <- name atEnd indents pure (Overloadable n) - <|> do pragma "language" + <|> do decorate fname Keyword $ pragma "language" e <- extension atEnd indents pure (Extension e) - <|> do pragma "default" - tot <- totalityOpt + <|> do decorate fname Keyword $ pragma "default" + tot <- totalityOpt fname atEnd indents pure (DefaultTotality tot) @@ -1113,7 +1219,7 @@ transformDecl fname indents = do b <- bounds (do pragma "transform" n <- simpleStr lhs <- expr plhs fname indents - symbol "=" + decoratedSymbol fname "=" rhs <- expr pnowith fname indents pure (n, lhs, rhs)) (n, lhs, rhs) <- pure b.val @@ -1126,29 +1232,30 @@ runElabDecl fname indents mutualDecls : FileName -> IndentInfo -> Rule PDecl mutualDecls fname indents - = do ds <- bounds (keyword "mutual" *> commit *> assert_total (nonEmptyBlock (topDecl fname))) + = do ds <- bounds (decoratedKeyword fname "mutual" *> commit *> assert_total (nonEmptyBlock (topDecl fname))) pure (PMutual (boundToFC fname ds) (concat ds.val)) usingDecls : FileName -> IndentInfo -> Rule PDecl usingDecls fname indents - = do b <- bounds (do keyword "using" + = do b <- bounds (do decoratedKeyword fname "using" commit - symbol "(" - us <- sepBy (symbol ",") + decoratedSymbol fname "(" + us <- sepBy (decoratedSymbol fname ",") (do n <- optional (do x <- unqualifiedName - symbol ":" + decoratedSymbol fname ":" pure (UN x)) ty <- typeExpr pdef fname indents pure (n, ty)) - symbol ")" + decoratedSymbol fname ")" ds <- assert_total (nonEmptyBlock (topDecl fname)) pure (us, ds)) (us, ds) <- pure b.val pure (PUsing (boundToFC fname b) us (collectDefs (concat ds))) -fnOpt : Rule PFnOpt -fnOpt = do x <- totalityOpt +fnOpt : FileName -> Rule PFnOpt +fnOpt fname + = do x <- totalityOpt fname pure $ IFnOpt (Totality x) fnDirectOpt : FileName -> Rule PFnOpt @@ -1170,7 +1277,7 @@ fnDirectOpt fname <|> do pragma "macro" pure $ IFnOpt Macro <|> do pragma "spec" - ns <- sepBy (symbol ",") name + ns <- sepBy (decoratedSymbol fname ",") name pure $ IFnOpt (SpecArgs ns) <|> do pragma "foreign" cs <- block (expr pdef fname) @@ -1192,15 +1299,15 @@ builtinDecl fname indents visOpt : FileName -> Rule (Either Visibility PFnOpt) visOpt fname - = do vis <- visOption + = do vis <- visOption fname pure (Left vis) - <|> do tot <- fnOpt + <|> do tot <- fnOpt fname pure (Right tot) <|> do opt <- fnDirectOpt fname pure (Right opt) getVisibility : Maybe Visibility -> List (Either Visibility PFnOpt) -> - SourceEmptyRule Visibility + EmptyRule Visibility getVisibility Nothing [] = pure Private getVisibility (Just vis) [] = pure vis getVisibility Nothing (Left x :: xs) = getVisibility (Just x) xs @@ -1208,67 +1315,67 @@ getVisibility (Just vis) (Left x :: xs) = fatalError "Multiple visibility modifiers" getVisibility v (_ :: xs) = getVisibility v xs -recordConstructor : Rule Name -recordConstructor +recordConstructor : FileName -> Rule Name +recordConstructor fname = do exactIdent "constructor" - n <- mustWork dataConstructorName + n <- mustWork $ decoratedDataConstructorName fname pure n -constraints : FileName -> IndentInfo -> SourceEmptyRule (List (Maybe Name, PTerm)) +constraints : FileName -> IndentInfo -> EmptyRule (List (Maybe Name, PTerm)) constraints fname indents = do tm <- appExpr pdef fname indents - symbol "=>" + decoratedSymbol fname "=>" more <- constraints fname indents pure ((Nothing, tm) :: more) - <|> do symbol "(" - n <- name - symbol ":" + <|> do decoratedSymbol fname "(" + n <- decorate fname Bound name + decoratedSymbol fname ":" tm <- expr pdef fname indents - symbol ")" - symbol "=>" + decoratedSymbol fname ")" + decoratedSymbol fname "=>" more <- constraints fname indents pure ((Just n, tm) :: more) <|> pure [] -implBinds : FileName -> IndentInfo -> SourceEmptyRule (List (Name, RigCount, PTerm)) +implBinds : FileName -> IndentInfo -> EmptyRule (List (Name, RigCount, PTerm)) implBinds fname indents - = do symbol "{" - rig <- multiplicity - n <- name - symbol ":" + = do decoratedSymbol fname "{" + rig <- multiplicity fname + n <- decorate fname Bound name + decoratedSymbol fname ":" tm <- expr pdef fname indents - symbol "}" - symbol "->" + decoratedSymbol fname "}" + decoratedSymbol fname "->" more <- implBinds fname indents pure ((n, rig, tm) :: more) <|> pure [] ifaceParam : FileName -> IndentInfo -> Rule (List Name, (RigCount, PTerm)) ifaceParam fname indents - = do symbol "(" - rig <- multiplicity - ns <- sepBy1 (symbol ",") name - symbol ":" + = do decoratedSymbol fname "(" + rig <- multiplicity fname + ns <- sepBy1 (decoratedSymbol fname ",") (decorate fname Bound name) + decoratedSymbol fname ":" tm <- expr pdef fname indents - symbol ")" + decoratedSymbol fname ")" pure (forget ns, (rig, tm)) - <|> do n <- bounds name + <|> do n <- bounds (decorate fname Bound name) pure ([n.val], (erased, PInfer (boundToFC fname n))) ifaceDecl : FileName -> IndentInfo -> Rule PDecl ifaceDecl fname indents = do b <- bounds (do doc <- option "" documentation - vis <- visibility + vis <- visibility fname col <- column - keyword "interface" + decoratedKeyword fname "interface" commit cons <- constraints fname indents - n <- name + n <- decorate fname Typ name paramss <- many (ifaceParam fname indents) let params = concatMap (\ (ns, rt) => map (\ n => (n, rt)) ns) paramss - det <- option [] $ symbol "|" *> sepBy (symbol ",") name - keyword "where" - dc <- optional recordConstructor + det <- option [] $ decoratedSymbol fname "|" *> sepBy (decoratedSymbol fname ",") (decorate fname Bound name) + decoratedKeyword fname "where" + dc <- optional (recordConstructor fname) body <- assert_total (blockAfter col (topDecl fname)) pure (\fc : FC => PInterface fc vis cons n doc params det dc (collectDefs (concat body)))) @@ -1281,14 +1388,17 @@ implDecl fname indents vis <- getVisibility Nothing visOpts let opts = mapMaybe getRight visOpts col <- column - option () (keyword "implementation") - iname <- optional $ symbol "[" *> name <* symbol "]" + option () (decoratedKeyword fname "implementation") + iname <- optional $ decoratedSymbol fname "[" + *> decorate fname Function name + <* decoratedSymbol fname "]" impls <- implBinds fname indents cons <- constraints fname indents - n <- name + n <- decorate fname Typ name params <- many (simpleExpr fname indents) - nusing <- option [] $ keyword "using" *> forget <$> some name - body <- optional $ keyword "where" *> blockAfter col (topDecl fname) + nusing <- option [] $ decoratedKeyword fname "using" + *> forget <$> some (decorate fname Function name) + body <- optional $ decoratedKeyword fname "where" *> blockAfter col (topDecl fname) pure $ \fc : FC => (PImplementation fc vis opts Single impls cons n params iname nusing (map (collectDefs . concat) body))) @@ -1298,11 +1408,11 @@ implDecl fname indents fieldDecl : FileName -> IndentInfo -> Rule (List PField) fieldDecl fname indents = do doc <- option "" documentation - symbol "{" + decoratedSymbol fname "{" commit - impl <- option Implicit (AutoImplicit <$ keyword "auto") + impl <- option Implicit (AutoImplicit <$ decoratedKeyword fname "auto") fs <- fieldBody doc impl - symbol "}" + decoratedSymbol fname "}" atEnd indents pure fs <|> do doc <- option "" documentation @@ -1312,27 +1422,29 @@ fieldDecl fname indents where fieldBody : String -> PiInfo PTerm -> Rule (List PField) fieldBody doc p - = do b <- bounds (do rig <- multiplicity - ns <- sepBy1 (symbol ",") name - symbol ":" + = do b <- bounds (do rig <- multiplicity fname + ns <- sepBy1 (decoratedSymbol fname ",") (do + n <- decorate fname Function name + pure n) + decoratedSymbol fname ":" ty <- expr pdef fname indents pure (\fc : FC => map (\n => MkField fc doc rig p n ty) (forget ns))) pure (b.val (boundToFC fname b)) typedArg : FileName -> IndentInfo -> Rule (List (Name, RigCount, PiInfo PTerm, PTerm)) typedArg fname indents - = do symbol "(" + = do decoratedSymbol fname "(" params <- pibindListName fname indents - symbol ")" + decoratedSymbol fname ")" pure $ map (\(c, n, tm) => (n.val, c, Explicit, tm)) params - <|> do symbol "{" + <|> do decoratedSymbol fname "{" commit - info <- the (SourceEmptyRule (PiInfo PTerm)) - (pure AutoImplicit <* keyword "auto" - <|> (keyword "default" *> DefImplicit <$> simpleExpr fname indents) + info <- the (EmptyRule (PiInfo PTerm)) + (pure AutoImplicit <* decoratedKeyword fname "auto" + <|> (decoratedKeyword fname "default" *> DefImplicit <$> simpleExpr fname indents) <|> pure Implicit) params <- pibindListName fname indents - symbol "}" + decoratedSymbol fname "}" pure $ map (\(c, n, tm) => (n.val, c, info, tm)) params recordParam : FileName -> IndentInfo -> Rule (List (Name, RigCount, PiInfo PTerm, PTerm)) @@ -1344,22 +1456,22 @@ recordParam fname indents recordDecl : FileName -> IndentInfo -> Rule PDecl recordDecl fname indents = do b <- bounds (do doc <- option "" documentation - vis <- visibility + vis <- visibility fname col <- column - keyword "record" - n <- mustWork dataTypeName + decoratedKeyword fname "record" + n <- mustWork (decoratedDataTypeName fname) paramss <- many (recordParam fname indents) let params = concat paramss - keyword "where" + decoratedKeyword fname "where" dcflds <- blockWithOptHeaderAfter col - (\ idt => recordConstructor <* atEnd idt) + (\ idt => recordConstructor fname <* atEnd idt) (fieldDecl fname) pure (\fc : FC => PRecord fc doc vis n params (fst dcflds) (concat (snd dcflds)))) pure (b.val (boundToFC fname b)) paramDecls : FileName -> IndentInfo -> Rule PDecl paramDecls fname indents - = do b1 <- bounds (keyword "parameters") + = do b1 <- bounds (decoratedKeyword fname "parameters") commit args <- bounds (newParamDecls fname indents <|> oldParamDecls fname indents) commit @@ -1370,13 +1482,13 @@ paramDecls fname indents where oldParamDecls : FileName -> IndentInfo -> Rule (List (Name, RigCount, PiInfo PTerm, PTerm)) oldParamDecls fname indents - = do symbol "(" - ps <- sepBy (symbol ",") - (do x <- unqualifiedName - symbol ":" + = do decoratedSymbol fname "(" + ps <- sepBy (decoratedSymbol fname ",") + (do x <- decoratedSimpleBinderName fname + decoratedSymbol fname ":" ty <- typeExpr pdef fname indents pure (UN x, top, Explicit, ty)) - symbol ")" + decoratedSymbol fname ")" pure ps newParamDecls : FileName -> IndentInfo -> Rule (List (Name, RigCount, PiInfo PTerm, PTerm)) @@ -1391,8 +1503,9 @@ claims fname indents visOpts <- many (visOpt fname) vis <- getVisibility Nothing visOpts let opts = mapMaybe getRight visOpts - rig <- multiplicity - cls <- tyDecls name doc fname indents + rig <- multiplicity fname + cls <- tyDecls (decorate fname Function name) + doc fname indents pure $ map (\cl => the (Pair _ _) (doc, vis, opts, rig, cl)) cls) pure $ map (\(doc, vis, opts, rig, cl) : Pair _ _ => PClaim (boundToFC fname bs) rig vis opts cl) @@ -1405,10 +1518,10 @@ definition fname indents fixDecl : FileName -> IndentInfo -> Rule (List PDecl) fixDecl fname indents - = do b <- bounds (do fixity <- fix + = do b <- bounds (do fixity <- decorate fname Keyword $ fix commit - prec <- intLit - ops <- sepBy1 (symbol ",") iOperator + prec <- decorate fname Keyword $ intLit + ops <- sepBy1 (decoratedSymbol fname ",") iOperator pure (fixity, prec, ops)) (fixity, prec, ops) <- pure b.val pure (map (PFixity (boundToFC fname b) fixity (fromInteger prec)) (forget ops)) @@ -1492,8 +1605,8 @@ collectDefs (d :: ds) export import_ : FileName -> IndentInfo -> Rule Import import_ fname indents - = do b <- bounds (do keyword "import" - reexp <- option False (keyword "public" $> True) + = do b <- bounds (do decoratedKeyword fname "import" + reexp <- option False (decoratedKeyword fname "public" $> True) ns <- mustWork moduleIdent nsAs <- option (miAsNamespace ns) (do exactIdent "as" @@ -1504,11 +1617,11 @@ import_ fname indents pure (MkImport (boundToFC fname b) reexp ns nsAs) export -prog : FileName -> SourceEmptyRule Module +prog : FileName -> EmptyRule Module prog fname = do b <- bounds (do doc <- option "" documentation nspace <- option (nsAsModuleIdent mainNS) - (do keyword "module" + (do decoratedKeyword fname "module" mustWork moduleIdent) imports <- block (import_ fname) pure (doc, nspace, imports)) @@ -1518,11 +1631,11 @@ prog fname nspace imports doc (collectDefs (concat ds))) export -progHdr : FileName -> SourceEmptyRule Module +progHdr : FileName -> EmptyRule Module progHdr fname = do b <- bounds (do doc <- option "" documentation nspace <- option (nsAsModuleIdent mainNS) - (do keyword "module" + (do decoratedKeyword fname "module" mustWork moduleIdent) imports <- block (import_ fname) pure (doc, nspace, imports)) @@ -1914,7 +2027,7 @@ eval = do pure (Eval tm) export -command : SourceEmptyRule REPLCmd +command : EmptyRule REPLCmd command = eoi $> NOP <|> nonEmptyCommand diff --git a/src/Idris/Parser/Let.idr b/src/Idris/Parser/Let.idr index 6326db84d..e15098118 100644 --- a/src/Idris/Parser/Let.idr +++ b/src/Idris/Parser/Let.idr @@ -56,7 +56,7 @@ mkLets : FileName -> List1 (WithBounds (Either LetBinder LetDecl)) -> PTerm -> PTerm mkLets fname = letFactory buildLets - (\ decls, scope => PLocal (boundToFC fname decls) decls.val scope) + (\ decls, scope => PLocal (virtualiseFC $ boundToFC fname decls) decls.val scope) where diff --git a/src/Idris/Pretty.idr b/src/Idris/Pretty.idr index 5b6723356..c467631a0 100644 --- a/src/Idris/Pretty.idr +++ b/src/Idris/Pretty.idr @@ -294,10 +294,10 @@ mutual go d (PDotted _ p) = dot <+> go d p go d (PImplicit _) = "_" go d (PInfer _) = "?" - go d (POp _ op x y) = parenthesise (d > appPrec) $ group $ go startPrec x <++> prettyOp op <++> go startPrec y - go d (PPrefixOp _ op x) = parenthesise (d > appPrec) $ pretty op <+> go startPrec x - go d (PSectionL _ op x) = parens (prettyOp op <++> go startPrec x) - go d (PSectionR _ x op) = parens (go startPrec x <++> prettyOp op) + go d (POp _ _ op x y) = parenthesise (d > appPrec) $ group $ go startPrec x <++> prettyOp op <++> go startPrec y + go d (PPrefixOp _ _ op x) = parenthesise (d > appPrec) $ pretty op <+> go startPrec x + go d (PSectionL _ _ op x) = parens (prettyOp op <++> go startPrec x) + go d (PSectionR _ _ x op) = parens (go startPrec x <++> prettyOp op) go d (PEq fc l r) = parenthesise (d > appPrec) $ go startPrec l <++> equals <++> go startPrec r go d (PBracketed _ tm) = parens (go startPrec tm) go d (PString _ xs) = parenthesise (d > appPrec) $ hsep $ punctuate "++" (prettyString <$> xs) @@ -305,10 +305,10 @@ mutual go d (PDoBlock _ ns ds) = parenthesise (d > appPrec) $ group $ align $ hang 2 $ do_ <++> (vsep $ punctuate semi (prettyDo <$> ds)) go d (PBang _ tm) = "!" <+> go d tm go d (PIdiom _ tm) = enclose (pretty "[|") (pretty "|]") (go startPrec tm) - go d (PList _ xs) = brackets (group $ align $ vsep $ punctuate comma (go startPrec <$> xs)) + go d (PList _ _ xs) = brackets (group $ align $ vsep $ punctuate comma (go startPrec . snd <$> xs)) go d (PPair _ l r) = group $ parens (go startPrec l <+> comma <+> line <+> go startPrec r) - go d (PDPair _ l (PImplicit _) r) = group $ parens (go startPrec l <++> pretty "**" <+> line <+> go startPrec r) - go d (PDPair _ l ty r) = group $ parens (go startPrec l <++> colon <++> go startPrec ty <++> pretty "**" <+> line <+> go startPrec r) + go d (PDPair _ _ l (PImplicit _) r) = group $ parens (go startPrec l <++> pretty "**" <+> line <+> go startPrec r) + go d (PDPair _ _ l ty r) = group $ parens (go startPrec l <++> colon <++> go startPrec ty <++> pretty "**" <+> line <+> go startPrec r) go d (PUnit _) = "()" go d (PIfThenElse _ x t e) = parenthesise (d > appPrec) $ group $ align $ hang 2 $ vsep diff --git a/src/Idris/ProcessIdr.idr b/src/Idris/ProcessIdr.idr index 5d9025f0a..ef958d597 100644 --- a/src/Idris/ProcessIdr.idr +++ b/src/Idris/ProcessIdr.idr @@ -197,7 +197,7 @@ readHeader path -- Stop at the first :, that's definitely not part of the header, to -- save lexing the whole file unnecessarily setCurrentElabSource res -- for error printing purposes - let Right mod = runParserTo path (isLitFile path) (is ':') res (progHdr path) + let Right (decor, mod) = runParserTo path (isLitFile path) (is ':') res (progHdr path) | Left err => throw err pure mod @@ -259,15 +259,16 @@ processMod srcf ttcf msg sourcecode pure Nothing else -- needs rebuilding do iputStrLn msg - Right mod <- logTime ("++ Parsing " ++ srcf) $ + Right (decor, mod) <- logTime ("++ Parsing " ++ srcf) $ pure (runParser srcf (isLitFile srcf) sourcecode (do p <- prog srcf; eoi; pure p)) | Left err => pure (Just [err]) + addSemanticDecorations decor initHash traverse_ addPublicHash (sort hs) resetNextVar when (ns /= nsAsModuleIdent mainNS) $ - do let MkFC fname _ _ = headerloc mod - | EmptyFC => throw (InternalError "No file name") + do let Just fname = map file (isNonEmptyFC $ headerloc mod) + | Nothing => throw (InternalError "No file name") d <- getDirs ns' <- pathToNS (working_dir d) (source_dir d) fname when (ns /= ns') $ diff --git a/src/Idris/REPL.idr b/src/Idris/REPL.idr index 8c09bc4f3..09f83aa22 100644 --- a/src/Idris/REPL.idr +++ b/src/Idris/REPL.idr @@ -18,6 +18,7 @@ import Core.Env import Core.InitPrimitives import Core.LinearCheck import Core.Metadata +import Core.FC import Core.Normalise import Core.Options import Core.Termination @@ -638,7 +639,7 @@ loadMainFile : {auto c : Ref Ctxt Defs} -> loadMainFile f = do opts <- get ROpts put ROpts (record { evalResultName = Nothing } opts) - resetContext + resetContext f Right res <- coreLift (readFile f) | Left err => do setSource "" pure (ErrorLoadingFile f err) @@ -970,16 +971,18 @@ processCatch cmd pure $ REPLError msg ) -parseEmptyCmd : SourceEmptyRule (Maybe REPLCmd) +parseEmptyCmd : EmptyRule (Maybe REPLCmd) parseEmptyCmd = eoi *> (pure Nothing) -parseCmd : SourceEmptyRule (Maybe REPLCmd) +parseCmd : EmptyRule (Maybe REPLCmd) parseCmd = do c <- command; eoi; pure $ Just c export parseRepl : String -> Either Error (Maybe REPLCmd) parseRepl inp - = runParser "(interactive)" Nothing inp (parseEmptyCmd <|> parseCmd) + = case runParser "(interactive)" Nothing inp (parseEmptyCmd <|> parseCmd) of + Left err => Left err + Right (decor, result) => Right result export interpret : {auto c : Ref Ctxt Defs} -> diff --git a/src/Idris/REPL/Common.idr b/src/Idris/REPL/Common.idr index 9000b4b10..1c4d764a2 100644 --- a/src/Idris/REPL/Common.idr +++ b/src/Idris/REPL/Common.idr @@ -29,7 +29,8 @@ import System.File -- Output informational messages, unless quiet flag is set export -iputStrLn : {auto o : Ref ROpts REPLOpts} -> +iputStrLn : {auto c : Ref Ctxt Defs} -> + {auto o : Ref ROpts REPLOpts} -> Doc IdrisAnn -> Core () iputStrLn msg = do opts <- get ROpts @@ -118,8 +119,7 @@ emitWarnings put Ctxt (record { warnings = [] } defs) getFCLine : FC -> Maybe Int -getFCLine (MkFC _ (line, _) _) = Just line -getFCLine EmptyFC = Nothing +getFCLine = map startLine . isNonEmptyFC export updateErrorLine : {auto o : Ref ROpts REPLOpts} -> @@ -136,14 +136,15 @@ resetContext : {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> {auto s : Ref Syn SyntaxInfo} -> {auto m : Ref MD Metadata} -> + (source : String) -> Core () -resetContext +resetContext fname = do defs <- get Ctxt put Ctxt (record { options = clearNames (options defs) } !initDefs) addPrimitives put UST initUState put Syn initSyntax - put MD initMetadata + put MD (initMetadata fname) public export data EditResult : Type where diff --git a/src/Idris/Resugar.idr b/src/Idris/Resugar.idr index 0700e36b6..180c57a3c 100644 --- a/src/Idris/Resugar.idr +++ b/src/Idris/Resugar.idr @@ -29,11 +29,11 @@ unbracketApp tm = tm -- TODO: Deal with precedences mkOp : {auto s : Ref Syn SyntaxInfo} -> PTerm -> Core PTerm -mkOp tm@(PApp fc (PApp _ (PRef _ n) x) y) +mkOp tm@(PApp fc (PApp _ (PRef opFC n) x) y) = do syn <- get Syn case StringMap.lookup (nameRoot n) (infixes syn) of Nothing => pure tm - Just _ => pure (POp fc n (unbracketApp x) (unbracketApp y)) + Just _ => pure (POp fc opFC n (unbracketApp x) (unbracketApp y)) mkOp tm = pure tm export @@ -44,10 +44,10 @@ addBracket fc tm = if needed tm then PBracketed fc tm else tm needed (PBracketed _ _) = False needed (PRef _ _) = False needed (PPair _ _ _) = False - needed (PDPair _ _ _ _) = False + needed (PDPair _ _ _ _ _) = False needed (PUnit _) = False needed (PComprehension _ _ _) = False - needed (PList _ _) = False + needed (PList _ _ _) = False needed (PPrimVal _ _) = False needed tm = True @@ -113,13 +113,13 @@ mutual ||| Put the special names (Nil, ::, Pair, Z, S, etc) back as syntax ||| Returns `Nothing` in case there was nothing to resugar. sugarAppM : PTerm -> Maybe PTerm - sugarAppM (PApp fc (PApp _ (PRef _ (NS ns nm)) l) r) = + sugarAppM (PApp fc (PApp _ (PRef opFC (NS ns nm)) l) r) = if builtinNS == ns then case nameRoot nm of "Pair" => pure $ PPair fc (unbracket l) (unbracket r) "MkPair" => pure $ PPair fc (unbracket l) (unbracket r) "DPair" => case unbracket r of - PLam _ _ _ n _ r' => pure $ PDPair fc n (unbracket l) (unbracket r') + PLam _ _ _ n _ r' => pure $ PDPair fc opFC n (unbracket l) (unbracket r') _ => Nothing "Equal" => pure $ PEq fc (unbracket l) (unbracket r) "===" => pure $ PEq fc (unbracket l) (unbracket r) @@ -127,8 +127,8 @@ mutual _ => Nothing else if nameRoot nm == "::" then case sugarApp (unbracket r) of - PList fc xs => pure $ PList fc (unbracketApp l :: xs) - _ => Nothing + PList fc nilFC xs => pure $ PList fc nilFC ((opFC, unbracketApp l) :: xs) + _ => Nothing else Nothing sugarAppM tm = -- refolding natural numbers if the expression is a constant @@ -142,7 +142,7 @@ mutual "MkUnit" => pure $ PUnit fc _ => Nothing else if nameRoot nm == "Nil" - then pure $ PList fc [] + then pure $ PList fc fc [] else Nothing _ => Nothing @@ -464,14 +464,14 @@ cleanPTerm ptm cleanNode : PTerm -> Core PTerm cleanNode (PRef fc nm) = PRef fc <$> cleanName nm - cleanNode (POp fc op x y) = - (\ op => POp fc op x y) <$> cleanName op - cleanNode (PPrefixOp fc op x) = - (\ op => PPrefixOp fc op x) <$> cleanName op - cleanNode (PSectionL fc op x) = - (\ op => PSectionL fc op x) <$> cleanName op - cleanNode (PSectionR fc x op) = - PSectionR fc x <$> cleanName op + cleanNode (POp fc opFC op x y) = + (\ op => POp fc opFC op x y) <$> cleanName op + cleanNode (PPrefixOp fc opFC op x) = + (\ op => PPrefixOp fc opFC op x) <$> cleanName op + cleanNode (PSectionL fc opFC op x) = + (\ op => PSectionL fc opFC op x) <$> cleanName op + cleanNode (PSectionR fc opFC x op) = + PSectionR fc opFC x <$> cleanName op cleanNode tm = pure tm toCleanPTerm : {auto c : Ref Ctxt Defs} -> diff --git a/src/Idris/Syntax.idr b/src/Idris/Syntax.idr index e1b734a56..aeb8198f4 100644 --- a/src/Idris/Syntax.idr +++ b/src/Idris/Syntax.idr @@ -82,10 +82,10 @@ mutual -- Operators - POp : FC -> OpStr -> PTerm -> PTerm -> PTerm - PPrefixOp : FC -> OpStr -> PTerm -> PTerm - PSectionL : FC -> OpStr -> PTerm -> PTerm - PSectionR : FC -> PTerm -> OpStr -> PTerm + POp : (full, opFC : FC) -> OpStr -> PTerm -> PTerm -> PTerm + PPrefixOp : (full, opFC : FC) -> OpStr -> PTerm -> PTerm + PSectionL : (full, opFC : FC) -> OpStr -> PTerm -> PTerm + PSectionR : (full, opFC : FC) -> PTerm -> OpStr -> PTerm PEq : FC -> PTerm -> PTerm -> PTerm PBracketed : FC -> PTerm -> PTerm @@ -95,9 +95,10 @@ mutual PDoBlock : FC -> Maybe Namespace -> List PDo -> PTerm PBang : FC -> PTerm -> PTerm PIdiom : FC -> PTerm -> PTerm - PList : FC -> List PTerm -> PTerm + PList : (full, nilFC : FC) -> List (FC, PTerm) -> PTerm + -- ^ location of the conses PPair : FC -> PTerm -> PTerm -> PTerm - PDPair : FC -> PTerm -> PTerm -> PTerm -> PTerm + PDPair : (full, opFC : FC) -> PTerm -> PTerm -> PTerm -> PTerm PUnit : FC -> PTerm PIfThenElse : FC -> PTerm -> PTerm -> PTerm -> PTerm PComprehension : FC -> PTerm -> List PDo -> PTerm @@ -146,10 +147,10 @@ mutual getPTermLoc (PDotted fc _) = fc getPTermLoc (PImplicit fc) = fc getPTermLoc (PInfer fc) = fc - getPTermLoc (POp fc _ _ _) = fc - getPTermLoc (PPrefixOp fc _ _) = fc - getPTermLoc (PSectionL fc _ _) = fc - getPTermLoc (PSectionR fc _ _) = fc + getPTermLoc (POp fc _ _ _ _) = fc + getPTermLoc (PPrefixOp fc _ _ _) = fc + getPTermLoc (PSectionL fc _ _ _) = fc + getPTermLoc (PSectionR fc _ _ _) = fc getPTermLoc (PEq fc _ _) = fc getPTermLoc (PBracketed fc _) = fc getPTermLoc (PString fc _) = fc @@ -157,9 +158,9 @@ mutual getPTermLoc (PDoBlock fc _ _) = fc getPTermLoc (PBang fc _) = fc getPTermLoc (PIdiom fc _) = fc - getPTermLoc (PList fc _) = fc + getPTermLoc (PList fc _ _) = fc getPTermLoc (PPair fc _ _) = fc - getPTermLoc (PDPair fc _ _ _) = fc + getPTermLoc (PDPair fc _ _ _ _) = fc getPTermLoc (PUnit fc) = fc getPTermLoc (PIfThenElse fc _ _ _) = fc getPTermLoc (PComprehension fc _ _) = fc @@ -360,10 +361,15 @@ mutual getPDeclLoc (PDirective fc _) = fc getPDeclLoc (PBuiltin fc _ _) = fc - export - isPDef : PDecl -> Maybe (FC, List PClause) - isPDef (PDef annot cs) = Just (annot, cs) - isPDef _ = Nothing +export +isStrInterp : PStr -> Maybe FC +isStrInterp (StrInterp fc _) = Just fc +isStrInterp (StrLiteral _ _) = Nothing + +export +isPDef : PDecl -> Maybe (FC, List PClause) +isPDef (PDef annot cs) = Just (annot, cs) +isPDef _ = Nothing definedInData : PDataDecl -> List Name @@ -604,10 +610,10 @@ mutual showPrec d (PDotted _ p) = "." ++ showPrec d p showPrec _ (PImplicit _) = "_" showPrec _ (PInfer _) = "?" - showPrec d (POp _ op x y) = showPrec d x ++ " " ++ showPrecOp d op ++ " " ++ showPrec d y - showPrec d (PPrefixOp _ op x) = showPrec d op ++ showPrec d x - showPrec d (PSectionL _ op x) = "(" ++ showPrecOp d op ++ " " ++ showPrec d x ++ ")" - showPrec d (PSectionR _ x op) = "(" ++ showPrec d x ++ " " ++ showPrecOp d op ++ ")" + showPrec d (POp _ _ op x y) = showPrec d x ++ " " ++ showPrecOp d op ++ " " ++ showPrec d y + showPrec d (PPrefixOp _ _ op x) = showPrec d op ++ showPrec d x + showPrec d (PSectionL _ _ op x) = "(" ++ showPrecOp d op ++ " " ++ showPrec d x ++ ")" + showPrec d (PSectionR _ _ x op) = "(" ++ showPrec d x ++ " " ++ showPrecOp d op ++ ")" showPrec d (PEq fc l r) = showPrec d l ++ " = " ++ showPrec d r showPrec d (PBracketed _ tm) = "(" ++ showPrec d tm ++ ")" showPrec d (PString _ xs) = join " ++ " $ show <$> xs @@ -616,11 +622,11 @@ mutual = "do " ++ showSep " ; " (map showDo ds) showPrec d (PBang _ tm) = "!" ++ showPrec d tm showPrec d (PIdiom _ tm) = "[|" ++ showPrec d tm ++ "|]" - showPrec d (PList _ xs) - = "[" ++ showSep ", " (map (showPrec d) xs) ++ "]" + showPrec d (PList _ _ xs) + = "[" ++ showSep ", " (map (showPrec d . snd) xs) ++ "]" showPrec d (PPair _ l r) = "(" ++ showPrec d l ++ ", " ++ showPrec d r ++ ")" - showPrec d (PDPair _ l (PImplicit _) r) = "(" ++ showPrec d l ++ " ** " ++ showPrec d r ++ ")" - showPrec d (PDPair _ l ty r) = "(" ++ showPrec d l ++ " : " ++ showPrec d ty ++ + showPrec d (PDPair _ _ l (PImplicit _) r) = "(" ++ showPrec d l ++ " ** " ++ showPrec d r ++ ")" + showPrec d (PDPair _ _ l ty r) = "(" ++ showPrec d l ++ " : " ++ showPrec d ty ++ " ** " ++ showPrec d r ++ ")" showPrec _ (PUnit _) = "()" showPrec d (PIfThenElse _ x t e) = "if " ++ showPrec d x ++ " then " ++ showPrec d t ++ @@ -941,18 +947,18 @@ mapPTermM f = goPTerm where >>= f goPTerm t@(PImplicit _) = f t goPTerm t@(PInfer _) = f t - goPTerm (POp fc x y z) = - POp fc x <$> goPTerm y - <*> goPTerm z + goPTerm (POp fc opFC x y z) = + POp fc opFC x <$> goPTerm y + <*> goPTerm z >>= f - goPTerm (PPrefixOp fc x y) = - PPrefixOp fc x <$> goPTerm y + goPTerm (PPrefixOp fc opFC x y) = + PPrefixOp fc opFC x <$> goPTerm y >>= f - goPTerm (PSectionL fc x y) = - PSectionL fc x <$> goPTerm y + goPTerm (PSectionL fc opFC x y) = + PSectionL fc opFC x <$> goPTerm y >>= f - goPTerm (PSectionR fc x y) = - PSectionR fc <$> goPTerm x + goPTerm (PSectionR fc opFC x y) = + PSectionR fc opFC <$> goPTerm x <*> pure y >>= f goPTerm (PEq fc x y) = @@ -977,15 +983,16 @@ mapPTermM f = goPTerm where goPTerm (PIdiom fc x) = PIdiom fc <$> goPTerm x >>= f - goPTerm (PList fc xs) = - PList fc <$> goPTerms xs + goPTerm (PList fc nilFC xs) = + PList fc nilFC <$> goPairedPTerms xs >>= f goPTerm (PPair fc x y) = PPair fc <$> goPTerm x <*> goPTerm y >>= f - goPTerm (PDPair fc x y z) = - PDPair fc <$> goPTerm x + goPTerm (PDPair fc opFC x y z) = + PDPair fc opFC + <$> goPTerm x <*> goPTerm y <*> goPTerm z >>= f diff --git a/src/Libraries/Data/PosMap.idr b/src/Libraries/Data/PosMap.idr index f0cd2a11d..a67c43abc 100644 --- a/src/Libraries/Data/PosMap.idr +++ b/src/Libraries/Data/PosMap.idr @@ -577,18 +577,27 @@ greater : FilePos -> Interval -> Bool greater k (MkInterval (MkRange low _)) = fst low > k greater k NoInterval = False --- Finds all the interval that overlaps with the given interval. --- takeUntil selects all the intervals within the given upper bound, --- however the remaining interval are not necessarily adjacent in --- the sequence, thus it drops elements until the next intersecting --- interval with dropUntil. +||| Finds all the intervals that overlap with the given interval. +export inRange : MeasureRM a => FilePos -> FilePos -> PosMap a -> List a inRange low high t = matches (takeUntil (greater high) t) + -- takeUntil selects all the intervals within the given upper bound, + -- however the remaining interval are not necessarily adjacent in + -- the sequence, thus it drops elements until the next intersecting + -- interval with dropUntil. where matches : PosMap a -> List a matches xs = case viewl (dropUntil (atleast low) xs) of EmptyL => [] x <: xs' => x :: assert_total (matches xs') +||| Finds the values matching the exact interval input +export +exactRange : MeasureRM a => FilePos -> FilePos -> PosMap a -> List a +exactRange low high t = flip mapMaybe (inRange low high t) $ \ a => + do let (MkRange rng _) = measureRM a + guard (rng == (low, high)) + pure a + ||| Returns all the interval that contains the given point. export searchPos : MeasureRM a => FilePos -> PosMap a -> List a diff --git a/src/Libraries/Text/Parser.idr b/src/Libraries/Text/Parser.idr index 56be98ce8..15c472690 100644 --- a/src/Libraries/Text/Parser.idr +++ b/src/Libraries/Text/Parser.idr @@ -15,7 +15,7 @@ import public Libraries.Text.Token export match : (Eq k, TokenKind k) => (kind : k) -> - Grammar (Token k) True (TokType kind) + Grammar state (Token k) True (TokType kind) match k = terminal "Unrecognised input" $ \t => if t.kind == k then Just $ tokValue k t.text @@ -25,8 +25,8 @@ match k = terminal "Unrecognised input" $ ||| match. May match the empty input. export option : {c : Bool} -> - (def : a) -> (p : Grammar tok c a) -> - Grammar tok False a + (def : a) -> (p : Grammar state tok c a) -> + Grammar state tok False a option {c = False} def p = p <|> pure def option {c = True} def p = p <|> pure def @@ -34,8 +34,8 @@ option {c = True} def p = p <|> pure def ||| To provide a default value, use `option` instead. export optional : {c : Bool} -> - (p : Grammar tok c a) -> - Grammar tok False (Maybe a) + (p : Grammar state tok c a) -> + Grammar state tok False (Maybe a) optional p = option Nothing (map Just p) ||| Try to parse one thing or the other, producing a value that indicates @@ -43,9 +43,9 @@ optional p = option Nothing (map Just p) ||| takes priority. export choose : {c1, c2 : Bool} -> - (l : Grammar tok c1 a) -> - (r : Grammar tok c2 b) -> - Grammar tok (c1 && c2) (Either a b) + (l : Grammar state tok c1 a) -> + (r : Grammar state tok c2 b) -> + Grammar state tok (c1 && c2) (Either a b) choose l r = map Left l <|> map Right r ||| Produce a grammar by applying a function to each element of a container, @@ -53,9 +53,9 @@ choose l r = map Left l <|> map Right r ||| container is empty. export choiceMap : {c : Bool} -> - (a -> Grammar tok c b) -> + (a -> Grammar state tok c b) -> Foldable t => t a -> - Grammar tok c b + Grammar state tok c b choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in f x <|> acc) (fail "No more options") xs @@ -67,28 +67,28 @@ choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in export choice : Foldable t => {c : Bool} -> - t (Grammar tok c a) -> - Grammar tok c a + t (Grammar state tok c a) -> + Grammar state tok c a choice = choiceMap id mutual ||| Parse one or more things export - some : Grammar tok True a -> - Grammar tok True (List1 a) + some : Grammar state tok True a -> + Grammar state tok True (List1 a) some p = pure (!p ::: !(many p)) ||| Parse zero or more things (may match the empty input) export - many : Grammar tok True a -> - Grammar tok False (List a) + many : Grammar state tok True a -> + Grammar state tok False (List a) many p = option [] (forget <$> some p) mutual private count1 : (q : Quantity) -> - (p : Grammar tok True a) -> - Grammar tok True (List a) + (p : Grammar state tok True a) -> + Grammar state tok True (List a) count1 q p = do x <- p seq (count q p) (\xs => pure (x :: xs)) @@ -96,8 +96,8 @@ mutual ||| Parse `p`, repeated as specified by `q`, returning the list of values. export count : (q : Quantity) -> - (p : Grammar tok True a) -> - Grammar tok (isSucc (min q)) (List a) + (p : Grammar state tok True a) -> + Grammar state tok (isSucc (min q)) (List a) count (Qty Z Nothing) p = many p count (Qty Z (Just Z)) _ = pure [] count (Qty Z (Just (S max))) p = option [] $ count1 (atMost max) p @@ -110,9 +110,9 @@ mutual ||| list of values from `p`. Guaranteed to consume input. export someTill : {c : Bool} -> - (end : Grammar tok c e) -> - (p : Grammar tok True a) -> - Grammar tok True (List1 a) + (end : Grammar state tok c e) -> + (p : Grammar state tok True a) -> + Grammar state tok True (List1 a) someTill {c} end p = do x <- p seq (manyTill end p) (\xs => pure (x ::: xs)) @@ -121,9 +121,9 @@ mutual ||| list of values from `p`. Guaranteed to consume input if `end` consumes. export manyTill : {c : Bool} -> - (end : Grammar tok c e) -> - (p : Grammar tok True a) -> - Grammar tok c (List a) + (end : Grammar state tok c e) -> + (p : Grammar state tok True a) -> + Grammar state tok c (List a) manyTill {c} end p = rewrite sym (andTrueNeutral c) in map (const []) end <|> (forget <$> someTill end p) @@ -132,9 +132,9 @@ mutual ||| returning its value. export afterSome : {c : Bool} -> - (skip : Grammar tok True s) -> - (p : Grammar tok c a) -> - Grammar tok True a + (skip : Grammar state tok True s) -> + (p : Grammar state tok c a) -> + Grammar state tok True a afterSome skip p = do ignore $ skip afterMany skip p @@ -142,18 +142,18 @@ mutual ||| returning its value. export afterMany : {c : Bool} -> - (skip : Grammar tok True s) -> - (p : Grammar tok c a) -> - Grammar tok c a + (skip : Grammar state tok True s) -> + (p : Grammar state tok c a) -> + Grammar state tok c a afterMany {c} skip p = rewrite sym (andTrueNeutral c) in p <|> afterSome skip p ||| Parse one or more things, each separated by another thing. export sepBy1 : {c : Bool} -> - (sep : Grammar tok True s) -> - (p : Grammar tok c a) -> - Grammar tok c (List1 a) + (sep : Grammar state tok True s) -> + (p : Grammar state tok c a) -> + Grammar state tok c (List1 a) sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in [| p ::: many (sep *> p) |] @@ -161,18 +161,18 @@ sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in ||| match the empty input. export sepBy : {c : Bool} -> - (sep : Grammar tok True s) -> - (p : Grammar tok c a) -> - Grammar tok False (List a) + (sep : Grammar state tok True s) -> + (p : Grammar state tok c a) -> + Grammar state tok False (List a) sepBy sep p = option [] $ forget <$> sepBy1 sep p ||| Parse one or more instances of `p` separated by and optionally terminated by ||| `sep`. export sepEndBy1 : {c : Bool} -> - (sep : Grammar tok True s) -> - (p : Grammar tok c a) -> - Grammar tok c (List1 a) + (sep : Grammar state tok True s) -> + (p : Grammar state tok c a) -> + Grammar state tok c (List1 a) sepEndBy1 {c} sep p = rewrite sym (orFalseNeutral c) in sepBy1 sep p <* optional sep @@ -180,32 +180,49 @@ sepEndBy1 {c} sep p = rewrite sym (orFalseNeutral c) in ||| by `sep`. Will not match a separator by itself. export sepEndBy : {c : Bool} -> - (sep : Grammar tok True s) -> - (p : Grammar tok c a) -> - Grammar tok False (List a) + (sep : Grammar state tok True s) -> + (p : Grammar state tok c a) -> + Grammar state tok False (List a) sepEndBy sep p = option [] $ forget <$> sepEndBy1 sep p ||| Parse one or more instances of `p`, separated and terminated by `sep`. export endBy1 : {c : Bool} -> - (sep : Grammar tok True s) -> - (p : Grammar tok c a) -> - Grammar tok True (List1 a) + (sep : Grammar state tok True s) -> + (p : Grammar state tok c a) -> + Grammar state tok True (List1 a) endBy1 {c} sep p = some $ rewrite sym (orTrueTrue c) in p <* sep export endBy : {c : Bool} -> - (sep : Grammar tok True s) -> - (p : Grammar tok c a) -> - Grammar tok False (List a) + (sep : Grammar state tok True s) -> + (p : Grammar state tok c a) -> + Grammar state tok False (List a) endBy sep p = option [] $ forget <$> endBy1 sep p ||| Parse an instance of `p` that is between `left` and `right`. export between : {c : Bool} -> - (left : Grammar tok True l) -> - (right : Grammar tok True r) -> - (p : Grammar tok c a) -> - Grammar tok True a + (left : Grammar state tok True l) -> + (right : Grammar state tok True r) -> + (p : Grammar state tok c a) -> + Grammar state tok True a between left right contents = left *> contents <* right + +export +location : Grammar state token False (Int, Int) +location = startBounds <$> position + +export +endLocation : Grammar state token False (Int, Int) +endLocation = endBounds <$> position + +export +column : Grammar state token False Int +column = snd <$> location + +public export +when : Bool -> Lazy (Grammar state token False ()) -> Grammar state token False () +when True y = y +when False y = pure () diff --git a/src/Libraries/Text/Parser/Core.idr b/src/Libraries/Text/Parser/Core.idr index 6e18f516c..6f6001f78 100644 --- a/src/Libraries/Text/Parser/Core.idr +++ b/src/Libraries/Text/Parser/Core.idr @@ -16,37 +16,39 @@ import public Libraries.Text.Bounded ||| to be non-empty - that is, successfully parsing the language is guaranteed ||| to consume some input. export -data Grammar : (tok : Type) -> (consumes : Bool) -> Type -> Type where - Empty : (val : ty) -> Grammar tok False ty - Terminal : String -> (tok -> Maybe a) -> Grammar tok True a - NextIs : String -> (tok -> Bool) -> Grammar tok False tok - EOF : Grammar tok False () +data Grammar : (state : Type) -> (tok : Type) -> (consumes : Bool) -> Type -> Type where + Empty : (val : ty) -> Grammar state tok False ty + Terminal : String -> (tok -> Maybe a) -> Grammar state tok True a + NextIs : String -> (tok -> Bool) -> Grammar state tok False tok + EOF : Grammar state tok False () - Fail : (location : Maybe Bounds) -> Bool -> String -> Grammar tok c ty - Try : Grammar tok c ty -> Grammar tok c ty + Fail : (location : Maybe Bounds) -> Bool -> String -> Grammar state tok c ty + Try : Grammar state tok c ty -> Grammar state tok c ty - Commit : Grammar tok False () - MustWork : Grammar tok c a -> Grammar tok c a + Commit : Grammar state tok False () + MustWork : Grammar state tok c a -> Grammar state tok c a SeqEat : {c2 : Bool} -> - Grammar tok True a -> Inf (a -> Grammar tok c2 b) -> - Grammar tok True b + Grammar state tok True a -> Inf (a -> Grammar state tok c2 b) -> + Grammar state tok True b SeqEmpty : {c1, c2 : Bool} -> - Grammar tok c1 a -> (a -> Grammar tok c2 b) -> - Grammar tok (c1 || c2) b + Grammar state tok c1 a -> (a -> Grammar state tok c2 b) -> + Grammar state tok (c1 || c2) b ThenEat : {c2 : Bool} -> - Grammar tok True () -> Inf (Grammar tok c2 a) -> - Grammar tok True a + Grammar state tok True () -> Inf (Grammar state tok c2 a) -> + Grammar state tok True a ThenEmpty : {c1, c2 : Bool} -> - Grammar tok c1 () -> Grammar tok c2 a -> - Grammar tok (c1 || c2) a + Grammar state tok c1 () -> Grammar state tok c2 a -> + Grammar state tok (c1 || c2) a Alt : {c1, c2 : Bool} -> - Grammar tok c1 ty -> Lazy (Grammar tok c2 ty) -> - Grammar tok (c1 && c2) ty - Bounds : Grammar tok c ty -> Grammar tok c (WithBounds ty) - Position : Grammar tok False Bounds + Grammar state tok c1 ty -> Lazy (Grammar state tok c2 ty) -> + Grammar state tok (c1 && c2) ty + Bounds : Grammar state tok c ty -> Grammar state tok c (WithBounds ty) + Position : Grammar state tok False Bounds + + Act : state -> Grammar state tok False () ||| Sequence two grammars. If either consumes some input, the sequence is ||| guaranteed to consume some input. If the first one consumes input, the @@ -54,9 +56,9 @@ data Grammar : (tok : Type) -> (consumes : Bool) -> Type -> Type where ||| consumed and therefore the input is smaller) export %inline (>>=) : {c1, c2 : Bool} -> - Grammar tok c1 a -> - inf c1 (a -> Grammar tok c2 b) -> - Grammar tok (c1 || c2) b + Grammar state tok c1 a -> + inf c1 (a -> Grammar state tok c2 b) -> + Grammar state tok (c1 || c2) b (>>=) {c1 = False} = SeqEmpty (>>=) {c1 = True} = SeqEat @@ -66,9 +68,9 @@ export %inline ||| consumed and therefore the input is smaller) public export %inline %tcinline (>>) : {c1, c2 : Bool} -> - Grammar tok c1 () -> - inf c1 (Grammar tok c2 a) -> - Grammar tok (c1 || c2) a + Grammar state tok c1 () -> + inf c1 (Grammar state tok c2 a) -> + Grammar state tok (c1 || c2) a (>>) {c1 = False} = ThenEmpty (>>) {c1 = True} = ThenEat @@ -77,23 +79,23 @@ public export %inline %tcinline ||| of `>>=`. export %inline seq : {c1,c2 : Bool} -> - Grammar tok c1 a -> - (a -> Grammar tok c2 b) -> - Grammar tok (c1 || c2) b + Grammar state tok c1 a -> + (a -> Grammar state tok c2 b) -> + Grammar state tok (c1 || c2) b seq = SeqEmpty ||| Sequence a grammar followed by the grammar it returns. export %inline join : {c1,c2 : Bool} -> - Grammar tok c1 (Grammar tok c2 a) -> - Grammar tok (c1 || c2) a + Grammar state tok c1 (Grammar state tok c2 a) -> + Grammar state tok (c1 || c2) a join {c1 = False} p = SeqEmpty p id join {c1 = True} p = SeqEat p id ||| Allows the result of a grammar to be mapped to a different value. export {c : _} -> -Functor (Grammar tok c) where +Functor (Grammar state tok c) where map f (Empty val) = Empty (f val) map f (Fail bd fatal msg) = Fail bd fatal msg map f (Try g) = Try (map f g) @@ -119,9 +121,9 @@ Functor (Grammar tok c) where ||| guaranteed to consume. export %inline (<|>) : {c1,c2 : Bool} -> - Grammar tok c1 ty -> - Lazy (Grammar tok c2 ty) -> - Grammar tok (c1 && c2) ty + Grammar state tok c1 ty -> + Lazy (Grammar state tok c2 ty) -> + Grammar state tok (c1 && c2) ty (<|>) = Alt infixr 2 <||> @@ -129,9 +131,9 @@ infixr 2 <||> ||| combination is guaranteed to consume. export (<||>) : {c1,c2 : Bool} -> - Grammar tok c1 a -> - Lazy (Grammar tok c2 b) -> - Grammar tok (c1 && c2) (Either a b) + Grammar state tok c1 a -> + Lazy (Grammar state tok c2 b) -> + Grammar state tok (c1 && c2) (Either a b) (<||>) p q = (Left <$> p) <|> (Right <$> q) ||| Sequence a grammar with value type `a -> b` and a grammar @@ -140,33 +142,37 @@ export ||| Guaranteed to consume if either grammar consumes. export %inline (<*>) : {c1, c2 : Bool} -> - Grammar tok c1 (a -> b) -> - Grammar tok c2 a -> - Grammar tok (c1 || c2) b + Grammar state tok c1 (a -> b) -> + Grammar state tok c2 a -> + Grammar state tok (c1 || c2) b (<*>) x y = SeqEmpty x (\f => map f y) ||| Sequence two grammars. If both succeed, use the value of the first one. ||| Guaranteed to consume if either grammar consumes. export %inline (<*) : {c1,c2 : Bool} -> - Grammar tok c1 a -> - Grammar tok c2 b -> - Grammar tok (c1 || c2) a + Grammar state tok c1 a -> + Grammar state tok c2 b -> + Grammar state tok (c1 || c2) a (<*) x y = map const x <*> y ||| Sequence two grammars. If both succeed, use the value of the second one. ||| Guaranteed to consume if either grammar consumes. export %inline (*>) : {c1,c2 : Bool} -> - Grammar tok c1 a -> - Grammar tok c2 b -> - Grammar tok (c1 || c2) b + Grammar state tok c1 a -> + Grammar state tok c2 b -> + Grammar state tok (c1 || c2) b (*>) x y = map (const id) x <*> y +export %inline +act : state -> Grammar state tok False () +act = Act + ||| Produce a grammar that can parse a different type of token by providing a ||| function converting the new token type into the original one. export -mapToken : (a -> b) -> Grammar b c ty -> Grammar a c ty +mapToken : (a -> b) -> Grammar state b c ty -> Grammar state a c ty mapToken f (Empty val) = Empty val mapToken f (Terminal msg g) = Terminal msg (g . f) mapToken f (NextIs msg g) = SeqEmpty (NextIs msg (g . f)) (Empty . f) @@ -186,149 +192,151 @@ mapToken f (ThenEmpty act next) mapToken f (Alt x y) = Alt (mapToken f x) (mapToken f y) mapToken f (Bounds act) = Bounds (mapToken f act) mapToken f Position = Position +mapToken f (Act action) = Act action ||| Always succeed with the given value. export %inline -pure : (val : ty) -> Grammar tok False ty +pure : (val : ty) -> Grammar state tok False ty pure = Empty ||| Check whether the next token satisfies a predicate export %inline -nextIs : String -> (tok -> Bool) -> Grammar tok False tok +nextIs : String -> (tok -> Bool) -> Grammar state tok False tok nextIs = NextIs ||| Look at the next token in the input export %inline -peek : Grammar tok False tok +peek : Grammar state tok False tok peek = nextIs "Unrecognised token" (const True) ||| Succeeds if running the predicate on the next token returns Just x, ||| returning x. Otherwise fails. export %inline -terminal : String -> (tok -> Maybe a) -> Grammar tok True a +terminal : String -> (tok -> Maybe a) -> Grammar state tok True a terminal = Terminal ||| Always fail with a message export %inline -fail : String -> Grammar tok c ty +fail : String -> Grammar state tok c ty fail = Fail Nothing False ||| Always fail with a message and a location export %inline -failLoc : Bounds -> String -> Grammar tok c ty +failLoc : Bounds -> String -> Grammar state tok c ty failLoc b = Fail (Just b) False export %inline -fatalError : String -> Grammar tok c ty +fatalError : String -> Grammar state tok c ty fatalError = Fail Nothing True export %inline -fatalLoc : Bounds -> String -> Grammar tok c ty +fatalLoc : Bounds -> String -> Grammar state tok c ty fatalLoc b = Fail (Just b) True ||| Catch a fatal error export %inline -try : Grammar tok c ty -> Grammar tok c ty +try : Grammar state tok c ty -> Grammar state tok c ty try = Try ||| Succeed if the input is empty export %inline -eof : Grammar tok False () +eof : Grammar state tok False () eof = EOF ||| Commit to an alternative; if the current branch of an alternative ||| fails to parse, no more branches will be tried export %inline -commit : Grammar tok False () +commit : Grammar state tok False () commit = Commit ||| If the parser fails, treat it as a fatal error export %inline -mustWork : {c : Bool} -> Grammar tok c ty -> Grammar tok c ty +mustWork : {c : Bool} -> Grammar state tok c ty -> Grammar state tok c ty mustWork = MustWork export %inline -bounds : Grammar tok c ty -> Grammar tok c (WithBounds ty) +bounds : Grammar state tok c ty -> Grammar state tok c (WithBounds ty) bounds = Bounds export %inline -position : Grammar tok False Bounds +position : Grammar state tok False Bounds position = Position -data ParseResult : Type -> Type -> Type where +data ParseResult : Type -> Type -> Type -> Type where Failure : (committed : Bool) -> (fatal : Bool) -> - (err : String) -> (location : Maybe Bounds) -> ParseResult tok ty - Res : (committed : Bool) -> - (val : WithBounds ty) -> (more : List (WithBounds tok)) -> ParseResult tok ty + (err : String) -> (location : Maybe Bounds) -> ParseResult state tok ty + Res : state -> (committed : Bool) -> + (val : WithBounds ty) -> (more : List (WithBounds tok)) -> ParseResult state tok ty -mergeWith : WithBounds ty -> ParseResult tok sy -> ParseResult tok sy -mergeWith x (Res committed val more) = Res committed (mergeBounds x val) more +mergeWith : WithBounds ty -> ParseResult state tok sy -> ParseResult state tok sy +mergeWith x (Res s committed val more) = Res s committed (mergeBounds x val) more mergeWith x v = v -doParse : (commit : Bool) -> - (act : Grammar tok c ty) -> +doParse : Semigroup state => state -> (commit : Bool) -> + (act : Grammar state tok c ty) -> (xs : List (WithBounds tok)) -> - ParseResult tok ty -doParse com (Empty val) xs = Res com (irrelevantBounds val) xs -doParse com (Fail location fatal str) xs = Failure com fatal str (maybe (bounds <$> head' xs) Just location) -doParse com (Try g) xs = case doParse com g xs of + ParseResult state tok ty +doParse s com (Empty val) xs = Res s com (irrelevantBounds val) xs +doParse s com (Fail location fatal str) xs = Failure com fatal str (maybe (bounds <$> head' xs) Just location) +doParse s com (Try g) xs = case doParse s com g xs of -- recover from fatal match but still propagate the 'commit' Failure com _ msg ts => Failure com False msg ts res => res -doParse com Commit xs = Res True (irrelevantBounds ()) xs -doParse com (MustWork g) xs = - case assert_total (doParse com g xs) of +doParse s com Commit xs = Res s True (irrelevantBounds ()) xs +doParse s com (MustWork g) xs = + case assert_total (doParse s com g xs) of Failure com' _ msg ts => Failure com' True msg ts res => res -doParse com (Terminal err f) [] = Failure com False "End of input" Nothing -doParse com (Terminal err f) (x :: xs) = +doParse s com (Terminal err f) [] = Failure com False "End of input" Nothing +doParse s com (Terminal err f) (x :: xs) = case f x.val of Nothing => Failure com False err (Just x.bounds) - Just a => Res com (const a <$> x) xs -doParse com EOF [] = Res com (irrelevantBounds ()) [] -doParse com EOF (x :: xs) = Failure com False "Expected end of input" (Just x.bounds) -doParse com (NextIs err f) [] = Failure com False "End of input" Nothing -doParse com (NextIs err f) (x :: xs) + Just a => Res s com (const a <$> x) xs +doParse s com EOF [] = Res s com (irrelevantBounds ()) [] +doParse s com EOF (x :: xs) = Failure com False "Expected end of input" (Just x.bounds) +doParse s com (NextIs err f) [] = Failure com False "End of input" Nothing +doParse s com (NextIs err f) (x :: xs) = if f x.val - then Res com (removeIrrelevance x) (x :: xs) + then Res s com (removeIrrelevance x) (x :: xs) else Failure com False err (Just x.bounds) -doParse com (Alt {c1} {c2} x y) xs - = case doParse False x xs of +doParse s com (Alt {c1} {c2} x y) xs + = case doParse s False x xs of Failure com' fatal msg ts => if com' || fatal -- If the alternative had committed, don't try the -- other branch (and reset commit flag) then Failure com fatal msg ts - else assert_total (doParse False y xs) + else assert_total (doParse s False y xs) -- Successfully parsed the first option, so use the outer commit flag - Res _ val xs => Res com val xs -doParse com (SeqEmpty act next) xs - = case assert_total (doParse com act xs) of + Res s _ val xs => Res s com val xs +doParse s com (SeqEmpty act next) xs + = case assert_total (doParse s com act xs) of Failure com fatal msg ts => Failure com fatal msg ts - Res com v xs => - mergeWith v (assert_total $ doParse com (next v.val) xs) -doParse com (SeqEat act next) xs - = case assert_total (doParse com act xs) of + Res s com v xs => + mergeWith v (assert_total $ doParse s com (next v.val) xs) +doParse s com (SeqEat act next) xs + = case assert_total (doParse s com act xs) of Failure com fatal msg ts => Failure com fatal msg ts - Res com v xs => - mergeWith v (assert_total $ doParse com (next v.val) xs) -doParse com (ThenEmpty act next) xs - = case assert_total (doParse com act xs) of + Res s com v xs => + mergeWith v (assert_total $ doParse s com (next v.val) xs) +doParse s com (ThenEmpty act next) xs + = case assert_total (doParse s com act xs) of Failure com fatal msg ts => Failure com fatal msg ts - Res com v xs => - mergeWith v (assert_total $ doParse com next xs) -doParse com (ThenEat act next) xs - = case assert_total (doParse com act xs) of + Res s com v xs => + mergeWith v (assert_total $ doParse s com next xs) +doParse s com (ThenEat act next) xs + = case assert_total (doParse s com act xs) of Failure com fatal msg ts => Failure com fatal msg ts - Res com v xs => - mergeWith v (assert_total $ doParse com next xs) -doParse com (Bounds act) xs - = case assert_total (doParse com act xs) of + Res s com v xs => + mergeWith v (assert_total $ doParse s com next xs) +doParse s com (Bounds act) xs + = case assert_total (doParse s com act xs) of Failure com fatal msg ts => Failure com fatal msg ts - Res com v xs => Res com (const v <$> v) xs -doParse com Position [] = Failure com False "End of input" Nothing -doParse com Position (x :: xs) - = Res com (irrelevantBounds x.bounds) (x :: xs) + Res s com v xs => Res s com (const v <$> v) xs +doParse s com Position [] = Failure com False "End of input" Nothing +doParse s com Position (x :: xs) + = Res s com (irrelevantBounds x.bounds) (x :: xs) +doParse s com (Act action) xs = Res (s <+> action) com (irrelevantBounds ()) xs public export data ParsingError tok = Error String (Maybe Bounds) @@ -337,9 +345,17 @@ data ParsingError tok = Error String (Maybe Bounds) ||| returns a pair of the parse result and the unparsed tokens (the remaining ||| input). export -parse : {c : Bool} -> (act : Grammar tok c ty) -> (xs : List (WithBounds tok)) -> +parse : {c : Bool} -> (act : Grammar () tok c ty) -> (xs : List (WithBounds tok)) -> Either (ParsingError tok) (ty, List (WithBounds tok)) parse act xs - = case doParse False act xs of + = case doParse neutral False act xs of Failure _ _ msg ts => Left (Error msg ts) - Res _ v rest => Right (v.val, rest) + Res _ _ v rest => Right (v.val, rest) + +export +parseWith : Monoid state => {c : Bool} -> (act : Grammar state tok c ty) -> (xs : List (WithBounds tok)) -> + Either (ParsingError tok) (state, ty, List (WithBounds tok)) +parseWith act xs + = case doParse neutral False act xs of + Failure _ _ msg ts => Left (Error msg ts) + Res s _ v rest => Right (s, v.val, rest) diff --git a/src/Libraries/Utils/Path.idr b/src/Libraries/Utils/Path.idr index f8cf57f4f..104c85200 100644 --- a/src/Libraries/Utils/Path.idr +++ b/src/Libraries/Utils/Path.idr @@ -136,6 +136,9 @@ Eq PathTokenKind where PathToken : Type PathToken = Token PathTokenKind +PathGrammar : Bool -> Type -> Type +PathGrammar = Grammar () PathToken + TokenKind PathTokenKind where TokType PTText = String TokType (PTPunct _) = () @@ -156,7 +159,7 @@ lexPath : String -> List (WithBounds PathToken) lexPath str = let (tokens, _, _, _) = lex pathTokenMap str in tokens -- match both '/' and '\\' regardless of the platform. -bodySeparator : Grammar PathToken True () +bodySeparator : PathGrammar True () bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/') -- Windows will automatically translate '/' to '\\'. And the verbatim prefix, @@ -164,7 +167,7 @@ bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/') -- However, we just parse it and ignore it. -- -- Example: \\?\ -verbatim : Grammar PathToken True () +verbatim : PathGrammar True () verbatim = do ignore $ count (exactly 2) $ match $ PTPunct '\\' @@ -173,7 +176,7 @@ verbatim = pure () -- Example: \\server\share -unc : Grammar PathToken True Volume +unc : PathGrammar True Volume unc = do ignore $ count (exactly 2) $ match $ PTPunct '\\' @@ -183,7 +186,7 @@ unc = pure $ UNC server share -- Example: \\?\server\share -verbatimUnc : Grammar PathToken True Volume +verbatimUnc : PathGrammar True Volume verbatimUnc = do verbatim @@ -193,7 +196,7 @@ verbatimUnc = pure $ UNC server share -- Example: C: -disk : Grammar PathToken True Volume +disk : PathGrammar True Volume disk = do text <- match PTText @@ -204,31 +207,31 @@ disk = pure $ Disk (toUpper disk) -- Example: \\?\C: -verbatimDisk : Grammar PathToken True Volume +verbatimDisk : PathGrammar True Volume verbatimDisk = do verbatim disk <- disk pure disk -parseVolume : Grammar PathToken True Volume +parseVolume : PathGrammar True Volume parseVolume = verbatimUnc <|> verbatimDisk <|> unc <|> disk -parseBody : Grammar PathToken True Body +parseBody : PathGrammar True Body parseBody = do text <- match PTText - the (Grammar _ False _) $ + the (PathGrammar False _) $ case text of ".." => pure ParentDir "." => pure CurDir normal => pure (Normal normal) -parsePath : Grammar PathToken False Path +parsePath : PathGrammar False Path parsePath = do vol <- optional parseVolume diff --git a/src/Libraries/Utils/Shunting.idr b/src/Libraries/Utils/Shunting.idr index c09866844..3f910a273 100644 --- a/src/Libraries/Utils/Shunting.idr +++ b/src/Libraries/Utils/Shunting.idr @@ -21,8 +21,10 @@ data OpPrec -- Tokens are either operators or already parsed expressions in some -- higher level language public export -data Tok op a = Op FC op OpPrec - | Expr a +data Tok op a + = ||| The second FC is for the operator alone + Op FC FC op OpPrec + | Expr a -- The result of shunting is a parse tree with the precedences made explicit -- in the tree. @@ -34,14 +36,14 @@ data Tok op a = Op FC op OpPrec -- there's a better way though! public export -data Tree op a = Infix FC op (Tree op a) (Tree op a) - | Pre FC op (Tree op a) +data Tree op a = Infix FC FC op (Tree op a) (Tree op a) + | Pre FC FC op (Tree op a) | Leaf a export (Show op, Show a) => Show (Tree op a) where - show (Infix _ op l r) = "(" ++ show op ++ " " ++ show l ++ " " ++ show r ++ ")" - show (Pre _ op l) = "(" ++ show op ++ " " ++ show l ++ ")" + show (Infix _ _ op l r) = "(" ++ show op ++ " " ++ show l ++ " " ++ show r ++ ")" + show (Pre _ _ op l) = "(" ++ show op ++ " " ++ show l ++ ")" show (Leaf val) = show val Show OpPrec where @@ -52,7 +54,7 @@ Show OpPrec where export (Show op, Show a) => Show (Tok op a) where - show (Op _ op p) = show op ++ " " ++ show p + show (Op _ _ op p) = show op ++ " " ++ show p show (Expr val) = show val -- Label for the output queue state @@ -60,9 +62,9 @@ data Out : Type where output : List (Tree op a) -> Tok op a -> Core (List (Tree op a)) -output [] (Op _ _ _) = throw (InternalError "Invalid input to shunting") -output (x :: stk) (Op loc str (Prefix _)) = pure $ Pre loc str x :: stk -output (x :: y :: stk) (Op loc str _) = pure $ Infix loc str y x :: stk +output [] (Op _ _ _ _) = throw (InternalError "Invalid input to shunting") +output (x :: stk) (Op loc opFC str (Prefix _)) = pure $ Pre loc opFC str x :: stk +output (x :: y :: stk) (Op loc opFC str _) = pure $ Infix loc opFC str y x :: stk output stk (Expr a) = pure $ Leaf a :: stk output _ _ = throw (InternalError "Invalid input to shunting") @@ -101,36 +103,32 @@ higher loc opl l opr r ((getPrec l == getPrec r) && isLAssoc l) processStack : Show op => {auto o : Ref Out (List (Tree op a))} -> - List (FC, op, OpPrec) -> op -> OpPrec -> - Core (List (FC, op, OpPrec)) + List (FC, FC, op, OpPrec) -> op -> OpPrec -> + Core (List (FC, FC, op, OpPrec)) processStack [] op prec = pure [] -processStack ((loc, opx, sprec) :: xs) opy prec +processStack (x@(loc, opFC, opx, sprec) :: xs) opy prec = if !(higher loc opx sprec opy prec) - then do emit (Op loc opx sprec) + then do emit (Op loc opFC opx sprec) processStack xs opy prec - else pure ((loc, opx, sprec) :: xs) + else pure (x :: xs) shunt : Show op => {auto o : Ref Out (List (Tree op a))} -> - (opstk : List (FC, op, OpPrec)) -> + (opstk : List (FC, FC, op, OpPrec)) -> List (Tok op a) -> Core (Tree op a) shunt stk (Expr x :: rest) = do emit (Expr x) shunt stk rest -shunt stk (Op loc op prec :: rest) +shunt stk (Op loc opFC op prec :: rest) = do stk' <- processStack stk op prec - shunt ((loc, op, prec) :: stk') rest + shunt ((loc, opFC, op, prec) :: stk') rest shunt stk [] - = do traverse_ (\s => emit (Op (sloc s) (sop s) (sprec s))) stk + = do traverse_ (emit . mkOp) stk [out] <- get Out | out => throw (InternalError "Invalid input to shunting") pure out where - sloc : (annot, b, c) -> annot - sloc (x, y, z) = x - sop : (annot, b, c) -> b - sop (x, y, z) = y - sprec : (annot, b, c) -> c - sprec (x, y, z) = z + mkOp : (FC, FC, op, OpPrec) -> Tok op a + mkOp (loc, opFC, op, prec) = Op loc opFC op prec export parseOps : Show op => List (Tok op a) -> Core (Tree op a) diff --git a/src/Parser/Rule/Common.idr b/src/Parser/Rule/Common.idr deleted file mode 100644 index 724064055..000000000 --- a/src/Parser/Rule/Common.idr +++ /dev/null @@ -1,26 +0,0 @@ -module Parser.Rule.Common - -import public Libraries.Text.Lexer -import public Libraries.Text.Parser - -%default total - -public export -Rule : Type -> Type -> Type -Rule token ty = Grammar token True ty - -public export -EmptyRule : Type -> Type -> Type -EmptyRule token ty = Grammar token False ty - -export -location : {token : _} -> EmptyRule token (Int, Int) -location = startBounds <$> position - -export -endLocation : {token : _} -> EmptyRule token (Int, Int) -endLocation = endBounds <$> position - -export -column : {token : _ } -> EmptyRule token Int -column = snd <$> location diff --git a/src/Parser/Rule/Package.idr b/src/Parser/Rule/Package.idr index 465153f28..dd7f5459d 100644 --- a/src/Parser/Rule/Package.idr +++ b/src/Parser/Rule/Package.idr @@ -1,7 +1,6 @@ module Parser.Rule.Package import public Parser.Lexer.Package -import public Parser.Rule.Common import Data.List import Data.List1 @@ -12,11 +11,11 @@ import Core.Name.Namespace public export Rule : Type -> Type -Rule = Rule Token +Rule = Grammar () Token True public export -PackageEmptyRule : Type -> Type -PackageEmptyRule = EmptyRule Token +EmptyRule : Type -> Type +EmptyRule = Grammar () Token False export equals : Rule () diff --git a/src/Parser/Rule/Source.idr b/src/Parser/Rule/Source.idr index f08d71d39..fcb7c5777 100644 --- a/src/Parser/Rule/Source.idr +++ b/src/Parser/Rule/Source.idr @@ -1,10 +1,10 @@ module Parser.Rule.Source import public Parser.Lexer.Source -import public Parser.Rule.Common import public Parser.Support import Core.TT +import Core.Metadata import Data.List1 import Data.Strings import Libraries.Data.List.Extra @@ -19,14 +19,14 @@ import Libraries.Data.String.Extra public export Rule : Type -> Type -Rule = Rule Token +Rule ty = Grammar SemanticDecorations Token True ty public export -SourceEmptyRule : Type -> Type -SourceEmptyRule = EmptyRule Token +EmptyRule : Type -> Type +EmptyRule ty = Grammar SemanticDecorations Token False ty export -eoi : SourceEmptyRule () +eoi : EmptyRule () eoi = ignore $ nextIs "Expected end of input" isEOI where isEOI : Token -> Bool @@ -208,11 +208,11 @@ namespacedIdent Ident i => Just (Nothing, i) _ => Nothing) -isCapitalisedIdent : WithBounds String -> SourceEmptyRule () +isCapitalisedIdent : WithBounds String -> EmptyRule () isCapitalisedIdent str = let val = str.val loc = str.bounds - err : SourceEmptyRule () + err : EmptyRule () = failLoc loc ("Expected a capitalised identifier, got: " ++ val) in case strM val of StrNil => err @@ -249,7 +249,7 @@ reservedNames , "String", "Char", "Double", "Lazy", "Inf", "Force", "Delay" ] -isNotReservedIdent : WithBounds String -> SourceEmptyRule () +isNotReservedIdent : WithBounds String -> EmptyRule () isNotReservedIdent x = if x.val `elem` reservedNames then failLoc x.bounds $ "can't use reserved name " ++ x.val @@ -260,7 +260,7 @@ opNonNS : Rule Name opNonNS = symbol "(" *> (operator <|> postfixProj) <* symbol ")" identWithCapital : (capitalised : Bool) -> WithBounds String -> - SourceEmptyRule () + EmptyRule () identWithCapital b x = if b then isCapitalisedIdent x else pure () nameWithCapital : (capitalised : Bool) -> Rule Name @@ -269,7 +269,7 @@ nameWithCapital b = opNonNS <|> do opNS nsx <|> nameNS nsx where - nameNS : WithBounds (Maybe Namespace, String) -> SourceEmptyRule Name + nameNS : WithBounds (Maybe Namespace, String) -> EmptyRule Name nameNS nsx = do let id = snd <$> nsx identWithCapital b id @@ -317,23 +317,22 @@ export init : IndentInfo init = 0 -continueF : SourceEmptyRule () -> (indent : IndentInfo) -> SourceEmptyRule () +continueF : EmptyRule () -> (indent : IndentInfo) -> EmptyRule () continueF err indent = do eoi; err <|> do keyword "where"; err - <|> do col <- Common.column - if col <= indent - then err - else pure () + <|> do col <- column + when (col <= indent) + err ||| Fail if this is the end of a block entry or end of file export -continue : (indent : IndentInfo) -> SourceEmptyRule () +continue : (indent : IndentInfo) -> EmptyRule () continue = continueF (fail "Unexpected end of expression") ||| As 'continue' but failing is fatal (i.e. entire parse fails) export -mustContinue : (indent : IndentInfo) -> Maybe String -> SourceEmptyRule () +mustContinue : (indent : IndentInfo) -> Maybe String -> EmptyRule () mustContinue indent Nothing = continueF (fatalError "Unexpected end of expression") indent mustContinue indent (Just req) @@ -355,7 +354,7 @@ Show ValidIndent where show (AfterPos i) = "[after " ++ show i ++ "]" show EndOfBlock = "[EOB]" -checkValid : ValidIndent -> Int -> SourceEmptyRule () +checkValid : ValidIndent -> Int -> EmptyRule () checkValid AnyIndent c = pure () checkValid (AtPos x) c = if c == x then pure () @@ -386,29 +385,27 @@ isTerminator _ = False ||| It's the end if we have a terminating token, or the next token starts ||| in or before indent. Works by looking ahead but not consuming. export -atEnd : (indent : IndentInfo) -> SourceEmptyRule () +atEnd : (indent : IndentInfo) -> EmptyRule () atEnd indent = eoi <|> do ignore $ nextIs "Expected end of block" isTerminator - <|> do col <- Common.column - if (col <= indent) - then pure () - else fail "Not the end of a block entry" + <|> do col <- column + when (not (col <= indent)) + $ fail "Not the end of a block entry" -- Check we're at the end, but only by looking at indentation export -atEndIndent : (indent : IndentInfo) -> SourceEmptyRule () +atEndIndent : (indent : IndentInfo) -> EmptyRule () atEndIndent indent = eoi - <|> do col <- Common.column - if col <= indent - then pure () - else fail "Not the end of a block entry" + <|> do col <- column + when (not (col <= indent)) + $ fail "Not the end of a block entry" -- Parse a terminator, return where the next block entry -- must start, given where the current block entry started -terminator : ValidIndent -> Int -> SourceEmptyRule ValidIndent +terminator : ValidIndent -> Int -> EmptyRule ValidIndent terminator valid laststart = do eoi pure EndOfBlock @@ -430,7 +427,7 @@ terminator valid laststart -- Expected indentation for the next token can either be anything (if -- we're inside a brace delimited block) or in exactly the initial column -- (if we're inside an indentation delimited block) - afterDedent : ValidIndent -> Int -> SourceEmptyRule ValidIndent + afterDedent : ValidIndent -> Int -> EmptyRule ValidIndent afterDedent AnyIndent col = if col <= laststart then pure AnyIndent @@ -456,7 +453,7 @@ blockEntry valid rule pure (p, valid') blockEntries : ValidIndent -> (IndentInfo -> Rule ty) -> - SourceEmptyRule (List ty) + EmptyRule (List ty) blockEntries valid rule = do eoi; pure [] <|> do res <- blockEntry valid rule @@ -465,7 +462,7 @@ blockEntries valid rule <|> pure [] export -block : (IndentInfo -> Rule ty) -> SourceEmptyRule (List ty) +block : (IndentInfo -> Rule ty) -> EmptyRule (List ty) block item = do symbol "{" commit @@ -481,35 +478,35 @@ block item ||| by curly braces). `rule` is a function of the actual indentation ||| level. export -blockAfter : Int -> (IndentInfo -> Rule ty) -> SourceEmptyRule (List ty) +blockAfter : Int -> (IndentInfo -> Rule ty) -> EmptyRule (List ty) blockAfter mincol item = do symbol "{" commit ps <- blockEntries AnyIndent item symbol "}" pure ps - <|> do col <- Common.column - if col <= mincol - then pure [] - else blockEntries (AtPos col) item + <|> do col <- column + ifThenElse (col <= mincol) + (pure []) + $ blockEntries (AtPos col) item export blockWithOptHeaderAfter : (column : Int) -> (header : IndentInfo -> Rule hd) -> (item : IndentInfo -> Rule ty) -> - SourceEmptyRule (Maybe hd, List ty) + EmptyRule (Maybe hd, List ty) blockWithOptHeaderAfter {ty} mincol header item = do symbol "{" commit hidt <- optional $ blockEntry AnyIndent header restOfBlock hidt - <|> do col <- Common.column - if col <= mincol - then pure (Nothing, []) - else do hidt <- optional $ blockEntry (AtPos col) header - ps <- blockEntries (AtPos col) item - pure (map fst hidt, ps) + <|> do col <- column + ifThenElse (col <= mincol) + (pure (Nothing, [])) + $ do hidt <- optional $ blockEntry (AtPos col) header + ps <- blockEntries (AtPos col) item + pure (map fst hidt, ps) where restOfBlock : Maybe (hd, ValidIndent) -> Rule (Maybe hd, List ty) restOfBlock (Just (h, idt)) = do ps <- blockEntries idt item diff --git a/src/Parser/Source.idr b/src/Parser/Source.idr index d609ff1b9..34ce9daa0 100644 --- a/src/Parser/Source.idr +++ b/src/Parser/Source.idr @@ -5,6 +5,9 @@ import public Parser.Rule.Source import public Parser.Unlit import Core.Core +import Core.Name +import Core.Metadata +import Core.FC import System.File import Libraries.Utils.Either @@ -14,21 +17,21 @@ export runParserTo : {e : _} -> (fname : String) -> Maybe LiterateStyle -> Lexer -> - String -> Grammar Token e ty -> Either Error ty + String -> Grammar SemanticDecorations Token e ty -> Either Error (SemanticDecorations, ty) runParserTo fname lit reject str p = do str <- mapError (fromLitError fname) $ unlit lit str toks <- mapError (fromLexError fname) $ lexTo reject str - parsed <- mapError (fromParsingError fname) $ parse p toks - Right (fst parsed) + (decs, (parsed, _)) <- mapError (fromParsingError fname) $ parseWith p toks + Right (decs, parsed) export runParser : {e : _} -> (fname : String) -> Maybe LiterateStyle -> String -> - Grammar Token e ty -> Either Error ty + Grammar SemanticDecorations Token e ty -> Either Error (SemanticDecorations, ty) runParser fname lit = runParserTo fname lit (pred $ const False) export covering -parseFile : (fname : String) -> Rule ty -> IO (Either Error ty) +parseFile : (fname : String) -> Rule ty -> IO (Either Error (SemanticDecorations, ty)) parseFile fname p = do Right str <- readFile fname | Left err => pure (Left (FileErr fname err)) diff --git a/src/TTImp/Elab/App.idr b/src/TTImp/Elab/App.idr index e3ec2b5a8..6b011b17d 100644 --- a/src/TTImp/Elab/App.idr +++ b/src/TTImp/Elab/App.idr @@ -55,6 +55,14 @@ getNameType rigc env fc x do est <- get EST put EST (record { linearUsed $= ((MkVar lv) :: ) } est) + log "ide-mode.highlight" 8 + $ "getNameType is trying to add Bound: " + ++ show x ++ " (" ++ show fc ++ ")" + when (isSourceName x) $ + whenJust (isConcreteFC fc) \nfc => do + log "ide-mode.highlight" 7 $ "getNameType is adding Bound: " ++ show x + addSemanticDecorations [(nfc, Bound, Just x)] + pure (Local fc (Just (isLet binder)) _ lv, gnf env bty) Nothing => do defs <- get Ctxt @@ -68,6 +76,18 @@ getNameType rigc env fc x DCon t a _ => DataCon t a TCon t a _ _ _ _ _ _ => TyCon t a _ => Func + + log "ide-mode.highlight" 8 + $ "getNameType is trying to add something for: " + ++ show def.fullname ++ " (" ++ show fc ++ ")" + + when (isSourceName def.fullname) $ + whenJust (isConcreteFC fc) \nfc => do + let decor = nameTypeDecoration nt + log "ide-mode.highlight" 7 + $ "getNameType is adding " ++ show decor ++ ": " ++ show def.fullname + addSemanticDecorations [(nfc, decor, Just def.fullname)] + pure (Ref fc nt (Resolved i), gnf env (embed (type def))) where rigSafe : RigCount -> RigCount -> Core () @@ -89,7 +109,7 @@ getVarType rigc nest env fc x Just (nestn, argns, tmf) => do defs <- get Ctxt let arglen = length argns - let n' = maybe x id nestn + let n' = fromMaybe x nestn case !(lookupCtxtExact n' (gamma defs)) of Nothing => undefinedName fc n' Just ndef => @@ -110,6 +130,14 @@ getVarType rigc nest env fc x log "metadata.names" 7 $ "getVarType is adding ↓" addNameType fc x env tyenv + when (isSourceName ndef.fullname) $ + whenJust (isConcreteFC fc) \nfc => do + let decor = nameTypeDecoration nt + log "ide-mode.highlight" 7 + $ "getNameType is adding "++ show decor ++": " + ++ show ndef.fullname + addSemanticDecorations [(nfc, decor, Just ndef.fullname)] + pure (tm, arglen, gnf env tyenv) where useVars : {vars : _} -> diff --git a/src/TTImp/Elab/Case.idr b/src/TTImp/Elab/Case.idr index 526ad333e..6669163a9 100644 --- a/src/TTImp/Elab/Case.idr +++ b/src/TTImp/Elab/Case.idr @@ -231,7 +231,7 @@ caseBlock {vars} rigc elabinfo fc nest env scr scrtm scrty caseRig alts expected -- This will be the case either if the scrutinee is a variable, in -- which case the duplication won't hurt, or if (TODO) none of the -- case patterns in alts are just a variable - maybe (pure ()) (const (setFlag fc casen Inline)) splitOn + whenJust splitOn $ \ _ => setFlag fc casen Inline let applyEnv = applyToFull fc caseRef env let appTm : Term vars diff --git a/src/TTImp/Elab/Check.idr b/src/TTImp/Elab/Check.idr index e466d0707..75b452b01 100644 --- a/src/TTImp/Elab/Check.idr +++ b/src/TTImp/Elab/Check.idr @@ -515,7 +515,7 @@ successful : {vars : _} -> Bool -> -- constraints allowed List (Maybe Name, Core a) -> Core (List (Either (Maybe Name, Error) - (Nat, a, Defs, UState, EState vars))) + (Nat, a, Defs, UState, EState vars, Metadata))) successful allowCons [] = pure [] successful allowCons ((tm, elab) :: elabs) = do ust <- get UST @@ -555,7 +555,7 @@ successful allowCons ((tm, elab) :: elabs) elabs' <- successful allowCons elabs -- Record success, and the state we ended at pure (Right (minus ncons' ncons, - res, defs', ust', est') :: elabs')) + res, defs', ust', est', md') :: elabs')) (\err => do put UST ust put EST est put MD md @@ -576,9 +576,10 @@ exactlyOne' allowCons fc env [(tm, elab)] = elab exactlyOne' {vars} allowCons fc env all = do elabs <- successful allowCons all case getRight elabs of - Right (res, defs, ust, est) => + Right (res, defs, ust, est, md) => do put UST ust put EST est + put MD md put Ctxt defs commit pure res diff --git a/src/TTImp/Elab/ImplicitBind.idr b/src/TTImp/Elab/ImplicitBind.idr index c1e6d8e63..4fec8cf34 100644 --- a/src/TTImp/Elab/ImplicitBind.idr +++ b/src/TTImp/Elab/ImplicitBind.idr @@ -416,6 +416,12 @@ checkBindVar rig elabinfo nest env fc str topexp noteLHSPatVar elabmode (UN str) notePatVar n est <- get EST + + whenJust (isConcreteFC fc) \nfc => do + log "ide-mode.highlight" 7 $ "getNameType is adding Bound: " ++ show n + addSemanticDecorations [(nfc, Bound, Just n)] + + case lookup n (boundNames est) of Nothing => do (tm, exp, bty) <- mkPatternHole fc rig n env diff --git a/src/TTImp/Elab/Record.idr b/src/TTImp/Elab/Record.idr index 9869f3fc1..a499b6f63 100644 --- a/src/TTImp/Elab/Record.idr +++ b/src/TTImp/Elab/Record.idr @@ -177,18 +177,21 @@ recUpdate : {vars : _} -> List IFieldUpdate -> (rec : RawImp) -> (grecty : Glued vars) -> Core RawImp -recUpdate rigc elabinfo loc nest env flds rec grecty +recUpdate rigc elabinfo iloc nest env flds rec grecty = do defs <- get Ctxt rectynf <- getNF grecty let Just rectyn = getRecordType env rectynf - | Nothing => throw (RecordTypeNeeded loc env) + | Nothing => throw (RecordTypeNeeded iloc env) fldn <- genFieldName "__fld" - sides <- getAllSides loc flds rectyn rec - (Field Nothing fldn (IVar loc (UN fldn))) - pure $ ICase loc rec (Implicit loc False) [mkClause sides] + sides <- getAllSides iloc flds rectyn rec + (Field Nothing fldn (IVar vloc (UN fldn))) + pure $ ICase vloc rec (Implicit vloc False) [mkClause sides] where + vloc : FC + vloc = virtualiseFC iloc + mkClause : Rec -> ImpClause - mkClause rec = PatClause loc (toLHS loc rec) (toRHS loc rec) + mkClause rec = PatClause vloc (toLHS vloc rec) (toRHS vloc rec) needType : Error -> Bool needType (RecordTypeNeeded _ _) = True diff --git a/src/TTImp/Elab/Rewrite.idr b/src/TTImp/Elab/Rewrite.idr index 94e9db9a0..34ca2fb93 100644 --- a/src/TTImp/Elab/Rewrite.idr +++ b/src/TTImp/Elab/Rewrite.idr @@ -107,19 +107,20 @@ checkRewrite : {vars : _} -> Core (Term vars, Glued vars) checkRewrite rigc elabinfo nest env fc rule tm Nothing = throw (GenericMsg fc "Can't infer a type for rewrite") -checkRewrite {vars} rigc elabinfo nest env fc rule tm (Just expected) - = delayOnFailure fc rigc env expected rewriteErr 10 (\delayed => - do (rulev, grulet) <- check erased elabinfo nest env rule Nothing +checkRewrite {vars} rigc elabinfo nest env ifc rule tm (Just expected) + = delayOnFailure ifc rigc env expected rewriteErr 10 $ \delayed => + do let vfc = virtualiseFC ifc + (rulev, grulet) <- check erased elabinfo nest env rule Nothing rulet <- getTerm grulet expTy <- getTerm expected when delayed $ log "elab.rewrite" 5 "Retrying rewrite" - (lemma, pred, predty) <- elabRewrite fc env expTy rulet + (lemma, pred, predty) <- elabRewrite vfc env expTy rulet rname <- genVarName "_" pname <- genVarName "_" - let pbind = Let fc erased pred predty - let rbind = Let fc erased (weaken rulev) (weaken rulet) + let pbind = Let vfc erased pred predty + let rbind = Let vfc erased (weaken rulev) (weaken rulet) let env' = rbind :: pbind :: env @@ -128,16 +129,15 @@ checkRewrite {vars} rigc elabinfo nest env fc rule tm (Just expected) -- implicits for the rewriting lemma are in the right place. But, -- we still need the right type for the EState, so weaken it once -- for each of the let bindings above. - (rwtm, grwty) <- inScope fc (pbind :: env) - (\e' => inScope {e=e'} fc env' - (\e'' => check {e = e''} {vars = rname :: pname :: vars} - rigc elabinfo (weaken (weaken nest)) env' - (apply (IVar fc lemma) [IVar fc pname, - IVar fc rname, - tm]) - (Just (gnf env' - (weakenNs (mkSizeOf [rname, pname]) expTy))) - )) + (rwtm, grwty) <- + inScope vfc (pbind :: env) $ \e' => + inScope {e=e'} vfc env' $ \e'' => + let offset = mkSizeOf [rname, pname] in + check {e = e''} rigc elabinfo (weakenNs offset nest) env' + (apply (IVar vfc lemma) [IVar vfc pname, + IVar vfc rname, + tm]) + (Just (gnf env' (weakenNs offset expTy))) rwty <- getTerm grwty - pure (Bind fc pname pbind (Bind fc rname rbind rwtm), - gnf env (Bind fc pname pbind (Bind fc rname rbind rwty)))) + let binding = Bind vfc pname pbind . Bind vfc rname rbind + pure (binding rwtm, gnf env (binding rwty)) diff --git a/src/TTImp/Interactive/CaseSplit.idr b/src/TTImp/Interactive/CaseSplit.idr index badd1acac..6d82507c7 100644 --- a/src/TTImp/Interactive/CaseSplit.idr +++ b/src/TTImp/Interactive/CaseSplit.idr @@ -330,7 +330,7 @@ mkCase : {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> Int -> RawImp -> RawImp -> Core ClauseUpdate mkCase {c} {u} fn orig lhs_raw - = do m <- newRef MD initMetadata + = do m <- newRef MD (initMetadata "(interactive)") defs <- get Ctxt ust <- get UST catch diff --git a/src/TTImp/Parser.idr b/src/TTImp/Parser.idr index d6c845919..41affcb17 100644 --- a/src/TTImp/Parser.idr +++ b/src/TTImp/Parser.idr @@ -2,6 +2,7 @@ module TTImp.Parser import Core.Context import Core.Core +import Core.Metadata import Core.Env import Core.TT import Parser.Source @@ -81,7 +82,7 @@ visOption <|> do keyword "private" pure Private -visibility : SourceEmptyRule Visibility +visibility : EmptyRule Visibility visibility = visOption <|> pure Private @@ -124,7 +125,7 @@ visOpt pure (Right opt) getVisibility : Maybe Visibility -> List (Either Visibility FnOpt) -> - SourceEmptyRule Visibility + EmptyRule Visibility getVisibility Nothing [] = pure Private getVisibility (Just vis) [] = pure vis getVisibility Nothing (Left x :: xs) = getVisibility (Just x) xs @@ -216,13 +217,13 @@ mutual symbol ")" pure e - multiplicity : SourceEmptyRule (Maybe Integer) + multiplicity : EmptyRule (Maybe Integer) multiplicity = do c <- intLit pure (Just c) <|> pure Nothing - getMult : Maybe Integer -> SourceEmptyRule RigCount + getMult : Maybe Integer -> EmptyRule RigCount getMult (Just 0) = pure erased getMult (Just 1) = pure linear getMult Nothing = pure top @@ -522,7 +523,7 @@ mutual let fc = MkFC fname start end pure (!(getFn lhs), ImpossibleClause fc lhs) where - getFn : RawImp -> SourceEmptyRule Name + getFn : RawImp -> EmptyRule Name getFn (IVar _ n) = pure n getFn (IApp _ f a) = getFn f getFn (IAutoApp _ f a) = getFn f @@ -593,7 +594,7 @@ recordParam fname indents <|> do symbol "{" commit start <- location - info <- the (SourceEmptyRule (PiInfo RawImp)) + info <- the (EmptyRule (PiInfo RawImp)) (pure AutoImplicit <* keyword "auto" <|>(do keyword "default" diff --git a/src/TTImp/ProcessDecls.idr b/src/TTImp/ProcessDecls.idr index de61b771b..98301fa1a 100644 --- a/src/TTImp/ProcessDecls.idr +++ b/src/TTImp/ProcessDecls.idr @@ -167,7 +167,7 @@ processTTImpFile : {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> String -> Core Bool processTTImpFile fname - = do Right tti <- logTime "Parsing" $ coreLift $ parseFile fname + = do Right (decor, tti) <- logTime "Parsing" $ coreLift $ parseFile fname (do decls <- prog fname eoi pure decls) diff --git a/src/TTImp/ProcessDef.idr b/src/TTImp/ProcessDef.idr index 7cc6b511c..ec00d0081 100644 --- a/src/TTImp/ProcessDef.idr +++ b/src/TTImp/ProcessDef.idr @@ -464,13 +464,13 @@ checkClause {vars} mult vis totreq hashit n opts nest env (PatClause fc lhs_in r pure (Right (MkClause env' lhstm' rhstm)) -- TODO: (to decide) With is complicated. Move this into its own module? checkClause {vars} mult vis totreq hashit n opts nest env - (WithClause fc lhs_in wval_raw mprf flags cs) + (WithClause ifc lhs_in wval_raw mprf flags cs) = do (lhs, (vars' ** (sub', env', nest', lhspat, reqty))) <- - checkLHS False mult hashit n opts nest env fc lhs_in + checkLHS False mult hashit n opts nest env ifc lhs_in let wmode = if isErased mult then InType else InExpr - (wval, gwvalTy) <- wrapErrorC opts (InRHS fc !(getFullName (Resolved n))) $ + (wval, gwvalTy) <- wrapErrorC opts (InRHS ifc !(getFullName (Resolved n))) $ elabTermSub n wmode opts nest' env' env sub' wval_raw Nothing clearHoleLHS @@ -498,7 +498,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env -- to get the 'magic with' behaviour (wargs ** (scenv, var, binder)) <- bindWithArgs wvalTy ((,wval) <$> mprf) wvalEnv - let bnr = bindNotReq fc 0 env' withSub [] reqty + let bnr = bindNotReq vfc 0 env' withSub [] reqty let notreqns = fst bnr let notreqty = snd bnr @@ -511,7 +511,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env (weakenNs (mkSizeOf wargs) notreqty)) let bNotReq = binder wtyScope - let Just (reqns, envns, wtype) = bindReq fc env' withSub [] bNotReq + let Just (reqns, envns, wtype) = bindReq vfc env' withSub [] bNotReq | Nothing => throw (InternalError "Impossible happened: With abstraction failure #4") -- list of argument names - 'Just' means we need to match the name @@ -526,11 +526,11 @@ checkClause {vars} mult vis totreq hashit n opts nest env wname <- genWithName !(prettyName !(toFullNames (Resolved n))) widx <- addDef wname (record {flags $= (SetTotal totreq ::)} - (newDef fc wname (if isErased mult then erased else top) + (newDef vfc wname (if isErased mult then erased else top) vars wtype vis None)) let toWarg : Maybe (PiInfo RawImp, Name) -> List (Maybe Name, RawImp) - := flip maybe (\pn => [(Nothing, IVar fc (snd pn))]) $ + := flip maybe (\pn => [(Nothing, IVar vfc (snd pn))]) $ (Nothing, wval_raw) :: case mprf of Nothing => [] @@ -539,12 +539,12 @@ checkClause {vars} mult vis totreq hashit n opts nest env let refl = IVar fc (NS builtinNS (UN "Refl")) in [(mprf, INamedApp fc refl (UN "x") wval_raw)] - let rhs_in = gapply (IVar fc wname) - $ map (\ nm => (Nothing, IVar fc nm)) envns + let rhs_in = gapply (IVar vfc wname) + $ map (\ nm => (Nothing, IVar vfc nm)) envns ++ concatMap toWarg wargNames log "declare.def.clause" 3 $ "Applying to with argument " ++ show rhs_in - rhs <- wrapErrorC opts (InRHS fc !(getFullName (Resolved n))) $ + rhs <- wrapErrorC opts (InRHS ifc !(getFullName (Resolved n))) $ checkTermSub n wmode opts nest' env' env sub' rhs_in (gnf env' reqty) @@ -556,11 +556,14 @@ checkClause {vars} mult vis totreq hashit n opts nest env nestname <- applyEnv env wname let nest'' = record { names $= (nestname ::) } nest - let wdef = IDef fc wname cs' + let wdef = IDef ifc wname cs' processDecl [] nest'' env wdef pure (Right (MkClause env' lhspat rhs)) where + vfc : FC + vfc = virtualiseFC ifc + bindWithArgs : (wvalTy : Term xs) -> Maybe (Name, Term xs) -> (wvalEnv : Env Term xs) -> @@ -576,13 +579,13 @@ checkClause {vars} mult vis totreq hashit n opts nest env wargs = [wargn] scenv : Env Term (wargs ++ xs) - := Pi fc top Explicit wvalTy :: wvalEnv + := Pi vfc top Explicit wvalTy :: wvalEnv var : Term (wargs ++ xs) - := Local fc (Just False) Z First + := Local vfc (Just False) Z First binder : Term (wargs ++ xs) -> Term xs - := Bind fc wargn (Pi fc top Explicit wvalTy) + := Bind vfc wargn (Pi vfc top Explicit wvalTy) in pure (wargs ** (scenv, var, binder)) @@ -592,7 +595,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env let eqName = NS builtinNS (UN "Equal") Just (TCon t ar _ _ _ _ _ _) <- lookupDefExact eqName (gamma defs) | _ => throw (InternalError "Cannot find builtin Equal") - let eqTyCon = Ref fc (TyCon t ar) eqName + let eqTyCon = Ref vfc (TyCon t ar) eqName let wargn : Name wargn = MN "warg" 0 @@ -601,24 +604,24 @@ checkClause {vars} mult vis totreq hashit n opts nest env wvalTy' := weaken wvalTy eqTy : Term (MN "warg" 0 :: xs) - := apply fc eqTyCon + := apply vfc eqTyCon [ wvalTy' , wvalTy' , weaken wval - , Local fc (Just False) Z First + , Local vfc (Just False) Z First ] scenv : Env Term (wargs ++ xs) - := Pi fc top Implicit eqTy - :: Pi fc top Explicit wvalTy + := Pi vfc top Implicit eqTy + :: Pi vfc top Explicit wvalTy :: wvalEnv var : Term (wargs ++ xs) - := Local fc (Just False) (S Z) (Later First) + := Local vfc (Just False) (S Z) (Later First) binder : Term (wargs ++ xs) -> Term xs - := \ t => Bind fc wargn (Pi fc top Explicit wvalTy) - $ Bind fc name (Pi fc top Implicit eqTy) t + := \ t => Bind vfc wargn (Pi vfc top Explicit wvalTy) + $ Bind vfc name (Pi vfc top Implicit eqTy) t pure (wargs ** (scenv, var, binder)) diff --git a/src/TTImp/ProcessRecord.idr b/src/TTImp/ProcessRecord.idr index 24d1ec88b..c348c43b7 100644 --- a/src/TTImp/ProcessRecord.idr +++ b/src/TTImp/ProcessRecord.idr @@ -77,7 +77,7 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields farg : IField -> (FC, Maybe Name, RigCount, PiInfo RawImp, RawImp) - farg (MkIField fc c p n ty) = (fc, Just n, c, p, ty) + farg (MkIField fc c p n ty) = (virtualiseFC fc, Just n, c, p, ty) mkTy : List (FC, Maybe Name, RigCount, PiInfo RawImp, RawImp) -> RawImp -> RawImp @@ -86,7 +86,7 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields = IPi fc c imp n argty (mkTy args ret) recTy : RawImp - recTy = apply (IVar fc tn) (map (\(n, c, p, tm) => (n, IVar EmptyFC n, p)) params) + recTy = apply (IVar (virtualiseFC fc) tn) (map (\(n, c, p, tm) => (n, IVar EmptyFC n, p)) params) where ||| Apply argument to list of explicit or implicit named arguments apply : RawImp -> List (Name, RawImp, PiInfo RawImp) -> RawImp @@ -96,7 +96,8 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields elabAsData : Name -> Core () elabAsData cname - = do let conty = mkTy paramTelescope $ + = do let fc = virtualiseFC fc + let conty = mkTy paramTelescope $ mkTy (map farg fields) recTy let con = MkImpTy EmptyFC EmptyFC cname !(bindTypeNames [] (map fst params ++ map fname fields ++ vars) conty) @@ -125,7 +126,9 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields Env Term vs -> Term vs -> Core () elabGetters con done upds tyenv (Bind bfc n b@(Pi _ rc imp ty_chk) sc) - = if (n `elem` map fst params) || (n `elem` vars) + = let rig = if isErased rc then erased else top + isVis = projVis vis + in if (n `elem` map fst params) || (n `elem` vars) then elabGetters con (if imp == Explicit && not (n `elem` vars) then S done else done) @@ -145,11 +148,14 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields (map fst params ++ map fname fields ++ vars) $ mkTy paramTelescope $ IPi bfc top Explicit (Just rname) recTy ty' - log "declare.record.projection" 5 $ "Projection " ++ show rfNameNS ++ " : " ++ show projTy - processDecl [] nest env - (IClaim bfc (if isErased rc - then erased - else top) (projVis vis) [Inline] (MkImpTy EmptyFC EmptyFC rfNameNS projTy)) + + let mkProjClaim = \ nm => + let ty = MkImpTy EmptyFC EmptyFC nm projTy + in IClaim bfc rig isVis [Inline] ty + + log "declare.record.projection" 5 $ + "Projection " ++ show rfNameNS ++ " : " ++ show projTy + processDecl [] nest env (mkProjClaim rfNameNS) -- Define the LHS and RHS let lhs_exp @@ -173,16 +179,15 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields when !isPrefixRecordProjections $ do -- beware: `!` is NOT boolean `not`! -- Claim the type. -- we just reuse `projTy` defined above - log "declare.record.projection.prefix" 5 $ "Prefix projection " ++ show unNameNS ++ " : " ++ show projTy - processDecl [] nest env - (IClaim bfc (if isErased rc - then erased - else top) (projVis vis) [Inline] (MkImpTy EmptyFC EmptyFC unNameNS projTy)) + log "declare.record.projection.prefix" 5 $ + "Prefix projection " ++ show unNameNS ++ " : " ++ show projTy + processDecl [] nest env (mkProjClaim unNameNS) -- Define the LHS and RHS let lhs = IVar bfc unNameNS let rhs = IVar bfc rfNameNS - log "declare.record.projection.prefix" 5 $ "Prefix projection " ++ show lhs ++ " = " ++ show rhs + log "declare.record.projection.prefix" 5 $ + "Prefix projection " ++ show lhs ++ " = " ++ show rhs processDecl [] nest env (IDef bfc unNameNS [PatClause bfc lhs rhs]) diff --git a/src/TTImp/ProcessType.idr b/src/TTImp/ProcessType.idr index 8c806bf39..38bf7d1be 100644 --- a/src/TTImp/ProcessType.idr +++ b/src/TTImp/ProcessType.idr @@ -264,6 +264,8 @@ processType : {vars : _} -> processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc nameFC n_in ty_raw) = do n <- inCurrentNS n_in + addNameLoc nameFC n + log "declare.type" 1 $ "Processing " ++ show n log "declare.type" 5 $ "Checking type decl " ++ show n ++ " : " ++ show ty_raw idx <- resolveName n diff --git a/src/TTImp/TTImp.idr b/src/TTImp/TTImp.idr index 129c64333..74c640244 100644 --- a/src/TTImp/TTImp.idr +++ b/src/TTImp/TTImp.idr @@ -542,7 +542,7 @@ implicitsAs n defs ns tm "\n In the type of " ++ show n ++ ": " ++ show ty ++ "\n Using locals: " ++ show ns ++ "\n Found implicits: " ++ show implicits - pure $ impAs loc implicits (IVar loc nm) + pure $ impAs (virtualiseFC loc) implicits (IVar loc nm) where -- If there's an @{c} in the list of given implicits, that's the next -- autoimplicit, so don't rewrite the LHS and update the list of given @@ -715,10 +715,28 @@ getFC (IAs x _ _ _ _) = x getFC (Implicit x _) = x getFC (IWithUnambigNames x _ _) = x +namespace ImpDecl + + public export + getFC : ImpDecl -> FC + getFC (IClaim fc _ _ _ _) = fc + getFC (IData fc _ _) = fc + getFC (IDef fc _ _) = fc + getFC (IParameters fc _ _) = fc + getFC (IRecord fc _ _ _ ) = fc + getFC (INamespace fc _ _) = fc + getFC (ITransform fc _ _ _) = fc + getFC (IRunElabDecl fc _) = fc + getFC (IPragma _ _) = EmptyFC + getFC (ILog _) = EmptyFC + getFC (IBuiltin fc _ _) = fc + export apply : RawImp -> List RawImp -> RawImp apply f [] = f -apply f (x :: xs) = apply (IApp (getFC f) f x) xs +apply f (x :: xs) = + let fFC = getFC f in + apply (IApp (fromMaybe fFC (mergeFC fFC (getFC x))) f x) xs export gapply : RawImp -> List (Maybe Name, RawImp) -> RawImp diff --git a/src/TTImp/WithClause.idr b/src/TTImp/WithClause.idr index 667d58305..3fee9e834 100644 --- a/src/TTImp/WithClause.idr +++ b/src/TTImp/WithClause.idr @@ -2,7 +2,9 @@ module TTImp.WithClause import Core.Context import Core.Context.Log +import Core.Metadata import Core.TT + import TTImp.BindImplicits import TTImp.TTImp import TTImp.Elab.Check @@ -15,30 +17,52 @@ import Data.Maybe matchFail : FC -> Core a matchFail loc = throw (GenericMsg loc "With clause does not match parent") +--- To be used on the lhs of a nested with clause to figure out a tight location +--- information to give to the generated LHS +getHeadLoc : RawImp -> Core FC +getHeadLoc (IVar fc _) = pure fc +getHeadLoc (IApp _ f _) = getHeadLoc f +getHeadLoc (IAutoApp _ f _) = getHeadLoc f +getHeadLoc (INamedApp _ f _ _) = getHeadLoc f +getHeadLoc t = throw (InternalError $ "Could not find head of LHS: " ++ show t) + +addAlias : {auto m : Ref MD Metadata} -> + {auto c : Ref Ctxt Defs} -> + FC -> FC -> Core () +addAlias from to = + whenJust (isConcreteFC from) $ \ from => + whenJust (isConcreteFC to) $ \ to => do + log "ide-mode.highlighting.alias" 25 $ + "Adding alias: " ++ show from ++ " -> " ++ show to + addSemanticAlias from to + mutual export - getMatch : (lhs : Bool) -> RawImp -> RawImp -> + getMatch : {auto m : Ref MD Metadata} -> + {auto c : Ref Ctxt Defs} -> + (lhs : Bool) -> RawImp -> RawImp -> Core (List (String, RawImp)) + getMatch lhs (IBindVar to n) tm@(IBindVar from _) + = [(n, tm)] <$ addAlias from to getMatch lhs (IBindVar _ n) tm = pure [(n, tm)] getMatch lhs (Implicit _ _) tm = pure [] - getMatch lhs (IVar _ (NS ns n)) (IVar loc (NS ns' n')) - = if n == n' && isParentOf ns' ns then pure [] else matchFail loc - getMatch lhs (IVar _ (NS ns n)) (IVar loc n') - = if n == n' then pure [] else matchFail loc - getMatch lhs (IVar _ n) (IVar loc n') - = if n == n' then pure [] else matchFail loc + getMatch lhs (IVar to (NS ns n)) (IVar from (NS ns' n')) + = if n == n' && isParentOf ns' ns + then [] <$ addAlias from to -- <$ decorateName loc nm + else matchFail from + getMatch lhs (IVar to (NS ns n)) (IVar from n') + = if n == n' + then [] <$ addAlias from to -- <$ decorateName loc (NS ns n') + else matchFail from + getMatch lhs (IVar to n) (IVar from n') + = if n == n' + then [] <$ addAlias from to -- <$ decorateName loc n' + else matchFail from getMatch lhs (IPi _ c p n arg ret) (IPi loc c' p' n' arg' ret') - = if c == c' && samePiInfo p p' && n == n' + = if c == c' && eqPiInfoBy (\_, _ => True) p p' && n == n' then matchAll lhs [(arg, arg'), (ret, ret')] else matchFail loc - where - samePiInfo : PiInfo RawImp -> PiInfo RawImp -> Bool - samePiInfo Explicit Explicit = True - samePiInfo Implicit Implicit = True - samePiInfo AutoImplicit AutoImplicit = True - samePiInfo (DefImplicit _) (DefImplicit _) = True - samePiInfo _ _ = False -- TODO: Lam, Let, Case, Local, Update getMatch lhs (IApp _ f a) (IApp loc f' a') = matchAll lhs [(f, f'), (a, a')] @@ -72,7 +96,7 @@ mutual -- one of them is okay getMatch lhs (IAlternative fc _ as) (IAlternative _ _ as') = matchAny fc lhs (zip as as') - getMatch lhs (IAs _ _ _ (UN n) p) (IAs fc _ _ (UN n') p') + getMatch lhs (IAs _ _ _ (UN n) p) (IAs _ fc _ (UN n') p') = do ms <- getMatch lhs p p' mergeMatches lhs ((n, IBindVar fc n') :: ms) getMatch lhs (IAs _ _ _ (UN n) p) p' @@ -87,14 +111,18 @@ mutual else matchFail fc getMatch lhs pat spec = matchFail (getFC pat) - matchAny : FC -> (lhs : Bool) -> List (RawImp, RawImp) -> + matchAny : {auto m : Ref MD Metadata} -> + {auto c : Ref Ctxt Defs} -> + FC -> (lhs : Bool) -> List (RawImp, RawImp) -> Core (List (String, RawImp)) matchAny fc lhs [] = matchFail fc matchAny fc lhs ((x, y) :: ms) = catch (getMatch lhs x y) (\err => matchAny fc lhs ms) - matchAll : (lhs : Bool) -> List (RawImp, RawImp) -> + matchAll : {auto m : Ref MD Metadata} -> + {auto c : Ref Ctxt Defs} -> + (lhs : Bool) -> List (RawImp, RawImp) -> Core (List (String, RawImp)) matchAll lhs [] = pure [] matchAll lhs ((x, y) :: ms) @@ -102,7 +130,9 @@ mutual mxy <- getMatch lhs x y mergeMatches lhs (mxy ++ matches) - mergeMatches : (lhs : Bool) -> List (String, RawImp) -> + mergeMatches : {auto m : Ref MD Metadata} -> + {auto c : Ref Ctxt Defs} -> + (lhs : Bool) -> List (String, RawImp) -> Core (List (String, RawImp)) mergeMatches lhs [] = pure [] mergeMatches lhs ((n, tm) :: rest) @@ -110,8 +140,9 @@ mutual case lookup n rest' of Nothing => pure ((n, tm) :: rest') Just tm' => - do ignore $ getMatch lhs tm tm' -- just need to know it succeeds - mergeMatches lhs rest + do ignore $ getMatch lhs tm tm' + -- ^ just need to know it succeeds + pure rest' -- Get the arguments for the rewritten pattern clause of a with by looking -- up how the argument names matched @@ -138,11 +169,13 @@ getArgMatch ploc mode search warg ms (Just (_, nm)) export getNewLHS : {auto c : Ref Ctxt Defs} -> + {auto m : Ref MD Metadata} -> FC -> (drop : Nat) -> NestedNames vars -> Name -> List (Maybe (PiInfo RawImp, Name)) -> RawImp -> RawImp -> Core RawImp -getNewLHS ploc drop nest wname wargnames lhs_raw patlhs - = do (mlhs_raw, wrest) <- dropWithArgs drop patlhs +getNewLHS iploc drop nest wname wargnames lhs_raw patlhs + = do let vploc = virtualiseFC iploc + (mlhs_raw, wrest) <- dropWithArgs drop patlhs autoimp <- isUnboundImplicits setUnboundImplicits True @@ -154,15 +187,16 @@ getNewLHS ploc drop nest wname wargnames lhs_raw patlhs log "declare.def.clause.with" 20 $ "Modified LHS (with implicits): " ++ show mlhs let (warg :: rest) = reverse wrest - | _ => throw (GenericMsg ploc "Badly formed 'with' clause") + | _ => throw (GenericMsg iploc "Badly formed 'with' clause") log "declare.def.clause.with" 5 $ show lhs ++ " against " ++ show mlhs ++ " dropping " ++ show (warg :: rest) ms <- getMatch True lhs mlhs log "declare.def.clause.with" 5 $ "Matches: " ++ show ms - let params = map (getArgMatch ploc (InLHS top) False warg ms) wargnames + let params = map (getArgMatch vploc (InLHS top) False warg ms) wargnames log "declare.def.clause.with" 5 $ "Parameters: " ++ show params - let newlhs = apply (IVar ploc wname) (params ++ rest) + hdloc <- getHeadLoc patlhs + let newlhs = apply (IVar hdloc wname) (params ++ rest) log "declare.def.clause.with" 5 $ "New LHS: " ++ show newlhs pure newlhs where @@ -174,11 +208,12 @@ getNewLHS ploc drop nest wname wargnames lhs_raw patlhs pure (tm, arg :: rest) -- Shouldn't happen if parsed correctly, but there's no guarantee that -- inputs come from parsed source so throw an error. - dropWithArgs _ _ = throw (GenericMsg ploc "Badly formed 'with' clause") + dropWithArgs _ _ = throw (GenericMsg iploc "Badly formed 'with' clause") -- Find a 'with' application on the RHS and update it export withRHS : {auto c : Ref Ctxt Defs} -> + {auto m : Ref MD Metadata} -> FC -> (drop : Nat) -> Name -> List (Maybe (PiInfo RawImp, Name)) -> RawImp -> RawImp -> Core RawImp @@ -196,8 +231,9 @@ withRHS fc drop wname wargnames tm toplhs updateWith fc tm (arg :: args) = do log "declare.def.clause.with" 10 $ "With-app: Matching " ++ show toplhs ++ " against " ++ show tm ms <- getMatch False toplhs tm + hdloc <- getHeadLoc tm log "declare.def.clause.with" 10 $ "Result: " ++ show ms - let newrhs = apply (IVar fc wname) + let newrhs = apply (IVar hdloc wname) (map (getArgMatch fc InExpr True arg ms) wargnames) log "declare.def.clause.with" 10 $ "With args for RHS: " ++ show wargnames log "declare.def.clause.with" 10 $ "New RHS: " ++ show newrhs diff --git a/src/Yaffle/Main.idr b/src/Yaffle/Main.idr index 6e2d611b4..d82e30e36 100644 --- a/src/Yaffle/Main.idr +++ b/src/Yaffle/Main.idr @@ -47,7 +47,7 @@ yaffleMain : String -> List String -> Core () yaffleMain fname args = do defs <- initDefs c <- newRef Ctxt defs - m <- newRef MD initMetadata + m <- newRef MD (initMetadata fname) u <- newRef UST initUState d <- getDirs t <- processArgs args diff --git a/src/Yaffle/REPL.idr b/src/Yaffle/REPL.idr index 341f9c319..56b90f6d3 100644 --- a/src/Yaffle/REPL.idr +++ b/src/Yaffle/REPL.idr @@ -153,4 +153,4 @@ repl case runParser "(interactive)" Nothing inp command of Left err => do coreLift_ (printLn err) repl - Right cmd => when !(processCatch cmd) repl + Right (decor, cmd) => when !(processCatch cmd) repl diff --git a/tests/Main.idr b/tests/Main.idr index 9dc52bd7a..ba909647a 100644 --- a/tests/Main.idr +++ b/tests/Main.idr @@ -246,7 +246,7 @@ nodeTests = MkTestPool [Node] ideModeTests : TestPool ideModeTests = MkTestPool [] - [ "ideMode001", "ideMode002", "ideMode003", "ideMode004" + [ "ideMode001", "ideMode002", "ideMode003", "ideMode004", "ideMode005" ] preludeTests : TestPool diff --git a/tests/ideMode/ideMode001/expected b/tests/ideMode/ideMode001/expected index 94d65eba4..a68b1a4b5 100644 --- a/tests/ideMode/ideMode001/expected +++ b/tests/ideMode/ideMode001/expected @@ -1,27 +1,68 @@ 000018(:protocol-version 2 0) 000038(:write-string "1/1: Building LocType (LocType.idr)" 1) -0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 24)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "a")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 7 40)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1) -0000df(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 37)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect {k:376} a)")))))) 1) -0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 10)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "?{_:377}_[]")))))) 1) -0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 16)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{k:376}_[] ?{_:377}_[])")))))) 1) -0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 20)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:378}_[] ?{_:377}_[])")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 6 18)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1) -0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 13)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:367}_[] ?{_:366}_[])")))))) 1) -0001ec(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 5 7)) ((:name "Main.append") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "{0 m : Prelude.Types.Nat} -> {0 a : Type} -> {0 n : Prelude.Types.Nat} -> ({arg:357} : (Main.Vect n[0] a[1])) -> ({arg:358} : (Main.Vect m[3] a[2])) -> (Main.Vect (Prelude.Num.+ Prelude.Types.Nat Prelude.Types.Num implementation at Prelude/Types.idr:64:1--69:33 n[2] m[4]) a[3]))))))")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 5 48)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 45)) ((:name "m") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 41)) ((:name "n") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 29) (:end 5 30)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 27) (:end 5 28)) ((:name "m") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 17) (:end 5 18)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 15) (:end 5 16)) ((:name "n") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 41) (:end 3 42)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 38) (:end 3 39)) ((:name "k") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 25) (:end 3 26)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 23) (:end 3 24)) ((:name "k") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 13) (:end 3 14)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 19) (:end 2 20)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 1) (:end 1 4)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 6) (:end 1 9)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 11)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 13) (:end 1 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 17) (:end 1 18)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 20) (:end 1 23)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 25) (:end 1 26)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 28) (:end 1 31)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 33) (:end 1 37)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 6) (:end 2 8)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 10)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 12) (:end 2 15)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 17) (:end 2 17)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 19) (:end 2 19)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 6) (:end 3 9)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 11) (:end 3 11)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 13) (:end 3 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 18) (:end 3 21)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 23) (:end 3 23)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 25) (:end 3 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 27) (:end 3 28)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 30) (:end 3 33)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 35) (:end 3 35)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 36) (:end 3 36)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 38) (:end 3 38)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 39) (:end 3 39)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 41) (:end 3 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 5 6)) ((:decor :function)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 8) (:end 5 8)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 10) (:end 5 13)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 15) (:end 5 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 17) (:end 5 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 19) (:end 5 20)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 22) (:end 5 25)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 27) (:end 5 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 29) (:end 5 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 31) (:end 5 32)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 34) (:end 5 37)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 39) (:end 5 39)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 42) (:end 5 42)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 44)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 45) (:end 5 45)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 5 47)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 1) (:end 6 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 8) (:end 6 9)) ((:name "Nil") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 12)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 14) (:end 6 14)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 6 17)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 1) (:end 7 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 8) (:end 7 8)) ((:decor :keyword)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 11) (:end 7 12)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 15)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 16) (:end 7 16)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 19)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 21) (:end 7 21)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 23)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 25) (:end 7 26)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 28) (:end 7 33)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 36)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 7 39)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000015(:return (:ok ()) 1) 000040(:return (:ok "Syntax highlight option changed to False" ()) 2) 000015(:return (:ok ()) 3) diff --git a/tests/ideMode/ideMode003/expected b/tests/ideMode/ideMode003/expected index 471dabbab..02770583b 100644 --- a/tests/ideMode/ideMode003/expected +++ b/tests/ideMode/ideMode003/expected @@ -1,27 +1,68 @@ 000018(:protocol-version 2 0) 000038(:write-string "1/1: Building LocType (LocType.idr)" 1) -0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 24)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "a")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 7 40)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1) -0000df(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 37)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect {k:376} a)")))))) 1) -0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 10)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "?{_:377}_[]")))))) 1) -0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 16)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{k:376}_[] ?{_:377}_[])")))))) 1) -0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 20)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:378}_[] ?{_:377}_[])")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 6 18)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect m a)")))))) 1) -0000ed(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 13)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "(Main.Vect ?{_:367}_[] ?{_:366}_[])")))))) 1) -0001ec(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 5 7)) ((:name "Main.append") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "{0 m : Prelude.Types.Nat} -> {0 a : Type} -> {0 n : Prelude.Types.Nat} -> ({arg:357} : (Main.Vect n[0] a[1])) -> ({arg:358} : (Main.Vect m[3] a[2])) -> (Main.Vect (Prelude.Num.+ Prelude.Types.Nat Prelude.Types.Num implementation at Prelude/Types.idr:64:1--69:33 n[2] m[4]) a[3]))))))")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 5 48)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 45)) ((:name "m") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 41)) ((:name "n") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 29) (:end 5 30)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 27) (:end 5 28)) ((:name "m") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 17) (:end 5 18)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 15) (:end 5 16)) ((:name "n") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 41) (:end 3 42)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 38) (:end 3 39)) ((:name "k") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 25) (:end 3 26)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000d9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 23) (:end 3 24)) ((:name "k") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Prelude.Types.Nat")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 13) (:end 3 14)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) -0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 19) (:end 2 20)) ((:name "a") (:namespace "") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "Type")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 1) (:end 1 4)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 6) (:end 1 9)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 11)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 13) (:end 1 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 17) (:end 1 18)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 20) (:end 1 23)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 25) (:end 1 26)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 28) (:end 1 31)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 33) (:end 1 37)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 6) (:end 2 8)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 10)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 12) (:end 2 15)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 17) (:end 2 17)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 19) (:end 2 19)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 6) (:end 3 9)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 11) (:end 3 11)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 13) (:end 3 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 18) (:end 3 21)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 23) (:end 3 23)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 25) (:end 3 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 27) (:end 3 28)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 30) (:end 3 33)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 35) (:end 3 35)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 36) (:end 3 36)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 38) (:end 3 38)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 39) (:end 3 39)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 41) (:end 3 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 5 6)) ((:decor :function)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 8) (:end 5 8)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 10) (:end 5 13)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 15) (:end 5 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 17) (:end 5 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 19) (:end 5 20)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 22) (:end 5 25)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 27) (:end 5 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 29) (:end 5 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 31) (:end 5 32)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 34) (:end 5 37)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 39) (:end 5 39)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 42) (:end 5 42)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 44)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 45) (:end 5 45)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 5 47)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 1) (:end 6 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 8) (:end 6 9)) ((:name "Nil") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 12)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 14) (:end 6 14)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 6 17)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 1) (:end 7 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 8) (:end 7 8)) ((:decor :keyword)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 11) (:end 7 12)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 15)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 16) (:end 7 16)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 19)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 21) (:end 7 21)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 23)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 25) (:end 7 26)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 28) (:end 7 33)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 36)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 7 39)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) 000015(:return (:ok ()) 1) 000037(:return (:ok "Main.Vect : Nat -> Type -> Type" ()) 2) Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/Ambiguity.idr b/tests/ideMode/ideMode005/Ambiguity.idr new file mode 100644 index 000000000..a0064f11f --- /dev/null +++ b/tests/ideMode/ideMode005/Ambiguity.idr @@ -0,0 +1,10 @@ +namespace X + export + OneX : Bool -> Bool + +namespace Y + export + data Foo = OneX + +foo : Bool +foo = OneX (id False) diff --git a/tests/ideMode/ideMode005/Case.idr b/tests/ideMode/ideMode005/Case.idr new file mode 100644 index 000000000..c75ac96d7 --- /dev/null +++ b/tests/ideMode/ideMode005/Case.idr @@ -0,0 +1,22 @@ +module Case + +listId : List a -> List a +listId xs = case xs of + [] => [] + (x :: xs) => x :: listId xs + +listRev : List a -> List a +listRev = \case + [] => [] + (x :: xs) => listRev xs ++ [x] + +listFilter2 : (p, q : a -> Bool) -> List a -> List a +listFilter2 p q xs + = do x <- xs + -- let pat + let True = p x + | False => [] + -- do pat + True <- pure (q x) + | False => [] + pure x diff --git a/tests/ideMode/ideMode005/Implementation.idr b/tests/ideMode/ideMode005/Implementation.idr new file mode 100644 index 000000000..7a273447c --- /dev/null +++ b/tests/ideMode/ideMode005/Implementation.idr @@ -0,0 +1,25 @@ +module Implementation + +import Data.String + +data OneTwoThree a + = One a + | Two a a + | Three a a a + +[OTT] Functor OneTwoThree where + map f = \case + One x => One (f x) + Two x y => Two (f x) (f y) + Three x y z => Three (f x) (f y) (f z) + +Show a => Show (OneTwoThree a) where + show (One x) = "One " ++ show x + show (Two x y) = unwords ["Two", show x, show y] + show (Three x y z) = unwords ["Three", show x, show y, show z] + +Eq a => Eq (OneTwoThree a) where + One a == One x = a == x + Two a b == Two x y = a == x && b == y + Three a b c == Three x y z = a == x && b == y && c == z + _ == _ = False diff --git a/tests/ideMode/ideMode005/Interface.idr b/tests/ideMode/ideMode005/Interface.idr new file mode 100644 index 000000000..aba897be5 --- /dev/null +++ b/tests/ideMode/ideMode005/Interface.idr @@ -0,0 +1,14 @@ +module Interface + +interface Show a => Pretty (0 a : Type) where + constructor MkPretty + + 0 Doc : Type + toDoc : String -> Doc + + pretty : a -> Doc + pretty n = toDoc (show n) + + prettys : List a -> List Doc + prettys [] = [] + prettys (a :: as) = pretty a :: prettys as diff --git a/tests/ideMode/ideMode005/README.md b/tests/ideMode/ideMode005/README.md new file mode 100644 index 000000000..640fab8cd --- /dev/null +++ b/tests/ideMode/ideMode005/README.md @@ -0,0 +1,17 @@ +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/Rainbow.idr b/tests/ideMode/ideMode005/Rainbow.idr new file mode 100644 index 000000000..26e53fc3d --- /dev/null +++ b/tests/ideMode/ideMode005/Rainbow.idr @@ -0,0 +1,30 @@ +namespace L0 + public export + data List0 : Type + + +namespace L1 + public export + data List1 : Type where + Nil : List1 + Cons1 : Nat -> List0 -> List1 + + public export + (::) : Nat -> List0 -> List1 + (::) = Cons1 + +namespace L2 + public export + data (::) : Nat -> List0 -> Type where + +namespace L0 + public export + data List0 : Type where + Nil : List0 + (::) : Nat ->Type -> List0 + +m : Nat +m = believe_me %MkWorld + +Rainbow : Nat -> Type +Rainbow n = [ n , m , n ] diff --git a/tests/ideMode/ideMode005/Ranges.idr b/tests/ideMode/ideMode005/Ranges.idr new file mode 100644 index 000000000..869a0c471 --- /dev/null +++ b/tests/ideMode/ideMode005/Ranges.idr @@ -0,0 +1,5 @@ +hours : List Nat +hours = [1..12] + +nats : Stream Nat +nats = [0,1..] diff --git a/tests/ideMode/ideMode005/RecordUpdate.idr b/tests/ideMode/ideMode005/RecordUpdate.idr new file mode 100644 index 000000000..0f475ef6d --- /dev/null +++ b/tests/ideMode/ideMode005/RecordUpdate.idr @@ -0,0 +1,23 @@ +module RecordUpdate + +record Attributes where + font : String + size : Nat + +bigMono : Attributes -> Attributes +bigMono = record { font $= (++ "Mono") + , size = 20 + } + +smallMono : Attributes -> Attributes +smallMono = { size := 5 } . bigMono + +-- Works for nested too +record Text where + attributes : Attributes + content : String + +setArial10 : Text -> Text +setArial10 = { attributes->font := "Arial" + , attributes.size := 10 + } diff --git a/tests/ideMode/ideMode005/Rewrite.idr b/tests/ideMode/ideMode005/Rewrite.idr new file mode 100644 index 000000000..8ea9f3d1f --- /dev/null +++ b/tests/ideMode/ideMode005/Rewrite.idr @@ -0,0 +1,9 @@ +transport : (0 eq : a === b) -> a -> b +transport eq x = rewrite sym eq in x + +nested : (0 _ : a === b) -> (0 _ : b === c) -> a === c +nested eq1 eq2 = + rewrite eq1 in + rewrite eq2 in + let prf : c === c := Refl in + prf diff --git a/tests/ideMode/ideMode005/SimpleData.idr b/tests/ideMode/ideMode005/SimpleData.idr new file mode 100644 index 000000000..f5fffe297 --- /dev/null +++ b/tests/ideMode/ideMode005/SimpleData.idr @@ -0,0 +1,8 @@ +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/StringLiterals.idr b/tests/ideMode/ideMode005/StringLiterals.idr new file mode 100644 index 000000000..17f3b40d2 --- /dev/null +++ b/tests/ideMode/ideMode005/StringLiterals.idr @@ -0,0 +1,19 @@ +module StringLiterals + +hello : FromString a => a +hello = "Hello" + +helloName : String -> String +helloName name = "\{hello {a = String}} \{name}" + +welcomeName : String -> String +welcomeName name = """ + \{helloName name} + and welcome! + """ + +scareQuotes : String +scareQuotes = #""hello""# + +test : StringLiterals.scareQuotes = "\"hello\"" +test = Refl diff --git a/tests/ideMode/ideMode005/Syntax.idr b/tests/ideMode/ideMode005/Syntax.idr new file mode 100644 index 000000000..1338a0637 --- /dev/null +++ b/tests/ideMode/ideMode005/Syntax.idr @@ -0,0 +1,52 @@ +module Syntax + +%default total + +{- -- interfaces don't work yet +{0 a : Type} -> (foo : Show a) => Show (Maybe a) where +-} + +showMaybe : {0 a : Type} -> (assumption : Show a) => Maybe a -> String +showMaybe x@ma = case map (id . id) ma of + Nothing => "Nothing" + Just a => "Just " ++ show a + +nats : List Nat +nats = + let n = 5 + m = 7 + xs = [id $ id 0] + ys = [1, 2, m, n, 3] ++ xs + in [n,id $ id m] ++ [1, 2, m, n, 3] ++ xs + + +record ANat where + constructor MkANat + aNat : Nat + +doBlock : Maybe ANat +doBlock + = do let a = 3 + let b = 5 + (c, _) <- Just (7, 9) + let (d, e) = (c, c) + f <- [| Nothing + Just d |] + pure $ MkANat $ sum [a,b,c,d,e,f] + +parameters (x, y, z : Nat) + + add3 : Nat + add3 = x + y + z + +parameters (x, y : Nat) (z, a : Nat) + + add4 : Nat + add4 = x + y + z + a + +anonLam : Maybe (m : Nat ** n : Nat ** m === n) +anonLam = map (\m => (m ** m ** Refl)) + $ map (uncurry $ \ m, n => m + n) + $ map (\ (m, n) => (n, m)) + $ map (\ m => (m, m)) + $ map (\ m => S m.aNat) + doBlock diff --git a/tests/ideMode/ideMode005/With.idr b/tests/ideMode/ideMode005/With.idr new file mode 100644 index 000000000..07c358660 --- /dev/null +++ b/tests/ideMode/ideMode005/With.idr @@ -0,0 +1,30 @@ +module With + +f : (n : Nat) -> (m : Nat ** n : Nat ** m = n + n) +f n with (n + n) proof eq + f n | Z = (Z ** n ** sym eq) + f n | (S m) = (S m ** n ** sym eq) + +g : List a -> Nat +g [] = Z +g (a :: as) with (as ++ as) + g (b :: bs) | asas = Z + +nested : Nat -> Nat +nested m with (m) + nested m | Z with (m + m) + nested m | Z | 0 = 1 + nested m | Z | S k with (k + k) + nested m | Z | S k | Z = 2 + nested m | Z | S k | S l with (l + l) + nested m | Z | S k | S l | Z = 3 + nested m | Z | S k | S l | S p = 4 + nested m | S k = 5 + +data ANat : Nat -> Type where + MkANat : (n : Nat) -> ANat n + +someNats : Nat -> Nat +someNats n with (MkANat n) + someNats n | m@(MkANat n) with (MkANat n) + someNats n | p@(MkANat n) | MkANat n = Z diff --git a/tests/ideMode/ideMode005/WithApp.idr b/tests/ideMode/ideMode005/WithApp.idr new file mode 100644 index 000000000..3cc42c336 --- /dev/null +++ b/tests/ideMode/ideMode005/WithApp.idr @@ -0,0 +1,12 @@ +data Natty : Nat -> Type where + Z : Natty Z + S : Natty n -> Natty (S n) + +view : (n: Nat) -> Natty n +view Z = Z +view (S n) = S (view n) + +id : Nat -> Nat +id n with (view n) + id _ | Z = Z + id _ | S n = id _ | n diff --git a/tests/ideMode/ideMode005/dummy.ipkg b/tests/ideMode/ideMode005/dummy.ipkg new file mode 100644 index 000000000..c31d6563c --- /dev/null +++ b/tests/ideMode/ideMode005/dummy.ipkg @@ -0,0 +1 @@ +package idemode005 diff --git a/tests/ideMode/ideMode005/expected b/tests/ideMode/ideMode005/expected new file mode 100644 index 000000000..e69de29bb diff --git a/tests/ideMode/ideMode005/expected1 b/tests/ideMode/ideMode005/expected1 new file mode 100644 index 000000000..1402cd144 --- /dev/null +++ b/tests/ideMode/ideMode005/expected1 @@ -0,0 +1,302 @@ +000018(:protocol-version 2 0) +000036(:write-string "1/1: Building Syntax (Syntax.idr)" 1) +000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 3 1) (:end 3 8)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 3 10) (:end 3 14)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 1) (:end 9 9)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 11) (:end 9 11)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 13) (:end 9 13)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 14) (:end 9 14)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 16) (:end 9 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 18) (:end 9 18)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 20) (:end 9 23)) ((:decor :type)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 24) (:end 9 24)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 26) (:end 9 27)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 29) (:end 9 29)) ((:decor :keyword)))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 30) (:end 9 39)) ((:name "assumption") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 41) (:end 9 41)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 43) (:end 9 46)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 48) (:end 9 48)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 49) (:end 9 49)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 51) (:end 9 52)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 54) (:end 9 58)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 60) (:end 9 60)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 62) (:end 9 63)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 65) (:end 9 70)) ((:decor :type)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 1) (:end 10 9)) ((:name "showMaybe") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 11) (:end 10 11)) ((:decor :bound)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 12) (:end 10 12)) ((:decor :keyword)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 13) (:end 10 14)) ((:name "ma") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 16) (:end 10 16)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 18) (:end 10 21)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 23) (:end 10 25)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 27) (:end 10 27)) ((:decor :keyword)))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 28) (:end 10 29)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 31) (:end 10 31)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 33) (:end 10 34)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 35) (:end 10 35)) ((:decor :keyword)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 37) (:end 10 38)) ((:name "ma") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 40) (:end 10 41)) ((:decor :keyword)))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 5) (:end 11 11)) ((:name "Nothing") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 13) (:end 11 14)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 16) (:end 11 24)) ((:decor :data)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 5) (:end 12 8)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 10) (:end 12 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 12) (:end 12 13)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 15) (:end 12 21)) ((:decor :data)))))) 1) +0000e2(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 23) (:end 12 24)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 26) (:end 12 29)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 31) (:end 12 31)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 1) (:end 14 4)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 6) (:end 14 6)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 8) (:end 14 11)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 13) (:end 14 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 1) (:end 15 4)) ((:name "nats") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 6) (:end 15 6)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 3) (:end 16 5)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 7) (:end 16 7)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 9) (:end 16 9)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 11) (:end 16 11)) ((:decor :data)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 17 7) (:end 17 7)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 17 9) (:end 17 9)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 17 11) (:end 17 11)) ((:decor :data)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 7) (:end 18 8)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 10) (:end 18 10)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 12) (:end 18 12)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 13) (:end 18 14)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 18) (:end 18 19)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 21) (:end 18 21)) ((:decor :data)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 22) (:end 18 22)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 7) (:end 19 8)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 10) (:end 19 10)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 12) (:end 19 12)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 13) (:end 19 13)) ((:decor :data)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 14) (:end 19 14)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 16) (:end 19 16)) ((:decor :data)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 17) (:end 19 17)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 19) (:end 19 19)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 20) (:end 19 20)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 22) (:end 19 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 23) (:end 19 23)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 25) (:end 19 25)) ((:decor :data)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 26) (:end 19 26)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 28) (:end 19 29)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 31) (:end 19 32)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 3) (:end 20 4)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 6) (:end 20 6)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 7) (:end 20 7)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 8) (:end 20 8)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 9) (:end 20 10)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 14) (:end 20 15)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 17) (:end 20 17)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 18) (:end 20 18)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 20) (:end 20 21)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 23) (:end 20 23)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 24) (:end 20 24)) ((:decor :data)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 25) (:end 20 25)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 27) (:end 20 27)) ((:decor :data)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 28) (:end 20 28)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 30) (:end 20 30)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 31) (:end 20 31)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 33) (:end 20 33)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 34) (:end 20 34)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 36) (:end 20 36)) ((:decor :data)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 37) (:end 20 37)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 39) (:end 20 40)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 42) (:end 20 43)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 23 1) (:end 23 6)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 23 8) (:end 23 11)) ((:decor :type)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 23 13) (:end 23 17)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 24 15) (:end 24 20)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 3) (:end 25 6)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 8) (:end 25 8)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 10) (:end 25 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 10) (:end 25 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 10) (:end 25 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 1) (:end 27 7)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 9) (:end 27 9)) ((:decor :keyword)))))) 1) +0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 11) (:end 27 15)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 17) (:end 27 20)) ((:name "ANat") (:namespace "Syntax") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 1) (:end 28 7)) ((:name "doBlock") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 3) (:end 29 3)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 8) (:end 29 10)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 12) (:end 29 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 14) (:end 29 14)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 16) (:end 29 16)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 8) (:end 30 10)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 12) (:end 30 12)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 14) (:end 30 14)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 16) (:end 30 16)) ((:decor :data)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 8) (:end 31 13)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 9) (:end 31 9)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 10) (:end 31 10)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 12) (:end 31 12)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 15) (:end 31 16)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 18) (:end 31 21)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 23) (:end 31 28)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 24) (:end 31 24)) ((:decor :data)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 25) (:end 31 25)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 27) (:end 31 27)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 8) (:end 32 10)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 12) (:end 32 17)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 13) (:end 32 13)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 14) (:end 32 14)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 16) (:end 32 16)) ((:name "e") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 19) (:end 32 19)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 21) (:end 32 26)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 22) (:end 32 22)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 23) (:end 32 23)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 25) (:end 32 25)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 8) (:end 33 8)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 10) (:end 33 11)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 13) (:end 33 14)) ((:decor :keyword)))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 16) (:end 33 22)) ((:name "Nothing") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 24) (:end 33 24)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 26) (:end 33 29)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 31) (:end 33 31)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 33) (:end 33 34)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 8) (:end 34 11)) ((:name "pure") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 15) (:end 34 20)) ((:name "MkANat") (:namespace "Syntax") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 24) (:end 34 26)) ((:name "sum") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 28) (:end 34 28)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 29) (:end 34 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 30) (:end 34 30)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 31) (:end 34 31)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 32) (:end 34 32)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 33) (:end 34 33)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 34) (:end 34 34)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 35) (:end 34 35)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 36) (:end 34 36)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 37) (:end 34 37)) ((:name "e") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 38) (:end 34 38)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 39) (:end 34 39)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 40) (:end 34 40)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 1) (:end 36 10)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 12) (:end 36 12)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 13) (:end 36 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 14) (:end 36 14)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 16) (:end 36 16)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 17) (:end 36 17)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 19) (:end 36 19)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 21) (:end 36 21)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 23) (:end 36 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 23) (:end 36 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 23) (:end 36 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 26) (:end 36 26)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 3) (:end 38 6)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 8) (:end 38 8)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 10) (: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 3) (:end 39 6)) ((:name "add3") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 8) (:end 39 8)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 10) (:end 39 10)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 12) (:end 39 12)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 14) (:end 39 14)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 16) (:end 39 16)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 18) (:end 39 18)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 1) (:end 41 10)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 12) (:end 41 12)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 13) (:end 41 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 14) (:end 41 14)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 16) (:end 41 16)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 18) (:end 41 18)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 20) (:end 41 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 20) (:end 41 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 23) (:end 41 23)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 25) (:end 41 25)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 26) (:end 41 26)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 27) (:end 41 27)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 29) (:end 41 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 31) (:end 41 31)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 33) (:end 41 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 33) (:end 41 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 36) (:end 41 36)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 3) (:end 43 6)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 8) (:end 43 8)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 10) (:end 43 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 3) (:end 44 6)) ((:name "add4") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 8) (:end 44 8)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 10) (:end 44 10)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 12) (:end 44 12)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 14) (:end 44 14)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 16) (:end 44 16)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 18) (:end 44 18)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 20) (:end 44 20)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 22) (:end 44 22)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 1) (:end 46 7)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 9) (:end 46 9)) ((:decor :keyword)))))) 1) +0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 11) (:end 46 15)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 17) (:end 46 17)) ((:decor :type)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 18) (:end 46 18)) ((:decor :bound)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 20) (:end 46 20)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 22) (:end 46 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 22) (:end 46 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 26) (:end 46 27)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 29) (:end 46 29)) ((:decor :bound)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 31) (:end 46 31)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 33) (:end 46 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 33) (:end 46 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 37) (:end 46 38)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 40) (:end 46 40)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 42) (:end 46 44)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 46) (:end 46 46)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 47) (:end 46 47)) ((:decor :type)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 1) (:end 47 7)) ((:name "anonLam") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 9) (:end 47 9)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 11) (:end 47 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 15) (:end 47 15)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 16) (:end 47 16)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 17) (:end 47 17)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 19) (:end 47 20)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 23) (:end 47 23)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 25) (:end 47 26)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 28) (:end 47 28)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 30) (:end 47 31)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 33) (:end 47 36)) ((:name "Refl") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 38) (:end 47 38)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 11) (:end 48 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 15) (:end 48 15)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 16) (:end 48 22)) ((:name "uncurry") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 26) (:end 48 26)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 28) (:end 48 28)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 29) (:end 48 29)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 31) (:end 48 31)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 33) (:end 48 34)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 36) (:end 48 36)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 38) (:end 48 38)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 40) (:end 48 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 41) (:end 48 41)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 11) (:end 49 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 15) (:end 49 15)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 16) (:end 49 16)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 18) (:end 49 23)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 19) (:end 49 19)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 20) (:end 49 20)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 22) (:end 49 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 25) (:end 49 26)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 28) (:end 49 33)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 29) (:end 49 29)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 30) (:end 49 30)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 32) (:end 49 32)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 49 34) (:end 49 34)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 11) (:end 50 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 15) (:end 50 15)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 16) (:end 50 16)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 18) (:end 50 18)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 20) (:end 50 21)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 23) (:end 50 28)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 24) (:end 50 24)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 25) (:end 50 25)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 27) (:end 50 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 50 29) (:end 50 29)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 11) (:end 51 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 15) (:end 51 15)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 16) (:end 51 16)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 18) (:end 51 18)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 20) (:end 51 21)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 23) (:end 51 23)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 25) (:end 51 25)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 25) (:end 51 30)) ((:name "aNat") (:namespace "Syntax.ANat") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 51 31) (:end 51 31)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 52 9) (:end 52 15)) ((:name "doBlock") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expected2 b/tests/ideMode/ideMode005/expected2 new file mode 100644 index 000000000..9729bc4ca --- /dev/null +++ b/tests/ideMode/ideMode005/expected2 @@ -0,0 +1,77 @@ +000018(:protocol-version 2 0) +00003c(:write-string "1/1: Building Interface (Interface.idr)" 1) +000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 1) (:end 3 9)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 11) (:end 3 14)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 11) (:end 3 14)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 16) (:end 3 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 16) (:end 3 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 18) (:end 3 19)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 21) (:end 3 26)) ((:decor :type)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 28) (:end 3 28)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 29) (:end 3 29)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 31) (:end 3 31)) ((:decor :bound)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 33) (:end 3 33)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 35) (:end 3 38)) ((:decor :type)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 39) (:end 3 39)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 41) (:end 3 45)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 4 15) (:end 4 22)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 3) (:end 6 3)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 5) (:end 6 7)) ((:name "Doc") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 5) (:end 6 7)) ((:decor :function)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 9) (:end 6 9)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 11) (:end 6 14)) ((:decor :type)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 7 3) (:end 7 7)) ((:name "toDoc") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 7 3) (:end 7 7)) ((:decor :function)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 7 9) (:end 7 9)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 7 11) (:end 7 16)) ((:decor :type)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 7 18) (:end 7 19)) ((:decor :keyword)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 7 21) (:end 7 23)) ((:name "Doc") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000da(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 3) (:end 9 8)) ((:name "pretty") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 3) (:end 9 8)) ((:decor :function)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 10) (:end 9 10)) ((:decor :keyword)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 12) (:end 9 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 12) (:end 9 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 12) (:end 9 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 14) (:end 9 15)) ((:decor :keyword)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 17) (:end 9 19)) ((:name "Doc") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 10 3) (:end 10 8)) ((:name "pretty") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 10 10) (:end 10 10)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 10 12) (:end 10 12)) ((:decor :keyword)))))) 1) +0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 10 14) (:end 10 18)) ((:name "toDoc") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 10 20) (:end 10 20)) ((:decor :keyword)))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 10 21) (:end 10 24)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 10 26) (:end 10 26)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 10 27) (:end 10 27)) ((:decor :keyword)))))) 1) +0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 3) (:end 12 9)) ((:name "prettys") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 3) (:end 12 9)) ((:decor :function)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 11) (:end 12 11)) ((:decor :keyword)))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 13) (:end 12 16)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 13) (:end 12 16)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 13) (:end 12 16)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 18) (:end 12 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 18) (:end 12 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 18) (:end 12 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 20) (:end 12 21)) ((:decor :keyword)))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 23) (:end 12 26)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 23) (:end 12 26)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 23) (:end 12 26)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 28) (:end 12 30)) ((:name "Doc") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 3) (:end 13 9)) ((:name "prettys") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 11) (:end 13 12)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 14) (:end 13 14)) ((:decor :keyword)))))) 1) +0000db(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 16) (:end 13 17)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 3) (:end 14 9)) ((:name "prettys") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 11) (:end 14 11)) ((:decor :keyword)))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 12) (:end 14 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000da(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 14) (:end 14 15)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 17) (:end 14 18)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 19) (:end 14 19)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 21) (:end 14 21)) ((:decor :keyword)))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 23) (:end 14 28)) ((:name "pretty") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 30) (:end 14 30)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000da(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 32) (:end 14 33)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 35) (:end 14 41)) ((:name "prettys") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 43) (:end 14 44)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expected3 b/tests/ideMode/ideMode005/expected3 new file mode 100644 index 000000000..080024d44 --- /dev/null +++ b/tests/ideMode/ideMode005/expected3 @@ -0,0 +1,63 @@ +000018(:protocol-version 2 0) +00003e(:write-string "1/1: Building SimpleData (SimpleData.idr)" 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 1) (:end 3 4)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 6) (:end 3 8)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 6) (:end 3 8)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 6) (:end 3 8)) ((:decor :type)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 10) (:end 3 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 10) (:end 3 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 10) (:end 3 10)) ((:decor :bound)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 12) (:end 3 12)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 14) (:end 3 17)) ((:decor :data)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 19) (:end 3 19)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 21) (:end 3 24)) ((:decor :data)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 26) (:end 3 26)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 28) (:end 3 28)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 29) (:end 3 31)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 33) (:end 3 33)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 34) (:end 3 34)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 1) (:end 4 4)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 6) (:end 4 8)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 6) (:end 4 8)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 6) (:end 4 8)) ((:decor :type)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 10) (:end 4 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 10) (:end 4 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 10) (:end 4 10)) ((:decor :bound)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 12) (:end 4 12)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 14) (:end 4 17)) ((:decor :data)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 19) (:end 4 19)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 21) (:end 4 24)) ((:decor :data)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 26) (:end 4 26)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 27) (:end 4 29)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 31) (:end 4 31)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 32) (:end 4 32)) ((:decor :keyword)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 4 34) (:end 4 34)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 1) (:end 6 4)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 6) (:end 6 9)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 6) (:end 6 9)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 6) (:end 6 9)) ((:decor :type)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 11) (:end 6 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 11) (:end 6 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 11) (:end 6 11)) ((:decor :bound)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 13) (:end 6 13)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 13) (:end 6 13)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 13) (:end 6 13)) ((:decor :bound)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 3) (:end 7 3)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 5) (:end 7 8)) ((:decor :data)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 10) (:end 7 10)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 3) (:end 8 3)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 5) (:end 8 8)) ((:decor :data)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 10) (:end 8 10)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 11) (:end 8 14)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 16) (:end 8 16)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 18) (:end 8 18)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 19) (:end 8 19)) ((:decor :keyword)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 21) (:end 8 21)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 23) (:end 8 23)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 24) (:end 8 27)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 29) (:end 8 29)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 31) (:end 8 31)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 8 32) (:end 8 32)) ((:decor :keyword)))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expected4 b/tests/ideMode/ideMode005/expected4 new file mode 100644 index 000000000..499c9c6be --- /dev/null +++ b/tests/ideMode/ideMode005/expected4 @@ -0,0 +1 @@ +000018(:protocol-version 2 0) diff --git a/tests/ideMode/ideMode005/expected5 b/tests/ideMode/ideMode005/expected5 new file mode 100644 index 000000000..8a4930c17 --- /dev/null +++ b/tests/ideMode/ideMode005/expected5 @@ -0,0 +1,89 @@ +000018(:protocol-version 2 0) +000038(:write-string "1/1: Building Rainbow (Rainbow.idr)" 1) +000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 3) (:end 2 8)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 10) (:end 2 15)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 3 3) (:end 3 6)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 3 8) (:end 3 12)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 3 14) (:end 3 14)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 3 16) (:end 3 19)) ((:decor :type)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 3) (:end 7 8)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 10) (:end 7 15)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 3) (:end 8 6)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 8) (:end 8 12)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 14) (:end 8 14)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 16) (:end 8 19)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 21) (:end 8 25)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 5) (:end 9 7)) ((:decor :data)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 9) (:end 9 9)) ((:decor :keyword)))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 11) (:end 9 15)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 10 5) (:end 10 9)) ((:decor :data)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 10 11) (:end 10 11)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 10 13) (:end 10 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 10 17) (:end 10 18)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 10 20) (:end 10 24)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 10 26) (:end 10 27)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 10 29) (:end 10 33)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 3) (:end 12 8)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 10) (:end 12 15)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 3) (:end 13 6)) ((:decor :function)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 8) (:end 13 8)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 10) (:end 13 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 14) (:end 13 15)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 17) (:end 13 21)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 23) (:end 13 24)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 26) (:end 13 30)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 14 3) (:end 14 6)) ((:name "::") (:namespace "Main.L1") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 14 8) (:end 14 8)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 14 10) (:end 14 14)) ((:name "Cons1") (:namespace "Main.L1") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 3) (:end 17 8)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 10) (:end 17 15)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 3) (:end 18 6)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 8) (:end 18 11)) ((:decor :type)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 13) (:end 18 13)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 15) (:end 18 17)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 19) (:end 18 20)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 22) (:end 18 26)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 28) (:end 18 29)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 31) (:end 18 34)) ((:decor :type)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 18 36) (:end 18 40)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 3) (:end 21 8)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 10) (:end 21 15)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 3) (:end 22 6)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 8) (:end 22 12)) ((:decor :type)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 14) (:end 22 14)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 16) (:end 22 19)) ((:decor :type)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 21) (:end 22 25)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 5) (:end 23 7)) ((:decor :data)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 9) (:end 23 9)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 11) (:end 23 15)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 24 5) (:end 24 8)) ((:decor :data)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 24 10) (:end 24 10)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 24 12) (:end 24 14)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 24 16) (:end 24 17)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 24 18) (:end 24 21)) ((:decor :type)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 24 23) (:end 24 24)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 24 26) (:end 24 30)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 1) (:end 26 1)) ((:decor :function)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 3) (:end 26 3)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 5) (:end 26 7)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 27 1) (:end 27 1)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 27 3) (:end 27 3)) ((:decor :keyword)))))) 1) +0000dd(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 27 5) (:end 27 14)) ((:name "believe_me") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 27 16) (:end 27 23)) ((:decor :data)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 1) (:end 29 7)) ((:decor :function)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 9) (:end 29 9)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 11) (:end 29 13)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 15) (:end 29 16)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 18) (:end 29 21)) ((:decor :type)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 1) (:end 30 7)) ((:name "Rainbow") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 9) (:end 30 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 11) (:end 30 11)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 14) (:end 30 14)) ((:name "::") (:namespace "Main.L2") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 16) (:end 30 16)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 18) (:end 30 18)) ((:name "::") (:namespace "Main.L0") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 20) (:end 30 20)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 22) (:end 30 22)) ((:name "::") (:namespace "Main.L2") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 24) (:end 30 24)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 30 26) (:end 30 26)) ((:name "Nil") (:namespace "Main.L0") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expected6 b/tests/ideMode/ideMode005/expected6 new file mode 100644 index 000000000..0d8334850 --- /dev/null +++ b/tests/ideMode/ideMode005/expected6 @@ -0,0 +1,215 @@ +000018(:protocol-version 2 0) +000046(:write-string "1/1: Building Implementation (Implementation.idr)" 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 3 1) (:end 3 6)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 1) (:end 5 4)) ((:decor :keyword)))))) 1) +0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 6) (:end 5 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 6) (:end 5 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 6) (:end 5 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 6) (:end 5 16)) ((:decor :type)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 18) (:end 5 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 18) (:end 5 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 18) (:end 5 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 18) (:end 5 18)) ((:decor :bound)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 3) (:end 6 3)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 5) (:end 6 7)) ((:decor :data)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 9) (:end 6 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 3) (:end 7 3)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 5) (:end 7 7)) ((:decor :data)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 9) (:end 7 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 11) (:end 7 11)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 8 3) (:end 8 3)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 8 5) (:end 8 9)) ((:decor :data)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 8 11) (:end 8 11)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 8 13) (:end 8 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 8 15) (:end 8 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 1) (:end 10 1)) ((:decor :keyword)))))) 1) +00007c(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 2) (:end 10 4)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 5) (:end 10 5)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 7) (:end 10 13)) ((:decor :type)))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 15) (:end 10 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 15) (:end 10 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 15) (:end 10 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 27) (:end 10 31)) ((:decor :keyword)))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 3) (:end 11 5)) ((:name "map") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 7) (:end 11 7)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 9) (:end 11 9)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 11) (:end 11 11)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 12) (:end 11 15)) ((:decor :keyword)))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 5) (:end 12 7)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 9) (:end 12 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 11) (:end 12 12)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 14) (:end 12 16)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 18) (:end 12 18)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 19) (:end 12 19)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 21) (:end 12 21)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 22) (:end 12 22)) ((:decor :keyword)))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 5) (:end 13 7)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 9) (:end 13 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 11) (:end 13 11)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 13) (:end 13 14)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 16) (:end 13 18)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 20) (:end 13 20)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 21) (:end 13 21)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 23) (:end 13 23)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 24) (:end 13 24)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 26) (:end 13 26)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 27) (:end 13 27)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 29) (:end 13 29)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 30) (:end 13 30)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 5) (:end 14 9)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 11) (:end 14 11)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 13) (:end 14 13)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 15) (:end 14 15)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 17) (:end 14 18)) ((:decor :keyword)))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 20) (:end 14 24)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 26) (:end 14 26)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 27) (:end 14 27)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 29) (:end 14 29)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 30) (:end 14 30)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 32) (:end 14 32)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 33) (:end 14 33)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 35) (:end 14 35)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 36) (:end 14 36)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 38) (:end 14 38)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 39) (:end 14 39)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 41) (:end 14 41)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 14 42) (:end 14 42)) ((:decor :keyword)))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 1) (:end 16 4)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 1) (:end 16 4)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 1) (:end 16 4)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 6) (:end 16 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 6) (:end 16 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 6) (:end 16 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 8) (:end 16 9)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 11) (:end 16 14)) ((:decor :type)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 16) (:end 16 16)) ((:decor :keyword)))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 17) (:end 16 27)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 17) (:end 16 27)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 17) (:end 16 27)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 29) (:end 16 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 29) (:end 16 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 29) (:end 16 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 30) (:end 16 30)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 32) (:end 16 36)) ((:decor :keyword)))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 3) (:end 17 6)) ((:name "show") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 8) (:end 17 8)) ((:decor :keyword)))))) 1) +0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 9) (:end 17 11)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 13) (:end 17 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 14) (:end 17 14)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 16) (:end 17 16)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 18) (:end 17 23)) ((:decor :data)))))) 1) +0000ea(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 25) (:end 17 26)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 28) (:end 17 31)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 33) (:end 17 33)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 3) (:end 18 6)) ((:name "show") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 8) (:end 18 8)) ((:decor :keyword)))))) 1) +0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 9) (:end 18 11)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 13) (:end 18 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 15) (:end 18 15)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 16) (:end 18 16)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 18) (:end 18 18)) ((:decor :keyword)))))) 1) +0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 20) (:end 18 26)) ((:name "unwords") (:namespace "Data.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 28) (:end 18 28)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 29) (:end 18 33)) ((:decor :data)))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 34) (:end 18 34)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 36) (:end 18 39)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 41) (:end 18 41)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 42) (:end 18 42)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 44) (:end 18 47)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 49) (:end 18 49)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 50) (:end 18 50)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 3) (:end 19 6)) ((:name "show") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 8) (:end 19 8)) ((:decor :keyword)))))) 1) +0000e2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 9) (:end 19 13)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 15) (:end 19 15)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 17) (:end 19 17)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 19) (:end 19 19)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 20) (:end 19 20)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 22) (:end 19 22)) ((:decor :keyword)))))) 1) +0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 24) (:end 19 30)) ((:name "unwords") (:namespace "Data.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 32) (:end 19 32)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 33) (:end 19 39)) ((:decor :data)))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 40) (:end 19 40)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 42) (:end 19 45)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 47) (:end 19 47)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 48) (:end 19 48)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 50) (:end 19 53)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 55) (:end 19 55)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 56) (:end 19 56)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 58) (:end 19 61)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 63) (:end 19 63)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 64) (:end 19 64)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dd(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 1) (:end 21 2)) ((:name "Eq") (:namespace "Prelude.EqOrd") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dd(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 1) (:end 21 2)) ((:name "Eq") (:namespace "Prelude.EqOrd") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dd(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 1) (:end 21 2)) ((:name "Eq") (:namespace "Prelude.EqOrd") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 4) (:end 21 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 4) (:end 21 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 4) (:end 21 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 6) (:end 21 7)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 9) (:end 21 10)) ((:decor :type)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 12) (:end 21 12)) ((:decor :keyword)))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 13) (:end 21 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 13) (:end 21 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 13) (:end 21 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 13) (:end 21 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 13) (:end 21 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 25) (:end 21 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 25) (:end 21 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 25) (:end 21 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 25) (:end 21 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 25) (:end 21 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 26) (:end 21 26)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 28) (:end 21 32)) ((:decor :keyword)))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 3) (:end 22 5)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 7) (:end 22 7)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 9) (:end 22 10)) ((:name "==") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 12) (:end 22 14)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 16) (:end 22 16)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 18) (:end 22 18)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 20) (:end 22 20)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 22) (:end 22 23)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 25) (:end 22 25)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 3) (:end 23 5)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 7) (:end 23 7)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 9) (:end 23 9)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 11) (:end 23 12)) ((:name "==") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 14) (:end 23 16)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 18) (:end 23 18)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 20) (:end 23 20)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 22) (:end 23 22)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 24) (:end 23 24)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 26) (:end 23 27)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 29) (:end 23 29)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 31) (:end 23 32)) ((:name "&&") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 34) (:end 23 34)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 36) (:end 23 37)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 39) (:end 23 39)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 3) (:end 24 7)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 9) (:end 24 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 11) (:end 24 11)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 13) (:end 24 13)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 15) (:end 24 16)) ((:name "==") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 18) (:end 24 22)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 24) (:end 24 24)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 26) (:end 24 26)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 28) (:end 24 28)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 30) (:end 24 30)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 32) (:end 24 32)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 34) (:end 24 35)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 37) (:end 24 37)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 39) (:end 24 40)) ((:name "&&") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 42) (:end 24 42)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 44) (:end 24 45)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 47) (:end 24 47)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 49) (:end 24 50)) ((:name "&&") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 52) (:end 24 52)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 54) (:end 24 55)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 57) (:end 24 57)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 25 3) (:end 25 3)) ((:decor :keyword)))))) 1) +0000e2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 25 5) (:end 25 6)) ((:name "==") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 25 8) (:end 25 8)) ((:decor :keyword)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 25 10) (:end 25 10)) ((:decor :keyword)))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 25 12) (:end 25 16)) ((:name "False") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expected7 b/tests/ideMode/ideMode005/expected7 new file mode 100644 index 000000000..620687d79 --- /dev/null +++ b/tests/ideMode/ideMode005/expected7 @@ -0,0 +1,106 @@ +000018(:protocol-version 2 0) +000032(:write-string "1/1: Building Case (Case.idr)" 1) +00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 1) (:end 3 6)) ((:decor :function)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 8) (:end 3 8)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 10) (:end 3 13)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 15) (:end 3 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 17) (:end 3 18)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 20) (:end 3 23)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 25) (:end 3 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 1) (:end 4 6)) ((:name "listId") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 8) (:end 4 9)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 11) (:end 4 11)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 13) (:end 4 16)) ((:decor :keyword)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 18) (:end 4 19)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 21) (:end 4 22)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 3) (:end 5 4)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 6) (:end 5 7)) ((:decor :keyword)))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 9) (:end 5 10)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 3) (:end 6 3)) ((:decor :keyword)))))) 1) +0000c4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 4) (:end 6 4)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 6) (:end 6 7)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 9) (:end 6 10)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 11) (:end 6 11)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 13) (:end 6 14)) ((:decor :keyword)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 16) (:end 6 16)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 18) (:end 6 19)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 21) (:end 6 26)) ((:name "listId") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 28) (:end 6 29)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 1) (:end 8 7)) ((:decor :function)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 9) (:end 8 9)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 11) (:end 8 14)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 16) (:end 8 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 18) (:end 8 19)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 21) (:end 8 24)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 26) (:end 8 26)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 9 1) (:end 9 7)) ((:name "listRev") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 9 9) (:end 9 9)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 9 11) (:end 9 11)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 9 12) (:end 9 15)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 3) (:end 10 4)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 6) (:end 10 7)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 9) (:end 10 10)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 3) (:end 11 3)) ((:decor :keyword)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 4) (:end 11 4)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 6) (:end 11 7)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 9) (:end 11 10)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 11) (:end 11 11)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 13) (:end 11 14)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 16) (:end 11 22)) ((:name "listRev") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 24) (:end 11 25)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 27) (:end 11 28)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 30) (:end 11 30)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 31) (:end 11 31)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 32) (:end 11 32)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 1) (:end 13 11)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 13) (:end 13 13)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 15) (:end 13 15)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 16) (:end 13 16)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 17) (:end 13 17)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 19) (:end 13 19)) ((:name "q") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 21) (:end 13 21)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 23) (:end 13 23)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 23) (:end 13 23)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 25) (:end 13 26)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 28) (:end 13 31)) ((:name "Bool") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 28) (:end 13 31)) ((:name "Bool") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 32) (:end 13 32)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 34) (:end 13 35)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 37) (:end 13 40)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 42) (:end 13 42)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 44) (:end 13 45)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 47) (:end 13 50)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 52) (:end 13 52)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 14 1) (:end 14 11)) ((:name "listFilter2") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 14 13) (:end 14 13)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 14 15) (:end 14 15)) ((:name "q") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 14 17) (:end 14 18)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 15 3) (:end 15 3)) ((:decor :keyword)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 15 8) (:end 15 8)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 15 10) (:end 15 11)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 15 13) (:end 15 14)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 17 8) (:end 17 10)) ((:decor :keyword)))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 17 12) (:end 17 15)) ((:name "True") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 17 17) (:end 17 17)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 17 19) (:end 17 19)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 17 21) (:end 17 21)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 18 10) (:end 18 10)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 18 12) (:end 18 16)) ((:name "False") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 18 18) (:end 18 19)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 18 21) (:end 18 22)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 8) (:end 20 11)) ((:name "True") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 13) (:end 20 14)) ((:decor :keyword)))))) 1) +0000e0(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 16) (:end 20 19)) ((:name "pure") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 21) (:end 20 21)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 22) (:end 20 22)) ((:name "q") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 24) (:end 20 24)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 25) (:end 20 25)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 21 10) (:end 21 10)) ((:decor :keyword)))))) 1) +0000d9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 21 12) (:end 21 16)) ((:name "False") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 21 18) (:end 21 19)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 21 21) (:end 21 22)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 22 8) (:end 22 11)) ((:name "pure") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 22 13) (:end 22 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expected8 b/tests/ideMode/ideMode005/expected8 new file mode 100644 index 000000000..fd82a147b --- /dev/null +++ b/tests/ideMode/ideMode005/expected8 @@ -0,0 +1,80 @@ +000018(:protocol-version 2 0) +000042(:write-string "1/1: Building RecordUpdate (RecordUpdate.idr)" 1) +000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 1) (:end 3 6)) ((:decor :keyword)))))) 1) +000075(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 8) (:end 3 17)) ((:decor :type)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 19) (:end 3 23)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 3) (:end 4 6)) ((:decor :function)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 8) (:end 4 8)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 10) (:end 4 15)) ((:decor :type)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 5 3) (:end 5 6)) ((:decor :function)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 5 8) (:end 5 8)) ((:decor :keyword)))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 5 10) (:end 5 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 5 10) (:end 5 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000dc(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 5 10) (:end 5 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 1) (:end 7 7)) ((:decor :function)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 9) (:end 7 9)) ((:decor :keyword)))))) 1) +0000e2(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 11) (:end 7 20)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 22) (:end 7 23)) ((:decor :keyword)))))) 1) +0000e2(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 25) (:end 7 34)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 1) (:end 8 7)) ((:name "bigMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 9) (:end 8 9)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 11) (:end 8 16)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 18) (:end 8 18)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 20) (:end 8 23)) ((:decor :function)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 25) (:end 8 26)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 28) (:end 8 28)) ((:decor :keyword)))))) 1) +0000e6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 29) (:end 8 30)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 32) (:end 8 37)) ((:decor :data)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 38) (:end 8 38)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 18) (:end 9 18)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 20) (:end 9 23)) ((:decor :function)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 25) (:end 9 25)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 27) (:end 9 28)) ((:decor :data)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 10 18) (:end 10 18)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 1) (:end 12 9)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 11) (:end 12 11)) ((:decor :keyword)))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 13) (:end 12 22)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 24) (:end 12 25)) ((:decor :keyword)))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 27) (:end 12 36)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e5(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 1) (:end 13 9)) ((:name "smallMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 11) (:end 13 11)) ((:decor :keyword)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 13) (:end 13 13)) ((:decor :keyword)))))) 1) +00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 15) (:end 13 18)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 20) (:end 13 21)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 23) (:end 13 23)) ((:decor :data)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 25) (:end 13 25)) ((:decor :keyword)))))) 1) +0000e1(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 27) (:end 13 27)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e5(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 13 29) (:end 13 35)) ((:name "bigMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 1) (:end 16 6)) ((:decor :keyword)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 8) (:end 16 11)) ((:decor :type)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 13) (:end 16 17)) ((:decor :keyword)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 3) (:end 17 12)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 14) (:end 17 14)) ((:decor :keyword)))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 16) (:end 17 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 16) (:end 17 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 16) (:end 17 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 18 3) (:end 18 9)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 18 11) (:end 18 11)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 18 13) (:end 18 18)) ((:decor :type)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 1) (:end 20 10)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 12) (:end 20 12)) ((:decor :keyword)))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 14) (:end 20 17)) ((:name "Text") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 19) (:end 20 20)) ((:decor :keyword)))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 22) (:end 20 25)) ((:name "Text") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e7(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 1) (:end 21 10)) ((:name "setArial10") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 12) (:end 21 12)) ((:decor :keyword)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 14) (:end 21 14)) ((:decor :keyword)))))) 1) +00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 16) (:end 21 25)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 26) (:end 21 27)) ((:decor :keyword)))))) 1) +00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 28) (:end 21 31)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 33) (:end 21 34)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 36) (:end 21 42)) ((:decor :data)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 22 14) (:end 22 14)) ((:decor :keyword)))))) 1) +00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 22 16) (:end 22 25)) ((:decor :function)))))) 1) +00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 22 26) (:end 22 30)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 22 32) (:end 22 33)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 22 35) (:end 22 36)) ((:decor :data)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 23 14) (:end 23 14)) ((:decor :keyword)))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expected9 b/tests/ideMode/ideMode005/expected9 new file mode 100644 index 000000000..2bea062c5 --- /dev/null +++ b/tests/ideMode/ideMode005/expected9 @@ -0,0 +1,327 @@ +000018(:protocol-version 2 0) +000032(:write-string "1/1: Building With (With.idr)" 1) +0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 3) (:end 5 3)) ((:name "f") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 5) (:end 5 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 3) (:end 6 3)) ((:name "f") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 5) (:end 6 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 3) (:end 11 3)) ((:name "g") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 6) (:end 11 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 8) (:end 11 9)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 11) (:end 11 12)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 3) (:end 15 8)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 10) (:end 15 10)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 5) (:end 16 10)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 12) (:end 16 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 12) (:end 16 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 16) (:end 16 16)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 5) (:end 17 10)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 12) (:end 17 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 12) (:end 17 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 16) (:end 17 16)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 7) (:end 18 12)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 14) (:end 18 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 14) (:end 18 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 14) (:end 18 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 18) (:end 18 18)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 22) (:end 18 22)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 24) (:end 18 24)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 7) (:end 19 12)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 14) (:end 19 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 14) (:end 19 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 14) (:end 19 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 18) (:end 19 18)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 22) (:end 19 22)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 24) (:end 19 24)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 9) (:end 20 14)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 16) (:end 20 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 16) (:end 20 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 16) (:end 20 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 16) (:end 20 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 20) (:end 20 20)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 24) (:end 20 24)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 26) (:end 20 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 26) (:end 20 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 30) (:end 20 30)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 32) (:end 20 32)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 9) (:end 21 14)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 16) (:end 21 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 16) (:end 21 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 16) (:end 21 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 16) (:end 21 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 20) (:end 21 20)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 24) (:end 21 24)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 26) (:end 21 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 26) (:end 21 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 30) (:end 21 30)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 32) (:end 21 32)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 22 3) (:end 22 8)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 22 10) (:end 22 10)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 3) (:end 29 10)) ((:name "someNats") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 12) (:end 29 12)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 5) (:end 30 12)) ((:name "someNats") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 14) (:end 30 14)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 14) (:end 30 14)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 21) (:end 30 26)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 28) (:end 30 28)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 1) (:end 3 1)) ((:decor :function)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 3) (:end 3 3)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 5) (:end 3 5)) ((:decor :keyword)))))) 1) +0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 6) (:end 3 6)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 8) (:end 3 8)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 10) (:end 3 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 13) (:end 3 13)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 1) +00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 18) (:end 3 18)) ((:decor :type)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 19) (:end 3 19)) ((:decor :bound)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 21) (:end 3 21)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 23) (:end 3 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 23) (:end 3 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 27) (:end 3 28)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 30) (:end 3 30)) ((:decor :bound)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 32) (:end 3 32)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 34) (:end 3 36)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 34) (:end 3 36)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 38) (:end 3 39)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 41) (:end 3 41)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 43) (:end 3 43)) ((:decor :keyword)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 45) (:end 3 45)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 47) (:end 3 47)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 49) (:end 3 49)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 50) (:end 3 50)) ((:decor :type)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 1) (:end 4 1)) ((:name "f") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 3) (:end 4 3)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 5) (:end 4 8)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 10) (:end 4 10)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 10) (:end 4 10)) ((:decor :keyword)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 11) (:end 4 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 11) (:end 4 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 11) (:end 4 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 13) (:end 4 13)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 13) (:end 4 13)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 13) (:end 4 13)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 15) (:end 4 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 15) (:end 4 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 15) (:end 4 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 16) (:end 4 16)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 18) (:end 4 22)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 24) (:end 4 25)) ((:decor :bound)))))) 1) +0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 5) (:end 5 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 7) (:end 5 7)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 9) (:end 5 9)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 15) (:end 5 15)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 18) (:end 5 18)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 20) (:end 5 21)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 23) (:end 5 23)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 25) (:end 5 26)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 28) (:end 5 30)) ((:name "sym") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 32) (:end 5 33)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 5) (:end 6 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 7) (:end 6 7)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 9) (:end 6 9)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 10) (:end 6 10)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 12) (:end 6 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 13) (:end 6 13)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 15) (:end 6 15)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 18) (:end 6 18)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 20) (:end 6 20)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 22) (:end 6 23)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 25) (:end 6 25)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 27) (:end 6 28)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 30) (:end 6 32)) ((:name "sym") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 6 34) (:end 6 35)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 1) (:end 8 1)) ((:decor :function)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 3) (:end 8 3)) ((:decor :keyword)))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 5) (:end 8 8)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 10) (:end 8 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 12) (:end 8 13)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 15) (:end 8 17)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 1) (:end 9 1)) ((:name "g") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 3) (:end 9 4)) ((:name "Nil") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 6) (:end 9 6)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 8) (:end 9 8)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cd(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 1) (:end 10 1)) ((:name "g") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 3) (:end 10 3)) ((:decor :keyword)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 4) (:end 10 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 6) (:end 10 7)) ((:name "::") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 9) (:end 10 10)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 11) (:end 10 11)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 13) (:end 10 16)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 18) (:end 10 18)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 18) (:end 10 18)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 19) (:end 10 20)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 19) (:end 10 20)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 22) (:end 10 23)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 22) (:end 10 23)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 25) (:end 10 26)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 25) (:end 10 26)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 27) (:end 10 27)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 5) (:end 11 5)) ((:decor :keyword)))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 6) (:end 11 6)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 11) (:end 11 12)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 13) (:end 11 13)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 15) (:end 11 15)) ((:decor :keyword)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 17) (:end 11 20)) ((:name "asas") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 22) (:end 11 22)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 24) (:end 11 24)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 1) (:end 13 6)) ((:decor :function)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 8) (:end 13 8)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 10) (:end 13 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 14) (:end 13 15)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 17) (:end 13 19)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 1) (:end 14 6)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 8) (:end 14 8)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 10) (:end 14 13)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 15) (:end 14 15)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 15) (:end 14 15)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 16) (:end 14 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 16) (:end 14 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 17) (:end 14 17)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 10) (:end 15 10)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 12) (:end 15 12)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 14) (:end 15 14)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 16) (:end 15 19)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 21) (:end 15 21)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 21) (:end 15 21)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 22) (:end 15 22)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 22) (:end 15 22)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 24) (:end 15 24)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 24) (:end 15 24)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 26) (:end 15 26)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 26) (:end 15 26)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 27) (:end 15 27)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 12) (:end 16 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 14) (:end 16 14)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 18) (:end 16 18)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 20) (:end 16 20)) ((:decor :data)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 22) (:end 16 22)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 24) (:end 16 24)) ((:decor :data)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 12) (:end 17 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 14) (:end 17 14)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 18) (:end 17 18)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 20) (:end 17 20)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 22) (:end 17 22)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 24) (:end 17 27)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 29) (:end 17 29)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 29) (:end 17 29)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 30) (:end 17 30)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 30) (:end 17 30)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 32) (:end 17 32)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 32) (:end 17 32)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 34) (:end 17 34)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 34) (:end 17 34)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 35) (:end 17 35)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 14) (:end 18 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 16) (:end 18 16)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 20) (:end 18 20)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 24) (:end 18 24)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 26) (:end 18 26)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 28) (:end 18 28)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 30) (:end 18 30)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 32) (:end 18 32)) ((:decor :data)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 14) (:end 19 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 16) (:end 19 16)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 20) (:end 19 20)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 24) (:end 19 24)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 26) (:end 19 26)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 28) (:end 19 28)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 30) (:end 19 30)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 32) (:end 19 35)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 37) (:end 19 37)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 37) (:end 19 37)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 38) (:end 19 38)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 38) (:end 19 38)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 40) (:end 19 40)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 40) (:end 19 40)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 42) (:end 19 42)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 42) (:end 19 42)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 43) (:end 19 43)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 16) (:end 20 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 18) (:end 20 18)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 22) (:end 20 22)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 26) (:end 20 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 28) (:end 20 28)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 32) (:end 20 32)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 34) (:end 20 34)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 36) (:end 20 36)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 38) (:end 20 38)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 40) (:end 20 40)) ((:decor :data)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 16) (:end 21 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 18) (:end 21 18)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 22) (:end 21 22)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 26) (:end 21 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 28) (:end 21 28)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 32) (:end 21 32)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 34) (:end 21 34)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 36) (:end 21 36)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 38) (:end 21 38)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 40) (:end 21 40)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 42) (:end 21 42)) ((:decor :data)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 22 10) (:end 22 10)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 22 12) (:end 22 12)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 22 14) (:end 22 14)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 22 16) (:end 22 16)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 22 18) (:end 22 18)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 22 20) (:end 22 20)) ((:decor :data)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 1) (:end 24 4)) ((:decor :keyword)))))) 1) +00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 6) (:end 24 9)) ((:decor :type)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 11) (:end 24 11)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 13) (:end 24 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 17) (:end 24 18)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 20) (:end 24 23)) ((:decor :type)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 25) (:end 24 29)) ((:decor :keyword)))))) 1) +00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 3) (:end 25 8)) ((:decor :data)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 10) (:end 25 10)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 12) (:end 25 12)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 13) (:end 25 13)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 15) (:end 25 15)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 17) (:end 25 19)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 20) (:end 25 20)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 22) (:end 25 23)) ((:decor :keyword)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 25) (:end 25 28)) ((:name "ANat") (:namespace "With") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 25 30) (:end 25 30)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 1) (:end 27 8)) ((:decor :function)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 10) (:end 27 10)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 12) (:end 27 14)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 16) (:end 27 17)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 19) (:end 27 21)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 1) (:end 28 8)) ((:name "someNats") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 10) (:end 28 10)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 12) (:end 28 15)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 17) (:end 28 17)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 17) (:end 28 17)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 18) (:end 28 23)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 18) (:end 28 23)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 25) (:end 28 25)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 25) (:end 28 25)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 26) (:end 28 26)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 12) (:end 29 12)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 14) (:end 29 14)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 16) (:end 29 16)) ((:decor :bound)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 17) (:end 29 17)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 18) (:end 29 18)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 19) (:end 29 24)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 26) (:end 29 26)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 27) (:end 29 27)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 29) (:end 29 32)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 34) (:end 29 34)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 34) (:end 29 34)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 35) (:end 29 40)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 35) (:end 29 40)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 42) (:end 29 42)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 42) (:end 29 42)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 43) (:end 29 43)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 16) (:end 30 16)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 18) (:end 30 18)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 18) (:end 30 18)) ((:decor :bound)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 19) (:end 30 19)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 20) (:end 30 20)) ((:decor :keyword)))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 28) (:end 30 28)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 29) (:end 30 29)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 31) (:end 30 31)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 33) (:end 30 38)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 40) (:end 30 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 42) (:end 30 42)) ((:decor :keyword)))))) 1) +0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 30 44) (:end 30 44)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expectedA b/tests/ideMode/ideMode005/expectedA new file mode 100644 index 000000000..f96ce1654 --- /dev/null +++ b/tests/ideMode/ideMode005/expectedA @@ -0,0 +1,27 @@ +000018(:protocol-version 2 0) +000036(:write-string "1/1: Building Ranges (Ranges.idr)" 1) +000072(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 1) (:end 1 5)) ((:decor :function)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 7) (:end 1 7)) ((:decor :keyword)))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 9) (:end 1 12)) ((:name "List") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 14) (:end 1 16)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 2 1) (:end 2 5)) ((:name "hours") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 2 7) (:end 2 7)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 2 9) (:end 2 9)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 2 10) (:end 2 10)) ((:decor :data)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 2 11) (:end 2 12)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 2 13) (:end 2 14)) ((:decor :data)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 2 15) (:end 2 15)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 1) (:end 4 4)) ((:decor :function)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 6) (:end 4 6)) ((:decor :keyword)))))) 1) +0000df(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 8) (:end 4 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 4 15) (:end 4 17)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 5 1) (:end 5 4)) ((:name "nats") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 5 6) (:end 5 6)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 5 8) (:end 5 8)) ((:decor :keyword)))))) 1) +00006e(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 5 9) (:end 5 9)) ((:decor :data)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 5 10) (:end 5 10)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 5 11) (:end 5 11)) ((:decor :data)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 5 12) (:end 5 13)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 5 14) (:end 5 14)) ((:decor :keyword)))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expectedB b/tests/ideMode/ideMode005/expectedB new file mode 100644 index 000000000..2a9dcd38a --- /dev/null +++ b/tests/ideMode/ideMode005/expectedB @@ -0,0 +1,55 @@ +000018(:protocol-version 2 0) +000046(:write-string "1/1: Building StringLiterals (StringLiterals.idr)" 1) +000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 1) (:end 3 5)) ((:decor :function)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 7) (:end 3 7)) ((:decor :keyword)))))) 1) +0000de(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 9) (:end 3 18)) ((:name "FromString") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 20) (:end 3 20)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 22) (:end 3 23)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 25) (:end 3 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000e3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 4 1) (:end 4 5)) ((:name "hello") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 4 7) (:end 4 7)) ((:decor :keyword)))))) 1) +000077(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 4 9) (:end 4 15)) ((:decor :data)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 1) (:end 6 9)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 11) (:end 6 11)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 13) (:end 6 18)) ((:decor :type)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 20) (:end 6 21)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 23) (:end 6 28)) ((:decor :type)))))) 1) +0000e7(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 1) (:end 7 9)) ((:name "helloName") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 11) (:end 7 14)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 16) (:end 7 16)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 18) (:end 7 48)) ((:decor :data)))))) 1) +0000e5(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 21) (:end 7 25)) ((:name "hello") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 27) (:end 7 27)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 28) (:end 7 28)) ((:decor :bound)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 30) (:end 7 30)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 32) (:end 7 37)) ((:decor :type)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 38) (:end 7 38)) ((:decor :keyword)))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 7 43) (:end 7 46)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 1) (:end 9 11)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 13) (:end 9 13)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 15) (:end 9 20)) ((:decor :type)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 22) (:end 9 23)) ((:decor :keyword)))))) 1) +000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 25) (:end 9 30)) ((:decor :type)))))) 1) +0000ec(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 10 1) (:end 10 11)) ((:name "welcomeName") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 10 13) (:end 10 16)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 10 18) (:end 10 18)) ((:decor :keyword)))))) 1) +000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 10 20) (:end 13 5)) ((:decor :data)))))) 1) +0000ea(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 11 5) (:end 11 13)) ((:name "helloName") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 11 15) (:end 11 18)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 1) (:end 15 11)) ((:decor :function)))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 13) (:end 15 13)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 15) (:end 15 20)) ((:decor :type)))))) 1) +0000ec(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 16 1) (:end 16 11)) ((:name "scareQuotes") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 16 13) (:end 16 13)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 16 15) (:end 16 25)) ((:decor :data)))))) 1) +00007c(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 18 1) (:end 18 4)) ((:decor :function)))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 18 6) (:end 18 6)) ((:decor :keyword)))))) 1) +0000ec(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 18 8) (:end 18 33)) ((:name "scareQuotes") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 18 35) (:end 18 35)) ((:decor :keyword)))))) 1) +00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 18 37) (:end 18 47)) ((:decor :data)))))) 1) +0000e4(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 19 1) (:end 19 4)) ((:name "test") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 19 6) (:end 19 6)) ((:decor :keyword)))))) 1) +0000da(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 19 8) (:end 19 11)) ((:name "Refl") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expectedC b/tests/ideMode/ideMode005/expectedC new file mode 100644 index 000000000..81297841e --- /dev/null +++ b/tests/ideMode/ideMode005/expectedC @@ -0,0 +1,81 @@ +000018(:protocol-version 2 0) +000038(:write-string "1/1: Building WithApp (WithApp.idr)" 1) +0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 3) (:end 11 4)) ((:name "id") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 3) (:end 12 4)) ((:name "id") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 16) (:end 12 17)) ((:name "id") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 1) (:end 1 4)) ((:decor :keyword)))))) 1) +000070(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 6) (:end 1 10)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 12) (:end 1 12)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 14) (:end 1 16)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 18) (:end 1 19)) ((:decor :keyword)))))) 1) +000071(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 21) (:end 1 24)) ((:decor :type)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 26) (:end 1 30)) ((:decor :keyword)))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 3) (:end 2 3)) ((:decor :data)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 5) (:end 2 5)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 7) (:end 2 11)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 13) (:end 2 13)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +00006f(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 3) (:end 3 3)) ((:decor :data)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 5) (:end 3 5)) ((:decor :keyword)))))) 1) +0000cf(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 7) (:end 3 11)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 13) (:end 3 13)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 18) (:end 3 22)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 24) (:end 3 24)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 25) (:end 3 25)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 27) (:end 3 27)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 3 28) (:end 3 28)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 1) (:end 5 4)) ((:decor :function)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 6) (:end 5 6)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 8) (:end 5 8)) ((:decor :keyword)))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 9) (:end 5 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 10) (:end 5 10)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 12) (:end 5 14)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 15) (:end 5 15)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 17) (:end 5 18)) ((:decor :keyword)))))) 1) +0000d0(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 20) (:end 5 24)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 26) (:end 5 26)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 1) (:end 6 4)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 6) (:end 6 6)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 8) (:end 6 8)) ((:decor :keyword)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 10) (:end 6 10)) ((:name "Z") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 1) (:end 7 4)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 6) (:end 7 6)) ((:decor :keyword)))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 7) (:end 7 7)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 9) (:end 7 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 10) (:end 7 10)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 12) (:end 7 12)) ((:decor :keyword)))))) 1) +0000cc(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 14) (:end 7 14)) ((:name "S") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 16) (:end 7 16)) ((:decor :keyword)))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 17) (:end 7 20)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 22) (:end 7 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 7 23) (:end 7 23)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 1) (:end 9 2)) ((:decor :function)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 4) (:end 9 4)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 6) (:end 9 8)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 13) (:end 9 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 1) (:end 10 2)) ((:name "id") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 4) (:end 10 4)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 6) (:end 10 9)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 11) (:end 10 11)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 11) (:end 10 11)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 12) (:end 10 15)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 12) (:end 10 15)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 17) (:end 10 17)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 17) (:end 10 17)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 18) (:end 10 18)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 6) (:end 11 6)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 8) (:end 11 8)) ((:decor :keyword)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 10) (:end 11 10)) ((:name "Z") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 12) (:end 11 12)) ((:decor :keyword)))))) 1) +0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 14) (:end 11 14)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 6) (:end 12 6)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 8) (:end 12 8)) ((:decor :keyword)))))) 1) +0000ce(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 10) (:end 12 10)) ((:name "S") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 12) (:end 12 12)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 14) (:end 12 14)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 19) (:end 12 19)) ((:decor :keyword)))))) 1) +000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 21) (:end 12 21)) ((:decor :keyword)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 12 23) (:end 12 23)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/expectedD b/tests/ideMode/ideMode005/expectedD new file mode 100644 index 000000000..fecc23758 --- /dev/null +++ b/tests/ideMode/ideMode005/expectedD @@ -0,0 +1,70 @@ +000018(:protocol-version 2 0) +000038(:write-string "1/1: Building Rewrite (Rewrite.idr)" 1) +000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 1) (:end 1 9)) ((:decor :function)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 11) (:end 1 11)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 13) (:end 1 13)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 14) (:end 1 14)) ((:decor :keyword)))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 16) (:end 1 17)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 19) (:end 1 19)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 21) (:end 1 21)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 23) (:end 1 25)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 27) (:end 1 27)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 28) (:end 1 28)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 30) (:end 1 31)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 33) (:end 1 33)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 35) (:end 1 36)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 38) (:end 1 38)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d6(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 1) (:end 2 9)) ((:name "transport") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 11) (:end 2 12)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 14) (:end 2 14)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 16) (:end 2 16)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 18) (:end 2 24)) ((:decor :keyword)))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 26) (:end 2 28)) ((:name "sym") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 30) (:end 2 31)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 33) (:end 2 34)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 2 36) (:end 2 36)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 1) (:end 4 6)) ((:decor :function)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 8) (:end 4 8)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 10) (:end 4 10)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 11) (:end 4 11)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 13) (:end 4 13)) ((:name "_") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 15) (:end 4 15)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 17) (:end 4 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 19) (:end 4 21)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 23) (:end 4 23)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 24) (:end 4 24)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 26) (:end 4 27)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 29) (:end 4 29)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 30) (:end 4 30)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 32) (:end 4 32)) ((:name "_") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 34) (:end 4 34)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 36) (:end 4 36)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 38) (:end 4 40)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 42) (:end 4 42)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 43) (:end 4 43)) ((:decor :keyword)))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 45) (:end 4 46)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 48) (:end 4 48)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 50) (:end 4 52)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 54) (:end 4 54)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d3(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 5 1) (:end 5 6)) ((:name "nested") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 5 8) (:end 5 10)) ((:name "eq1") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 5 12) (:end 5 14)) ((:name "eq2") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 5 16) (:end 5 16)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 6 3) (:end 6 9)) ((:decor :keyword)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 6 11) (:end 6 13)) ((:name "eq1") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 6 15) (:end 6 16)) ((:decor :keyword)))))) 1) +000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 5) (:end 7 11)) ((:decor :keyword)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 13) (:end 7 15)) ((:name "eq2") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 17) (:end 7 18)) ((:decor :keyword)))))) 1) +000072(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 7) (:end 8 9)) ((:decor :keyword)))))) 1) +0000cb(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 11) (:end 8 13)) ((:name "prf") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 15) (:end 8 15)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 17) (:end 8 17)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 19) (:end 8 21)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 23) (:end 8 23)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 25) (:end 8 26)) ((:decor :keyword)))))) 1) +0000d2(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 28) (:end 8 31)) ((:name "Refl") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 33) (:end 8 34)) ((:decor :keyword)))))) 1) +0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 9 7) (:end 9 9)) ((:name "prf") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1) +000015(:return (:ok ()) 1) +Alas the file is done, aborting diff --git a/tests/ideMode/ideMode005/input1 b/tests/ideMode/ideMode005/input1 new file mode 100644 index 000000000..45e58d8d4 --- /dev/null +++ b/tests/ideMode/ideMode005/input1 @@ -0,0 +1 @@ +00001d((:load-file "Syntax.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input2 b/tests/ideMode/ideMode005/input2 new file mode 100644 index 000000000..4ece153e2 --- /dev/null +++ b/tests/ideMode/ideMode005/input2 @@ -0,0 +1 @@ +000020((:load-file "Interface.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input3 b/tests/ideMode/ideMode005/input3 new file mode 100644 index 000000000..2fd7d3b40 --- /dev/null +++ b/tests/ideMode/ideMode005/input3 @@ -0,0 +1 @@ +000021((:load-file "SimpleData.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input4 b/tests/ideMode/ideMode005/input4 new file mode 100644 index 000000000..6ae6c86aa --- /dev/null +++ b/tests/ideMode/ideMode005/input4 @@ -0,0 +1 @@ +000021((:load-file "Ambiguity.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input5 b/tests/ideMode/ideMode005/input5 new file mode 100644 index 000000000..169294bae --- /dev/null +++ b/tests/ideMode/ideMode005/input5 @@ -0,0 +1 @@ +00001e((:load-file "Rainbow.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input6 b/tests/ideMode/ideMode005/input6 new file mode 100644 index 000000000..25b911d28 --- /dev/null +++ b/tests/ideMode/ideMode005/input6 @@ -0,0 +1 @@ +000025((:load-file "Implementation.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input7 b/tests/ideMode/ideMode005/input7 new file mode 100644 index 000000000..c5616c1f5 --- /dev/null +++ b/tests/ideMode/ideMode005/input7 @@ -0,0 +1 @@ +00001b((:load-file "Case.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input8 b/tests/ideMode/ideMode005/input8 new file mode 100644 index 000000000..1c4edf5f3 --- /dev/null +++ b/tests/ideMode/ideMode005/input8 @@ -0,0 +1 @@ +000023((:load-file "RecordUpdate.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/input9 b/tests/ideMode/ideMode005/input9 new file mode 100644 index 000000000..675b4d8d3 --- /dev/null +++ b/tests/ideMode/ideMode005/input9 @@ -0,0 +1 @@ +00001b((:load-file "With.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/inputA b/tests/ideMode/ideMode005/inputA new file mode 100644 index 000000000..a297280f6 --- /dev/null +++ b/tests/ideMode/ideMode005/inputA @@ -0,0 +1 @@ +00001d((:load-file "Ranges.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/inputB b/tests/ideMode/ideMode005/inputB new file mode 100644 index 000000000..185c188da --- /dev/null +++ b/tests/ideMode/ideMode005/inputB @@ -0,0 +1 @@ +000025((:load-file "StringLiterals.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/inputC b/tests/ideMode/ideMode005/inputC new file mode 100644 index 000000000..9db790bfb --- /dev/null +++ b/tests/ideMode/ideMode005/inputC @@ -0,0 +1 @@ +00001e((:load-file "WithApp.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/inputD b/tests/ideMode/ideMode005/inputD new file mode 100644 index 000000000..482429e86 --- /dev/null +++ b/tests/ideMode/ideMode005/inputD @@ -0,0 +1 @@ +00001e((:load-file "Rewrite.idr") 1) \ No newline at end of file diff --git a/tests/ideMode/ideMode005/regenerate b/tests/ideMode/ideMode005/regenerate new file mode 100755 index 000000000..02dca7926 --- /dev/null +++ b/tests/ideMode/ideMode005/regenerate @@ -0,0 +1,19 @@ +rm -rf build + +$1 --no-color --console-width 0 --ide-mode < input1 > expected1 +$1 --no-color --console-width 0 --ide-mode < input2 > expected2 +$1 --no-color --console-width 0 --ide-mode < input3 > expected3 +$1 --no-color --console-width 0 --ide-mode < input4 > expected4 +$1 --no-color --console-width 0 --ide-mode < input5 > expected5 +$1 --no-color --console-width 0 --ide-mode < input6 > expected6 +$1 --no-color --console-width 0 --ide-mode < input7 > expected7 +$1 --no-color --console-width 0 --ide-mode < input8 > expected8 +$1 --no-color --console-width 0 --ide-mode < input9 > expected9 +$1 --no-color --console-width 0 --ide-mode < inputA > expectedA +$1 --no-color --console-width 0 --ide-mode < inputB > expectedB +$1 --no-color --console-width 0 --ide-mode < inputC > expectedC +$1 --no-color --console-width 0 --ide-mode < inputD > expectedD + +rm -f expected output* +touch expected +rm -rf build \ No newline at end of file diff --git a/tests/ideMode/ideMode005/run b/tests/ideMode/ideMode005/run new file mode 100755 index 000000000..433706c21 --- /dev/null +++ b/tests/ideMode/ideMode005/run @@ -0,0 +1,42 @@ +rm -f output* + +$1 --no-color --console-width 0 --ide-mode < input1 > output1 +diff expected1 output1 >> output + +$1 --no-color --console-width 0 --ide-mode < input2 > output2 +diff expected2 output2 >> output + +$1 --no-color --console-width 0 --ide-mode < input3 > output3 +diff expected3 output3 >> output + +$1 --no-color --console-width 0 --ide-mode < input4 > output4 +diff expected4 output4 >> output + +$1 --no-color --console-width 0 --ide-mode < input5 > output5 +diff expected5 output5 >> output + +$1 --no-color --console-width 0 --ide-mode < input6 > output6 +diff expected6 output6 >> output + +$1 --no-color --console-width 0 --ide-mode < input7 > output7 +diff expected7 output7 >> output + +$1 --no-color --console-width 0 --ide-mode < input8 > output8 +diff expected8 output8 >> output + +$1 --no-color --console-width 0 --ide-mode < input9 > output9 +diff expected9 output9 >> output + +$1 --no-color --console-width 0 --ide-mode < inputA > outputA +diff expectedA outputA >> output + +$1 --no-color --console-width 0 --ide-mode < inputB > outputB +diff expectedB outputB >> output + +$1 --no-color --console-width 0 --ide-mode < inputC > outputC +diff expectedC outputC >> output + +$1 --no-color --console-width 0 --ide-mode < inputD > outputD +diff expectedD outputD >> output + +rm -rf build diff --git a/tests/idris2/reflection001/expected b/tests/idris2/reflection001/expected index 962820b8e..a75cfd9a8 100644 --- a/tests/idris2/reflection001/expected +++ b/tests/idris2/reflection001/expected @@ -30,9 +30,9 @@ quote.idr:37:1--37:21 37 | %runElab noExtension ^^^^^^^^^^^^^^^^^^^^ -Main> IApp (MkFC "quote.idr" (3, 12) (3, 23)) (IApp (MkFC "quote.idr" (3, 12) (3, 23)) (IVar (MkFC "quote.idr" (3, 12) (3, 23)) (UN "+")) (IApp (MkFC "quote.idr" (6, 13) (6, 14)) (IVar (MkFC "quote.idr" (6, 13) (6, 14)) (UN "fromInteger")) (IPrimVal (MkFC "quote.idr" (6, 13) (6, 14)) (BI 3)))) (IApp (MkFC "quote.idr" (6, 18) (6, 19)) (IVar (MkFC "quote.idr" (6, 18) (6, 19)) (UN "fromInteger")) (IPrimVal (MkFC "quote.idr" (6, 18) (6, 19)) (BI 4))) -Main> IApp (MkFC "quote.idr" (3, 12) (3, 23)) (IApp (MkFC "quote.idr" (3, 12) (3, 23)) (IVar (MkFC "quote.idr" (3, 12) (3, 23)) (UN "+")) (IVar (MkFC "(interactive)" (0, 6) (0, 10)) (UN "True"))) (IVar (MkFC "(interactive)" (0, 14) (0, 19)) (UN "False")) -Main> ILocal (MkFC "quote.idr" (10, 12) (11, 29)) [IClaim (MkFC "quote.idr" (10, 12) (10, 28)) MW Private [] (MkTy (MkFC "quote.idr" (10, 12) (10, 15)) (MkFC "quote.idr" (10, 12) (10, 15)) (UN "xfn") (IPi (MkFC "quote.idr" (10, 18) (10, 21)) MW ExplicitArg Nothing (IPrimVal (MkFC "quote.idr" (10, 18) (10, 21)) IntType) (IPrimVal (MkFC "quote.idr" (10, 25) (10, 28)) IntType))), IDef (MkFC "quote.idr" (11, 12) (11, 29)) (UN "xfn") [PatClause (MkFC "quote.idr" (11, 12) (11, 29)) (IApp (MkFC "quote.idr" (11, 12) (11, 19)) (IVar (MkFC "quote.idr" (11, 12) (11, 15)) (UN "xfn")) (IBindVar (MkFC "quote.idr" (11, 16) (11, 19)) "var")) (IApp (MkFC "quote.idr" (11, 22) (11, 29)) (IApp (MkFC "quote.idr" (11, 22) (11, 29)) (IVar (MkFC "quote.idr" (11, 22) (11, 29)) (UN "*")) (IVar (MkFC "quote.idr" (11, 22) (11, 25)) (UN "var"))) (IApp (MkFC "quote.idr" (11, 28) (11, 29)) (IVar (MkFC "quote.idr" (11, 28) (11, 29)) (UN "fromInteger")) (IPrimVal (MkFC "quote.idr" (11, 28) (11, 29)) (BI 2))))]] (IApp (MkFC "quote.idr" (12, 12) (12, 22)) (IVar (MkFC "quote.idr" (12, 12) (12, 15)) (UN "xfn")) (IApp (MkFC "(interactive)" (0, 9) (0, 22)) (IApp (MkFC "(interactive)" (0, 9) (0, 22)) (IVar (MkFC "(interactive)" (0, 9) (0, 12)) (UN "the")) (IPrimVal (MkFC "(interactive)" (0, 13) (0, 16)) IntType)) (IApp (MkFC "(interactive)" (0, 17) (0, 22)) (IVar (MkFC "(interactive)" (0, 17) (0, 22)) (UN "fromInteger")) (IPrimVal (MkFC "(interactive)" (0, 17) (0, 22)) (BI 99994))))) -Main> ILocal (MkFC "quote.idr" (16, 12) (17, 29)) [IClaim (MkFC "quote.idr" (16, 12) (16, 28)) MW Private [] (MkTy (MkFC "quote.idr" (16, 12) (16, 15)) (MkFC "quote.idr" (16, 12) (16, 15)) (UN "xfn") (IPi (MkFC "quote.idr" (16, 18) (16, 21)) MW ExplicitArg Nothing (IPrimVal (MkFC "quote.idr" (16, 18) (16, 21)) IntType) (IPrimVal (MkFC "quote.idr" (16, 25) (16, 28)) IntType))), IDef (MkFC "quote.idr" (17, 12) (17, 29)) (UN "xfn") [PatClause (MkFC "quote.idr" (17, 12) (17, 29)) (IApp (MkFC "quote.idr" (17, 12) (17, 19)) (IVar (MkFC "quote.idr" (17, 12) (17, 15)) (UN "xfn")) (IBindVar (MkFC "quote.idr" (17, 16) (17, 19)) "var")) (IApp (MkFC "quote.idr" (17, 22) (17, 29)) (IApp (MkFC "quote.idr" (17, 22) (17, 29)) (IVar (MkFC "quote.idr" (17, 22) (17, 29)) (UN "*")) (IVar (MkFC "quote.idr" (17, 22) (17, 25)) (UN "var"))) (IApp (MkFC "quote.idr" (17, 28) (17, 29)) (IVar (MkFC "quote.idr" (17, 28) (17, 29)) (UN "fromInteger")) (IPrimVal (MkFC "quote.idr" (17, 28) (17, 29)) (BI 2))))]] (IApp (MkFC "quote.idr" (18, 12) (18, 43)) (IVar (MkFC "quote.idr" (18, 12) (18, 15)) (UN "xfn")) (IPrimVal EmptyFC (I 99994))) +Main> IApp (MkFC "quote.idr" (3, 12) (3, 23)) (IApp (MkFC "quote.idr" (3, 12) (3, 23)) (IVar (MkFC "quote.idr" (3, 17) (3, 18)) (UN "+")) (IApp (MkFC "quote.idr" (6, 13) (6, 14)) (IVar (MkFC "quote.idr" (6, 13) (6, 14)) (UN "fromInteger")) (IPrimVal (MkFC "quote.idr" (6, 13) (6, 14)) (BI 3)))) (IApp (MkFC "quote.idr" (6, 18) (6, 19)) (IVar (MkFC "quote.idr" (6, 18) (6, 19)) (UN "fromInteger")) (IPrimVal (MkFC "quote.idr" (6, 18) (6, 19)) (BI 4))) +Main> IApp (MkFC "quote.idr" (3, 12) (3, 23)) (IApp (MkFC "quote.idr" (3, 12) (3, 23)) (IVar (MkFC "quote.idr" (3, 17) (3, 18)) (UN "+")) (IVar (MkFC "(interactive)" (0, 6) (0, 10)) (UN "True"))) (IVar (MkFC "(interactive)" (0, 14) (0, 19)) (UN "False")) +Main> ILocal (MkFC "quote.idr" (10, 12) (11, 29)) [IClaim (MkFC "quote.idr" (10, 12) (10, 28)) MW Private [] (MkTy (MkFC "quote.idr" (10, 12) (10, 15)) (MkFC "quote.idr" (10, 12) (10, 15)) (UN "xfn") (IPi (MkFC "quote.idr" (10, 18) (10, 21)) MW ExplicitArg Nothing (IPrimVal (MkFC "quote.idr" (10, 18) (10, 21)) IntType) (IPrimVal (MkFC "quote.idr" (10, 25) (10, 28)) IntType))), IDef (MkFC "quote.idr" (11, 12) (11, 29)) (UN "xfn") [PatClause (MkFC "quote.idr" (11, 12) (11, 29)) (IApp (MkFC "quote.idr" (11, 12) (11, 19)) (IVar (MkFC "quote.idr" (11, 12) (11, 15)) (UN "xfn")) (IBindVar (MkFC "quote.idr" (11, 16) (11, 19)) "var")) (IApp (MkFC "quote.idr" (11, 22) (11, 29)) (IApp (MkFC "quote.idr" (11, 22) (11, 29)) (IVar (MkFC "quote.idr" (11, 26) (11, 27)) (UN "*")) (IVar (MkFC "quote.idr" (11, 22) (11, 25)) (UN "var"))) (IApp (MkFC "quote.idr" (11, 28) (11, 29)) (IVar (MkFC "quote.idr" (11, 28) (11, 29)) (UN "fromInteger")) (IPrimVal (MkFC "quote.idr" (11, 28) (11, 29)) (BI 2))))]] (IApp (MkFC "quote.idr" (12, 12) (12, 22)) (IVar (MkFC "quote.idr" (12, 12) (12, 15)) (UN "xfn")) (IApp (MkFC "(interactive)" (0, 9) (0, 22)) (IApp (MkFC "(interactive)" (0, 9) (0, 22)) (IVar (MkFC "(interactive)" (0, 9) (0, 12)) (UN "the")) (IPrimVal (MkFC "(interactive)" (0, 13) (0, 16)) IntType)) (IApp (MkFC "(interactive)" (0, 17) (0, 22)) (IVar (MkFC "(interactive)" (0, 17) (0, 22)) (UN "fromInteger")) (IPrimVal (MkFC "(interactive)" (0, 17) (0, 22)) (BI 99994))))) +Main> ILocal (MkFC "quote.idr" (16, 12) (17, 29)) [IClaim (MkFC "quote.idr" (16, 12) (16, 28)) MW Private [] (MkTy (MkFC "quote.idr" (16, 12) (16, 15)) (MkFC "quote.idr" (16, 12) (16, 15)) (UN "xfn") (IPi (MkFC "quote.idr" (16, 18) (16, 21)) MW ExplicitArg Nothing (IPrimVal (MkFC "quote.idr" (16, 18) (16, 21)) IntType) (IPrimVal (MkFC "quote.idr" (16, 25) (16, 28)) IntType))), IDef (MkFC "quote.idr" (17, 12) (17, 29)) (UN "xfn") [PatClause (MkFC "quote.idr" (17, 12) (17, 29)) (IApp (MkFC "quote.idr" (17, 12) (17, 19)) (IVar (MkFC "quote.idr" (17, 12) (17, 15)) (UN "xfn")) (IBindVar (MkFC "quote.idr" (17, 16) (17, 19)) "var")) (IApp (MkFC "quote.idr" (17, 22) (17, 29)) (IApp (MkFC "quote.idr" (17, 22) (17, 29)) (IVar (MkFC "quote.idr" (17, 26) (17, 27)) (UN "*")) (IVar (MkFC "quote.idr" (17, 22) (17, 25)) (UN "var"))) (IApp (MkFC "quote.idr" (17, 28) (17, 29)) (IVar (MkFC "quote.idr" (17, 28) (17, 29)) (UN "fromInteger")) (IPrimVal (MkFC "quote.idr" (17, 28) (17, 29)) (BI 2))))]] (IApp (MkFC "quote.idr" (18, 12) (18, 43)) (IVar (MkFC "quote.idr" (18, 12) (18, 15)) (UN "xfn")) (IPrimVal EmptyFC (I 99994))) Main> [UN "names", NS (MkNS ["Prelude"]) (UN "+")] Main> Bye for now! diff --git a/tests/idris2/reg007/expected b/tests/idris2/reg007/expected index 5cfb4c238..9d9ce8004 100644 --- a/tests/idris2/reg007/expected +++ b/tests/idris2/reg007/expected @@ -2,11 +2,11 @@ Error: While processing right hand side of dpairWithExtraInfoBad. When unifying [MN 0, MN 0] and [MN 0]. Mismatch between: [MN 0] and []. -Main.idr:27:26--27:73 +Main.idr:27:27--27:36 23 | dpairWithExtraInfoWorks : List (vars : List Name ** Expr vars) 24 | dpairWithExtraInfoWorks = [([MN 0] ** CLocal {x=MN 0} (First {ns=[]}))] 25 | 26 | dpairWithExtraInfoBad : List (vars : List Name ** Expr vars) 27 | dpairWithExtraInfoBad = [([MN 0] ** CLocal {x=MN 0} (First {ns=[MN 0]}))] - ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ + ^^^^^^^^^ diff --git a/tests/idris2/with003/expected b/tests/idris2/with003/expected index 525110f8e..841925794 100644 --- a/tests/idris2/with003/expected +++ b/tests/idris2/with003/expected @@ -26,9 +26,9 @@ Main> [] : List ?a Main> Error: When unifying Vect 0 ?elem and List ?a. Mismatch between: Vect 0 ?elem and List ?a. -(interactive):1:34--1:41 +(interactive):1:40--1:41 1 | :t with [Vect.Nil, Prelude.(::)] [1,2,3] - ^^^^^^^ + ^ Main> the (Maybe Integer) (pure 4) : Maybe Integer Main> Expected 'case', 'if', 'do', application or operator expression. diff --git a/tests/typedd-book/chapter10/expected b/tests/typedd-book/chapter10/expected index 220df863f..a56e3fd42 100644 --- a/tests/typedd-book/chapter10/expected +++ b/tests/typedd-book/chapter10/expected @@ -1,11 +1,11 @@ 1/1: Building DLFail (DLFail.idr) Error: While processing left hand side of describe_list_end. Can't match on ?xs ++ [?x] (Not a constructor application or primitive). -DLFail.idr:3:20--3:29 +DLFail.idr:3:23--3:25 1 | describe_list_end : List Int -> String 2 | describe_list_end [] = "Empty" 3 | describe_list_end (xs ++ [x]) = "Non-empty, initial portion = " ++ show xs - ^^^^^^^^^ + ^^ 1/13: Building DataStore (DataStore.idr) 2/13: Building DescribeList (DescribeList.idr)