Semantic highlighting (#1335)

Co-authored-by: Guillaume ALLAIS <guillaume.allais@ens-lyon.org>
This commit is contained in:
Ohad Kammar 2021-05-10 09:05:43 +01:00 committed by GitHub
parent 416cf6af3e
commit e58bcfc7ef
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
108 changed files with 3618 additions and 1127 deletions

15
.gitignore vendored
View File

@ -17,7 +17,7 @@ idris2docs_venv
/libs/**/build /libs/**/build
/tests/**/build /tests/**/build
/tests/**/output /tests/**/output*
/tests/**/*.so /tests/**/*.so
/tests/**/*.dylib /tests/**/*.dylib
/tests/**/*.dll /tests/**/*.dll
@ -39,9 +39,16 @@ idris2docs_venv
/result /result
# Editor/IDE Related # Editor/IDE Related
*~ # Vim swap file # WARNING:
.\#* # Emacs swap file # do not put comments on the same line as a regex
.vscode/* # VS Code # git seems to ignore the pattern in this case
# Vim swap file
*~
# Emacs swap file
.\#*
# VS Code
.vscode/*
# macOS # macOS
.DS_Store .DS_Store

View File

@ -203,7 +203,17 @@ The following keys are available:
``implicit`` ``implicit``
provides a Boolean value that is True if the region is the name of an implicit argument provides a Boolean value that is True if the region is the name of an implicit argument
``decor`` ``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`` ``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: states that the region refers to a source code location. Its body is a collection of key-value pairs, with the following possibilities:

View File

@ -169,7 +169,6 @@ modules =
Parser.Lexer.Package, Parser.Lexer.Package,
Parser.Lexer.Source, Parser.Lexer.Source,
Parser.Rule.Common,
Parser.Rule.Package, Parser.Rule.Package,
Parser.Rule.Source, Parser.Rule.Source,

View File

@ -14,8 +14,8 @@ FilePos = (Int, Int)
-- the second 'FilePos' indicates the start of the next term. -- the second 'FilePos' indicates the start of the next term.
public export public export
data FC : Type where data FC : Type where
MkFC : String -> FilePos -> FilePos -> FC MkFC : String -> FilePos -> FilePos -> FC
EmptyFC : FC EmptyFC : FC
public export public export
emptyFC : FC emptyFC : FC

View File

@ -463,6 +463,10 @@ export %inline
(<$>) : (a -> b) -> Core a -> Core b (<$>) : (a -> b) -> Core a -> Core b
(<$>) f (MkCore a) = MkCore (map (map f) a) (<$>) f (MkCore a) = MkCore (map (map f) a)
export %inline
(<$) : b -> Core a -> Core b
(<$) = (<$>) . const
export %inline export %inline
ignore : Core a -> Core () ignore : Core a -> Core ()
ignore = map (\ _ => ()) ignore = map (\ _ => ())

View File

@ -25,7 +25,10 @@ FileName = String
||| file or by the compiler. That makes it useful to have the notion of ||| file or by the compiler. That makes it useful to have the notion of
||| `EmptyFC` as part of the type. ||| `EmptyFC` as part of the type.
public export 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 | EmptyFC
||| A version of a file context that cannot be empty ||| A version of a file context that cannot be empty
@ -45,8 +48,22 @@ justFC (fname, start, end) = MkFC fname start end
export export
isNonEmptyFC : FC -> Maybe NonEmptyFC isNonEmptyFC : FC -> Maybe NonEmptyFC
isNonEmptyFC (MkFC fn start end) = Just (fn, start, end) isNonEmptyFC (MkFC fn start end) = Just (fn, start, end)
isNonEmptyFC (MkVirtualFC fn start end) = Just (fn, start, end)
isNonEmptyFC EmptyFC = Nothing 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 export
defaultFC : NonEmptyFC defaultFC : NonEmptyFC
defaultFC = ("", (0, 0), (0, 0)) defaultFC = ("", (0, 0), (0, 0))
@ -140,6 +157,7 @@ mergeFC _ _ = Nothing
export export
Eq FC where Eq FC where
(==) (MkFC n s e) (MkFC n' s' e') = n == n' && s == s' && e == e' (==) (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 (==) EmptyFC EmptyFC = True
(==) _ _ = False (==) _ _ = False
@ -149,6 +167,12 @@ Show FC where
show (MkFC file startPos endPos) = file ++ ":" ++ show (MkFC file startPos endPos) = file ++ ":" ++
showPos startPos ++ "--" ++ showPos startPos ++ "--" ++
showPos endPos 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 export
Pretty FC where Pretty FC where
@ -156,6 +180,6 @@ Pretty FC where
pretty (MkFC file startPos endPos) = pretty file <+> colon pretty (MkFC file startPos endPos) = pretty file <+> colon
<+> prettyPos startPos <+> pretty "--" <+> prettyPos startPos <+> pretty "--"
<+> prettyPos endPos <+> prettyPos endPos
where pretty (MkVirtualFC file startPos endPos) = pretty file <+> colon
prettyPos : FilePos -> Doc ann <+> prettyPos startPos <+> pretty "--"
prettyPos (l, c) = pretty (l + 1) <+> colon <+> pretty (c + 1) <+> prettyPos endPos

View File

@ -17,7 +17,60 @@ import Libraries.Utils.Binary
%default covering %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 public export
record Metadata where record Metadata where
@ -43,26 +96,41 @@ record Metadata where
currentLHS : Maybe ClosedTerm currentLHS : Maybe ClosedTerm
holeLHS : List (Name, ClosedTerm) holeLHS : List (Name, ClosedTerm)
nameLocMap : PosMap (NonEmptyFC, Name) 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 Metadata where
show (MkMetadata apps names tydecls currentLHS holeLHS nameLocMap) show (MkMetadata apps names tydecls currentLHS holeLHS nameLocMap
fname semanticHighlighting semanticAliases)
= "Metadata:\n" ++ = "Metadata:\n" ++
" lhsApps: " ++ show apps ++ "\n" ++ " lhsApps: " ++ show apps ++ "\n" ++
" names: " ++ show names ++ "\n" ++ " names: " ++ show names ++ "\n" ++
" type declarations: " ++ show tydecls ++ "\n" ++ " type declarations: " ++ show tydecls ++ "\n" ++
" current LHS: " ++ show currentLHS ++ "\n" ++ " current LHS: " ++ show currentLHS ++ "\n" ++
" holes: " ++ show holeLHS ++ "\n" ++ " holes: " ++ show holeLHS ++ "\n" ++
" nameLocMap: " ++ show nameLocMap " nameLocMap: " ++ show nameLocMap ++ "\n" ++
" sourcefile: " ++ fname ++
" semanticHighlighting: " ++ show semanticHighlighting ++
" semanticAliases: " ++ show semanticAliases
export export
initMetadata : Metadata initMetadata : String -> Metadata
initMetadata = MkMetadata initMetadata fname = MkMetadata
{ lhsApps = [] { lhsApps = []
, names = [] , names = []
, tydecls = [] , tydecls = []
, currentLHS = Nothing , currentLHS = Nothing
, holeLHS = [] , holeLHS = []
, nameLocMap = empty , nameLocMap = empty
, sourcefile = fname
, semanticHighlighting = empty
, semanticAliases = empty
} }
-- A label for metadata in the global state -- A label for metadata in the global state
@ -76,6 +144,9 @@ TTC Metadata where
toBuf b (tydecls m) toBuf b (tydecls m)
toBuf b (holeLHS m) toBuf b (holeLHS m)
toBuf b (nameLocMap m) toBuf b (nameLocMap m)
toBuf b (sourcefile m)
toBuf b (semanticHighlighting m)
toBuf b (semanticAliases m)
fromBuf b fromBuf b
= do apps <- fromBuf b = do apps <- fromBuf b
@ -83,7 +154,10 @@ TTC Metadata where
tys <- fromBuf b tys <- fromBuf b
hlhs <- fromBuf b hlhs <- fromBuf b
dlocs <- 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 export
addLHS : {vars : _} -> addLHS : {vars : _} ->
@ -213,6 +287,31 @@ findHoleLHS hn
= do meta <- get MD = do meta <- get MD
pure (lookupBy (\x, y => dropNS x == dropNS y) hn (holeLHS meta)) 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 -- Normalise all the types of the names, since they might have had holes
-- when added and the holes won't necessarily get saved -- when added and the holes won't necessarily get saved
normaliseTypes : {auto m : Ref MD Metadata} -> normaliseTypes : {auto m : Ref MD Metadata} ->
@ -249,13 +348,14 @@ TTC TTMFile where
pure (MkTTMFile ver md) pure (MkTTMFile ver md)
HasNames Metadata where HasNames Metadata where
full gam (MkMetadata lhs ns tys clhs hlhs dlocs) full gam md
= pure $ MkMetadata !(traverse fullLHS lhs) = pure $ record { lhsApps = !(traverse fullLHS $ md.lhsApps)
!(traverse fullTy ns) , names = !(traverse fullTy $ md.names)
!(traverse fullTy tys) , tydecls = !(traverse fullTy $ md.tydecls)
Nothing , currentLHS = Nothing
!(traverse fullHLHS hlhs) , holeLHS = !(traverse fullHLHS $ md.holeLHS)
(fromList !(traverse fullDecls (toList dlocs))) , nameLocMap = fromList !(traverse fullDecls (toList $ md.nameLocMap))
} md
where where
fullLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm)) fullLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm))
fullLHS (fc, (i, tm)) = pure (fc, (i, !(full gam tm))) fullLHS (fc, (i, tm)) = pure (fc, (i, !(full gam tm)))
@ -269,13 +369,16 @@ HasNames Metadata where
fullDecls : (NonEmptyFC, Name) -> Core (NonEmptyFC, Name) fullDecls : (NonEmptyFC, Name) -> Core (NonEmptyFC, Name)
fullDecls (fc, n) = pure (fc, !(full gam n)) 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) = pure $ MkMetadata !(traverse resolvedLHS lhs)
!(traverse resolvedTy ns) !(traverse resolvedTy ns)
!(traverse resolvedTy tys) !(traverse resolvedTy tys)
Nothing Nothing
!(traverse resolvedHLHS hlhs) !(traverse resolvedHLHS hlhs)
(fromList !(traverse resolvedDecls (toList dlocs))) (fromList !(traverse resolvedDecls (toList dlocs)))
fname
semhl
semal
where where
resolvedLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm)) resolvedLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm))
resolvedLHS (fc, (i, tm)) = pure (fc, (i, !(resolved gam tm))) resolvedLHS (fc, (i, tm)) = pure (fc, (i, !(resolved gam tm)))

View File

@ -2,6 +2,7 @@ module Core.Name
import Data.List import Data.List
import Data.Strings import Data.Strings
import Data.Maybe
import Decidable.Equality import Decidable.Equality
import Libraries.Text.PrettyPrint.Prettyprinter import Libraries.Text.PrettyPrint.Prettyprinter
import Libraries.Text.PrettyPrint.Prettyprinter.Util import Libraries.Text.PrettyPrint.Prettyprinter.Util
@ -75,6 +76,21 @@ isUserName (NS _ n) = isUserName n
isUserName (DN _ n) = isUserName n isUserName (DN _ n) = isUserName n
isUserName _ = True 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 export
isUN : Name -> Maybe String isUN : Name -> Maybe String
isUN (UN str) = Just str isUN (UN str) = Just str
@ -93,6 +109,19 @@ nameRoot (CaseBlock n _) = "$" ++ show n
nameRoot (WithBlock n _) = "$" ++ show n nameRoot (WithBlock n _) = "$" ++ show n
nameRoot (Resolved i) = "$" ++ show i 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 --- Drop a namespace from a name
export export
dropNS : Name -> Name dropNS : Name -> Name
@ -129,7 +158,8 @@ Pretty Name where
pretty (PV n d) = braces (pretty 'P' <+> colon <+> pretty n <+> colon <+> pretty d) pretty (PV n d) = braces (pretty 'P' <+> colon <+> pretty n <+> colon <+> pretty d)
pretty (DN str _) = pretty str pretty (DN str _) = pretty str
pretty (RF n) = "." <+> pretty n 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 (CaseBlock outer _) = reflow "case block in" <++> pretty outer
pretty (WithBlock outer _) = reflow "with block in" <++> pretty outer pretty (WithBlock outer _) = reflow "with block in" <++> pretty outer
pretty (Resolved x) = pretty "$resolved" <+> pretty x pretty (Resolved x) = pretty "$resolved" <+> pretty x

View File

@ -594,6 +594,11 @@ Reflect FC where
start' <- reflect fc defs lhs env start start' <- reflect fc defs lhs env start
end' <- reflect fc defs lhs env end end' <- reflect fc defs lhs env end
appCon fc defs (reflectiontt "MkFC") [fn', start', 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") reflect fc defs lhs env EmptyFC = getCon fc defs (reflectiontt "EmptyFC")
{- {-

View File

@ -82,6 +82,38 @@ isConstantType (UN n) = case n of
_ => Nothing _ => Nothing
isConstantType _ = 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 export
constantEq : (x, y : Constant) -> Maybe (x = y) constantEq : (x, y : Constant) -> Maybe (x = y)
constantEq (I x) (I y) = case decEq x y of constantEq (I x) (I y) = case decEq x y of

View File

@ -22,6 +22,8 @@ export
TTC FC where TTC FC where
toBuf b (MkFC file startPos endPos) toBuf b (MkFC file startPos endPos)
= do tag 0; toBuf b file; toBuf b startPos; toBuf b 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 toBuf b EmptyFC = tag 1
fromBuf b fromBuf b
@ -30,6 +32,9 @@ TTC FC where
s <- fromBuf b; e <- fromBuf b s <- fromBuf b; e <- fromBuf b
pure (MkFC f s e) pure (MkFC f s e)
1 => pure EmptyFC 1 => pure EmptyFC
2 => do f <- fromBuf b;
s <- fromBuf b; e <- fromBuf b
pure (MkVirtualFC f s e)
_ => corrupt "FC" _ => corrupt "FC"
export export
TTC Namespace where TTC Namespace where

View File

@ -87,7 +87,7 @@ mkPrec Prefix p = Prefix p
toTokList : {auto s : Ref Syn SyntaxInfo} -> toTokList : {auto s : Ref Syn SyntaxInfo} ->
PTerm -> Core (List (Tok OpStr PTerm)) PTerm -> Core (List (Tok OpStr PTerm))
toTokList (POp fc opn l r) toTokList (POp fc opFC opn l r)
= do syn <- get Syn = do syn <- get Syn
let op = nameRoot opn let op = nameRoot opn
case lookup op (infixes syn) of case lookup op (infixes syn) of
@ -95,16 +95,16 @@ toTokList (POp fc opn l r)
if any isOpChar (fastUnpack op) if any isOpChar (fastUnpack op)
then throw (GenericMsg fc $ "Unknown operator '" ++ op ++ "'") then throw (GenericMsg fc $ "Unknown operator '" ++ op ++ "'")
else do rtoks <- toTokList r else do rtoks <- toTokList r
pure (Expr l :: Op fc opn backtickPrec :: rtoks) pure (Expr l :: Op fc opFC opn backtickPrec :: rtoks)
Just (Prefix, _) => Just (Prefix, _) =>
throw (GenericMsg fc $ "'" ++ op ++ "' is a prefix operator") throw (GenericMsg fc $ "'" ++ op ++ "' is a prefix operator")
Just (fix, prec) => Just (fix, prec) =>
do rtoks <- toTokList r 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 where
backtickPrec : OpPrec backtickPrec : OpPrec
backtickPrec = NonAssoc 1 backtickPrec = NonAssoc 1
toTokList (PPrefixOp fc opn arg) toTokList (PPrefixOp fc opFC opn arg)
= do syn <- get Syn = do syn <- get Syn
let op = nameRoot opn let op = nameRoot opn
case lookup op (prefixes syn) of case lookup op (prefixes syn) of
@ -112,7 +112,7 @@ toTokList (PPrefixOp fc opn arg)
throw (GenericMsg fc $ "'" ++ op ++ "' is not a prefix operator") throw (GenericMsg fc $ "'" ++ op ++ "' is not a prefix operator")
Just prec => Just prec =>
do rtoks <- toTokList arg do rtoks <- toTokList arg
pure (Op fc opn (Prefix prec) :: rtoks) pure (Op fc opFC opn (Prefix prec) :: rtoks)
toTokList t = pure [Expr t] toTokList t = pure [Expr t]
record BangData where record BangData where
@ -123,21 +123,39 @@ record BangData where
initBangs : BangData initBangs : BangData
initBangs = MkBangData 0 [] 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 : List (Name, FC, RawImp) -> RawImp -> RawImp
bindBangs [] tm = tm bindBangs [] tm = tm
bindBangs ((n, fc, btm) :: bs) tm bindBangs ((n, fc, btm) :: bs) tm
= bindBangs bs $ IApp fc (IApp fc (IVar fc (UN ">>=")) btm) = bindBangs bs
(ILam EmptyFC top Explicit (Just n) $ bindFun fc Nothing btm
(Implicit fc False) tm) $ ILam EmptyFC top Explicit (Just n) (Implicit fc False) tm
idiomise : FC -> RawImp -> RawImp idiomise : FC -> RawImp -> RawImp
idiomise fc (IAlternative afc u alts) idiomise fc (IAlternative afc u alts)
= IAlternative afc (mapAltType (idiomise afc) u) (idiomise afc <$> alts) = IAlternative afc (mapAltType (idiomise afc) u) (idiomise afc <$> alts)
idiomise fc (IApp afc f a) idiomise fc (IApp afc f a)
= IApp fc (IApp fc (IVar fc (UN "<*>")) = let fc = virtualiseFC fc in
(idiomise afc f)) IApp fc (IApp fc (IVar fc (UN "<*>"))
a (idiomise afc f))
idiomise fc fn = IApp fc (IVar fc (UN "pure")) fn a
idiomise fc fn
= let fc = virtualiseFC fc in
IApp fc (IVar fc (UN "pure")) fn
pairname : Name pairname : Name
pairname = NS builtinNS (UN "Pair") pairname = NS builtinNS (UN "Pair")
@ -166,9 +184,11 @@ mutual
pure $ IPi fc rig !(traverse (desugar side ps') p) pure $ IPi fc rig !(traverse (desugar side ps') p)
mn !(desugarB side ps argTy) mn !(desugarB side ps argTy)
!(desugarB side ps' retTy) !(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 == "_" = 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) (Just n) !(desugarB side ps argTy)
!(desugar side (n :: ps) scope) !(desugar side (n :: ps) scope)
else pure $ ILam EmptyFC rig !(traverse (desugar side ps) p) 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) ICase fc (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False)
[snd !(desugarClause ps True (MkPatClause fc pat scope []))] [snd !(desugarClause ps True (MkPatClause fc pat scope []))]
desugarB side ps (PLet fc rig (PRef prefFC n) nTy nVal 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) = do whenJust (isConcreteFC prefFC) \nfc =>
!(desugar side (n :: ps) scope) 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) desugarB side ps (PLet fc rig pat nTy nVal scope alts)
= pure $ ICase fc !(desugarB side ps nVal) !(desugarB side ps nTy) = pure $ ICase fc !(desugarB side ps nVal) !(desugarB side ps nTy)
!(traverse (map snd . desugarClause ps True) !(traverse (map snd . desugarClause ps True)
(MkPatClause fc pat scope [] :: alts)) (MkPatClause fc pat scope [] :: alts))
desugarB side ps (PCase fc x xs) desugarB side ps (PCase fc x xs)
= pure $ ICase fc !(desugarB side ps x) = pure $ ICase fc !(desugarB side ps x)
(Implicit fc False) (Implicit (virtualiseFC fc) False)
!(traverse (map snd . desugarClause ps True) xs) !(traverse (map snd . desugarClause ps True) xs)
desugarB side ps (PLocal fc xs scope) desugarB side ps (PLocal fc xs scope)
= let ps' = definedIn xs ++ ps in = let ps' = definedIn xs ++ ps in
@ -207,8 +229,10 @@ mutual
= pure $ IUpdate pfc !(traverse (desugarUpdate side ps) fs) = pure $ IUpdate pfc !(traverse (desugarUpdate side ps) fs)
!(desugarB side ps rec) !(desugarB side ps rec)
desugarB side ps (PUpdate fc fs) desugarB side ps (PUpdate fc fs)
= desugarB side ps (PLam fc top Explicit (PRef fc (MN "rec" 0)) (PImplicit fc) = desugarB side ps
(PApp fc (PUpdate fc fs) (PRef fc (MN "rec" 0)))) $ 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) desugarB side ps (PApp fc x y)
= pure $ IApp fc !(desugarB side ps x) !(desugarB side ps y) = pure $ IApp fc !(desugarB side ps x) !(desugarB side ps y)
desugarB side ps (PAutoApp fc x 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'],
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 (PBracketed fc e) = desugarB side ps e
desugarB side ps (POp fc op l r) desugarB side ps (POp fc opFC op l r)
= do ts <- toTokList (POp fc op l r) = do ts <- toTokList (POp fc opFC op l r)
desugarTree side ps !(parseOps ts) desugarTree side ps !(parseOps ts)
desugarB side ps (PPrefixOp fc op arg) desugarB side ps (PPrefixOp fc opFC op arg)
= do ts <- toTokList (PPrefixOp fc op arg) = do ts <- toTokList (PPrefixOp fc opFC op arg)
desugarTree side ps !(parseOps ts) desugarTree side ps !(parseOps ts)
desugarB side ps (PSectionL fc op arg) desugarB side ps (PSectionL fc opFC op arg)
= do syn <- get Syn = do syn <- get Syn
-- It might actually be a prefix argument rather than a section -- It might actually be a prefix argument rather than a section
-- so check that first, otherwise desugar as a lambda -- so check that first, otherwise desugar as a lambda
case lookup (nameRoot op) (prefixes syn) of case lookup (nameRoot op) (prefixes syn) of
Nothing => Nothing =>
desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc)
(POp fc op (PRef fc (MN "arg" 0)) arg)) (POp fc opFC op (PRef fc (MN "arg" 0)) arg))
Just prec => desugarB side ps (PPrefixOp fc op arg) Just prec => desugarB side ps (PPrefixOp fc opFC op arg)
desugarB side ps (PSectionR fc arg op) desugarB side ps (PSectionR fc opFC arg op)
= desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) = 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 (PSearch fc depth) = pure $ ISearch fc depth
desugarB side ps (PPrimVal fc (BI x)) desugarB side ps (PPrimVal fc (BI x))
= case !fromIntegerName of = case !fromIntegerName of
@ -255,20 +279,23 @@ mutual
pure $ IAlternative fc (UniqueDefault (IPrimVal fc (BI x))) pure $ IAlternative fc (UniqueDefault (IPrimVal fc (BI x)))
[IPrimVal fc (BI x), [IPrimVal fc (BI x),
IPrimVal fc (I (fromInteger x))] IPrimVal fc (I (fromInteger x))]
Just fi => pure $ IApp fc (IVar fc fi) Just fi =>
(IPrimVal fc (BI x)) let vfc = virtualiseFC fc in
pure $ IApp vfc (IVar vfc fi) (IPrimVal fc (BI x))
desugarB side ps (PPrimVal fc (Ch x)) desugarB side ps (PPrimVal fc (Ch x))
= case !fromCharName of = case !fromCharName of
Nothing => Nothing =>
pure $ IPrimVal fc (Ch x) pure $ IPrimVal fc (Ch x)
Just f => pure $ IApp fc (IVar fc f) Just f =>
(IPrimVal fc (Ch x)) let vfc = virtualiseFC fc in
pure $ IApp vfc (IVar vfc f) (IPrimVal fc (Ch x))
desugarB side ps (PPrimVal fc (Db x)) desugarB side ps (PPrimVal fc (Db x))
= case !fromDoubleName of = case !fromDoubleName of
Nothing => Nothing =>
pure $ IPrimVal fc (Db x) pure $ IPrimVal fc (Db x)
Just f => pure $ IApp fc (IVar fc f) Just f =>
(IPrimVal fc (Db x)) 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 (PPrimVal fc x) = pure $ IPrimVal fc x
desugarB side ps (PQuote fc tm) desugarB side ps (PQuote fc tm)
= pure $ IQuote fc !(desugarB side ps tm) = pure $ IQuote fc !(desugarB side ps tm)
@ -313,40 +340,41 @@ mutual
let val = idiomise fc itm let val = idiomise fc itm
logRaw "desugar.idiom" 10 "Desugared to" val logRaw "desugar.idiom" 10 "Desugared to" val
pure val pure val
desugarB side ps (PList fc args) desugarB side ps (PList fc nilFC args)
= expandList side ps fc args = expandList side ps nilFC args
desugarB side ps (PPair fc l r) desugarB side ps (PPair fc l r)
= do l' <- desugarB side ps l = do l' <- desugarB side ps l
r' <- desugarB side ps r r' <- desugarB side ps r
let pval = apply (IVar fc mkpairname) [l', r'] let pval = apply (IVar fc mkpairname) [l', r']
pure $ IAlternative fc (UniqueDefault pval) pure $ IAlternative fc (UniqueDefault pval)
[apply (IVar fc pairname) [l', r'], 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 = 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) pure $ IAlternative fc (UniqueDefault pval)
[apply (IVar fc dpairname) [apply (IVar fc dpairname)
[Implicit nfc False, [Implicit nfc False,
ILam nfc top Explicit (Just (UN n)) (Implicit nfc False) r'], ILam nfc top Explicit (Just (UN n)) (Implicit nfc False) r'],
pval] 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 = do ty' <- desugarB side ps ty
r' <- desugarB side ps r r' <- desugarB side ps r
pure $ apply (IVar fc dpairname) pure $ apply (IVar opFC dpairname)
[ty', [ty',
ILam nfc top Explicit (Just (UN n)) ty' r'] 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 = do l' <- desugarB side ps l
r' <- desugarB side ps r r' <- desugarB side ps r
pure $ apply (IVar fc mkdpairname) [l', r'] pure $ apply (IVar opFC mkdpairname) [l', r']
desugarB side ps (PDPair fc l ty r) desugarB side ps (PDPair fc opFC l ty r)
= throw (GenericMsg fc "Invalid dependent pair type") = throw (GenericMsg fc "Invalid dependent pair type")
desugarB side ps (PUnit fc) desugarB side ps (PUnit fc)
= pure $ IAlternative fc (UniqueDefault (IVar fc (UN "MkUnit"))) = pure $ IAlternative fc (UniqueDefault (IVar fc (UN "MkUnit")))
[IVar fc (UN "Unit"), [IVar fc (UN "Unit"),
IVar fc (UN "MkUnit")] IVar fc (UN "MkUnit")]
desugarB side ps (PIfThenElse fc x t e) 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 "True")) !(desugar side ps t),
PatClause fc (IVar fc (UN "False")) !(desugar side ps e)] PatClause fc (IVar fc (UN "False")) !(desugar side ps e)]
desugarB side ps (PComprehension fc ret conds) desugarB side ps (PComprehension fc ret conds)
@ -361,22 +389,15 @@ mutual
desugarB side ps (PRewrite fc rule tm) desugarB side ps (PRewrite fc rule tm)
= pure $ IRewrite fc !(desugarB side ps rule) !(desugarB side ps tm) = pure $ IRewrite fc !(desugarB side ps rule) !(desugarB side ps tm)
desugarB side ps (PRange fc start next end) desugarB side ps (PRange fc start next end)
= case next of = let fc = virtualiseFC fc in
Nothing => desugarB side ps $ case next of
desugarB side ps (PApp fc Nothing => papply fc (PRef fc (UN "rangeFromTo")) [start,end]
(PApp fc (PRef fc (UN "rangeFromTo")) Just n => papply fc (PRef fc (UN "rangeFromThenTo")) [start, n, end]
start) end)
Just n =>
desugarB side ps (PApp fc
(PApp fc
(PApp fc (PRef fc (UN "rangeFromThenTo"))
start) n) end)
desugarB side ps (PRangeStream fc start next) desugarB side ps (PRangeStream fc start next)
= case next of = let fc = virtualiseFC fc in
Nothing => desugarB side ps $ case next of
desugarB side ps (PApp fc (PRef fc (UN "rangeFrom")) start) Nothing => papply fc (PRef fc (UN "rangeFrom")) [start]
Just n => Just n => papply fc (PRef fc (UN "rangeFromThen")) [start, n]
desugarB side ps (PApp fc (PApp fc (PRef fc (UN "rangeFromThen")) start) n)
desugarB side ps (PUnifyLog fc lvl tm) desugarB side ps (PUnifyLog fc lvl tm)
= pure $ IUnifyLog fc lvl !(desugarB side ps tm) = pure $ IUnifyLog fc lvl !(desugarB side ps tm)
desugarB side ps (PPostfixApp fc rec projs) desugarB side ps (PPostfixApp fc rec projs)
@ -404,23 +425,21 @@ mutual
{auto c : Ref Ctxt Defs} -> {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
{auto m : Ref MD Metadata} -> {auto m : Ref MD Metadata} ->
Side -> List Name -> FC -> List PTerm -> Core RawImp Side -> List Name ->
expandList side ps fc [] = pure (IVar fc (UN "Nil")) (nilFC : FC) -> List (FC, PTerm) -> Core RawImp
expandList side ps fc (x :: xs) expandList side ps nilFC [] = pure (IVar nilFC (UN "Nil"))
= pure $ apply (IVar fc (UN "::")) expandList side ps nilFC ((consFC, x) :: xs)
[!(desugarB side ps x), !(expandList side ps fc xs)] = pure $ apply (IVar consFC (UN "::"))
[!(desugarB side ps x), !(expandList side ps nilFC xs)]
addNS : Maybe Namespace -> Name -> Name
addNS (Just ns) n@(NS _ _) = n
addNS (Just ns) n = NS ns n
addNS _ n = n
addFromString : {auto c : Ref Ctxt Defs} -> addFromString : {auto c : Ref Ctxt Defs} ->
FC -> RawImp -> Core RawImp FC -> RawImp -> Core RawImp
addFromString fc tm addFromString fc tm
= pure $ case !fromStringName of = pure $ case !fromStringName of
Nothing => tm 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} -> expandString : {auto s : Ref Syn SyntaxInfo} ->
{auto b : Ref Bang BangData} -> {auto b : Ref Bang BangData} ->
@ -428,9 +447,11 @@ mutual
{auto m : Ref MD Metadata} -> {auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
Side -> List Name -> FC -> List PStr -> Core RawImp Side -> List Name -> FC -> List PStr -> Core RawImp
expandString side ps fc xs = pure $ case !(traverse toRawImp (filter notEmpty $ mergeStrLit xs)) of expandString side ps fc xs
[] => IPrimVal fc (Str "") = do xs <- traverse toRawImp (filter notEmpty $ mergeStrLit xs)
xs@(_::_) => foldr1 concatStr xs pure $ case xs of
[] => IPrimVal fc (Str "")
(_ :: _) => foldr1 concatStr xs
where where
toRawImp : PStr -> Core RawImp toRawImp : PStr -> Core RawImp
toRawImp (StrLiteral fc str) = pure $ IPrimVal fc (Str str) toRawImp (StrLiteral fc str) = pure $ IPrimVal fc (Str str)
@ -449,7 +470,10 @@ mutual
notEmpty (StrInterp _ _) = True notEmpty (StrInterp _ _) = True
concatStr : RawImp -> RawImp -> RawImp 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 -> Nat -> List (List PStr) -> Core (List PStr)
trimMultiline fc indent lines 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" then throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters"
else pure initLines else pure initLines
trimLast _ (initLines `snoc` xs) | Snoc xs initLines _ trimLast _ (initLines `snoc` xs) | Snoc xs initLines _
= let fc = fromMaybe fc $ findBy (\case StrInterp fc _ => Just fc; = let fc = fromMaybe fc $ findBy isStrInterp xs in
StrLiteral _ _ => Nothing) xs in
throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters" throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters"
dropLastNL : List PStr -> List PStr dropLastNL : List PStr -> List PStr
@ -508,13 +531,14 @@ mutual
= do tm' <- desugar side ps tm = do tm' <- desugar side ps tm
rest' <- expandDo side ps topfc ns rest rest' <- expandDo side ps topfc ns rest
gam <- get Ctxt 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) expandDo side ps topfc ns (DoBind fc nameFC n tm :: rest)
= do tm' <- desugar side ps tm = do tm' <- desugar side ps tm
rest' <- expandDo side ps topfc ns rest rest' <- expandDo side ps topfc ns rest
pure $ IApp fc (IApp fc (IVar fc (addNS ns (UN ">>="))) tm') whenJust (isConcreteFC nameFC) \nfc => addSemanticDecorations [(nfc, Bound, Just n)]
(ILam nameFC top Explicit (Just n) pure $ bindFun fc ns tm'
(Implicit fc False) rest') $ ILam nameFC top Explicit (Just n)
(Implicit (virtualiseFC fc) False) rest'
expandDo side ps topfc ns (DoBindPat fc pat exp alts :: rest) expandDo side ps topfc ns (DoBindPat fc pat exp alts :: rest)
= do pat' <- desugar LHS ps pat = do pat' <- desugar LHS ps pat
(newps, bpat) <- bindNames False pat' (newps, bpat) <- bindNames False pat'
@ -522,19 +546,22 @@ mutual
alts' <- traverse (map snd . desugarClause ps True) alts alts' <- traverse (map snd . desugarClause ps True) alts
let ps' = newps ++ ps let ps' = newps ++ ps
rest' <- expandDo side ps' topfc ns rest rest' <- expandDo side ps' topfc ns rest
pure $ IApp fc (IApp fc (IVar fc (addNS ns (UN ">>="))) exp') let fcOriginal = fc
(ILam EmptyFC top Explicit (Just (MN "_" 0)) let fc = virtualiseFC fc
pure $ bindFun fc ns exp'
$ ILam EmptyFC top Explicit (Just (MN "_" 0))
(Implicit fc False) (Implicit fc False)
(ICase fc (IVar EmptyFC (MN "_" 0)) (ICase fc (IVar EmptyFC (MN "_" 0))
(Implicit fc False) (Implicit fc False)
(PatClause fc bpat rest' (PatClause fcOriginal bpat rest'
:: alts'))) :: alts'))
expandDo side ps topfc ns (DoLet fc lhsFC n rig ty tm :: rest) expandDo side ps topfc ns (DoLet fc lhsFC n rig ty tm :: rest)
= do b <- newRef Bang initBangs = do b <- newRef Bang initBangs
tm' <- desugarB side ps tm tm' <- desugarB side ps tm
ty' <- desugar side ps ty ty' <- desugar side ps ty
rest' <- expandDo side ps topfc ns rest 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 bd <- get Bang
pure $ bindBangs (bangNames bd) bind pure $ bindBangs (bangNames bd) bind
expandDo side ps topfc ns (DoLetPat fc pat ty tm alts :: rest) expandDo side ps topfc ns (DoLetPat fc pat ty tm alts :: rest)
@ -547,6 +574,7 @@ mutual
let ps' = newps ++ ps let ps' = newps ++ ps
rest' <- expandDo side ps' topfc ns rest rest' <- expandDo side ps' topfc ns rest
bd <- get Bang bd <- get Bang
let fc = virtualiseFC fc
pure $ bindBangs (bangNames bd) $ pure $ bindBangs (bangNames bd) $
ICase fc tm' ty' ICase fc tm' ty'
(PatClause fc bpat rest' (PatClause fc bpat rest'
@ -566,20 +594,20 @@ mutual
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
{auto m : Ref MD Metadata} -> {auto m : Ref MD Metadata} ->
Side -> List Name -> Tree OpStr PTerm -> Core RawImp 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 = do l' <- desugarTree side ps l
r' <- desugarTree side ps r r' <- desugarTree side ps r
pure (IAlternative loc FirstSuccess pure (IAlternative loc FirstSuccess
[apply (IVar loc (UN "===")) [l', r'], [apply (IVar eqFC (UN "===")) [l', r'],
apply (IVar loc (UN "~=~")) [l', r']]) apply (IVar eqFC (UN "~=~")) [l', r']])
desugarTree side ps (Infix loc (UN "$") l r) -- special case since '$' is special syntax desugarTree side ps (Infix loc _ (UN "$") l r) -- special case since '$' is special syntax
= do l' <- desugarTree side ps l = do l' <- desugarTree side ps l
r' <- desugarTree side ps r r' <- desugarTree side ps r
pure (IApp loc l' 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 = do l' <- desugarTree side ps l
r' <- desugarTree side ps r 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 -- negation is a special case, since we can't have an operator with
-- two meanings otherwise -- two meanings otherwise
@ -587,7 +615,7 @@ mutual
-- Note: In case of negated signed integer literals, we apply the -- Note: In case of negated signed integer literals, we apply the
-- negation directly. Otherwise, the literal might be -- negation directly. Otherwise, the literal might be
-- truncated to 0 before being passed on to `negate`. -- 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) = let newFC = fromMaybe EmptyFC (mergeFC loc fc)
continue = desugarTree side ps . Leaf . PPrimVal newFC continue = desugarTree side ps . Leaf . PPrimVal newFC
in case c of in case c of
@ -601,15 +629,15 @@ mutual
-- not a signed integer literal. proceed by desugaring -- not a signed integer literal. proceed by desugaring
-- and applying to `negate`. -- and applying to `negate`.
_ => do arg' <- desugarTree side ps (Leaf $ PPrimVal fc c) _ => 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 = 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 = 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 desugarTree side ps (Leaf t) = desugarB side ps t
desugarType : {auto s : Ref Syn SyntaxInfo} -> desugarType : {auto s : Ref Syn SyntaxInfo} ->

View File

@ -286,9 +286,9 @@ getDocsForPTerm (PRef fc name) = pure $ [!(render styleAnn !(getDocsForName fc n
getDocsForPTerm (PPrimVal _ constant) = getDocsForPrimitive constant 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 (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 (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 (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 (PUnit _) = pure ["Unit Literal\n\tDesugars to MkUnit or Unit"]
getDocsForPTerm pterm = pure ["Docs not implemented for " ++ show pterm ++ " yet"] getDocsForPTerm pterm = pure ["Docs not implemented for " ++ show pterm ++ " yet"]

View File

@ -178,7 +178,7 @@ stMain cgs opts
when (checkVerbose opts) $ -- override Quiet if implicitly set when (checkVerbose opts) $ -- override Quiet if implicitly set
setOutput (REPL False) setOutput (REPL False)
u <- newRef UST initUState u <- newRef UST initUState
m <- newRef MD initMetadata m <- newRef MD (initMetadata $ fromMaybe "(interactive)" fname)
updateREPLOpts updateREPLOpts
session <- getSession session <- getSession
when (not $ nobanner session) $ do when (not $ nobanner session) $ do

View File

@ -124,7 +124,7 @@ elabImplementation : {vars : _} ->
Maybe (List ImpDecl) -> Maybe (List ImpDecl) ->
Core () Core ()
-- TODO: Refactor all these steps into separate functions -- 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 = do -- let impName_in = maybe (mkImplName fc iname ps) id impln
-- If we're in a nested block, update the name -- If we're in a nested block, update the name
let impName_nest = case lookup impName_in (names nest) of 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) inames <- lookupCtxtName iname (gamma defs)
let [cndata] = concatMap (\n => lookupName n (ifaces syn)) let [cndata] = concatMap (\n => lookupName n (ifaces syn))
(map fst inames) (map fst inames)
| [] => undefinedName fc iname | [] => undefinedName vfc iname
| ns => throw (AmbiguousName fc (map fst ns)) | ns => throw (AmbiguousName vfc (map fst ns))
let cn : Name = fst cndata let cn : Name = fst cndata
let cdata : IFaceInfo = snd cndata let cdata : IFaceInfo = snd cndata
Just ity <- lookupTyExact cn (gamma defs) Just ity <- lookupTyExact cn (gamma defs)
| Nothing => undefinedName fc cn | Nothing => undefinedName vfc cn
Just conty <- lookupTyExact (iconstructor cdata) (gamma defs) Just conty <- lookupTyExact (iconstructor cdata) (gamma defs)
| Nothing => undefinedName fc (iconstructor cdata) | Nothing => undefinedName vfc (iconstructor cdata)
let impsp = nub (concatMap findIBinds ps ++ let impsp = nub (concatMap findIBinds ps ++
concatMap findIBinds (map snd cons)) 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] then [Inline]
else [Inline, Hint True] else [Inline, Hint True]
let initTy = bindImpls fc is $ bindConstraints fc AutoImplicit cons let initTy = bindImpls vfc is $ bindConstraints vfc AutoImplicit cons
(apply (IVar fc iname) ps) (apply (IVar vfc iname) ps)
let paramBinds = if !isUnboundImplicits let paramBinds = if !isUnboundImplicits
then findBindableNames True vars [] initTy then findBindableNames True vars [] initTy
else [] else []
let impTy = doBind paramBinds initTy 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 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 -- If the body is empty, we're done for now (just declaring that
-- the implementation exists and define it later) -- the implementation exists and define it later)
when (defPass pass) $ maybe (pure ()) when (defPass pass) $
(\body_in => do whenJust mbody $ \body_in => do
defs <- get Ctxt defs <- get Ctxt
Just impTyc <- lookupTyExact impName (gamma defs) Just impTyc <- lookupTyExact impName (gamma defs)
| Nothing => throw (InternalError ("Can't happen, can't find type of " ++ show impName)) | Nothing => throw (InternalError ("Can't happen, can't find type of " ++ show impName))
methImps <- getMethImps [] impTyc methImps <- getMethImps [] impTyc
log "elab.implementation" 3 $ "Bind implicits to each method: " ++ show methImps 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) let (body, missing)
= addDefaults fc impName = addDefaults vfc impName
(zip (params cdata) ps) (zip (params cdata) ps)
(map (dropNS . name) (methods cdata)) (map (dropNS . name) (methods cdata))
(defaults cdata) body_in (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 defs <- get Ctxt
let hs = openHints defs -- snapshot open hint state let hs = openHints defs -- snapshot open hint state
log "elab.implementation" 10 $ "Open hints: " ++ (show (impName :: nusing)) 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 addOpenHint n') nusing
-- 2. Elaborate top level function types for this interface -- 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 mtops = map (Builtin.fst . snd) fns
let con = iconstructor cdata let con = iconstructor cdata
let ilhs = impsApply (IVar EmptyFC impName) let ilhs = impsApply (IVar EmptyFC impName)
(map (\x => (x, IBindVar fc (show x))) (map (\x => (x, IBindVar vfc (show x)))
(map fst methImps)) (map fst methImps))
-- RHS is the constructor applied to a search for the necessary -- RHS is the constructor applied to a search for the necessary
-- parent constraints, then the method implementations -- parent constraints, then the method implementations
defs <- get Ctxt defs <- get Ctxt
let fldTys = getFieldArgs !(normaliseHoles defs [] conty) let fldTys = getFieldArgs !(normaliseHoles defs [] conty)
log "elab.implementation" 5 $ "Field types " ++ show fldTys 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) (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 log "elab.implementation" 5 $ "Implementation record: " ++ show impFn
-- If it's a named implementation, add it as a global hint while -- 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 -- Make sure we don't use this name to solve parent constraints
-- when elaborating the record, or we'll end up in a cycle! -- 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 -- Update nested names so we elaborate the body in the right
-- environment -- 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 let nest' = record { names $= (names' ++) } nest
traverse_ (processDecl [] nest' env) [impFn] 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 -- 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 -- 4. (TODO: Order method bodies to be in declaration order, in
-- case of dependencies) -- 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 -- inline flag has done its job, and outside the interface
-- it can hurt, so unset it now -- it can hurt, so unset it now
unsetFlag fc impName TCInline unsetFlag vfc impName TCInline
-- Reset the open hints (remove the named implementation) -- Reset the open hints (remove the named implementation)
setOpenHints hs setOpenHints hs
pure ()) mbody
where where
vfc : FC
vfc = virtualiseFC ifc
applyEnv : Name -> applyEnv : Name ->
Core (Name, (Maybe Name, List (Var vars), FC -> NameType -> Term vars)) Core (Name, (Maybe Name, List (Var vars), FC -> NameType -> Term vars))
applyEnv n applyEnv n
= do n' <- resolveName n = do n' <- resolveName n
pure (Resolved n', (Nothing, reverse (allVars env), pure (Resolved n', (Nothing, reverse (allVars env),
\fn, nt => applyToFull fc \fn, nt => applyToFull vfc
(Ref fc nt (Resolved n')) env)) (Ref vfc nt (Resolved n')) env))
-- For the method fields in the record, get the arguments we need to abstract -- For the method fields in the record, get the arguments we need to abstract
-- over -- 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 : RawImp -> List (Name, RawImp) -> RawImp
impsApply fn [] = fn impsApply fn [] = fn
impsApply fn ((n, arg) :: ns) 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 : RawImp -> List RawImp -> RawImp
autoImpsApply f [] = f 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 : List (Name, RigCount, PiInfo RawImp) -> RawImp -> RawImp
mkLam [] tm = tm mkLam [] tm = tm
mkLam ((x, c, p) :: xs) 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 : RawImp -> List (Name, RigCount, PiInfo RawImp) -> RawImp
applyTo tm [] = tm applyTo tm [] = tm
@ -335,7 +338,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
mkLam argns mkLam argns
(impsApply (impsApply
(applyTo (IVar EmptyFC n) argns) (applyTo (IVar EmptyFC n) argns)
(map (\n => (n, IVar fc (UN (show n)))) imps)) (map (\n => (n, IVar vfc (UN (show n)))) imps))
where where
applyUpdate : (Name, RigCount, PiInfo RawImp) -> applyUpdate : (Name, RigCount, PiInfo RawImp) ->
(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 : Name -> Name -> Core (Name, RawImp)
applyCon impl n applyCon impl n
= do mn <- inCurrentNS (methName 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 : List (Name, RigCount, RawImp) -> RawImp -> RawImp
bindImps [] ty = ty bindImps [] ty = ty
bindImps ((n, c, t) :: ts) 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, -- Return method name, specialised method name, implicit name updates,
-- and method type. Also return how the method name should be updated -- 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. -- substitute in the explicit parameters.
let mty_iparams let mty_iparams
= substBindVars vars = substBindVars vars
(map (\n => (n, Implicit fc False)) imppnames) (map (\n => (n, Implicit vfc False)) imppnames)
mty_in mty_in
let mty_params let mty_params
= substNames vars (zip pnames ps) mty_iparams = 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 show mty_params
let mbase = bindImps methImps $ let mbase = bindImps methImps $
bindConstraints fc AutoImplicit cons $ bindConstraints vfc AutoImplicit cons $
mty_params mty_params
let ibound = findImplicits mbase 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 log "elab.implementation" 10 $ "Used names " ++ show ibound
let ibinds = map fst methImps let ibinds = map fst methImps
let methupds' = if isNil ibinds then [] let methupds' = if isNil ibinds then []
else [(n, impsApply (IVar fc n) else [(n, impsApply (IVar vfc n)
(map (\x => (x, IBindVar fc (show x))) ibinds))] (map (\x => (x, IBindVar vfc (show x))) ibinds))]
pure ((meth.name, n, upds, meth.count, meth.totalReq, mty), methupds') 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 : (Name, Name, List (String, String), RigCount, Maybe TotalReq, RawImp) -> ImpDecl
mkTopMethDecl (mn, n, upds, c, treq, mty) mkTopMethDecl (mn, n, upds, c, treq, mty)
= let opts = maybe opts_in (\t => Totality t :: opts_in) treq in = 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 -- Given the method type (result of topMethType) return the mapping from
-- top level method name to current implementation's method name -- 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) updateBody ns (IDef fc n cs)
= do cs' <- traverse (updateClause ns) cs = do cs' <- traverse (updateClause ns) cs
n' <- findMethName ns fc n n' <- findMethName ns fc n
log "ide-mode.highlight" 1 $ show (n, n', fc)
pure (IDef fc n' cs') pure (IDef fc n' cs')
updateBody ns _ updateBody ns e
= throw (GenericMsg fc = throw (GenericMsg (getFC e)
"Implementation body can only contain definitions") "Implementation body can only contain definitions")
addTransform : Name -> List (Name, Name) -> 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 $ = do log "elab.implementation" 3 $
"Adding transform for " ++ show meth.name ++ " : " ++ show meth.type ++ "Adding transform for " ++ show meth.name ++ " : " ++ show meth.type ++
"\n\tfor " ++ show iname ++ " in " ++ show ns "\n\tfor " ++ show iname ++ " in " ++ show ns
let lhs = INamedApp fc (IVar fc meth.name) let lhs = INamedApp vfc (IVar vfc meth.name)
(UN "__con") (UN "__con")
(IVar fc iname) (IVar vfc iname)
let Just mname = lookup (dropNS meth.name) ns let Just mname = lookup (dropNS meth.name) ns
| Nothing => pure () | Nothing => pure ()
let rhs = IVar fc mname let rhs = IVar vfc mname
log "elab.implementation" 5 $ show lhs ++ " ==> " ++ show rhs log "elab.implementation" 5 $ show lhs ++ " ==> " ++ show rhs
handleUnify handleUnify
(processDecl [] nest env (processDecl [] nest env
(ITransform fc (UN (show meth.name ++ " " ++ show iname)) lhs rhs)) (ITransform vfc (UN (show meth.name ++ " " ++ show iname)) lhs rhs))
(\err => (\err =>
log "elab.implementation" 5 $ "Can't add transform " ++ log "elab.implementation" 5 $ "Can't add transform " ++
show lhs ++ " ==> " ++ show rhs ++ show lhs ++ " ==> " ++ show rhs ++

View File

@ -4,6 +4,7 @@ import Core.Binary
import Core.Context import Core.Context
import Core.Context.Log import Core.Context.Log
import Core.Core import Core.Core
import Core.Name
import Core.Env import Core.Env
import Core.Metadata import Core.Metadata
import Core.TT import Core.TT
@ -111,20 +112,24 @@ mkIfaceData : {vars : _} ->
List (Maybe Name, RigCount, RawImp) -> List (Maybe Name, RigCount, RawImp) ->
Name -> Name -> List (Name, (RigCount, RawImp)) -> Name -> Name -> List (Name, (RigCount, RawImp)) ->
List Name -> List (Name, RigCount, RawImp) -> Core ImpDecl 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 = let opts = if isNil dets
then [NoHints, UniqueSearch] then [NoHints, UniqueSearch]
else [NoHints, UniqueSearch, SearchBy dets] else [NoHints, UniqueSearch, SearchBy dets]
pNames = map fst ps 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) $ conty = mkTy Implicit (map jname ps) $
mkTy AutoImplicit (map bhere constraints) (mkTy Explicit (map bname meths) retty) mkTy AutoImplicit (map bhere constraints) (mkTy Explicit (map bname meths) retty)
con = MkImpTy EmptyFC EmptyFC conName !(bindTypeNames [] (pNames ++ map fst meths ++ vars) conty) in 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) !(bindTypeNames [] (pNames ++ map fst meths ++ vars)
(mkDataTy fc ps)) (mkDataTy vfc ps))
opts [con]) opts [con])
where where
vfc : FC
vfc = virtualiseFC ifc
jname : (Name, (RigCount, RawImp)) -> (Maybe Name, RigCount, RawImp) jname : (Name, (RigCount, RawImp)) -> (Maybe Name, RigCount, RawImp)
jname (n, rig, t) = (Just n, rig, t) 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 List (Maybe Name, RigCount, RawImp) -> RawImp -> RawImp
mkTy imp [] ret = ret mkTy imp [] ret = ret
mkTy imp ((n, c, argty) :: args) 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 -- Get the implicit arguments for a method declaration or constraint hint
-- to allow us to build the data declaration -- to allow us to build the data declaration
@ -181,31 +186,31 @@ getMethToplevel : {vars : _} ->
Core (List ImpDecl) Core (List ImpDecl)
getMethToplevel {vars} env vis iname cname constraints allmeths params sig getMethToplevel {vars} env vis iname cname constraints allmeths params sig
= do let paramNames = map fst params = 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 -- Make the constraint application explicit for any method names
-- which appear in other method types -- which appear in other method types
let ty_constr = let ty_constr =
substNames vars (map applyCon allmeths) sig.type 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 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]) else [Inline])
(MkImpTy fc sig.nameLoc cn ty_imp) (MkImpTy vfc sig.nameLoc cn ty_imp)
let conapp = apply (IVar fc cname) let conapp = apply (IVar vfc cname)
(map (IBindVar EmptyFC) (map bindName allmeths)) (map (IBindVar EmptyFC) (map bindName allmeths))
let argns = getExplicitArgs 0 sig.type let argns = getExplicitArgs 0 sig.type
-- eta expand the RHS so that we put implicits in the right place -- 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") (UN "__con")
conapp) conapp)
(mkLam argns (mkLam argns
(apply (IVar EmptyFC (methName sig.name)) (apply (IVar EmptyFC (methName sig.name))
(map (IVar EmptyFC) argns))) (map (IVar EmptyFC) argns)))
let fndef = IDef fc cn [fnclause] let fndef = IDef vfc cn [fnclause]
pure [tydecl, fndef] pure [tydecl, fndef]
where where
fc : FC vfc : FC
fc = sig.location vfc = virtualiseFC sig.location
-- Bind the type parameters given explicitly - there might be information -- Bind the type parameters given explicitly - there might be information
-- in there that we can't infer after all -- 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 : Name -> (Name, RawImp)
applyCon n = let name = UN "__con" in 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 : Int -> RawImp -> List Name
getExplicitArgs i (IPi _ _ Explicit n _ sc) 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 : List Name -> RawImp -> RawImp
mkLam [] tm = tm mkLam [] tm = tm
mkLam (x :: xs) 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 : Name -> String
bindName (UN n) = "__bind_" ++ n bindName (UN n) = "__bind_" ++ n
@ -337,10 +342,10 @@ elabInterface : {vars : _} ->
(conName : Maybe Name) -> (conName : Maybe Name) ->
List ImpDecl -> List ImpDecl ->
Core () 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 = do fullIName <- getFullName iname
ns_iname <- inCurrentNS fullIName 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 -- Machine generated names need to be qualified when looking them up
conName <- inCurrentNS conName_in conName <- inCurrentNS conName_in
let meth_sigs = mapMaybe getSig body 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 pure (record { name = n } mt)) meth_decls
defs <- get Ctxt defs <- get Ctxt
Just ty <- lookupTyExact ns_iname (gamma defs) Just ty <- lookupTyExact ns_iname (gamma defs)
| Nothing => undefinedName fc iname | Nothing => undefinedName ifc iname
let implParams = getImplParams ty let implParams = getImplParams ty
updateIfaceSyn ns_iname conName updateIfaceSyn ns_iname conName
implParams paramNames (map snd constraints) implParams paramNames (map snd constraints)
ns_meths ds ns_meths ds
where where
vfc : FC
vfc = virtualiseFC ifc
paramNames : List Name paramNames : List Name
paramNames = map fst params 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 consts <- traverse (getMethDecl env nest params meth_names . (top,)) constraints
log "elab.interface" 5 $ "Constraints: " ++ show consts 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 dets meths
log "elab.interface" 10 $ "Methods: " ++ show meths log "elab.interface" 10 $ "Methods: " ++ show meths
log "elab.interface" 5 $ "Making interface data type " ++ show dt 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 log "elab.interface" 5 $ "Top level methods: " ++ show fns
traverse_ (processDecl [] nest env) fns traverse_ (processDecl [] nest env) fns
traverse_ (\n => do mn <- inCurrentNS n traverse_ (\n => do mn <- inCurrentNS n
setFlag fc mn Inline setFlag vfc mn Inline
setFlag fc mn TCInline setFlag vfc mn TCInline
setFlag fc mn Overloadable) meth_names setFlag vfc mn Overloadable) meth_names
-- Check that a default definition is correct. We just discard it here once -- 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 -- 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 -> elabDefault : List Declaration ->
(FC, List FnOpt, Name, List ImpClause) -> (FC, List FnOpt, Name, List ImpClause) ->
Core (Name, List ImpClause) Core (Name, List ImpClause)
elabDefault tydecls (fc, opts, n, cs) elabDefault tydecls (dfc, opts, n, cs)
= do -- orig <- branch = do -- orig <- branch
let dn_in = MN ("Default implementation of " ++ show n) 0 let dn_in = MN ("Default implementation of " ++ show n) 0
dn <- inCurrentNS dn_in dn <- inCurrentNS dn_in
@ -425,37 +434,40 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
the (Core (RigCount, RawImp)) $ the (Core (RigCount, RawImp)) $
case findBy (\ d => d <$ guard (n == d.name)) tydecls of case findBy (\ d => d <$ guard (n == d.name)) tydecls of
Just d => pure (d.count, d.type) 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 -- Substitute the method names with their top level function
-- name, so they don't get implicitly bound in the name -- name, so they don't get implicitly bound in the name
methNameMap <- traverse (\d => methNameMap <- traverse (\d =>
do let n = d.name do let n = d.name
cn <- inCurrentNS n cn <- inCurrentNS n
pure (n, applyParams (IVar fc cn) paramNames)) pure (n, applyParams (IVar vdfc cn) paramNames))
tydecls tydecls
let dty = bindPs params -- bind parameters let dty = bindPs params -- bind parameters
$ bindIFace fc ity -- bind interface (?!) $ bindIFace vdfc ity -- bind interface (?!)
$ substNames vars methNameMap dty $ substNames vars methNameMap dty
dty_imp <- bindTypeNames [] (map name tydecls ++ vars) dty dty_imp <- bindTypeNames [] (map name tydecls ++ vars) dty
log "elab.interface.default" 5 $ "Default method " ++ show dn ++ " : " ++ show dty_imp 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 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' 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 -- Reset the original context, we don't need to keep the definition
-- Actually we do for the metadata and name map! -- Actually we do for the metadata and name map!
-- put Ctxt orig -- put Ctxt orig
pure (n, cs) pure (n, cs)
where where
vdfc : FC
vdfc = virtualiseFC dfc
-- Bind the type parameters given explicitly - there might be information -- Bind the type parameters given explicitly - there might be information
-- in there that we can't infer after all -- in there that we can't infer after all
bindPs : List (Name, (RigCount, RawImp)) -> RawImp -> RawImp 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 : RawImp -> List Name -> RawImp
applyParams tm [] = tm applyParams tm [] = tm
applyParams tm (UN n :: ns) 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 applyParams tm (_ :: ns) = applyParams tm ns
changeNameTerm : Name -> RawImp -> RawImp changeNameTerm : Name -> RawImp -> Core RawImp
changeNameTerm dn (IVar fc n') 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) 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) 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) changeNameTerm dn (INamedApp fc f x arg)
= INamedApp fc (changeNameTerm dn f) x arg = INamedApp fc <$> changeNameTerm dn f <*> pure x <*> pure arg
changeNameTerm dn tm = tm changeNameTerm dn tm = pure tm
changeName : Name -> ImpClause -> ImpClause changeName : Name -> ImpClause -> Core ImpClause
changeName dn (PatClause fc lhs rhs) 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) changeName dn (WithClause fc lhs wval prf flags cs)
= WithClause fc (changeNameTerm dn lhs) wval = WithClause fc
prf flags (map (changeName dn) cs) <$> changeNameTerm dn lhs
<*> pure wval
<*> pure prf
<*> pure flags
<*> traverse (changeName dn) cs
changeName dn (ImpossibleClause fc lhs) changeName dn (ImpossibleClause fc lhs)
= ImpossibleClause fc (changeNameTerm dn lhs) = ImpossibleClause fc <$> changeNameTerm dn lhs
elabConstraintHints : (conName : Name) -> List Name -> elabConstraintHints : (conName : Name) -> List Name ->
Core () Core ()
elabConstraintHints conName meth_names elabConstraintHints conName meth_names
= do let nconstraints = nameCons 0 constraints = 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) (map fst nconstraints)
meth_names meth_names
paramNames) nconstraints paramNames) nconstraints
log "elab.interface" 5 $ "Constraint hints from " ++ show constraints ++ ": " ++ show chints log "elab.interface" 5 $ "Constraint hints from " ++ show constraints ++ ": " ++ show chints
traverse_ (processDecl [] nest env) (concatMap snd chints) traverse_ (processDecl [] nest env) (concatMap snd chints)
traverse_ (\n => do mn <- inCurrentNS n traverse_ (\n => do mn <- inCurrentNS n
setFlag fc mn TCInline) (map fst chints) setFlag vfc mn TCInline) (map fst chints)

View File

@ -66,8 +66,9 @@ pshowNoNorm env tm
ploc : {auto o : Ref ROpts REPLOpts} -> ploc : {auto o : Ref ROpts REPLOpts} ->
FC -> Core (Doc IdrisAnn) FC -> Core (Doc IdrisAnn)
ploc EmptyFC = pure emptyDoc ploc fc = do
ploc fc@(MkFC fn s e) = do let Just (fn, s, e) = isNonEmptyFC fc
| Nothing => pure emptyDoc
let (sr, sc) = mapHom (fromInteger . cast) s let (sr, sc) = mapHom (fromInteger . cast) s
let (er, ec) = mapHom (fromInteger . cast) e let (er, ec) = mapHom (fromInteger . cast) e
let nsize = length $ show (er + 1) let nsize = length $ show (er + 1)
@ -91,10 +92,12 @@ ploc fc@(MkFC fn s e) = do
-- Assumes the two FCs are sorted -- Assumes the two FCs are sorted
ploc2 : {auto o : Ref ROpts REPLOpts} -> ploc2 : {auto o : Ref ROpts REPLOpts} ->
FC -> FC -> Core (Doc IdrisAnn) FC -> FC -> Core (Doc IdrisAnn)
ploc2 fc EmptyFC = ploc fc ploc2 fc1 fc2 =
ploc2 EmptyFC fc = ploc fc do let Just (fn1, s1, e1) = isNonEmptyFC fc1
ploc2 (MkFC fn1 s1 e1) (MkFC fn2 s2 e2) = | Nothing => ploc fc2
do let (sr1, sc1) = mapHom (fromInteger . cast) s1 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 (sr2, sc2) = mapHom (fromInteger . cast) s2
let (er1, ec1) = mapHom (fromInteger . cast) e1 let (er1, ec1) = mapHom (fromInteger . cast) e1
let (er2, ec2) = mapHom (fromInteger . cast) e2 let (er2, ec2) = mapHom (fromInteger . cast) e2
@ -179,10 +182,12 @@ perror (PatternVariableUnifies fc env n tm)
prettyVar (PV n _) = prettyVar n prettyVar (PV n _) = prettyVar n
prettyVar n = pretty n prettyVar n = pretty n
order : FC -> FC -> (FC, FC) order : FC -> FC -> (FC, FC)
order EmptyFC fc2 = (EmptyFC, fc2) order fc1 fc2 =
order fc1 EmptyFC = (fc1, EmptyFC) let Just (_, sr1, sc1) = isNonEmptyFC fc1
order fc1@(MkFC _ sr1 sc1) fc2@(MkFC _ sr2 sc2) = | Nothing => (EmptyFC, fc2)
if sr1 < sr2 then (fc1, fc2) else if sr1 == sr2 && sc1 < sc2 then (fc1, fc2) else (fc2, fc1) 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) perror (CyclicMeta fc env n tm)
= pure $ errorDesc (reflow "Cycle detected in solution of metavariable" <++> meta (pretty !(prettyName n)) <++> equals = pure $ errorDesc (reflow "Cycle detected in solution of metavariable" <++> meta (pretty !(prettyName n)) <++> equals
<++> code !(pshow env tm)) <+> line <+> !(ploc fc) <++> code !(pshow env tm)) <+> line <+> !(ploc fc)

View File

@ -48,7 +48,7 @@ toStrUpdate (UN n, term)
where where
bracket : PTerm -> PTerm bracket : PTerm -> PTerm
bracket tm@(PRef _ _) = tm bracket tm@(PRef _ _) = tm
bracket tm@(PList _ _) = tm bracket tm@(PList _ _ _) = tm
bracket tm@(PPair _ _ _) = tm bracket tm@(PPair _ _ _) = tm
bracket tm@(PUnit _) = tm bracket tm@(PUnit _) = tm
bracket tm@(PComprehension _ _ _) = tm bracket tm@(PComprehension _ _ _) = tm

View File

@ -1,6 +1,8 @@
module Idris.IDEMode.Commands module Idris.IDEMode.Commands
import Core.Core import Core.Core
import Core.Context
import Core.Context.Log
import Core.Name import Core.Name
import public Idris.REPL.Opts import public Idris.REPL.Opts
import Libraries.Utils.Hex import Libraries.Utils.Hex
@ -281,9 +283,10 @@ sendStr f st =
map (const ()) (fPutStr f st) map (const ()) (fPutStr f st)
export export
send : SExpable a => File -> a -> Core () send : {auto c : Ref Ctxt Defs} -> SExpable a => File -> a -> Core ()
send f resp send f resp
= do let r = show (toSExp resp) ++ "\n" = 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 $ leftPad '0' 6 (asHex (cast (length r)))
coreLift $ sendStr f r coreLift $ sendStr f r
coreLift $ fflush f coreLift $ fflush f

View File

@ -5,6 +5,9 @@ module Idris.IDEMode.Parser
import Idris.IDEMode.Commands import Idris.IDEMode.Commands
import Core.Core import Core.Core
import Core.Name
import Core.Metadata
import Core.FC
import Data.Maybe import Data.Maybe
import Data.List import Data.List
@ -76,11 +79,11 @@ sexp
pure (SExpList xs) pure (SExpList xs)
ideParser : {e : _} -> 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 ideParser fname str p
= do toks <- mapError (fromLexError fname) $ idelex str = do toks <- mapError (fromLexError fname) $ idelex str
parsed <- mapError (fromParsingError fname) $ parse p toks (decor, (parsed, _)) <- mapError (fromParsingError fname) $ parseWith p toks
Right (fst parsed) Right parsed
export export

View File

@ -133,7 +133,8 @@ getInput f
pure (pack inp) pure (pack inp)
||| Do nothing and tell the user to wait for us to implmement this (or join the effort!) ||| 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 () String -> Core ()
todoCmd cmdName = iputStrLn $ reflow $ cmdName ++ ": command not yet implemented. Hopefully soon!" todoCmd cmdName = iputStrLn $ reflow $ cmdName ++ ": command not yet implemented. Hopefully soon!"
@ -260,24 +261,24 @@ processCatch cmd
msg <- perror err msg <- perror err
pure $ REPL $ REPLError msg) pure $ REPL $ REPLError msg)
idePutStrLn : File -> Integer -> String -> Core () idePutStrLn : {auto c : Ref Ctxt Defs} -> File -> Integer -> String -> Core ()
idePutStrLn outf i msg idePutStrLn outf i msg
= send outf (SExpList [SymbolAtom "write-string", = send outf (SExpList [SymbolAtom "write-string",
toSExp msg, toSExp i]) toSExp msg, toSExp i])
returnFromIDE : File -> Integer -> SExp -> Core () returnFromIDE : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core ()
returnFromIDE outf i msg returnFromIDE outf i msg
= do send outf (SExpList [SymbolAtom "return", msg, toSExp i]) = 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]) 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 printIDEResultWithHighlight outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "ok", toSExp msg
-- TODO return syntax highlighted result -- TODO return syntax highlighted result
, SExpList []]) , 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) ]) printIDEError outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "error", toSExp !(renderWithoutColor msg) ])
SExpable REPLEval where SExpable REPLEval where

View File

@ -1,35 +1,34 @@
module Idris.IDEMode.SyntaxHighlight module Idris.IDEMode.SyntaxHighlight
import Core.Context import Core.Context
import Core.Context.Log
import Core.InitPrimitives import Core.InitPrimitives
import Core.Metadata import Core.Metadata
import Core.TT import Core.TT
import Idris.REPL import Idris.REPL
import Idris.Syntax
import Idris.DocString
import Idris.IDEMode.Commands import Idris.IDEMode.Commands
import Data.List import Data.List
import Data.Maybe
import Libraries.Data.PosMap
%default covering %default covering
data Decoration : Type where
Typ : Decoration
Function : Decoration
Data : Decoration
Keyword : Decoration
Bound : Decoration
SExpable Decoration where SExpable Decoration where
toSExp Typ = SymbolAtom "type" toSExp Typ = SExpList [ SymbolAtom "decor", SymbolAtom "type"]
toSExp Function = SymbolAtom "function" toSExp Function = SExpList [ SymbolAtom "decor", SymbolAtom "function"]
toSExp Data = SymbolAtom "data" toSExp Data = SExpList [ SymbolAtom "decor", SymbolAtom "data"]
toSExp Keyword = SymbolAtom "keyword" toSExp Keyword = SExpList [ SymbolAtom "decor", SymbolAtom "keyword"]
toSExp Bound = SymbolAtom "bound" toSExp Bound = SExpList [ SymbolAtom "decor", SymbolAtom "bound"]
record Highlight where record Highlight where
constructor MkHighlight constructor MkHighlight
location : NonEmptyFC location : NonEmptyFC
name : Name name : String
isImplicit : Bool isImplicit : Bool
key : String key : String
decor : Decoration decor : Decoration
@ -38,19 +37,26 @@ record Highlight where
ns : String ns : String
SExpable FC where SExpable FC where
toSExp (MkFC fname (startLine, startCol) (endLine, endCol)) toSExp fc = case isNonEmptyFC fc of
= SExpList [ SExpList [ SymbolAtom "filename", StringAtom fname ] Just (fname , (startLine, startCol), (endLine, endCol)) =>
, SExpList [ SymbolAtom "start", IntegerAtom (cast startLine + 1), IntegerAtom (cast startCol + 1) ] SExpList [ SExpList [ SymbolAtom "filename", StringAtom fname ]
, SExpList [ SymbolAtom "end", IntegerAtom (cast endLine + 1), IntegerAtom (cast endCol + 1) ] , 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 SExpable Highlight where
toSExp (MkHighlight loc nam impl k dec doc t ns) toSExp (MkHighlight loc nam impl k dec doc t ns)
= SExpList [ toSExp $ justFC loc = SExpList [ toSExp $ justFC loc
, SExpList [ SExpList [ SymbolAtom "name", StringAtom (show nam) ] , SExpList [ SExpList [ SymbolAtom "name", StringAtom nam ]
, SExpList [ SymbolAtom "namespace", StringAtom ns ] , SExpList [ SymbolAtom "namespace", StringAtom ns ]
, SExpList [ SymbolAtom "decor", toSExp dec ] , toSExp dec
, SExpList [ SymbolAtom "implicit", toSExp impl ] , SExpList [ SymbolAtom "implicit", toSExp impl ]
, SExpList [ SymbolAtom "key", StringAtom k ] , SExpList [ SymbolAtom "key", StringAtom k ]
, SExpList [ SymbolAtom "doc-overview", StringAtom doc ] , 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 inFile fname ((file, _, _), _) = file == fname
||| Output some data using current dialog index ||| Output some data using current dialog index
export export
printOutput : {auto o : Ref ROpts REPLOpts} -> printOutput : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} ->
SExp -> Core () SExp -> Core ()
printOutput msg printOutput msg
= do opts <- get ROpts = do opts <- get ROpts
@ -74,7 +81,8 @@ printOutput msg
msg, toSExp i]) msg, toSExp i])
outputHighlight : {auto opts : Ref ROpts REPLOpts} -> outputHighlight : {auto c : Ref Ctxt Defs} ->
{auto opts : Ref ROpts REPLOpts} ->
Highlight -> Core () Highlight -> Core ()
outputHighlight h = outputHighlight h =
printOutput $ SExpList [ SymbolAtom "ok" printOutput $ SExpList [ SymbolAtom "ok"
@ -86,37 +94,46 @@ outputHighlight h =
hlt : List Highlight hlt : List Highlight
hlt = [h] hlt = [h]
outputNameSyntax : {auto opts : Ref ROpts REPLOpts} -> lwOutputHighlight :
(NonEmptyFC, (Name, Nat, ClosedTerm)) -> Core () {auto c : Ref Ctxt Defs} ->
outputNameSyntax (fc, (name, _, term)) = {auto opts : Ref ROpts REPLOpts} ->
let dec = case term of (NonEmptyFC, Decoration) -> Core ()
(Local fc x idx y) => Just Bound 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 outputNameSyntax : {auto c : Ref Ctxt Defs} ->
-- Func : NameType {auto s : Ref Syn SyntaxInfo} ->
-- DataCon : (tag : Int) -> (arity : Nat) -> NameType {auto opts : Ref ROpts REPLOpts} ->
-- TyCon : (tag : Int) -> (arity : Nat) -> NameType (NonEmptyFC, Decoration, Name) -> Core ()
(Ref fc Bound name) => Just Bound outputNameSyntax (nfc, decor, nm) = do
(Ref fc Func name) => Just Function defs <- get Ctxt
(Ref fc (DataCon tag arity) name) => Just Data log "ide-mode.highlight" 20 $ "highlighting at " ++ show nfc
(Ref fc (TyCon tag arity) name) => Just Typ ++ ": " ++ show nm
(Meta fc x y xs) => Just Bound ++ "\nAs: " ++ show decor
(Bind fc x b scope) => Just Bound let fc = justFC nfc
(App fc fn arg) => Just Bound let (mns, name) = displayName nm
(As fc x as pat) => Just Bound outputHighlight $ MkHighlight
(TDelayed fc x y) => Nothing { location = nfc
(TDelay fc x ty arg) => Nothing , name
(TForce fc x y) => Nothing , isImplicit = False
(PrimVal fc c) => Just Typ , key = ""
(Erased fc imp) => Just Bound , decor
(TType fc) => Just Typ , docOverview = "" --!(getDocsForName fc nm)
hilite = Prelude.map (\ d => MkHighlight fc name False "" d "" (show term) "") dec , typ = "" -- TODO: extract type maybe "" show !(lookupTyExact nm (gamma defs))
in maybe (pure ()) outputHighlight hilite , ns = maybe "" show mns
}
export 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} -> {auto opts : Ref ROpts REPLOpts} ->
String -> String ->
REPLResult -> REPLResult ->
@ -124,9 +141,28 @@ outputSyntaxHighlighting : {auto m : Ref MD Metadata} ->
outputSyntaxHighlighting fname loadResult = do outputSyntaxHighlighting fname loadResult = do
opts <- get ROpts opts <- get ROpts
when (opts.synHighlightOn) $ do when (opts.synHighlightOn) $ do
allNames <- the (Core ?) $ filter (inFile fname) . names <$> get MD meta <- get MD
-- decls <- filter (inFile fname) . tydecls <$> get MD let allNames = filter (inFile fname) $ toList meta.nameLocMap
_ <- traverse outputNameSyntax allNames -- ++ decls) --decls <- filter (inFile fname) . tydecls <$> get MD
pure () --_ <- 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 pure loadResult

View File

@ -178,7 +178,7 @@ buildMod loc num len mod
Nothing => True Nothing => True
Just t => any (\x => x > t) (srcTime :: map snd depTimes) Just t => any (\x => x > t) (srcTime :: map snd depTimes)
u <- newRef UST initUState u <- newRef UST initUState
m <- newRef MD initMetadata m <- newRef MD (initMetadata src)
put Syn initSyntax put Syn initSyntax
if needsBuilding if needsBuilding
@ -220,7 +220,7 @@ buildDeps fname
case ok of case ok of
[] => do -- On success, reload the main ttc in a clean context [] => do -- On success, reload the main ttc in a clean context
clearCtxt; addPrimitives clearCtxt; addPrimitives
put MD initMetadata put MD (initMetadata fname)
mainttc <- getTTCFileName fname "ttc" mainttc <- getTTCFileName fname "ttc"
log "import" 10 $ "Reloading " ++ show mainttc ++ " from " ++ fname log "import" 10 $ "Reloading " ++ show mainttc ++ " from " ++ fname
readAsMain mainttc readAsMain mainttc

View File

@ -159,7 +159,7 @@ field fname
pure [LT (MkPkgVersion (fromInteger <$> vs)) True, pure [LT (MkPkgVersion (fromInteger <$> vs)) True,
GT (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 mkBound (LT b i :: bs) pkgbs
= maybe (mkBound bs (record { upperBound = Just b, = maybe (mkBound bs (record { upperBound = Just b,
upperInclusive = i } pkgbs)) upperInclusive = i } pkgbs))
@ -285,7 +285,7 @@ compileMain : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} -> {auto o : Ref ROpts REPLOpts} ->
Name -> String -> String -> Core () Name -> String -> String -> Core ()
compileMain mainn mmod exec compileMain mainn mmod exec
= do m <- newRef MD initMetadata = do m <- newRef MD (initMetadata mmod)
u <- newRef UST initUState u <- newRef UST initUState
ignore $ loadMainFile mmod ignore $ loadMainFile mmod
ignore $ compileExp (PRef replFC mainn) exec ignore $ compileExp (PRef replFC mainn) exec
@ -559,7 +559,7 @@ runRepl : {auto c : Ref Ctxt Defs} ->
Core () Core ()
runRepl fname = do runRepl fname = do
u <- newRef UST initUState u <- newRef UST initUState
m <- newRef MD initMetadata m <- newRef MD (initMetadata $ fromMaybe "(interactive)" fname)
the (Core ()) $ the (Core ()) $
case fname of case fname of
Nothing => pure () Nothing => pure ()

File diff suppressed because it is too large Load Diff

View File

@ -56,7 +56,7 @@ mkLets : FileName ->
List1 (WithBounds (Either LetBinder LetDecl)) -> List1 (WithBounds (Either LetBinder LetDecl)) ->
PTerm -> PTerm PTerm -> PTerm
mkLets fname = letFactory buildLets mkLets fname = letFactory buildLets
(\ decls, scope => PLocal (boundToFC fname decls) decls.val scope) (\ decls, scope => PLocal (virtualiseFC $ boundToFC fname decls) decls.val scope)
where where

View File

@ -294,10 +294,10 @@ mutual
go d (PDotted _ p) = dot <+> go d p go d (PDotted _ p) = dot <+> go d p
go d (PImplicit _) = "_" go d (PImplicit _) = "_"
go d (PInfer _) = "?" go d (PInfer _) = "?"
go d (POp _ op x y) = parenthesise (d > appPrec) $ group $ go startPrec x <++> prettyOp op <++> go startPrec y 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 (PPrefixOp _ _ op x) = parenthesise (d > appPrec) $ pretty op <+> go startPrec x
go d (PSectionL _ op x) = parens (prettyOp 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 (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 (PEq fc l r) = parenthesise (d > appPrec) $ go startPrec l <++> equals <++> go startPrec r
go d (PBracketed _ tm) = parens (go startPrec tm) go d (PBracketed _ tm) = parens (go startPrec tm)
go d (PString _ xs) = parenthesise (d > appPrec) $ hsep $ punctuate "++" (prettyString <$> xs) 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 (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 (PBang _ tm) = "!" <+> go d tm
go d (PIdiom _ tm) = enclose (pretty "[|") (pretty "|]") (go startPrec 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 (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 (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 ty r) = group $ parens (go startPrec l <++> colon <++> go startPrec ty <++> pretty "**" <+> line <+> go startPrec r)
go d (PUnit _) = "()" go d (PUnit _) = "()"
go d (PIfThenElse _ x t e) = go d (PIfThenElse _ x t e) =
parenthesise (d > appPrec) $ group $ align $ hang 2 $ vsep parenthesise (d > appPrec) $ group $ align $ hang 2 $ vsep

View File

@ -197,7 +197,7 @@ readHeader path
-- Stop at the first :, that's definitely not part of the header, to -- Stop at the first :, that's definitely not part of the header, to
-- save lexing the whole file unnecessarily -- save lexing the whole file unnecessarily
setCurrentElabSource res -- for error printing purposes 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 | Left err => throw err
pure mod pure mod
@ -259,15 +259,16 @@ processMod srcf ttcf msg sourcecode
pure Nothing pure Nothing
else -- needs rebuilding else -- needs rebuilding
do iputStrLn msg 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)) pure (runParser srcf (isLitFile srcf) sourcecode (do p <- prog srcf; eoi; pure p))
| Left err => pure (Just [err]) | Left err => pure (Just [err])
addSemanticDecorations decor
initHash initHash
traverse_ addPublicHash (sort hs) traverse_ addPublicHash (sort hs)
resetNextVar resetNextVar
when (ns /= nsAsModuleIdent mainNS) $ when (ns /= nsAsModuleIdent mainNS) $
do let MkFC fname _ _ = headerloc mod do let Just fname = map file (isNonEmptyFC $ headerloc mod)
| EmptyFC => throw (InternalError "No file name") | Nothing => throw (InternalError "No file name")
d <- getDirs d <- getDirs
ns' <- pathToNS (working_dir d) (source_dir d) fname ns' <- pathToNS (working_dir d) (source_dir d) fname
when (ns /= ns') $ when (ns /= ns') $

View File

@ -18,6 +18,7 @@ import Core.Env
import Core.InitPrimitives import Core.InitPrimitives
import Core.LinearCheck import Core.LinearCheck
import Core.Metadata import Core.Metadata
import Core.FC
import Core.Normalise import Core.Normalise
import Core.Options import Core.Options
import Core.Termination import Core.Termination
@ -638,7 +639,7 @@ loadMainFile : {auto c : Ref Ctxt Defs} ->
loadMainFile f loadMainFile f
= do opts <- get ROpts = do opts <- get ROpts
put ROpts (record { evalResultName = Nothing } opts) put ROpts (record { evalResultName = Nothing } opts)
resetContext resetContext f
Right res <- coreLift (readFile f) Right res <- coreLift (readFile f)
| Left err => do setSource "" | Left err => do setSource ""
pure (ErrorLoadingFile f err) pure (ErrorLoadingFile f err)
@ -970,16 +971,18 @@ processCatch cmd
pure $ REPLError msg pure $ REPLError msg
) )
parseEmptyCmd : SourceEmptyRule (Maybe REPLCmd) parseEmptyCmd : EmptyRule (Maybe REPLCmd)
parseEmptyCmd = eoi *> (pure Nothing) parseEmptyCmd = eoi *> (pure Nothing)
parseCmd : SourceEmptyRule (Maybe REPLCmd) parseCmd : EmptyRule (Maybe REPLCmd)
parseCmd = do c <- command; eoi; pure $ Just c parseCmd = do c <- command; eoi; pure $ Just c
export export
parseRepl : String -> Either Error (Maybe REPLCmd) parseRepl : String -> Either Error (Maybe REPLCmd)
parseRepl inp 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 export
interpret : {auto c : Ref Ctxt Defs} -> interpret : {auto c : Ref Ctxt Defs} ->

View File

@ -29,7 +29,8 @@ import System.File
-- Output informational messages, unless quiet flag is set -- Output informational messages, unless quiet flag is set
export export
iputStrLn : {auto o : Ref ROpts REPLOpts} -> iputStrLn : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} ->
Doc IdrisAnn -> Core () Doc IdrisAnn -> Core ()
iputStrLn msg iputStrLn msg
= do opts <- get ROpts = do opts <- get ROpts
@ -118,8 +119,7 @@ emitWarnings
put Ctxt (record { warnings = [] } defs) put Ctxt (record { warnings = [] } defs)
getFCLine : FC -> Maybe Int getFCLine : FC -> Maybe Int
getFCLine (MkFC _ (line, _) _) = Just line getFCLine = map startLine . isNonEmptyFC
getFCLine EmptyFC = Nothing
export export
updateErrorLine : {auto o : Ref ROpts REPLOpts} -> updateErrorLine : {auto o : Ref ROpts REPLOpts} ->
@ -136,14 +136,15 @@ resetContext : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
{auto s : Ref Syn SyntaxInfo} -> {auto s : Ref Syn SyntaxInfo} ->
{auto m : Ref MD Metadata} -> {auto m : Ref MD Metadata} ->
(source : String) ->
Core () Core ()
resetContext resetContext fname
= do defs <- get Ctxt = do defs <- get Ctxt
put Ctxt (record { options = clearNames (options defs) } !initDefs) put Ctxt (record { options = clearNames (options defs) } !initDefs)
addPrimitives addPrimitives
put UST initUState put UST initUState
put Syn initSyntax put Syn initSyntax
put MD initMetadata put MD (initMetadata fname)
public export public export
data EditResult : Type where data EditResult : Type where

View File

@ -29,11 +29,11 @@ unbracketApp tm = tm
-- TODO: Deal with precedences -- TODO: Deal with precedences
mkOp : {auto s : Ref Syn SyntaxInfo} -> mkOp : {auto s : Ref Syn SyntaxInfo} ->
PTerm -> Core PTerm 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 = do syn <- get Syn
case StringMap.lookup (nameRoot n) (infixes syn) of case StringMap.lookup (nameRoot n) (infixes syn) of
Nothing => pure tm 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 mkOp tm = pure tm
export export
@ -44,10 +44,10 @@ addBracket fc tm = if needed tm then PBracketed fc tm else tm
needed (PBracketed _ _) = False needed (PBracketed _ _) = False
needed (PRef _ _) = False needed (PRef _ _) = False
needed (PPair _ _ _) = False needed (PPair _ _ _) = False
needed (PDPair _ _ _ _) = False needed (PDPair _ _ _ _ _) = False
needed (PUnit _) = False needed (PUnit _) = False
needed (PComprehension _ _ _) = False needed (PComprehension _ _ _) = False
needed (PList _ _) = False needed (PList _ _ _) = False
needed (PPrimVal _ _) = False needed (PPrimVal _ _) = False
needed tm = True needed tm = True
@ -113,13 +113,13 @@ mutual
||| Put the special names (Nil, ::, Pair, Z, S, etc) back as syntax ||| Put the special names (Nil, ::, Pair, Z, S, etc) back as syntax
||| Returns `Nothing` in case there was nothing to resugar. ||| Returns `Nothing` in case there was nothing to resugar.
sugarAppM : PTerm -> Maybe PTerm 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 if builtinNS == ns
then case nameRoot nm of then case nameRoot nm of
"Pair" => pure $ PPair fc (unbracket l) (unbracket r) "Pair" => pure $ PPair fc (unbracket l) (unbracket r)
"MkPair" => pure $ PPair fc (unbracket l) (unbracket r) "MkPair" => pure $ PPair fc (unbracket l) (unbracket r)
"DPair" => case unbracket r of "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 _ => Nothing
"Equal" => pure $ PEq fc (unbracket l) (unbracket r) "Equal" => pure $ PEq fc (unbracket l) (unbracket r)
"===" => pure $ PEq fc (unbracket l) (unbracket r) "===" => pure $ PEq fc (unbracket l) (unbracket r)
@ -127,8 +127,8 @@ mutual
_ => Nothing _ => Nothing
else if nameRoot nm == "::" else if nameRoot nm == "::"
then case sugarApp (unbracket r) of then case sugarApp (unbracket r) of
PList fc xs => pure $ PList fc (unbracketApp l :: xs) PList fc nilFC xs => pure $ PList fc nilFC ((opFC, unbracketApp l) :: xs)
_ => Nothing _ => Nothing
else Nothing else Nothing
sugarAppM tm = sugarAppM tm =
-- refolding natural numbers if the expression is a constant -- refolding natural numbers if the expression is a constant
@ -142,7 +142,7 @@ mutual
"MkUnit" => pure $ PUnit fc "MkUnit" => pure $ PUnit fc
_ => Nothing _ => Nothing
else if nameRoot nm == "Nil" else if nameRoot nm == "Nil"
then pure $ PList fc [] then pure $ PList fc fc []
else Nothing else Nothing
_ => Nothing _ => Nothing
@ -464,14 +464,14 @@ cleanPTerm ptm
cleanNode : PTerm -> Core PTerm cleanNode : PTerm -> Core PTerm
cleanNode (PRef fc nm) = cleanNode (PRef fc nm) =
PRef fc <$> cleanName nm PRef fc <$> cleanName nm
cleanNode (POp fc op x y) = cleanNode (POp fc opFC op x y) =
(\ op => POp fc op x y) <$> cleanName op (\ op => POp fc opFC op x y) <$> cleanName op
cleanNode (PPrefixOp fc op x) = cleanNode (PPrefixOp fc opFC op x) =
(\ op => PPrefixOp fc op x) <$> cleanName op (\ op => PPrefixOp fc opFC op x) <$> cleanName op
cleanNode (PSectionL fc op x) = cleanNode (PSectionL fc opFC op x) =
(\ op => PSectionL fc op x) <$> cleanName op (\ op => PSectionL fc opFC op x) <$> cleanName op
cleanNode (PSectionR fc x op) = cleanNode (PSectionR fc opFC x op) =
PSectionR fc x <$> cleanName op PSectionR fc opFC x <$> cleanName op
cleanNode tm = pure tm cleanNode tm = pure tm
toCleanPTerm : {auto c : Ref Ctxt Defs} -> toCleanPTerm : {auto c : Ref Ctxt Defs} ->

View File

@ -82,10 +82,10 @@ mutual
-- Operators -- Operators
POp : FC -> OpStr -> PTerm -> PTerm -> PTerm POp : (full, opFC : FC) -> OpStr -> PTerm -> PTerm -> PTerm
PPrefixOp : FC -> OpStr -> PTerm -> PTerm PPrefixOp : (full, opFC : FC) -> OpStr -> PTerm -> PTerm
PSectionL : FC -> OpStr -> PTerm -> PTerm PSectionL : (full, opFC : FC) -> OpStr -> PTerm -> PTerm
PSectionR : FC -> PTerm -> OpStr -> PTerm PSectionR : (full, opFC : FC) -> PTerm -> OpStr -> PTerm
PEq : FC -> PTerm -> PTerm -> PTerm PEq : FC -> PTerm -> PTerm -> PTerm
PBracketed : FC -> PTerm -> PTerm PBracketed : FC -> PTerm -> PTerm
@ -95,9 +95,10 @@ mutual
PDoBlock : FC -> Maybe Namespace -> List PDo -> PTerm PDoBlock : FC -> Maybe Namespace -> List PDo -> PTerm
PBang : FC -> PTerm -> PTerm PBang : FC -> PTerm -> PTerm
PIdiom : 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 PPair : FC -> PTerm -> PTerm -> PTerm
PDPair : FC -> PTerm -> PTerm -> PTerm -> PTerm PDPair : (full, opFC : FC) -> PTerm -> PTerm -> PTerm -> PTerm
PUnit : FC -> PTerm PUnit : FC -> PTerm
PIfThenElse : FC -> PTerm -> PTerm -> PTerm -> PTerm PIfThenElse : FC -> PTerm -> PTerm -> PTerm -> PTerm
PComprehension : FC -> PTerm -> List PDo -> PTerm PComprehension : FC -> PTerm -> List PDo -> PTerm
@ -146,10 +147,10 @@ mutual
getPTermLoc (PDotted fc _) = fc getPTermLoc (PDotted fc _) = fc
getPTermLoc (PImplicit fc) = fc getPTermLoc (PImplicit fc) = fc
getPTermLoc (PInfer fc) = fc getPTermLoc (PInfer fc) = fc
getPTermLoc (POp fc _ _ _) = fc getPTermLoc (POp fc _ _ _ _) = fc
getPTermLoc (PPrefixOp fc _ _) = fc getPTermLoc (PPrefixOp fc _ _ _) = fc
getPTermLoc (PSectionL fc _ _) = fc getPTermLoc (PSectionL fc _ _ _) = fc
getPTermLoc (PSectionR fc _ _) = fc getPTermLoc (PSectionR fc _ _ _) = fc
getPTermLoc (PEq fc _ _) = fc getPTermLoc (PEq fc _ _) = fc
getPTermLoc (PBracketed fc _) = fc getPTermLoc (PBracketed fc _) = fc
getPTermLoc (PString fc _) = fc getPTermLoc (PString fc _) = fc
@ -157,9 +158,9 @@ mutual
getPTermLoc (PDoBlock fc _ _) = fc getPTermLoc (PDoBlock fc _ _) = fc
getPTermLoc (PBang fc _) = fc getPTermLoc (PBang fc _) = fc
getPTermLoc (PIdiom fc _) = fc getPTermLoc (PIdiom fc _) = fc
getPTermLoc (PList fc _) = fc getPTermLoc (PList fc _ _) = fc
getPTermLoc (PPair fc _ _) = fc getPTermLoc (PPair fc _ _) = fc
getPTermLoc (PDPair fc _ _ _) = fc getPTermLoc (PDPair fc _ _ _ _) = fc
getPTermLoc (PUnit fc) = fc getPTermLoc (PUnit fc) = fc
getPTermLoc (PIfThenElse fc _ _ _) = fc getPTermLoc (PIfThenElse fc _ _ _) = fc
getPTermLoc (PComprehension fc _ _) = fc getPTermLoc (PComprehension fc _ _) = fc
@ -360,10 +361,15 @@ mutual
getPDeclLoc (PDirective fc _) = fc getPDeclLoc (PDirective fc _) = fc
getPDeclLoc (PBuiltin fc _ _) = fc getPDeclLoc (PBuiltin fc _ _) = fc
export export
isPDef : PDecl -> Maybe (FC, List PClause) isStrInterp : PStr -> Maybe FC
isPDef (PDef annot cs) = Just (annot, cs) isStrInterp (StrInterp fc _) = Just fc
isPDef _ = Nothing isStrInterp (StrLiteral _ _) = Nothing
export
isPDef : PDecl -> Maybe (FC, List PClause)
isPDef (PDef annot cs) = Just (annot, cs)
isPDef _ = Nothing
definedInData : PDataDecl -> List Name definedInData : PDataDecl -> List Name
@ -604,10 +610,10 @@ mutual
showPrec d (PDotted _ p) = "." ++ showPrec d p showPrec d (PDotted _ p) = "." ++ showPrec d p
showPrec _ (PImplicit _) = "_" showPrec _ (PImplicit _) = "_"
showPrec _ (PInfer _) = "?" showPrec _ (PInfer _) = "?"
showPrec d (POp _ op x y) = showPrec d x ++ " " ++ showPrecOp d op ++ " " ++ showPrec d y 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 (PPrefixOp _ _ op x) = showPrec d op ++ showPrec d x
showPrec d (PSectionL _ op x) = "(" ++ showPrecOp 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 (PSectionR _ _ x op) = "(" ++ showPrec d x ++ " " ++ showPrecOp d op ++ ")"
showPrec d (PEq fc l r) = showPrec d l ++ " = " ++ showPrec d r showPrec d (PEq fc l r) = showPrec d l ++ " = " ++ showPrec d r
showPrec d (PBracketed _ tm) = "(" ++ showPrec d tm ++ ")" showPrec d (PBracketed _ tm) = "(" ++ showPrec d tm ++ ")"
showPrec d (PString _ xs) = join " ++ " $ show <$> xs showPrec d (PString _ xs) = join " ++ " $ show <$> xs
@ -616,11 +622,11 @@ mutual
= "do " ++ showSep " ; " (map showDo ds) = "do " ++ showSep " ; " (map showDo ds)
showPrec d (PBang _ tm) = "!" ++ showPrec d tm showPrec d (PBang _ tm) = "!" ++ showPrec d tm
showPrec d (PIdiom _ tm) = "[|" ++ showPrec d tm ++ "|]" showPrec d (PIdiom _ tm) = "[|" ++ showPrec d tm ++ "|]"
showPrec d (PList _ xs) showPrec d (PList _ _ xs)
= "[" ++ showSep ", " (map (showPrec d) xs) ++ "]" = "[" ++ showSep ", " (map (showPrec d . snd) xs) ++ "]"
showPrec d (PPair _ l r) = "(" ++ showPrec d l ++ ", " ++ showPrec d r ++ ")" 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 (PImplicit _) r) = "(" ++ showPrec d l ++ " ** " ++ showPrec d r ++ ")"
showPrec d (PDPair _ l ty r) = "(" ++ showPrec d l ++ " : " ++ showPrec d ty ++ showPrec d (PDPair _ _ l ty r) = "(" ++ showPrec d l ++ " : " ++ showPrec d ty ++
" ** " ++ showPrec d r ++ ")" " ** " ++ showPrec d r ++ ")"
showPrec _ (PUnit _) = "()" showPrec _ (PUnit _) = "()"
showPrec d (PIfThenElse _ x t e) = "if " ++ showPrec d x ++ " then " ++ showPrec d t ++ showPrec d (PIfThenElse _ x t e) = "if " ++ showPrec d x ++ " then " ++ showPrec d t ++
@ -941,18 +947,18 @@ mapPTermM f = goPTerm where
>>= f >>= f
goPTerm t@(PImplicit _) = f t goPTerm t@(PImplicit _) = f t
goPTerm t@(PInfer _) = f t goPTerm t@(PInfer _) = f t
goPTerm (POp fc x y z) = goPTerm (POp fc opFC x y z) =
POp fc x <$> goPTerm y POp fc opFC x <$> goPTerm y
<*> goPTerm z <*> goPTerm z
>>= f >>= f
goPTerm (PPrefixOp fc x y) = goPTerm (PPrefixOp fc opFC x y) =
PPrefixOp fc x <$> goPTerm y PPrefixOp fc opFC x <$> goPTerm y
>>= f >>= f
goPTerm (PSectionL fc x y) = goPTerm (PSectionL fc opFC x y) =
PSectionL fc x <$> goPTerm y PSectionL fc opFC x <$> goPTerm y
>>= f >>= f
goPTerm (PSectionR fc x y) = goPTerm (PSectionR fc opFC x y) =
PSectionR fc <$> goPTerm x PSectionR fc opFC <$> goPTerm x
<*> pure y <*> pure y
>>= f >>= f
goPTerm (PEq fc x y) = goPTerm (PEq fc x y) =
@ -977,15 +983,16 @@ mapPTermM f = goPTerm where
goPTerm (PIdiom fc x) = goPTerm (PIdiom fc x) =
PIdiom fc <$> goPTerm x PIdiom fc <$> goPTerm x
>>= f >>= f
goPTerm (PList fc xs) = goPTerm (PList fc nilFC xs) =
PList fc <$> goPTerms xs PList fc nilFC <$> goPairedPTerms xs
>>= f >>= f
goPTerm (PPair fc x y) = goPTerm (PPair fc x y) =
PPair fc <$> goPTerm x PPair fc <$> goPTerm x
<*> goPTerm y <*> goPTerm y
>>= f >>= f
goPTerm (PDPair fc x y z) = goPTerm (PDPair fc opFC x y z) =
PDPair fc <$> goPTerm x PDPair fc opFC
<$> goPTerm x
<*> goPTerm y <*> goPTerm y
<*> goPTerm z <*> goPTerm z
>>= f >>= f

View File

@ -577,18 +577,27 @@ greater : FilePos -> Interval -> Bool
greater k (MkInterval (MkRange low _)) = fst low > k greater k (MkInterval (MkRange low _)) = fst low > k
greater k NoInterval = False greater k NoInterval = False
-- Finds all the interval that overlaps with the given interval. ||| Finds all the intervals that overlap with the given interval.
-- takeUntil selects all the intervals within the given upper bound, export
-- however the remaining interval are not necessarily adjacent in
-- the sequence, thus it drops elements until the next intersecting
-- interval with dropUntil.
inRange : MeasureRM a => FilePos -> FilePos -> PosMap a -> List a inRange : MeasureRM a => FilePos -> FilePos -> PosMap a -> List a
inRange low high t = matches (takeUntil (greater high) t) 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 where matches : PosMap a -> List a
matches xs = case viewl (dropUntil (atleast low) xs) of matches xs = case viewl (dropUntil (atleast low) xs) of
EmptyL => [] EmptyL => []
x <: xs' => x :: assert_total (matches xs') 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. ||| Returns all the interval that contains the given point.
export export
searchPos : MeasureRM a => FilePos -> PosMap a -> List a searchPos : MeasureRM a => FilePos -> PosMap a -> List a

View File

@ -15,7 +15,7 @@ import public Libraries.Text.Token
export export
match : (Eq k, TokenKind k) => match : (Eq k, TokenKind k) =>
(kind : k) -> (kind : k) ->
Grammar (Token k) True (TokType kind) Grammar state (Token k) True (TokType kind)
match k = terminal "Unrecognised input" $ match k = terminal "Unrecognised input" $
\t => if t.kind == k \t => if t.kind == k
then Just $ tokValue k t.text then Just $ tokValue k t.text
@ -25,8 +25,8 @@ match k = terminal "Unrecognised input" $
||| match. May match the empty input. ||| match. May match the empty input.
export export
option : {c : Bool} -> option : {c : Bool} ->
(def : a) -> (p : Grammar tok c a) -> (def : a) -> (p : Grammar state tok c a) ->
Grammar tok False a Grammar state tok False a
option {c = False} def p = p <|> pure def option {c = False} def p = p <|> pure def
option {c = True} 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. ||| To provide a default value, use `option` instead.
export export
optional : {c : Bool} -> optional : {c : Bool} ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok False (Maybe a) Grammar state tok False (Maybe a)
optional p = option Nothing (map Just p) optional p = option Nothing (map Just p)
||| Try to parse one thing or the other, producing a value that indicates ||| 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. ||| takes priority.
export export
choose : {c1, c2 : Bool} -> choose : {c1, c2 : Bool} ->
(l : Grammar tok c1 a) -> (l : Grammar state tok c1 a) ->
(r : Grammar tok c2 b) -> (r : Grammar state tok c2 b) ->
Grammar tok (c1 && c2) (Either a b) Grammar state tok (c1 && c2) (Either a b)
choose l r = map Left l <|> map Right r choose l r = map Left l <|> map Right r
||| Produce a grammar by applying a function to each element of a container, ||| 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. ||| container is empty.
export export
choiceMap : {c : Bool} -> choiceMap : {c : Bool} ->
(a -> Grammar tok c b) -> (a -> Grammar state tok c b) ->
Foldable t => t a -> 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 choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in
f x <|> acc) f x <|> acc)
(fail "No more options") xs (fail "No more options") xs
@ -67,28 +67,28 @@ choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in
export export
choice : Foldable t => choice : Foldable t =>
{c : Bool} -> {c : Bool} ->
t (Grammar tok c a) -> t (Grammar state tok c a) ->
Grammar tok c a Grammar state tok c a
choice = choiceMap id choice = choiceMap id
mutual mutual
||| Parse one or more things ||| Parse one or more things
export export
some : Grammar tok True a -> some : Grammar state tok True a ->
Grammar tok True (List1 a) Grammar state tok True (List1 a)
some p = pure (!p ::: !(many p)) some p = pure (!p ::: !(many p))
||| Parse zero or more things (may match the empty input) ||| Parse zero or more things (may match the empty input)
export export
many : Grammar tok True a -> many : Grammar state tok True a ->
Grammar tok False (List a) Grammar state tok False (List a)
many p = option [] (forget <$> some p) many p = option [] (forget <$> some p)
mutual mutual
private private
count1 : (q : Quantity) -> count1 : (q : Quantity) ->
(p : Grammar tok True a) -> (p : Grammar state tok True a) ->
Grammar tok True (List a) Grammar state tok True (List a)
count1 q p = do x <- p count1 q p = do x <- p
seq (count q p) seq (count q p)
(\xs => pure (x :: xs)) (\xs => pure (x :: xs))
@ -96,8 +96,8 @@ mutual
||| Parse `p`, repeated as specified by `q`, returning the list of values. ||| Parse `p`, repeated as specified by `q`, returning the list of values.
export export
count : (q : Quantity) -> count : (q : Quantity) ->
(p : Grammar tok True a) -> (p : Grammar state tok True a) ->
Grammar tok (isSucc (min q)) (List a) Grammar state tok (isSucc (min q)) (List a)
count (Qty Z Nothing) p = many p count (Qty Z Nothing) p = many p
count (Qty Z (Just Z)) _ = pure [] count (Qty Z (Just Z)) _ = pure []
count (Qty Z (Just (S max))) p = option [] $ count1 (atMost max) p 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. ||| list of values from `p`. Guaranteed to consume input.
export export
someTill : {c : Bool} -> someTill : {c : Bool} ->
(end : Grammar tok c e) -> (end : Grammar state tok c e) ->
(p : Grammar tok True a) -> (p : Grammar state tok True a) ->
Grammar tok True (List1 a) Grammar state tok True (List1 a)
someTill {c} end p = do x <- p someTill {c} end p = do x <- p
seq (manyTill end p) seq (manyTill end p)
(\xs => pure (x ::: xs)) (\xs => pure (x ::: xs))
@ -121,9 +121,9 @@ mutual
||| list of values from `p`. Guaranteed to consume input if `end` consumes. ||| list of values from `p`. Guaranteed to consume input if `end` consumes.
export export
manyTill : {c : Bool} -> manyTill : {c : Bool} ->
(end : Grammar tok c e) -> (end : Grammar state tok c e) ->
(p : Grammar tok True a) -> (p : Grammar state tok True a) ->
Grammar tok c (List a) Grammar state tok c (List a)
manyTill {c} end p = rewrite sym (andTrueNeutral c) in manyTill {c} end p = rewrite sym (andTrueNeutral c) in
map (const []) end <|> (forget <$> someTill end p) map (const []) end <|> (forget <$> someTill end p)
@ -132,9 +132,9 @@ mutual
||| returning its value. ||| returning its value.
export export
afterSome : {c : Bool} -> afterSome : {c : Bool} ->
(skip : Grammar tok True s) -> (skip : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok True a Grammar state tok True a
afterSome skip p = do ignore $ skip afterSome skip p = do ignore $ skip
afterMany skip p afterMany skip p
@ -142,18 +142,18 @@ mutual
||| returning its value. ||| returning its value.
export export
afterMany : {c : Bool} -> afterMany : {c : Bool} ->
(skip : Grammar tok True s) -> (skip : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok c a Grammar state tok c a
afterMany {c} skip p = rewrite sym (andTrueNeutral c) in afterMany {c} skip p = rewrite sym (andTrueNeutral c) in
p <|> afterSome skip p p <|> afterSome skip p
||| Parse one or more things, each separated by another thing. ||| Parse one or more things, each separated by another thing.
export export
sepBy1 : {c : Bool} -> sepBy1 : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok c (List1 a) Grammar state tok c (List1 a)
sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
[| p ::: many (sep *> p) |] [| p ::: many (sep *> p) |]
@ -161,18 +161,18 @@ sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
||| match the empty input. ||| match the empty input.
export export
sepBy : {c : Bool} -> sepBy : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok False (List a) Grammar state tok False (List a)
sepBy sep p = option [] $ forget <$> sepBy1 sep p sepBy sep p = option [] $ forget <$> sepBy1 sep p
||| Parse one or more instances of `p` separated by and optionally terminated by ||| Parse one or more instances of `p` separated by and optionally terminated by
||| `sep`. ||| `sep`.
export export
sepEndBy1 : {c : Bool} -> sepEndBy1 : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok c (List1 a) Grammar state tok c (List1 a)
sepEndBy1 {c} sep p = rewrite sym (orFalseNeutral c) in sepEndBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
sepBy1 sep p <* optional sep 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. ||| by `sep`. Will not match a separator by itself.
export export
sepEndBy : {c : Bool} -> sepEndBy : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok False (List a) Grammar state tok False (List a)
sepEndBy sep p = option [] $ forget <$> sepEndBy1 sep p sepEndBy sep p = option [] $ forget <$> sepEndBy1 sep p
||| Parse one or more instances of `p`, separated and terminated by `sep`. ||| Parse one or more instances of `p`, separated and terminated by `sep`.
export export
endBy1 : {c : Bool} -> endBy1 : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok True (List1 a) Grammar state tok True (List1 a)
endBy1 {c} sep p = some $ rewrite sym (orTrueTrue c) in endBy1 {c} sep p = some $ rewrite sym (orTrueTrue c) in
p <* sep p <* sep
export export
endBy : {c : Bool} -> endBy : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok False (List a) Grammar state tok False (List a)
endBy sep p = option [] $ forget <$> endBy1 sep p endBy sep p = option [] $ forget <$> endBy1 sep p
||| Parse an instance of `p` that is between `left` and `right`. ||| Parse an instance of `p` that is between `left` and `right`.
export export
between : {c : Bool} -> between : {c : Bool} ->
(left : Grammar tok True l) -> (left : Grammar state tok True l) ->
(right : Grammar tok True r) -> (right : Grammar state tok True r) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok True a Grammar state tok True a
between left right contents = left *> contents <* right 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 ()

View File

@ -16,37 +16,39 @@ import public Libraries.Text.Bounded
||| to be non-empty - that is, successfully parsing the language is guaranteed ||| to be non-empty - that is, successfully parsing the language is guaranteed
||| to consume some input. ||| to consume some input.
export export
data Grammar : (tok : Type) -> (consumes : Bool) -> Type -> Type where data Grammar : (state : Type) -> (tok : Type) -> (consumes : Bool) -> Type -> Type where
Empty : (val : ty) -> Grammar tok False ty Empty : (val : ty) -> Grammar state tok False ty
Terminal : String -> (tok -> Maybe a) -> Grammar tok True a Terminal : String -> (tok -> Maybe a) -> Grammar state tok True a
NextIs : String -> (tok -> Bool) -> Grammar tok False tok NextIs : String -> (tok -> Bool) -> Grammar state tok False tok
EOF : Grammar tok False () EOF : Grammar state tok False ()
Fail : (location : Maybe Bounds) -> Bool -> String -> Grammar tok c ty Fail : (location : Maybe Bounds) -> Bool -> String -> Grammar state tok c ty
Try : Grammar tok c ty -> Grammar tok c ty Try : Grammar state tok c ty -> Grammar state tok c ty
Commit : Grammar tok False () Commit : Grammar state tok False ()
MustWork : Grammar tok c a -> Grammar tok c a MustWork : Grammar state tok c a -> Grammar state tok c a
SeqEat : {c2 : Bool} -> SeqEat : {c2 : Bool} ->
Grammar tok True a -> Inf (a -> Grammar tok c2 b) -> Grammar state tok True a -> Inf (a -> Grammar state tok c2 b) ->
Grammar tok True b Grammar state tok True b
SeqEmpty : {c1, c2 : Bool} -> SeqEmpty : {c1, c2 : Bool} ->
Grammar tok c1 a -> (a -> Grammar tok c2 b) -> Grammar state tok c1 a -> (a -> Grammar state tok c2 b) ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
ThenEat : {c2 : Bool} -> ThenEat : {c2 : Bool} ->
Grammar tok True () -> Inf (Grammar tok c2 a) -> Grammar state tok True () -> Inf (Grammar state tok c2 a) ->
Grammar tok True a Grammar state tok True a
ThenEmpty : {c1, c2 : Bool} -> ThenEmpty : {c1, c2 : Bool} ->
Grammar tok c1 () -> Grammar tok c2 a -> Grammar state tok c1 () -> Grammar state tok c2 a ->
Grammar tok (c1 || c2) a Grammar state tok (c1 || c2) a
Alt : {c1, c2 : Bool} -> Alt : {c1, c2 : Bool} ->
Grammar tok c1 ty -> Lazy (Grammar tok c2 ty) -> Grammar state tok c1 ty -> Lazy (Grammar state tok c2 ty) ->
Grammar tok (c1 && c2) ty Grammar state tok (c1 && c2) ty
Bounds : Grammar tok c ty -> Grammar tok c (WithBounds ty) Bounds : Grammar state tok c ty -> Grammar state tok c (WithBounds ty)
Position : Grammar tok False Bounds Position : Grammar state tok False Bounds
Act : state -> Grammar state tok False ()
||| Sequence two grammars. If either consumes some input, the sequence is ||| Sequence two grammars. If either consumes some input, the sequence is
||| guaranteed to consume some input. If the first one consumes input, the ||| 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) ||| consumed and therefore the input is smaller)
export %inline export %inline
(>>=) : {c1, c2 : Bool} -> (>>=) : {c1, c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
inf c1 (a -> Grammar tok c2 b) -> inf c1 (a -> Grammar state tok c2 b) ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
(>>=) {c1 = False} = SeqEmpty (>>=) {c1 = False} = SeqEmpty
(>>=) {c1 = True} = SeqEat (>>=) {c1 = True} = SeqEat
@ -66,9 +68,9 @@ export %inline
||| consumed and therefore the input is smaller) ||| consumed and therefore the input is smaller)
public export %inline %tcinline public export %inline %tcinline
(>>) : {c1, c2 : Bool} -> (>>) : {c1, c2 : Bool} ->
Grammar tok c1 () -> Grammar state tok c1 () ->
inf c1 (Grammar tok c2 a) -> inf c1 (Grammar state tok c2 a) ->
Grammar tok (c1 || c2) a Grammar state tok (c1 || c2) a
(>>) {c1 = False} = ThenEmpty (>>) {c1 = False} = ThenEmpty
(>>) {c1 = True} = ThenEat (>>) {c1 = True} = ThenEat
@ -77,23 +79,23 @@ public export %inline %tcinline
||| of `>>=`. ||| of `>>=`.
export %inline export %inline
seq : {c1,c2 : Bool} -> seq : {c1,c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
(a -> Grammar tok c2 b) -> (a -> Grammar state tok c2 b) ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
seq = SeqEmpty seq = SeqEmpty
||| Sequence a grammar followed by the grammar it returns. ||| Sequence a grammar followed by the grammar it returns.
export %inline export %inline
join : {c1,c2 : Bool} -> join : {c1,c2 : Bool} ->
Grammar tok c1 (Grammar tok c2 a) -> Grammar state tok c1 (Grammar state tok c2 a) ->
Grammar tok (c1 || c2) a Grammar state tok (c1 || c2) a
join {c1 = False} p = SeqEmpty p id join {c1 = False} p = SeqEmpty p id
join {c1 = True} p = SeqEat p id join {c1 = True} p = SeqEat p id
||| Allows the result of a grammar to be mapped to a different value. ||| Allows the result of a grammar to be mapped to a different value.
export export
{c : _} -> {c : _} ->
Functor (Grammar tok c) where Functor (Grammar state tok c) where
map f (Empty val) = Empty (f val) map f (Empty val) = Empty (f val)
map f (Fail bd fatal msg) = Fail bd fatal msg map f (Fail bd fatal msg) = Fail bd fatal msg
map f (Try g) = Try (map f g) map f (Try g) = Try (map f g)
@ -119,9 +121,9 @@ Functor (Grammar tok c) where
||| guaranteed to consume. ||| guaranteed to consume.
export %inline export %inline
(<|>) : {c1,c2 : Bool} -> (<|>) : {c1,c2 : Bool} ->
Grammar tok c1 ty -> Grammar state tok c1 ty ->
Lazy (Grammar tok c2 ty) -> Lazy (Grammar state tok c2 ty) ->
Grammar tok (c1 && c2) ty Grammar state tok (c1 && c2) ty
(<|>) = Alt (<|>) = Alt
infixr 2 <||> infixr 2 <||>
@ -129,9 +131,9 @@ infixr 2 <||>
||| combination is guaranteed to consume. ||| combination is guaranteed to consume.
export export
(<||>) : {c1,c2 : Bool} -> (<||>) : {c1,c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
Lazy (Grammar tok c2 b) -> Lazy (Grammar state tok c2 b) ->
Grammar tok (c1 && c2) (Either a b) Grammar state tok (c1 && c2) (Either a b)
(<||>) p q = (Left <$> p) <|> (Right <$> q) (<||>) p q = (Left <$> p) <|> (Right <$> q)
||| Sequence a grammar with value type `a -> b` and a grammar ||| Sequence a grammar with value type `a -> b` and a grammar
@ -140,33 +142,37 @@ export
||| Guaranteed to consume if either grammar consumes. ||| Guaranteed to consume if either grammar consumes.
export %inline export %inline
(<*>) : {c1, c2 : Bool} -> (<*>) : {c1, c2 : Bool} ->
Grammar tok c1 (a -> b) -> Grammar state tok c1 (a -> b) ->
Grammar tok c2 a -> Grammar state tok c2 a ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
(<*>) x y = SeqEmpty x (\f => map f y) (<*>) x y = SeqEmpty x (\f => map f y)
||| Sequence two grammars. If both succeed, use the value of the first one. ||| Sequence two grammars. If both succeed, use the value of the first one.
||| Guaranteed to consume if either grammar consumes. ||| Guaranteed to consume if either grammar consumes.
export %inline export %inline
(<*) : {c1,c2 : Bool} -> (<*) : {c1,c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
Grammar tok c2 b -> Grammar state tok c2 b ->
Grammar tok (c1 || c2) a Grammar state tok (c1 || c2) a
(<*) x y = map const x <*> y (<*) x y = map const x <*> y
||| Sequence two grammars. If both succeed, use the value of the second one. ||| Sequence two grammars. If both succeed, use the value of the second one.
||| Guaranteed to consume if either grammar consumes. ||| Guaranteed to consume if either grammar consumes.
export %inline export %inline
(*>) : {c1,c2 : Bool} -> (*>) : {c1,c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
Grammar tok c2 b -> Grammar state tok c2 b ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
(*>) x y = map (const id) x <*> y (*>) 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 ||| Produce a grammar that can parse a different type of token by providing a
||| function converting the new token type into the original one. ||| function converting the new token type into the original one.
export 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 (Empty val) = Empty val
mapToken f (Terminal msg g) = Terminal msg (g . f) mapToken f (Terminal msg g) = Terminal msg (g . f)
mapToken f (NextIs msg g) = SeqEmpty (NextIs msg (g . f)) (Empty . 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 (Alt x y) = Alt (mapToken f x) (mapToken f y)
mapToken f (Bounds act) = Bounds (mapToken f act) mapToken f (Bounds act) = Bounds (mapToken f act)
mapToken f Position = Position mapToken f Position = Position
mapToken f (Act action) = Act action
||| Always succeed with the given value. ||| Always succeed with the given value.
export %inline export %inline
pure : (val : ty) -> Grammar tok False ty pure : (val : ty) -> Grammar state tok False ty
pure = Empty pure = Empty
||| Check whether the next token satisfies a predicate ||| Check whether the next token satisfies a predicate
export %inline export %inline
nextIs : String -> (tok -> Bool) -> Grammar tok False tok nextIs : String -> (tok -> Bool) -> Grammar state tok False tok
nextIs = NextIs nextIs = NextIs
||| Look at the next token in the input ||| Look at the next token in the input
export %inline export %inline
peek : Grammar tok False tok peek : Grammar state tok False tok
peek = nextIs "Unrecognised token" (const True) peek = nextIs "Unrecognised token" (const True)
||| Succeeds if running the predicate on the next token returns Just x, ||| Succeeds if running the predicate on the next token returns Just x,
||| returning x. Otherwise fails. ||| returning x. Otherwise fails.
export %inline export %inline
terminal : String -> (tok -> Maybe a) -> Grammar tok True a terminal : String -> (tok -> Maybe a) -> Grammar state tok True a
terminal = Terminal terminal = Terminal
||| Always fail with a message ||| Always fail with a message
export %inline export %inline
fail : String -> Grammar tok c ty fail : String -> Grammar state tok c ty
fail = Fail Nothing False fail = Fail Nothing False
||| Always fail with a message and a location ||| Always fail with a message and a location
export %inline export %inline
failLoc : Bounds -> String -> Grammar tok c ty failLoc : Bounds -> String -> Grammar state tok c ty
failLoc b = Fail (Just b) False failLoc b = Fail (Just b) False
export %inline export %inline
fatalError : String -> Grammar tok c ty fatalError : String -> Grammar state tok c ty
fatalError = Fail Nothing True fatalError = Fail Nothing True
export %inline export %inline
fatalLoc : Bounds -> String -> Grammar tok c ty fatalLoc : Bounds -> String -> Grammar state tok c ty
fatalLoc b = Fail (Just b) True fatalLoc b = Fail (Just b) True
||| Catch a fatal error ||| Catch a fatal error
export %inline export %inline
try : Grammar tok c ty -> Grammar tok c ty try : Grammar state tok c ty -> Grammar state tok c ty
try = Try try = Try
||| Succeed if the input is empty ||| Succeed if the input is empty
export %inline export %inline
eof : Grammar tok False () eof : Grammar state tok False ()
eof = EOF eof = EOF
||| Commit to an alternative; if the current branch of an alternative ||| Commit to an alternative; if the current branch of an alternative
||| fails to parse, no more branches will be tried ||| fails to parse, no more branches will be tried
export %inline export %inline
commit : Grammar tok False () commit : Grammar state tok False ()
commit = Commit commit = Commit
||| If the parser fails, treat it as a fatal error ||| If the parser fails, treat it as a fatal error
export %inline 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 mustWork = MustWork
export %inline 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 bounds = Bounds
export %inline export %inline
position : Grammar tok False Bounds position : Grammar state tok False Bounds
position = Position position = Position
data ParseResult : Type -> Type -> Type where data ParseResult : Type -> Type -> Type -> Type where
Failure : (committed : Bool) -> (fatal : Bool) -> Failure : (committed : Bool) -> (fatal : Bool) ->
(err : String) -> (location : Maybe Bounds) -> ParseResult tok ty (err : String) -> (location : Maybe Bounds) -> ParseResult state tok ty
Res : (committed : Bool) -> Res : state -> (committed : Bool) ->
(val : WithBounds ty) -> (more : List (WithBounds tok)) -> ParseResult tok ty (val : WithBounds ty) -> (more : List (WithBounds tok)) -> ParseResult state tok ty
mergeWith : WithBounds ty -> ParseResult tok sy -> ParseResult tok sy mergeWith : WithBounds ty -> ParseResult state tok sy -> ParseResult state tok sy
mergeWith x (Res committed val more) = Res committed (mergeBounds x val) more mergeWith x (Res s committed val more) = Res s committed (mergeBounds x val) more
mergeWith x v = v mergeWith x v = v
doParse : (commit : Bool) -> doParse : Semigroup state => state -> (commit : Bool) ->
(act : Grammar tok c ty) -> (act : Grammar state tok c ty) ->
(xs : List (WithBounds tok)) -> (xs : List (WithBounds tok)) ->
ParseResult tok ty ParseResult state tok ty
doParse com (Empty val) xs = Res com (irrelevantBounds val) xs doParse s com (Empty val) xs = Res s com (irrelevantBounds val) xs
doParse com (Fail location fatal str) xs = Failure com fatal str (maybe (bounds <$> head' xs) Just location) doParse s 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 doParse s com (Try g) xs = case doParse s com g xs of
-- recover from fatal match but still propagate the 'commit' -- recover from fatal match but still propagate the 'commit'
Failure com _ msg ts => Failure com False msg ts Failure com _ msg ts => Failure com False msg ts
res => res res => res
doParse com Commit xs = Res True (irrelevantBounds ()) xs doParse s com Commit xs = Res s True (irrelevantBounds ()) xs
doParse com (MustWork g) xs = doParse s com (MustWork g) xs =
case assert_total (doParse com g xs) of case assert_total (doParse s com g xs) of
Failure com' _ msg ts => Failure com' True msg ts Failure com' _ msg ts => Failure com' True msg ts
res => res res => res
doParse com (Terminal err f) [] = Failure com False "End of input" Nothing doParse s com (Terminal err f) [] = Failure com False "End of input" Nothing
doParse com (Terminal err f) (x :: xs) = doParse s com (Terminal err f) (x :: xs) =
case f x.val of case f x.val of
Nothing => Failure com False err (Just x.bounds) Nothing => Failure com False err (Just x.bounds)
Just a => Res com (const a <$> x) xs Just a => Res s com (const a <$> x) xs
doParse com EOF [] = Res com (irrelevantBounds ()) [] doParse s com EOF [] = Res s com (irrelevantBounds ()) []
doParse com EOF (x :: xs) = Failure com False "Expected end of input" (Just x.bounds) doParse s 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 s com (NextIs err f) [] = Failure com False "End of input" Nothing
doParse com (NextIs err f) (x :: xs) doParse s com (NextIs err f) (x :: xs)
= if f x.val = 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) else Failure com False err (Just x.bounds)
doParse com (Alt {c1} {c2} x y) xs doParse s com (Alt {c1} {c2} x y) xs
= case doParse False x xs of = case doParse s False x xs of
Failure com' fatal msg ts Failure com' fatal msg ts
=> if com' || fatal => if com' || fatal
-- If the alternative had committed, don't try the -- If the alternative had committed, don't try the
-- other branch (and reset commit flag) -- other branch (and reset commit flag)
then Failure com fatal msg ts 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 -- Successfully parsed the first option, so use the outer commit flag
Res _ val xs => Res com val xs Res s _ val xs => Res s com val xs
doParse com (SeqEmpty act next) xs doParse s com (SeqEmpty act next) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res s com v xs =>
mergeWith v (assert_total $ doParse com (next v.val) xs) mergeWith v (assert_total $ doParse s com (next v.val) xs)
doParse com (SeqEat act next) xs doParse s com (SeqEat act next) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res s com v xs =>
mergeWith v (assert_total $ doParse com (next v.val) xs) mergeWith v (assert_total $ doParse s com (next v.val) xs)
doParse com (ThenEmpty act next) xs doParse s com (ThenEmpty act next) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res s com v xs =>
mergeWith v (assert_total $ doParse com next xs) mergeWith v (assert_total $ doParse s com next xs)
doParse com (ThenEat act next) xs doParse s com (ThenEat act next) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res s com v xs =>
mergeWith v (assert_total $ doParse com next xs) mergeWith v (assert_total $ doParse s com next xs)
doParse com (Bounds act) xs doParse s com (Bounds act) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res com (const v <$> v) xs Res s com v xs => Res s com (const v <$> v) xs
doParse com Position [] = Failure com False "End of input" Nothing doParse s com Position [] = Failure com False "End of input" Nothing
doParse com Position (x :: xs) doParse s com Position (x :: xs)
= Res com (irrelevantBounds x.bounds) (x :: xs) = Res s com (irrelevantBounds x.bounds) (x :: xs)
doParse s com (Act action) xs = Res (s <+> action) com (irrelevantBounds ()) xs
public export public export
data ParsingError tok = Error String (Maybe Bounds) 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 ||| returns a pair of the parse result and the unparsed tokens (the remaining
||| input). ||| input).
export 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)) Either (ParsingError tok) (ty, List (WithBounds tok))
parse act xs parse act xs
= case doParse False act xs of = case doParse neutral False act xs of
Failure _ _ msg ts => Left (Error msg ts) 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)

View File

@ -136,6 +136,9 @@ Eq PathTokenKind where
PathToken : Type PathToken : Type
PathToken = Token PathTokenKind PathToken = Token PathTokenKind
PathGrammar : Bool -> Type -> Type
PathGrammar = Grammar () PathToken
TokenKind PathTokenKind where TokenKind PathTokenKind where
TokType PTText = String TokType PTText = String
TokType (PTPunct _) = () TokType (PTPunct _) = ()
@ -156,7 +159,7 @@ lexPath : String -> List (WithBounds PathToken)
lexPath str = let (tokens, _, _, _) = lex pathTokenMap str in tokens lexPath str = let (tokens, _, _, _) = lex pathTokenMap str in tokens
-- match both '/' and '\\' regardless of the platform. -- match both '/' and '\\' regardless of the platform.
bodySeparator : Grammar PathToken True () bodySeparator : PathGrammar True ()
bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/') bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/')
-- Windows will automatically translate '/' to '\\'. And the verbatim prefix, -- 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. -- However, we just parse it and ignore it.
-- --
-- Example: \\?\ -- Example: \\?\
verbatim : Grammar PathToken True () verbatim : PathGrammar True ()
verbatim = verbatim =
do do
ignore $ count (exactly 2) $ match $ PTPunct '\\' ignore $ count (exactly 2) $ match $ PTPunct '\\'
@ -173,7 +176,7 @@ verbatim =
pure () pure ()
-- Example: \\server\share -- Example: \\server\share
unc : Grammar PathToken True Volume unc : PathGrammar True Volume
unc = unc =
do do
ignore $ count (exactly 2) $ match $ PTPunct '\\' ignore $ count (exactly 2) $ match $ PTPunct '\\'
@ -183,7 +186,7 @@ unc =
pure $ UNC server share pure $ UNC server share
-- Example: \\?\server\share -- Example: \\?\server\share
verbatimUnc : Grammar PathToken True Volume verbatimUnc : PathGrammar True Volume
verbatimUnc = verbatimUnc =
do do
verbatim verbatim
@ -193,7 +196,7 @@ verbatimUnc =
pure $ UNC server share pure $ UNC server share
-- Example: C: -- Example: C:
disk : Grammar PathToken True Volume disk : PathGrammar True Volume
disk = disk =
do do
text <- match PTText text <- match PTText
@ -204,31 +207,31 @@ disk =
pure $ Disk (toUpper disk) pure $ Disk (toUpper disk)
-- Example: \\?\C: -- Example: \\?\C:
verbatimDisk : Grammar PathToken True Volume verbatimDisk : PathGrammar True Volume
verbatimDisk = verbatimDisk =
do do
verbatim verbatim
disk <- disk disk <- disk
pure disk pure disk
parseVolume : Grammar PathToken True Volume parseVolume : PathGrammar True Volume
parseVolume = parseVolume =
verbatimUnc verbatimUnc
<|> verbatimDisk <|> verbatimDisk
<|> unc <|> unc
<|> disk <|> disk
parseBody : Grammar PathToken True Body parseBody : PathGrammar True Body
parseBody = parseBody =
do do
text <- match PTText text <- match PTText
the (Grammar _ False _) $ the (PathGrammar False _) $
case text of case text of
".." => pure ParentDir ".." => pure ParentDir
"." => pure CurDir "." => pure CurDir
normal => pure (Normal normal) normal => pure (Normal normal)
parsePath : Grammar PathToken False Path parsePath : PathGrammar False Path
parsePath = parsePath =
do do
vol <- optional parseVolume vol <- optional parseVolume

View File

@ -21,8 +21,10 @@ data OpPrec
-- Tokens are either operators or already parsed expressions in some -- Tokens are either operators or already parsed expressions in some
-- higher level language -- higher level language
public export public export
data Tok op a = Op FC op OpPrec data Tok op a
| Expr 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 -- The result of shunting is a parse tree with the precedences made explicit
-- in the tree. -- in the tree.
@ -34,14 +36,14 @@ data Tok op a = Op FC op OpPrec
-- there's a better way though! -- there's a better way though!
public export public export
data Tree op a = Infix FC op (Tree op a) (Tree op a) data Tree op a = Infix FC FC op (Tree op a) (Tree op a)
| Pre FC op (Tree op a) | Pre FC FC op (Tree op a)
| Leaf a | Leaf a
export export
(Show op, Show a) => Show (Tree op a) where (Show op, Show a) => Show (Tree op a) where
show (Infix _ op l r) = "(" ++ show op ++ " " ++ show l ++ " " ++ show r ++ ")" show (Infix _ _ op l r) = "(" ++ show op ++ " " ++ show l ++ " " ++ show r ++ ")"
show (Pre _ op l) = "(" ++ show op ++ " " ++ show l ++ ")" show (Pre _ _ op l) = "(" ++ show op ++ " " ++ show l ++ ")"
show (Leaf val) = show val show (Leaf val) = show val
Show OpPrec where Show OpPrec where
@ -52,7 +54,7 @@ Show OpPrec where
export export
(Show op, Show a) => Show (Tok op a) where (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 show (Expr val) = show val
-- Label for the output queue state -- Label for the output queue state
@ -60,9 +62,9 @@ data Out : Type where
output : List (Tree op a) -> Tok op a -> output : List (Tree op a) -> Tok op a ->
Core (List (Tree op a)) Core (List (Tree op a))
output [] (Op _ _ _) = throw (InternalError "Invalid input to shunting") output [] (Op _ _ _ _) = throw (InternalError "Invalid input to shunting")
output (x :: stk) (Op loc str (Prefix _)) = pure $ Pre loc str x :: stk output (x :: stk) (Op loc opFC str (Prefix _)) = pure $ Pre loc opFC str x :: stk
output (x :: y :: stk) (Op loc str _) = pure $ Infix loc str y 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 stk (Expr a) = pure $ Leaf a :: stk
output _ _ = throw (InternalError "Invalid input to shunting") output _ _ = throw (InternalError "Invalid input to shunting")
@ -101,36 +103,32 @@ higher loc opl l opr r
((getPrec l == getPrec r) && isLAssoc l) ((getPrec l == getPrec r) && isLAssoc l)
processStack : Show op => {auto o : Ref Out (List (Tree op a))} -> processStack : Show op => {auto o : Ref Out (List (Tree op a))} ->
List (FC, op, OpPrec) -> op -> OpPrec -> List (FC, FC, op, OpPrec) -> op -> OpPrec ->
Core (List (FC, op, OpPrec)) Core (List (FC, FC, op, OpPrec))
processStack [] op prec = pure [] 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) = 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 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))} -> 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) List (Tok op a) -> Core (Tree op a)
shunt stk (Expr x :: rest) shunt stk (Expr x :: rest)
= do emit (Expr x) = do emit (Expr x)
shunt stk rest shunt stk rest
shunt stk (Op loc op prec :: rest) shunt stk (Op loc opFC op prec :: rest)
= do stk' <- processStack stk op prec = do stk' <- processStack stk op prec
shunt ((loc, op, prec) :: stk') rest shunt ((loc, opFC, op, prec) :: stk') rest
shunt stk [] shunt stk []
= do traverse_ (\s => emit (Op (sloc s) (sop s) (sprec s))) stk = do traverse_ (emit . mkOp) stk
[out] <- get Out [out] <- get Out
| out => throw (InternalError "Invalid input to shunting") | out => throw (InternalError "Invalid input to shunting")
pure out pure out
where where
sloc : (annot, b, c) -> annot mkOp : (FC, FC, op, OpPrec) -> Tok op a
sloc (x, y, z) = x mkOp (loc, opFC, op, prec) = Op loc opFC op prec
sop : (annot, b, c) -> b
sop (x, y, z) = y
sprec : (annot, b, c) -> c
sprec (x, y, z) = z
export export
parseOps : Show op => List (Tok op a) -> Core (Tree op a) parseOps : Show op => List (Tok op a) -> Core (Tree op a)

View File

@ -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

View File

@ -1,7 +1,6 @@
module Parser.Rule.Package module Parser.Rule.Package
import public Parser.Lexer.Package import public Parser.Lexer.Package
import public Parser.Rule.Common
import Data.List import Data.List
import Data.List1 import Data.List1
@ -12,11 +11,11 @@ import Core.Name.Namespace
public export public export
Rule : Type -> Type Rule : Type -> Type
Rule = Rule Token Rule = Grammar () Token True
public export public export
PackageEmptyRule : Type -> Type EmptyRule : Type -> Type
PackageEmptyRule = EmptyRule Token EmptyRule = Grammar () Token False
export export
equals : Rule () equals : Rule ()

View File

@ -1,10 +1,10 @@
module Parser.Rule.Source module Parser.Rule.Source
import public Parser.Lexer.Source import public Parser.Lexer.Source
import public Parser.Rule.Common
import public Parser.Support import public Parser.Support
import Core.TT import Core.TT
import Core.Metadata
import Data.List1 import Data.List1
import Data.Strings import Data.Strings
import Libraries.Data.List.Extra import Libraries.Data.List.Extra
@ -19,14 +19,14 @@ import Libraries.Data.String.Extra
public export public export
Rule : Type -> Type Rule : Type -> Type
Rule = Rule Token Rule ty = Grammar SemanticDecorations Token True ty
public export public export
SourceEmptyRule : Type -> Type EmptyRule : Type -> Type
SourceEmptyRule = EmptyRule Token EmptyRule ty = Grammar SemanticDecorations Token False ty
export export
eoi : SourceEmptyRule () eoi : EmptyRule ()
eoi = ignore $ nextIs "Expected end of input" isEOI eoi = ignore $ nextIs "Expected end of input" isEOI
where where
isEOI : Token -> Bool isEOI : Token -> Bool
@ -208,11 +208,11 @@ namespacedIdent
Ident i => Just (Nothing, i) Ident i => Just (Nothing, i)
_ => Nothing) _ => Nothing)
isCapitalisedIdent : WithBounds String -> SourceEmptyRule () isCapitalisedIdent : WithBounds String -> EmptyRule ()
isCapitalisedIdent str = isCapitalisedIdent str =
let val = str.val let val = str.val
loc = str.bounds loc = str.bounds
err : SourceEmptyRule () err : EmptyRule ()
= failLoc loc ("Expected a capitalised identifier, got: " ++ val) = failLoc loc ("Expected a capitalised identifier, got: " ++ val)
in case strM val of in case strM val of
StrNil => err StrNil => err
@ -249,7 +249,7 @@ reservedNames
, "String", "Char", "Double", "Lazy", "Inf", "Force", "Delay" , "String", "Char", "Double", "Lazy", "Inf", "Force", "Delay"
] ]
isNotReservedIdent : WithBounds String -> SourceEmptyRule () isNotReservedIdent : WithBounds String -> EmptyRule ()
isNotReservedIdent x isNotReservedIdent x
= if x.val `elem` reservedNames = if x.val `elem` reservedNames
then failLoc x.bounds $ "can't use reserved name " ++ x.val then failLoc x.bounds $ "can't use reserved name " ++ x.val
@ -260,7 +260,7 @@ opNonNS : Rule Name
opNonNS = symbol "(" *> (operator <|> postfixProj) <* symbol ")" opNonNS = symbol "(" *> (operator <|> postfixProj) <* symbol ")"
identWithCapital : (capitalised : Bool) -> WithBounds String -> identWithCapital : (capitalised : Bool) -> WithBounds String ->
SourceEmptyRule () EmptyRule ()
identWithCapital b x = if b then isCapitalisedIdent x else pure () identWithCapital b x = if b then isCapitalisedIdent x else pure ()
nameWithCapital : (capitalised : Bool) -> Rule Name nameWithCapital : (capitalised : Bool) -> Rule Name
@ -269,7 +269,7 @@ nameWithCapital b = opNonNS <|> do
opNS nsx <|> nameNS nsx opNS nsx <|> nameNS nsx
where where
nameNS : WithBounds (Maybe Namespace, String) -> SourceEmptyRule Name nameNS : WithBounds (Maybe Namespace, String) -> EmptyRule Name
nameNS nsx = do nameNS nsx = do
let id = snd <$> nsx let id = snd <$> nsx
identWithCapital b id identWithCapital b id
@ -317,23 +317,22 @@ export
init : IndentInfo init : IndentInfo
init = 0 init = 0
continueF : SourceEmptyRule () -> (indent : IndentInfo) -> SourceEmptyRule () continueF : EmptyRule () -> (indent : IndentInfo) -> EmptyRule ()
continueF err indent continueF err indent
= do eoi; err = do eoi; err
<|> do keyword "where"; err <|> do keyword "where"; err
<|> do col <- Common.column <|> do col <- column
if col <= indent when (col <= indent)
then err err
else pure ()
||| Fail if this is the end of a block entry or end of file ||| Fail if this is the end of a block entry or end of file
export export
continue : (indent : IndentInfo) -> SourceEmptyRule () continue : (indent : IndentInfo) -> EmptyRule ()
continue = continueF (fail "Unexpected end of expression") continue = continueF (fail "Unexpected end of expression")
||| As 'continue' but failing is fatal (i.e. entire parse fails) ||| As 'continue' but failing is fatal (i.e. entire parse fails)
export export
mustContinue : (indent : IndentInfo) -> Maybe String -> SourceEmptyRule () mustContinue : (indent : IndentInfo) -> Maybe String -> EmptyRule ()
mustContinue indent Nothing mustContinue indent Nothing
= continueF (fatalError "Unexpected end of expression") indent = continueF (fatalError "Unexpected end of expression") indent
mustContinue indent (Just req) mustContinue indent (Just req)
@ -355,7 +354,7 @@ Show ValidIndent where
show (AfterPos i) = "[after " ++ show i ++ "]" show (AfterPos i) = "[after " ++ show i ++ "]"
show EndOfBlock = "[EOB]" show EndOfBlock = "[EOB]"
checkValid : ValidIndent -> Int -> SourceEmptyRule () checkValid : ValidIndent -> Int -> EmptyRule ()
checkValid AnyIndent c = pure () checkValid AnyIndent c = pure ()
checkValid (AtPos x) c = if c == x checkValid (AtPos x) c = if c == x
then pure () then pure ()
@ -386,29 +385,27 @@ isTerminator _ = False
||| It's the end if we have a terminating token, or the next token starts ||| 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. ||| in or before indent. Works by looking ahead but not consuming.
export export
atEnd : (indent : IndentInfo) -> SourceEmptyRule () atEnd : (indent : IndentInfo) -> EmptyRule ()
atEnd indent atEnd indent
= eoi = eoi
<|> do ignore $ nextIs "Expected end of block" isTerminator <|> do ignore $ nextIs "Expected end of block" isTerminator
<|> do col <- Common.column <|> do col <- column
if (col <= indent) when (not (col <= indent))
then pure () $ fail "Not the end of a block entry"
else fail "Not the end of a block entry"
-- Check we're at the end, but only by looking at indentation -- Check we're at the end, but only by looking at indentation
export export
atEndIndent : (indent : IndentInfo) -> SourceEmptyRule () atEndIndent : (indent : IndentInfo) -> EmptyRule ()
atEndIndent indent atEndIndent indent
= eoi = eoi
<|> do col <- Common.column <|> do col <- column
if col <= indent when (not (col <= indent))
then pure () $ fail "Not the end of a block entry"
else fail "Not the end of a block entry"
-- Parse a terminator, return where the next block entry -- Parse a terminator, return where the next block entry
-- must start, given where the current block entry started -- must start, given where the current block entry started
terminator : ValidIndent -> Int -> SourceEmptyRule ValidIndent terminator : ValidIndent -> Int -> EmptyRule ValidIndent
terminator valid laststart terminator valid laststart
= do eoi = do eoi
pure EndOfBlock pure EndOfBlock
@ -430,7 +427,7 @@ terminator valid laststart
-- Expected indentation for the next token can either be anything (if -- Expected indentation for the next token can either be anything (if
-- we're inside a brace delimited block) or in exactly the initial column -- we're inside a brace delimited block) or in exactly the initial column
-- (if we're inside an indentation delimited block) -- (if we're inside an indentation delimited block)
afterDedent : ValidIndent -> Int -> SourceEmptyRule ValidIndent afterDedent : ValidIndent -> Int -> EmptyRule ValidIndent
afterDedent AnyIndent col afterDedent AnyIndent col
= if col <= laststart = if col <= laststart
then pure AnyIndent then pure AnyIndent
@ -456,7 +453,7 @@ blockEntry valid rule
pure (p, valid') pure (p, valid')
blockEntries : ValidIndent -> (IndentInfo -> Rule ty) -> blockEntries : ValidIndent -> (IndentInfo -> Rule ty) ->
SourceEmptyRule (List ty) EmptyRule (List ty)
blockEntries valid rule blockEntries valid rule
= do eoi; pure [] = do eoi; pure []
<|> do res <- blockEntry valid rule <|> do res <- blockEntry valid rule
@ -465,7 +462,7 @@ blockEntries valid rule
<|> pure [] <|> pure []
export export
block : (IndentInfo -> Rule ty) -> SourceEmptyRule (List ty) block : (IndentInfo -> Rule ty) -> EmptyRule (List ty)
block item block item
= do symbol "{" = do symbol "{"
commit commit
@ -481,35 +478,35 @@ block item
||| by curly braces). `rule` is a function of the actual indentation ||| by curly braces). `rule` is a function of the actual indentation
||| level. ||| level.
export export
blockAfter : Int -> (IndentInfo -> Rule ty) -> SourceEmptyRule (List ty) blockAfter : Int -> (IndentInfo -> Rule ty) -> EmptyRule (List ty)
blockAfter mincol item blockAfter mincol item
= do symbol "{" = do symbol "{"
commit commit
ps <- blockEntries AnyIndent item ps <- blockEntries AnyIndent item
symbol "}" symbol "}"
pure ps pure ps
<|> do col <- Common.column <|> do col <- column
if col <= mincol ifThenElse (col <= mincol)
then pure [] (pure [])
else blockEntries (AtPos col) item $ blockEntries (AtPos col) item
export export
blockWithOptHeaderAfter : blockWithOptHeaderAfter :
(column : Int) -> (column : Int) ->
(header : IndentInfo -> Rule hd) -> (header : IndentInfo -> Rule hd) ->
(item : IndentInfo -> Rule ty) -> (item : IndentInfo -> Rule ty) ->
SourceEmptyRule (Maybe hd, List ty) EmptyRule (Maybe hd, List ty)
blockWithOptHeaderAfter {ty} mincol header item blockWithOptHeaderAfter {ty} mincol header item
= do symbol "{" = do symbol "{"
commit commit
hidt <- optional $ blockEntry AnyIndent header hidt <- optional $ blockEntry AnyIndent header
restOfBlock hidt restOfBlock hidt
<|> do col <- Common.column <|> do col <- column
if col <= mincol ifThenElse (col <= mincol)
then pure (Nothing, []) (pure (Nothing, []))
else do hidt <- optional $ blockEntry (AtPos col) header $ do hidt <- optional $ blockEntry (AtPos col) header
ps <- blockEntries (AtPos col) item ps <- blockEntries (AtPos col) item
pure (map fst hidt, ps) pure (map fst hidt, ps)
where where
restOfBlock : Maybe (hd, ValidIndent) -> Rule (Maybe hd, List ty) restOfBlock : Maybe (hd, ValidIndent) -> Rule (Maybe hd, List ty)
restOfBlock (Just (h, idt)) = do ps <- blockEntries idt item restOfBlock (Just (h, idt)) = do ps <- blockEntries idt item

View File

@ -5,6 +5,9 @@ import public Parser.Rule.Source
import public Parser.Unlit import public Parser.Unlit
import Core.Core import Core.Core
import Core.Name
import Core.Metadata
import Core.FC
import System.File import System.File
import Libraries.Utils.Either import Libraries.Utils.Either
@ -14,21 +17,21 @@ export
runParserTo : {e : _} -> runParserTo : {e : _} ->
(fname : String) -> (fname : String) ->
Maybe LiterateStyle -> Lexer -> 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 runParserTo fname lit reject str p
= do str <- mapError (fromLitError fname) $ unlit lit str = do str <- mapError (fromLitError fname) $ unlit lit str
toks <- mapError (fromLexError fname) $ lexTo reject str toks <- mapError (fromLexError fname) $ lexTo reject str
parsed <- mapError (fromParsingError fname) $ parse p toks (decs, (parsed, _)) <- mapError (fromParsingError fname) $ parseWith p toks
Right (fst parsed) Right (decs, parsed)
export export
runParser : {e : _} -> runParser : {e : _} ->
(fname : String) -> Maybe LiterateStyle -> String -> (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) runParser fname lit = runParserTo fname lit (pred $ const False)
export covering export covering
parseFile : (fname : String) -> Rule ty -> IO (Either Error ty) parseFile : (fname : String) -> Rule ty -> IO (Either Error (SemanticDecorations, ty))
parseFile fname p parseFile fname p
= do Right str <- readFile fname = do Right str <- readFile fname
| Left err => pure (Left (FileErr fname err)) | Left err => pure (Left (FileErr fname err))

View File

@ -55,6 +55,14 @@ getNameType rigc env fc x
do est <- get EST do est <- get EST
put EST put EST
(record { linearUsed $= ((MkVar lv) :: ) } 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) pure (Local fc (Just (isLet binder)) _ lv, gnf env bty)
Nothing => Nothing =>
do defs <- get Ctxt do defs <- get Ctxt
@ -68,6 +76,18 @@ getNameType rigc env fc x
DCon t a _ => DataCon t a DCon t a _ => DataCon t a
TCon t a _ _ _ _ _ _ => TyCon t a TCon t a _ _ _ _ _ _ => TyCon t a
_ => Func _ => 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))) pure (Ref fc nt (Resolved i), gnf env (embed (type def)))
where where
rigSafe : RigCount -> RigCount -> Core () rigSafe : RigCount -> RigCount -> Core ()
@ -89,7 +109,7 @@ getVarType rigc nest env fc x
Just (nestn, argns, tmf) => Just (nestn, argns, tmf) =>
do defs <- get Ctxt do defs <- get Ctxt
let arglen = length argns let arglen = length argns
let n' = maybe x id nestn let n' = fromMaybe x nestn
case !(lookupCtxtExact n' (gamma defs)) of case !(lookupCtxtExact n' (gamma defs)) of
Nothing => undefinedName fc n' Nothing => undefinedName fc n'
Just ndef => Just ndef =>
@ -110,6 +130,14 @@ getVarType rigc nest env fc x
log "metadata.names" 7 $ "getVarType is adding ↓" log "metadata.names" 7 $ "getVarType is adding ↓"
addNameType fc x env tyenv 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) pure (tm, arglen, gnf env tyenv)
where where
useVars : {vars : _} -> useVars : {vars : _} ->

View File

@ -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 -- 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 -- which case the duplication won't hurt, or if (TODO) none of the
-- case patterns in alts are just a variable -- 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 applyEnv = applyToFull fc caseRef env
let appTm : Term vars let appTm : Term vars

View File

@ -515,7 +515,7 @@ successful : {vars : _} ->
Bool -> -- constraints allowed Bool -> -- constraints allowed
List (Maybe Name, Core a) -> List (Maybe Name, Core a) ->
Core (List (Either (Maybe Name, Error) Core (List (Either (Maybe Name, Error)
(Nat, a, Defs, UState, EState vars))) (Nat, a, Defs, UState, EState vars, Metadata)))
successful allowCons [] = pure [] successful allowCons [] = pure []
successful allowCons ((tm, elab) :: elabs) successful allowCons ((tm, elab) :: elabs)
= do ust <- get UST = do ust <- get UST
@ -555,7 +555,7 @@ successful allowCons ((tm, elab) :: elabs)
elabs' <- successful allowCons elabs elabs' <- successful allowCons elabs
-- Record success, and the state we ended at -- Record success, and the state we ended at
pure (Right (minus ncons' ncons, pure (Right (minus ncons' ncons,
res, defs', ust', est') :: elabs')) res, defs', ust', est', md') :: elabs'))
(\err => do put UST ust (\err => do put UST ust
put EST est put EST est
put MD md put MD md
@ -576,9 +576,10 @@ exactlyOne' allowCons fc env [(tm, elab)] = elab
exactlyOne' {vars} allowCons fc env all exactlyOne' {vars} allowCons fc env all
= do elabs <- successful allowCons all = do elabs <- successful allowCons all
case getRight elabs of case getRight elabs of
Right (res, defs, ust, est) => Right (res, defs, ust, est, md) =>
do put UST ust do put UST ust
put EST est put EST est
put MD md
put Ctxt defs put Ctxt defs
commit commit
pure res pure res

View File

@ -416,6 +416,12 @@ checkBindVar rig elabinfo nest env fc str topexp
noteLHSPatVar elabmode (UN str) noteLHSPatVar elabmode (UN str)
notePatVar n notePatVar n
est <- get EST 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 case lookup n (boundNames est) of
Nothing => Nothing =>
do (tm, exp, bty) <- mkPatternHole fc rig n env do (tm, exp, bty) <- mkPatternHole fc rig n env

View File

@ -177,18 +177,21 @@ recUpdate : {vars : _} ->
List IFieldUpdate -> List IFieldUpdate ->
(rec : RawImp) -> (grecty : Glued vars) -> (rec : RawImp) -> (grecty : Glued vars) ->
Core RawImp Core RawImp
recUpdate rigc elabinfo loc nest env flds rec grecty recUpdate rigc elabinfo iloc nest env flds rec grecty
= do defs <- get Ctxt = do defs <- get Ctxt
rectynf <- getNF grecty rectynf <- getNF grecty
let Just rectyn = getRecordType env rectynf let Just rectyn = getRecordType env rectynf
| Nothing => throw (RecordTypeNeeded loc env) | Nothing => throw (RecordTypeNeeded iloc env)
fldn <- genFieldName "__fld" fldn <- genFieldName "__fld"
sides <- getAllSides loc flds rectyn rec sides <- getAllSides iloc flds rectyn rec
(Field Nothing fldn (IVar loc (UN fldn))) (Field Nothing fldn (IVar vloc (UN fldn)))
pure $ ICase loc rec (Implicit loc False) [mkClause sides] pure $ ICase vloc rec (Implicit vloc False) [mkClause sides]
where where
vloc : FC
vloc = virtualiseFC iloc
mkClause : Rec -> ImpClause 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 : Error -> Bool
needType (RecordTypeNeeded _ _) = True needType (RecordTypeNeeded _ _) = True

View File

@ -107,19 +107,20 @@ checkRewrite : {vars : _} ->
Core (Term vars, Glued vars) Core (Term vars, Glued vars)
checkRewrite rigc elabinfo nest env fc rule tm Nothing checkRewrite rigc elabinfo nest env fc rule tm Nothing
= throw (GenericMsg fc "Can't infer a type for rewrite") = throw (GenericMsg fc "Can't infer a type for rewrite")
checkRewrite {vars} rigc elabinfo nest env fc rule tm (Just expected) checkRewrite {vars} rigc elabinfo nest env ifc rule tm (Just expected)
= delayOnFailure fc rigc env expected rewriteErr 10 (\delayed => = delayOnFailure ifc rigc env expected rewriteErr 10 $ \delayed =>
do (rulev, grulet) <- check erased elabinfo nest env rule Nothing do let vfc = virtualiseFC ifc
(rulev, grulet) <- check erased elabinfo nest env rule Nothing
rulet <- getTerm grulet rulet <- getTerm grulet
expTy <- getTerm expected expTy <- getTerm expected
when delayed $ log "elab.rewrite" 5 "Retrying rewrite" 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 "_" rname <- genVarName "_"
pname <- genVarName "_" pname <- genVarName "_"
let pbind = Let fc erased pred predty let pbind = Let vfc erased pred predty
let rbind = Let fc erased (weaken rulev) (weaken rulet) let rbind = Let vfc erased (weaken rulev) (weaken rulet)
let env' = rbind :: pbind :: env 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, -- implicits for the rewriting lemma are in the right place. But,
-- we still need the right type for the EState, so weaken it once -- we still need the right type for the EState, so weaken it once
-- for each of the let bindings above. -- for each of the let bindings above.
(rwtm, grwty) <- inScope fc (pbind :: env) (rwtm, grwty) <-
(\e' => inScope {e=e'} fc env' inScope vfc (pbind :: env) $ \e' =>
(\e'' => check {e = e''} {vars = rname :: pname :: vars} inScope {e=e'} vfc env' $ \e'' =>
rigc elabinfo (weaken (weaken nest)) env' let offset = mkSizeOf [rname, pname] in
(apply (IVar fc lemma) [IVar fc pname, check {e = e''} rigc elabinfo (weakenNs offset nest) env'
IVar fc rname, (apply (IVar vfc lemma) [IVar vfc pname,
tm]) IVar vfc rname,
(Just (gnf env' tm])
(weakenNs (mkSizeOf [rname, pname]) expTy))) (Just (gnf env' (weakenNs offset expTy)))
))
rwty <- getTerm grwty rwty <- getTerm grwty
pure (Bind fc pname pbind (Bind fc rname rbind rwtm), let binding = Bind vfc pname pbind . Bind vfc rname rbind
gnf env (Bind fc pname pbind (Bind fc rname rbind rwty)))) pure (binding rwtm, gnf env (binding rwty))

View File

@ -330,7 +330,7 @@ mkCase : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
Int -> RawImp -> RawImp -> Core ClauseUpdate Int -> RawImp -> RawImp -> Core ClauseUpdate
mkCase {c} {u} fn orig lhs_raw mkCase {c} {u} fn orig lhs_raw
= do m <- newRef MD initMetadata = do m <- newRef MD (initMetadata "(interactive)")
defs <- get Ctxt defs <- get Ctxt
ust <- get UST ust <- get UST
catch catch

View File

@ -2,6 +2,7 @@ module TTImp.Parser
import Core.Context import Core.Context
import Core.Core import Core.Core
import Core.Metadata
import Core.Env import Core.Env
import Core.TT import Core.TT
import Parser.Source import Parser.Source
@ -81,7 +82,7 @@ visOption
<|> do keyword "private" <|> do keyword "private"
pure Private pure Private
visibility : SourceEmptyRule Visibility visibility : EmptyRule Visibility
visibility visibility
= visOption = visOption
<|> pure Private <|> pure Private
@ -124,7 +125,7 @@ visOpt
pure (Right opt) pure (Right opt)
getVisibility : Maybe Visibility -> List (Either Visibility FnOpt) -> getVisibility : Maybe Visibility -> List (Either Visibility FnOpt) ->
SourceEmptyRule Visibility EmptyRule Visibility
getVisibility Nothing [] = pure Private getVisibility Nothing [] = pure Private
getVisibility (Just vis) [] = pure vis getVisibility (Just vis) [] = pure vis
getVisibility Nothing (Left x :: xs) = getVisibility (Just x) xs getVisibility Nothing (Left x :: xs) = getVisibility (Just x) xs
@ -216,13 +217,13 @@ mutual
symbol ")" symbol ")"
pure e pure e
multiplicity : SourceEmptyRule (Maybe Integer) multiplicity : EmptyRule (Maybe Integer)
multiplicity multiplicity
= do c <- intLit = do c <- intLit
pure (Just c) pure (Just c)
<|> pure Nothing <|> pure Nothing
getMult : Maybe Integer -> SourceEmptyRule RigCount getMult : Maybe Integer -> EmptyRule RigCount
getMult (Just 0) = pure erased getMult (Just 0) = pure erased
getMult (Just 1) = pure linear getMult (Just 1) = pure linear
getMult Nothing = pure top getMult Nothing = pure top
@ -522,7 +523,7 @@ mutual
let fc = MkFC fname start end let fc = MkFC fname start end
pure (!(getFn lhs), ImpossibleClause fc lhs) pure (!(getFn lhs), ImpossibleClause fc lhs)
where where
getFn : RawImp -> SourceEmptyRule Name getFn : RawImp -> EmptyRule Name
getFn (IVar _ n) = pure n getFn (IVar _ n) = pure n
getFn (IApp _ f a) = getFn f getFn (IApp _ f a) = getFn f
getFn (IAutoApp _ f a) = getFn f getFn (IAutoApp _ f a) = getFn f
@ -593,7 +594,7 @@ recordParam fname indents
<|> do symbol "{" <|> do symbol "{"
commit commit
start <- location start <- location
info <- the (SourceEmptyRule (PiInfo RawImp)) info <- the (EmptyRule (PiInfo RawImp))
(pure AutoImplicit <* keyword "auto" (pure AutoImplicit <* keyword "auto"
<|>(do <|>(do
keyword "default" keyword "default"

View File

@ -167,7 +167,7 @@ processTTImpFile : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
String -> Core Bool String -> Core Bool
processTTImpFile fname processTTImpFile fname
= do Right tti <- logTime "Parsing" $ coreLift $ parseFile fname = do Right (decor, tti) <- logTime "Parsing" $ coreLift $ parseFile fname
(do decls <- prog fname (do decls <- prog fname
eoi eoi
pure decls) pure decls)

View File

@ -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)) pure (Right (MkClause env' lhstm' rhstm))
-- TODO: (to decide) With is complicated. Move this into its own module? -- TODO: (to decide) With is complicated. Move this into its own module?
checkClause {vars} mult vis totreq hashit n opts nest env 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))) <- = 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 let wmode
= if isErased mult then InType else InExpr = 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 elabTermSub n wmode opts nest' env' env sub' wval_raw Nothing
clearHoleLHS clearHoleLHS
@ -498,7 +498,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
-- to get the 'magic with' behaviour -- to get the 'magic with' behaviour
(wargs ** (scenv, var, binder)) <- bindWithArgs wvalTy ((,wval) <$> mprf) wvalEnv (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 notreqns = fst bnr
let notreqty = snd bnr let notreqty = snd bnr
@ -511,7 +511,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
(weakenNs (mkSizeOf wargs) notreqty)) (weakenNs (mkSizeOf wargs) notreqty))
let bNotReq = binder wtyScope 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") | Nothing => throw (InternalError "Impossible happened: With abstraction failure #4")
-- list of argument names - 'Just' means we need to match the name -- 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))) wname <- genWithName !(prettyName !(toFullNames (Resolved n)))
widx <- addDef wname (record {flags $= (SetTotal totreq ::)} 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)) vars wtype vis None))
let toWarg : Maybe (PiInfo RawImp, Name) -> List (Maybe Name, RawImp) 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) :: (Nothing, wval_raw) ::
case mprf of case mprf of
Nothing => [] Nothing => []
@ -539,12 +539,12 @@ checkClause {vars} mult vis totreq hashit n opts nest env
let refl = IVar fc (NS builtinNS (UN "Refl")) in let refl = IVar fc (NS builtinNS (UN "Refl")) in
[(mprf, INamedApp fc refl (UN "x") wval_raw)] [(mprf, INamedApp fc refl (UN "x") wval_raw)]
let rhs_in = gapply (IVar fc wname) let rhs_in = gapply (IVar vfc wname)
$ map (\ nm => (Nothing, IVar fc nm)) envns $ map (\ nm => (Nothing, IVar vfc nm)) envns
++ concatMap toWarg wargNames ++ concatMap toWarg wargNames
log "declare.def.clause" 3 $ "Applying to with argument " ++ show rhs_in 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 checkTermSub n wmode opts nest' env' env sub' rhs_in
(gnf env' reqty) (gnf env' reqty)
@ -556,11 +556,14 @@ checkClause {vars} mult vis totreq hashit n opts nest env
nestname <- applyEnv env wname nestname <- applyEnv env wname
let nest'' = record { names $= (nestname ::) } nest let nest'' = record { names $= (nestname ::) } nest
let wdef = IDef fc wname cs' let wdef = IDef ifc wname cs'
processDecl [] nest'' env wdef processDecl [] nest'' env wdef
pure (Right (MkClause env' lhspat rhs)) pure (Right (MkClause env' lhspat rhs))
where where
vfc : FC
vfc = virtualiseFC ifc
bindWithArgs : bindWithArgs :
(wvalTy : Term xs) -> Maybe (Name, Term xs) -> (wvalTy : Term xs) -> Maybe (Name, Term xs) ->
(wvalEnv : Env Term xs) -> (wvalEnv : Env Term xs) ->
@ -576,13 +579,13 @@ checkClause {vars} mult vis totreq hashit n opts nest env
wargs = [wargn] wargs = [wargn]
scenv : Env Term (wargs ++ xs) scenv : Env Term (wargs ++ xs)
:= Pi fc top Explicit wvalTy :: wvalEnv := Pi vfc top Explicit wvalTy :: wvalEnv
var : Term (wargs ++ xs) var : Term (wargs ++ xs)
:= Local fc (Just False) Z First := Local vfc (Just False) Z First
binder : Term (wargs ++ xs) -> Term xs 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)) 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") let eqName = NS builtinNS (UN "Equal")
Just (TCon t ar _ _ _ _ _ _) <- lookupDefExact eqName (gamma defs) Just (TCon t ar _ _ _ _ _ _) <- lookupDefExact eqName (gamma defs)
| _ => throw (InternalError "Cannot find builtin Equal") | _ => 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 let wargn : Name
wargn = MN "warg" 0 wargn = MN "warg" 0
@ -601,24 +604,24 @@ checkClause {vars} mult vis totreq hashit n opts nest env
wvalTy' := weaken wvalTy wvalTy' := weaken wvalTy
eqTy : Term (MN "warg" 0 :: xs) eqTy : Term (MN "warg" 0 :: xs)
:= apply fc eqTyCon := apply vfc eqTyCon
[ wvalTy' [ wvalTy'
, wvalTy' , wvalTy'
, weaken wval , weaken wval
, Local fc (Just False) Z First , Local vfc (Just False) Z First
] ]
scenv : Env Term (wargs ++ xs) scenv : Env Term (wargs ++ xs)
:= Pi fc top Implicit eqTy := Pi vfc top Implicit eqTy
:: Pi fc top Explicit wvalTy :: Pi vfc top Explicit wvalTy
:: wvalEnv :: wvalEnv
var : Term (wargs ++ xs) 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 binder : Term (wargs ++ xs) -> Term xs
:= \ t => Bind fc wargn (Pi fc top Explicit wvalTy) := \ t => Bind vfc wargn (Pi vfc top Explicit wvalTy)
$ Bind fc name (Pi fc top Implicit eqTy) t $ Bind vfc name (Pi vfc top Implicit eqTy) t
pure (wargs ** (scenv, var, binder)) pure (wargs ** (scenv, var, binder))

View File

@ -77,7 +77,7 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields
farg : IField -> farg : IField ->
(FC, Maybe Name, RigCount, PiInfo RawImp, RawImp) (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) -> mkTy : List (FC, Maybe Name, RigCount, PiInfo RawImp, RawImp) ->
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) = IPi fc c imp n argty (mkTy args ret)
recTy : RawImp 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 where
||| Apply argument to list of explicit or implicit named arguments ||| Apply argument to list of explicit or implicit named arguments
apply : RawImp -> List (Name, RawImp, PiInfo RawImp) -> RawImp 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 : Name -> Core ()
elabAsData cname elabAsData cname
= do let conty = mkTy paramTelescope $ = do let fc = virtualiseFC fc
let conty = mkTy paramTelescope $
mkTy (map farg fields) recTy mkTy (map farg fields) recTy
let con = MkImpTy EmptyFC EmptyFC cname !(bindTypeNames [] (map fst params ++ let con = MkImpTy EmptyFC EmptyFC cname !(bindTypeNames [] (map fst params ++
map fname fields ++ vars) conty) 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 -> Env Term vs -> Term vs ->
Core () Core ()
elabGetters con done upds tyenv (Bind bfc n b@(Pi _ rc imp ty_chk) sc) 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 then elabGetters con
(if imp == Explicit && not (n `elem` vars) (if imp == Explicit && not (n `elem` vars)
then S done else done) 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) $ (map fst params ++ map fname fields ++ vars) $
mkTy paramTelescope $ mkTy paramTelescope $
IPi bfc top Explicit (Just rname) recTy ty' IPi bfc top Explicit (Just rname) recTy ty'
log "declare.record.projection" 5 $ "Projection " ++ show rfNameNS ++ " : " ++ show projTy
processDecl [] nest env let mkProjClaim = \ nm =>
(IClaim bfc (if isErased rc let ty = MkImpTy EmptyFC EmptyFC nm projTy
then erased in IClaim bfc rig isVis [Inline] ty
else top) (projVis vis) [Inline] (MkImpTy EmptyFC EmptyFC rfNameNS projTy))
log "declare.record.projection" 5 $
"Projection " ++ show rfNameNS ++ " : " ++ show projTy
processDecl [] nest env (mkProjClaim rfNameNS)
-- Define the LHS and RHS -- Define the LHS and RHS
let lhs_exp 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`! when !isPrefixRecordProjections $ do -- beware: `!` is NOT boolean `not`!
-- Claim the type. -- Claim the type.
-- we just reuse `projTy` defined above -- we just reuse `projTy` defined above
log "declare.record.projection.prefix" 5 $ "Prefix projection " ++ show unNameNS ++ " : " ++ show projTy log "declare.record.projection.prefix" 5 $
processDecl [] nest env "Prefix projection " ++ show unNameNS ++ " : " ++ show projTy
(IClaim bfc (if isErased rc processDecl [] nest env (mkProjClaim unNameNS)
then erased
else top) (projVis vis) [Inline] (MkImpTy EmptyFC EmptyFC unNameNS projTy))
-- Define the LHS and RHS -- Define the LHS and RHS
let lhs = IVar bfc unNameNS let lhs = IVar bfc unNameNS
let rhs = IVar bfc rfNameNS 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 processDecl [] nest env
(IDef bfc unNameNS [PatClause bfc lhs rhs]) (IDef bfc unNameNS [PatClause bfc lhs rhs])

View File

@ -264,6 +264,8 @@ processType : {vars : _} ->
processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc nameFC n_in ty_raw) processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc nameFC n_in ty_raw)
= do n <- inCurrentNS n_in = do n <- inCurrentNS n_in
addNameLoc nameFC n
log "declare.type" 1 $ "Processing " ++ show n log "declare.type" 1 $ "Processing " ++ show n
log "declare.type" 5 $ "Checking type decl " ++ show n ++ " : " ++ show ty_raw log "declare.type" 5 $ "Checking type decl " ++ show n ++ " : " ++ show ty_raw
idx <- resolveName n idx <- resolveName n

View File

@ -542,7 +542,7 @@ implicitsAs n defs ns tm
"\n In the type of " ++ show n ++ ": " ++ show ty ++ "\n In the type of " ++ show n ++ ": " ++ show ty ++
"\n Using locals: " ++ show ns ++ "\n Using locals: " ++ show ns ++
"\n Found implicits: " ++ show implicits "\n Found implicits: " ++ show implicits
pure $ impAs loc implicits (IVar loc nm) pure $ impAs (virtualiseFC loc) implicits (IVar loc nm)
where where
-- If there's an @{c} in the list of given implicits, that's the next -- 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 -- 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 (Implicit x _) = x
getFC (IWithUnambigNames 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 export
apply : RawImp -> List RawImp -> RawImp apply : RawImp -> List RawImp -> RawImp
apply f [] = f 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 export
gapply : RawImp -> List (Maybe Name, RawImp) -> RawImp gapply : RawImp -> List (Maybe Name, RawImp) -> RawImp

View File

@ -2,7 +2,9 @@ module TTImp.WithClause
import Core.Context import Core.Context
import Core.Context.Log import Core.Context.Log
import Core.Metadata
import Core.TT import Core.TT
import TTImp.BindImplicits import TTImp.BindImplicits
import TTImp.TTImp import TTImp.TTImp
import TTImp.Elab.Check import TTImp.Elab.Check
@ -15,30 +17,52 @@ import Data.Maybe
matchFail : FC -> Core a matchFail : FC -> Core a
matchFail loc = throw (GenericMsg loc "With clause does not match parent") 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 mutual
export 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)) 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 (IBindVar _ n) tm = pure [(n, tm)]
getMatch lhs (Implicit _ _) tm = pure [] getMatch lhs (Implicit _ _) tm = pure []
getMatch lhs (IVar _ (NS ns n)) (IVar loc (NS ns' n')) getMatch lhs (IVar to (NS ns n)) (IVar from (NS ns' n'))
= if n == n' && isParentOf ns' ns then pure [] else matchFail loc = if n == n' && isParentOf ns' ns
getMatch lhs (IVar _ (NS ns n)) (IVar loc n') then [] <$ addAlias from to -- <$ decorateName loc nm
= if n == n' then pure [] else matchFail loc else matchFail from
getMatch lhs (IVar _ n) (IVar loc n') getMatch lhs (IVar to (NS ns n)) (IVar from n')
= if n == n' then pure [] else matchFail loc = 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') 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')] then matchAll lhs [(arg, arg'), (ret, ret')]
else matchFail loc 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 -- TODO: Lam, Let, Case, Local, Update
getMatch lhs (IApp _ f a) (IApp loc f' a') getMatch lhs (IApp _ f a) (IApp loc f' a')
= matchAll lhs [(f, f'), (a, a')] = matchAll lhs [(f, f'), (a, a')]
@ -72,7 +96,7 @@ mutual
-- one of them is okay -- one of them is okay
getMatch lhs (IAlternative fc _ as) (IAlternative _ _ as') getMatch lhs (IAlternative fc _ as) (IAlternative _ _ as')
= matchAny fc lhs (zip as 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' = do ms <- getMatch lhs p p'
mergeMatches lhs ((n, IBindVar fc n') :: ms) mergeMatches lhs ((n, IBindVar fc n') :: ms)
getMatch lhs (IAs _ _ _ (UN n) p) p' getMatch lhs (IAs _ _ _ (UN n) p) p'
@ -87,14 +111,18 @@ mutual
else matchFail fc else matchFail fc
getMatch lhs pat spec = matchFail (getFC pat) 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)) Core (List (String, RawImp))
matchAny fc lhs [] = matchFail fc matchAny fc lhs [] = matchFail fc
matchAny fc lhs ((x, y) :: ms) matchAny fc lhs ((x, y) :: ms)
= catch (getMatch lhs x y) = catch (getMatch lhs x y)
(\err => matchAny fc lhs ms) (\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)) Core (List (String, RawImp))
matchAll lhs [] = pure [] matchAll lhs [] = pure []
matchAll lhs ((x, y) :: ms) matchAll lhs ((x, y) :: ms)
@ -102,7 +130,9 @@ mutual
mxy <- getMatch lhs x y mxy <- getMatch lhs x y
mergeMatches lhs (mxy ++ matches) 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)) Core (List (String, RawImp))
mergeMatches lhs [] = pure [] mergeMatches lhs [] = pure []
mergeMatches lhs ((n, tm) :: rest) mergeMatches lhs ((n, tm) :: rest)
@ -110,8 +140,9 @@ mutual
case lookup n rest' of case lookup n rest' of
Nothing => pure ((n, tm) :: rest') Nothing => pure ((n, tm) :: rest')
Just tm' => Just tm' =>
do ignore $ getMatch lhs tm tm' -- just need to know it succeeds do ignore $ getMatch lhs tm tm'
mergeMatches lhs rest -- ^ just need to know it succeeds
pure rest'
-- Get the arguments for the rewritten pattern clause of a with by looking -- Get the arguments for the rewritten pattern clause of a with by looking
-- up how the argument names matched -- up how the argument names matched
@ -138,11 +169,13 @@ getArgMatch ploc mode search warg ms (Just (_, nm))
export export
getNewLHS : {auto c : Ref Ctxt Defs} -> getNewLHS : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
FC -> (drop : Nat) -> NestedNames vars -> FC -> (drop : Nat) -> NestedNames vars ->
Name -> List (Maybe (PiInfo RawImp, Name)) -> Name -> List (Maybe (PiInfo RawImp, Name)) ->
RawImp -> RawImp -> Core RawImp RawImp -> RawImp -> Core RawImp
getNewLHS ploc drop nest wname wargnames lhs_raw patlhs getNewLHS iploc drop nest wname wargnames lhs_raw patlhs
= do (mlhs_raw, wrest) <- dropWithArgs drop patlhs = do let vploc = virtualiseFC iploc
(mlhs_raw, wrest) <- dropWithArgs drop patlhs
autoimp <- isUnboundImplicits autoimp <- isUnboundImplicits
setUnboundImplicits True 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 log "declare.def.clause.with" 20 $ "Modified LHS (with implicits): " ++ show mlhs
let (warg :: rest) = reverse wrest 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 ++ log "declare.def.clause.with" 5 $ show lhs ++ " against " ++ show mlhs ++
" dropping " ++ show (warg :: rest) " dropping " ++ show (warg :: rest)
ms <- getMatch True lhs mlhs ms <- getMatch True lhs mlhs
log "declare.def.clause.with" 5 $ "Matches: " ++ show ms 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 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 log "declare.def.clause.with" 5 $ "New LHS: " ++ show newlhs
pure newlhs pure newlhs
where where
@ -174,11 +208,12 @@ getNewLHS ploc drop nest wname wargnames lhs_raw patlhs
pure (tm, arg :: rest) pure (tm, arg :: rest)
-- Shouldn't happen if parsed correctly, but there's no guarantee that -- Shouldn't happen if parsed correctly, but there's no guarantee that
-- inputs come from parsed source so throw an error. -- 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 -- Find a 'with' application on the RHS and update it
export export
withRHS : {auto c : Ref Ctxt Defs} -> withRHS : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
FC -> (drop : Nat) -> Name -> List (Maybe (PiInfo RawImp, Name)) -> FC -> (drop : Nat) -> Name -> List (Maybe (PiInfo RawImp, Name)) ->
RawImp -> RawImp -> RawImp -> RawImp ->
Core RawImp Core RawImp
@ -196,8 +231,9 @@ withRHS fc drop wname wargnames tm toplhs
updateWith fc tm (arg :: args) updateWith fc tm (arg :: args)
= do log "declare.def.clause.with" 10 $ "With-app: Matching " ++ show toplhs ++ " against " ++ show tm = do log "declare.def.clause.with" 10 $ "With-app: Matching " ++ show toplhs ++ " against " ++ show tm
ms <- getMatch False toplhs tm ms <- getMatch False toplhs tm
hdloc <- getHeadLoc tm
log "declare.def.clause.with" 10 $ "Result: " ++ show ms 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) (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 $ "With args for RHS: " ++ show wargnames
log "declare.def.clause.with" 10 $ "New RHS: " ++ show newrhs log "declare.def.clause.with" 10 $ "New RHS: " ++ show newrhs

View File

@ -47,7 +47,7 @@ yaffleMain : String -> List String -> Core ()
yaffleMain fname args yaffleMain fname args
= do defs <- initDefs = do defs <- initDefs
c <- newRef Ctxt defs c <- newRef Ctxt defs
m <- newRef MD initMetadata m <- newRef MD (initMetadata fname)
u <- newRef UST initUState u <- newRef UST initUState
d <- getDirs d <- getDirs
t <- processArgs args t <- processArgs args

View File

@ -153,4 +153,4 @@ repl
case runParser "(interactive)" Nothing inp command of case runParser "(interactive)" Nothing inp command of
Left err => do coreLift_ (printLn err) Left err => do coreLift_ (printLn err)
repl repl
Right cmd => when !(processCatch cmd) repl Right (decor, cmd) => when !(processCatch cmd) repl

View File

@ -246,7 +246,7 @@ nodeTests = MkTestPool [Node]
ideModeTests : TestPool ideModeTests : TestPool
ideModeTests = MkTestPool [] ideModeTests = MkTestPool []
[ "ideMode001", "ideMode002", "ideMode003", "ideMode004" [ "ideMode001", "ideMode002", "ideMode003", "ideMode004", "ideMode005"
] ]
preludeTests : TestPool preludeTests : TestPool

View File

@ -1,27 +1,68 @@
000018(:protocol-version 2 0) 000018(:protocol-version 2 0)
000038(:write-string "1/1: Building LocType (LocType.idr)" 1) 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) 000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 1) (:end 1 4)) ((:decor :keyword)))))) 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) 00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 6) (:end 1 9)) ((:decor :type)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 11)) ((:decor :keyword)))))) 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) 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)
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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 17) (:end 1 18)) ((:decor :keyword)))))) 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) 000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 20) (:end 1 23)) ((:decor :type)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 25) (:end 1 26)) ((:decor :keyword)))))) 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) 000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 28) (:end 1 31)) ((:decor :type)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 33) (:end 1 37)) ((:decor :keyword)))))) 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) 00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 6) (:end 2 8)) ((:decor :data)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 10)) ((:decor :keyword)))))) 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) 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)
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) 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)
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) 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)
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) 00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 6) (:end 3 9)) ((:decor :data)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 11) (:end 3 11)) ((:decor :keyword)))))) 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) 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)
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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 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) 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)
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) 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)
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) 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)
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) 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) 000015(:return (:ok ()) 1)
000040(:return (:ok "Syntax highlight option changed to False" ()) 2) 000040(:return (:ok "Syntax highlight option changed to False" ()) 2)
000015(:return (:ok ()) 3) 000015(:return (:ok ()) 3)

View File

@ -1,27 +1,68 @@
000018(:protocol-version 2 0) 000018(:protocol-version 2 0)
000038(:write-string "1/1: Building LocType (LocType.idr)" 1) 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) 000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 1) (:end 1 4)) ((:decor :keyword)))))) 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) 00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 6) (:end 1 9)) ((:decor :type)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 11)) ((:decor :keyword)))))) 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) 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)
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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 17) (:end 1 18)) ((:decor :keyword)))))) 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) 000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 20) (:end 1 23)) ((:decor :type)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 25) (:end 1 26)) ((:decor :keyword)))))) 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) 000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 28) (:end 1 31)) ((:decor :type)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 33) (:end 1 37)) ((:decor :keyword)))))) 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) 00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 6) (:end 2 8)) ((:decor :data)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 10)) ((:decor :keyword)))))) 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) 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)
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) 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)
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) 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)
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) 00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 6) (:end 3 9)) ((:decor :data)))))) 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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 11) (:end 3 11)) ((:decor :keyword)))))) 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) 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)
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) 000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 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) 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)
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) 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)
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) 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)
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) 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) 000015(:return (:ok ()) 1)
000037(:return (:ok "Main.Vect : Nat -> Type -> Type" ()) 2) 000037(:return (:ok "Main.Vect : Nat -> Type -> Type" ()) 2)
Alas the file is done, aborting Alas the file is done, aborting

View File

@ -0,0 +1,10 @@
namespace X
export
OneX : Bool -> Bool
namespace Y
export
data Foo = OneX
foo : Bool
foo = OneX (id False)

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -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`

View File

@ -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 ]

View File

@ -0,0 +1,5 @@
hours : List Nat
hours = [1..12]
nats : Stream Nat
nats = [0,1..]

View File

@ -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
}

View File

@ -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

View File

@ -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)

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -0,0 +1 @@
package idemode005

View File

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -0,0 +1 @@
000018(:protocol-version 2 0)

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -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

View File

@ -0,0 +1 @@
00001d((:load-file "Syntax.idr") 1)

View File

@ -0,0 +1 @@
000020((:load-file "Interface.idr") 1)

View File

@ -0,0 +1 @@
000021((:load-file "SimpleData.idr") 1)

View File

@ -0,0 +1 @@
000021((:load-file "Ambiguity.idr") 1)

View File

@ -0,0 +1 @@
00001e((:load-file "Rainbow.idr") 1)

View File

@ -0,0 +1 @@
000025((:load-file "Implementation.idr") 1)

View File

@ -0,0 +1 @@
00001b((:load-file "Case.idr") 1)

View File

@ -0,0 +1 @@
000023((:load-file "RecordUpdate.idr") 1)

View File

@ -0,0 +1 @@
00001b((:load-file "With.idr") 1)

View File

@ -0,0 +1 @@
00001d((:load-file "Ranges.idr") 1)

View File

@ -0,0 +1 @@
000025((:load-file "StringLiterals.idr") 1)

Some files were not shown because too many files have changed in this diff Show More