mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-21 10:41:59 +03:00
[ refactor ] IDE protocol as a separate module hierarchy (#2171)
This commit is contained in:
parent
a09c5082c5
commit
3c532ea35d
@ -129,6 +129,7 @@ modules =
|
||||
Idris.IDEMode.Holes,
|
||||
Idris.IDEMode.MakeClause,
|
||||
Idris.IDEMode.Parser,
|
||||
Idris.IDEMode.Pretty,
|
||||
Idris.IDEMode.REPL,
|
||||
Idris.IDEMode.SyntaxHighlight,
|
||||
Idris.IDEMode.TokenLine,
|
||||
@ -159,6 +160,7 @@ modules =
|
||||
Libraries.Data.NameMap.Traversable,
|
||||
Libraries.Data.PosMap,
|
||||
Libraries.Data.Primitives,
|
||||
Libraries.Data.Span,
|
||||
Libraries.Data.SortedMap,
|
||||
Libraries.Data.SortedSet,
|
||||
Libraries.Data.String.Extra,
|
||||
@ -190,7 +192,6 @@ modules =
|
||||
Libraries.Text.Token,
|
||||
|
||||
Libraries.Utils.Binary,
|
||||
Libraries.Utils.Hex,
|
||||
Libraries.Utils.Octal,
|
||||
Libraries.Utils.Path,
|
||||
Libraries.Utils.Scheme,
|
||||
@ -210,6 +211,17 @@ modules =
|
||||
Parser.Rule.Package,
|
||||
Parser.Rule.Source,
|
||||
|
||||
Protocol.IDE,
|
||||
Protocol.IDE.Command,
|
||||
Protocol.IDE.Decoration,
|
||||
Protocol.IDE.FileContext,
|
||||
Protocol.IDE.Formatting,
|
||||
Protocol.IDE.Holes,
|
||||
Protocol.IDE.Result,
|
||||
Protocol.IDE.Highlight,
|
||||
Protocol.Hex,
|
||||
Protocol.SExp,
|
||||
|
||||
TTImp.BindImplicits,
|
||||
TTImp.Elab,
|
||||
TTImp.Impossible,
|
||||
|
20
idris2protocols.ipkg.hide
Normal file
20
idris2protocols.ipkg.hide
Normal file
@ -0,0 +1,20 @@
|
||||
package idris-protocols
|
||||
version = 0.0.1
|
||||
|
||||
modules
|
||||
= Protocol.Hex
|
||||
, Protocol.IDE
|
||||
, Protocol.IDE.Command
|
||||
, Protocol.IDE.Decoration
|
||||
, Protocol.SExp
|
||||
, Libraries.Data.Span
|
||||
, Protocol.IDE.FileContext
|
||||
, Protocol.IDE.Formatting
|
||||
, Protocol.IDE.Holes
|
||||
, Protocol.IDE.Result
|
||||
, Protocol.IDE.Highlight
|
||||
|
||||
|
||||
--depends =
|
||||
|
||||
sourcedir = "src"
|
@ -14,7 +14,7 @@ import Compiler.ES.TailRec
|
||||
import Compiler.ES.State
|
||||
import Compiler.NoMangle
|
||||
import Libraries.Data.SortedMap
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
import Libraries.Data.String.Extra
|
||||
|
||||
import Data.Vect
|
||||
|
@ -20,7 +20,7 @@ import Data.Vect
|
||||
import System
|
||||
import System.File
|
||||
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
import Libraries.Utils.Path
|
||||
|
||||
%default covering
|
||||
|
@ -11,7 +11,7 @@ import Core.Directory
|
||||
import Core.Name
|
||||
import Core.Options
|
||||
import Core.TT
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
import Libraries.Utils.Path
|
||||
|
||||
import Data.List
|
||||
|
@ -10,7 +10,7 @@ import Core.Directory
|
||||
import Core.Name
|
||||
import Core.Options
|
||||
import Core.TT
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
import Libraries.Utils.Path
|
||||
|
||||
import Data.List
|
||||
|
@ -11,7 +11,7 @@ import Core.Context.Log
|
||||
import Core.Directory
|
||||
import Core.Name
|
||||
import Core.TT
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
import Libraries.Utils.Path
|
||||
|
||||
import Data.List
|
||||
|
@ -15,19 +15,9 @@ import System.File
|
||||
import Libraries.Data.PosMap
|
||||
import Libraries.Utils.Binary
|
||||
|
||||
%default covering
|
||||
import public Protocol.IDE.Decoration as Protocol.IDE
|
||||
|
||||
public export
|
||||
data Decoration : Type where
|
||||
Comment : Decoration
|
||||
Typ : Decoration
|
||||
Function : Decoration
|
||||
Data : Decoration
|
||||
Keyword : Decoration
|
||||
Bound : Decoration
|
||||
Namespace : Decoration
|
||||
Postulate : Decoration
|
||||
Module : Decoration
|
||||
%default covering
|
||||
|
||||
export
|
||||
nameDecoration : Name -> NameType -> Decoration
|
||||
@ -51,34 +41,6 @@ public export
|
||||
SemanticDecorations : Type
|
||||
SemanticDecorations = List ASemanticDecoration
|
||||
|
||||
public export
|
||||
Eq Decoration where
|
||||
Comment == Comment = True
|
||||
Typ == Typ = True
|
||||
Function == Function = True
|
||||
Data == Data = True
|
||||
Keyword == Keyword = True
|
||||
Bound == Bound = True
|
||||
Namespace == Namespace = True
|
||||
Postulate == Postulate = True
|
||||
Module == Module = True
|
||||
_ == _ = False
|
||||
|
||||
-- CAREFUL: this instance is used in SExpable Decoration. If you change
|
||||
-- it then you need to fix the SExpable implementation in order not to
|
||||
-- break the IDE mode.
|
||||
public export
|
||||
Show Decoration where
|
||||
show Comment = "comment"
|
||||
show Typ = "type"
|
||||
show Function = "function"
|
||||
show Data = "data"
|
||||
show Keyword = "keyword"
|
||||
show Bound = "bound"
|
||||
show Namespace = "namespace"
|
||||
show Postulate = "postulate"
|
||||
show Module = "module"
|
||||
|
||||
TTC Decoration where
|
||||
toBuf b Typ = tag 0
|
||||
toBuf b Function = tag 1
|
||||
|
@ -204,9 +204,7 @@ DecEq Namespace where
|
||||
-- TODO: move somewhere more appropriate
|
||||
export
|
||||
showSep : String -> List String -> String
|
||||
showSep sep [] = ""
|
||||
showSep sep [x] = x
|
||||
showSep sep (x :: xs) = x ++ sep ++ showSep sep xs
|
||||
showSep sep = Libraries.Data.String.Extra.join sep
|
||||
|
||||
export
|
||||
showNSWithSep : String -> Namespace -> String
|
||||
|
@ -5,277 +5,40 @@ import Core.Context
|
||||
import Core.Context.Log
|
||||
import Core.Name
|
||||
import public Idris.REPL.Opts
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
|
||||
import System.File
|
||||
|
||||
import public Protocol.IDE
|
||||
import public Protocol.SExp
|
||||
|
||||
%default total
|
||||
|
||||
public export
|
||||
data SExp = SExpList (List SExp)
|
||||
| StringAtom String
|
||||
| BoolAtom Bool
|
||||
| IntegerAtom Integer
|
||||
| SymbolAtom String
|
||||
|
||||
public export
|
||||
data DocMode = Overview | Full
|
||||
|
||||
public export
|
||||
data IDECommand
|
||||
= Interpret String
|
||||
| LoadFile String (Maybe Integer)
|
||||
| TypeOf String (Maybe (Integer, Integer))
|
||||
| NameAt String (Maybe (Integer, Integer))
|
||||
| CaseSplit Integer Integer String
|
||||
| AddClause Integer String
|
||||
-- deprecated: | AddProofClause
|
||||
| AddMissing Integer String
|
||||
| ExprSearch Integer String (List String) Bool
|
||||
| ExprSearchNext
|
||||
| GenerateDef Integer String
|
||||
| GenerateDefNext
|
||||
| MakeLemma Integer String
|
||||
| MakeCase Integer String
|
||||
| MakeWith Integer String
|
||||
| DocsFor String (Maybe DocMode)
|
||||
| Directive String
|
||||
| Apropos String
|
||||
| Metavariables Integer
|
||||
| WhoCalls String
|
||||
| CallsWho String
|
||||
| BrowseNamespace String
|
||||
| NormaliseTerm String -- TODO: implement a Binary lib
|
||||
| ShowTermImplicits String -- and implement Binary (Term)
|
||||
| HideTermImplicits String -- for these four defintions,
|
||||
| ElaborateTerm String -- then change String to Term, as in idris1
|
||||
| PrintDefinition String
|
||||
| ReplCompletions String
|
||||
| EnableSyntax Bool
|
||||
| Version
|
||||
| GetOptions
|
||||
|
||||
readHints : List SExp -> Maybe (List String)
|
||||
readHints [] = Just []
|
||||
readHints (StringAtom s :: rest)
|
||||
= do rest' <- readHints rest
|
||||
pure (s :: rest')
|
||||
readHints _ = Nothing
|
||||
|
||||
export
|
||||
getIDECommand : SExp -> Maybe IDECommand
|
||||
getIDECommand (SExpList [SymbolAtom "interpret", StringAtom cmd])
|
||||
= Just $ Interpret cmd
|
||||
getIDECommand (SExpList [SymbolAtom "load-file", StringAtom fname])
|
||||
= Just $ LoadFile fname Nothing
|
||||
getIDECommand (SExpList [SymbolAtom "load-file", StringAtom fname, IntegerAtom l])
|
||||
= Just $ LoadFile fname (Just l)
|
||||
getIDECommand (SExpList [SymbolAtom "type-of", StringAtom n])
|
||||
= Just $ TypeOf n Nothing
|
||||
getIDECommand (SExpList [SymbolAtom "type-of", StringAtom n,
|
||||
IntegerAtom l, IntegerAtom c])
|
||||
= Just $ TypeOf n (Just (l, c))
|
||||
getIDECommand (SExpList [SymbolAtom "name-at", StringAtom n])
|
||||
= Just $ NameAt n Nothing
|
||||
getIDECommand (SExpList [SymbolAtom "name-at", StringAtom n,
|
||||
IntegerAtom l, IntegerAtom c])
|
||||
= Just $ NameAt n (Just (l, c))
|
||||
getIDECommand (SExpList [SymbolAtom "case-split", IntegerAtom l, IntegerAtom c,
|
||||
StringAtom n])
|
||||
= Just $ CaseSplit l c n
|
||||
getIDECommand (SExpList [SymbolAtom "case-split", IntegerAtom l, StringAtom n])
|
||||
= Just $ CaseSplit l 0 n
|
||||
getIDECommand (SExpList [SymbolAtom "add-clause", IntegerAtom l, StringAtom n])
|
||||
= Just $ AddClause l n
|
||||
getIDECommand (SExpList [SymbolAtom "add-missing", IntegerAtom line, StringAtom n])
|
||||
= Just $ AddMissing line n
|
||||
getIDECommand (SExpList [SymbolAtom "proof-search", IntegerAtom l, StringAtom n])
|
||||
= Just $ ExprSearch l n [] False
|
||||
getIDECommand (SExpList [SymbolAtom "proof-search", IntegerAtom l, StringAtom n, SExpList hs])
|
||||
= (\hs' => ExprSearch l n hs' False) <$> readHints hs
|
||||
getIDECommand (SExpList [SymbolAtom "proof-search", IntegerAtom l, StringAtom n, SExpList hs, SymbolAtom mode])
|
||||
= (\hs' => ExprSearch l n hs' (getMode mode)) <$> readHints hs
|
||||
where
|
||||
getMode : String -> Bool
|
||||
getMode m = m == "all"
|
||||
getIDECommand (SymbolAtom "proof-search-next") = Just ExprSearchNext
|
||||
getIDECommand (SExpList [SymbolAtom "generate-def", IntegerAtom l, StringAtom n])
|
||||
= Just $ GenerateDef l n
|
||||
getIDECommand (SymbolAtom "generate-def-next") = Just GenerateDefNext
|
||||
getIDECommand (SExpList [SymbolAtom "make-lemma", IntegerAtom l, StringAtom n])
|
||||
= Just $ MakeLemma l n
|
||||
getIDECommand (SExpList [SymbolAtom "make-case", IntegerAtom l, StringAtom n])
|
||||
= Just $ MakeCase l n
|
||||
getIDECommand (SExpList [SymbolAtom "make-with", IntegerAtom l, StringAtom n])
|
||||
= Just $ MakeWith l n
|
||||
getIDECommand (SExpList (SymbolAtom "docs-for" :: StringAtom n :: modeTail))
|
||||
= do -- Maybe monad
|
||||
modeOpt <- case modeTail of
|
||||
[] => Just Nothing
|
||||
[SymbolAtom "overview"] => Just $ Just Overview
|
||||
[SymbolAtom "full" ] => Just $ Just Full
|
||||
_ => Nothing
|
||||
Just $ DocsFor n modeOpt
|
||||
getIDECommand (SExpList [SymbolAtom "apropos", StringAtom n])
|
||||
= Just $ Apropos n
|
||||
getIDECommand (SExpList [SymbolAtom "directive", StringAtom n])
|
||||
= Just $ Directive n
|
||||
getIDECommand (SExpList [SymbolAtom "metavariables", IntegerAtom n])
|
||||
= Just $ Metavariables n
|
||||
getIDECommand (SExpList [SymbolAtom "who-calls", StringAtom n])
|
||||
= Just $ WhoCalls n
|
||||
getIDECommand (SExpList [SymbolAtom "calls-who", StringAtom n])
|
||||
= Just $ CallsWho n
|
||||
getIDECommand (SExpList [SymbolAtom "browse-namespace", StringAtom ns])
|
||||
= Just $ BrowseNamespace ns
|
||||
getIDECommand (SExpList [SymbolAtom "normalise-term", StringAtom tm])
|
||||
= Just $ NormaliseTerm tm
|
||||
getIDECommand (SExpList [SymbolAtom "show-term-implicits", StringAtom tm])
|
||||
= Just $ ShowTermImplicits tm
|
||||
getIDECommand (SExpList [SymbolAtom "hide-term-implicits", StringAtom tm])
|
||||
= Just $ HideTermImplicits tm
|
||||
getIDECommand (SExpList [SymbolAtom "elaborate-term", StringAtom tm])
|
||||
= Just $ ElaborateTerm tm
|
||||
getIDECommand (SExpList [SymbolAtom "print-definition", StringAtom n])
|
||||
= Just $ PrintDefinition n
|
||||
getIDECommand (SExpList [SymbolAtom "repl-completions", StringAtom n])
|
||||
= Just $ ReplCompletions n
|
||||
getIDECommand (SExpList [SymbolAtom "enable-syntax", BoolAtom b])
|
||||
= Just $ EnableSyntax b
|
||||
getIDECommand (SymbolAtom "version") = Just Version
|
||||
getIDECommand (SExpList [SymbolAtom "get-options"]) = Just GetOptions
|
||||
getIDECommand _ = Nothing
|
||||
|
||||
export
|
||||
putIDECommand : IDECommand -> SExp
|
||||
putIDECommand (Interpret cmd) = (SExpList [SymbolAtom "interpret", StringAtom cmd])
|
||||
putIDECommand (LoadFile fname Nothing) = (SExpList [SymbolAtom "load-file", StringAtom fname])
|
||||
putIDECommand (LoadFile fname (Just line)) = (SExpList [SymbolAtom "load-file", StringAtom fname, IntegerAtom line])
|
||||
putIDECommand (TypeOf cmd Nothing) = (SExpList [SymbolAtom "type-of", StringAtom cmd])
|
||||
putIDECommand (TypeOf cmd (Just (line, col))) = (SExpList [SymbolAtom "type-of", StringAtom cmd, IntegerAtom line, IntegerAtom col])
|
||||
putIDECommand (NameAt cmd Nothing) = (SExpList [SymbolAtom "name-at", StringAtom cmd])
|
||||
putIDECommand (NameAt cmd (Just (line, col))) = (SExpList [SymbolAtom "name-at", StringAtom cmd, IntegerAtom line, IntegerAtom col])
|
||||
putIDECommand (CaseSplit line col n) = (SExpList [SymbolAtom "case-split", IntegerAtom line, IntegerAtom col, StringAtom n])
|
||||
putIDECommand (AddClause line n) = (SExpList [SymbolAtom "add-clause", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (AddMissing line n) = (SExpList [SymbolAtom "add-missing", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (ExprSearch line n exprs mode) = (SExpList [SymbolAtom "proof-search", IntegerAtom line, StringAtom n, SExpList $ map StringAtom exprs, getMode mode])
|
||||
where
|
||||
getMode : Bool -> SExp
|
||||
getMode True = SymbolAtom "all"
|
||||
getMode False = SymbolAtom "other"
|
||||
putIDECommand ExprSearchNext = SymbolAtom "proof-search-next"
|
||||
putIDECommand (GenerateDef line n) = (SExpList [SymbolAtom "generate-def", IntegerAtom line, StringAtom n])
|
||||
putIDECommand GenerateDefNext = SymbolAtom "generate-def-next"
|
||||
putIDECommand (MakeLemma line n) = (SExpList [SymbolAtom "make-lemma", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (MakeCase line n) = (SExpList [SymbolAtom "make-case", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (MakeWith line n) = (SExpList [SymbolAtom "make-with", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (DocsFor n modeOpt) = let modeTail = case modeOpt of
|
||||
Nothing => []
|
||||
Just Overview => [SymbolAtom "overview"]
|
||||
Just Full => [SymbolAtom "full"] in
|
||||
(SExpList (SymbolAtom "docs-for" :: StringAtom n :: modeTail))
|
||||
putIDECommand (Apropos n) = (SExpList [SymbolAtom "apropos", StringAtom n])
|
||||
putIDECommand (Metavariables n) = (SExpList [SymbolAtom "metavariables", IntegerAtom n])
|
||||
putIDECommand (WhoCalls n) = (SExpList [SymbolAtom "who-calls", StringAtom n])
|
||||
putIDECommand (CallsWho n) = (SExpList [SymbolAtom "calls-who", StringAtom n])
|
||||
putIDECommand (BrowseNamespace ns) = (SExpList [SymbolAtom "browse-namespace", StringAtom ns])
|
||||
putIDECommand (NormaliseTerm tm) = (SExpList [SymbolAtom "normalise-term", StringAtom tm])
|
||||
putIDECommand (ShowTermImplicits tm) = (SExpList [SymbolAtom "show-term-implicits", StringAtom tm])
|
||||
putIDECommand (HideTermImplicits tm) = (SExpList [SymbolAtom "hide-term-implicits", StringAtom tm])
|
||||
putIDECommand (ElaborateTerm tm) = (SExpList [SymbolAtom "elaborate-term", StringAtom tm])
|
||||
putIDECommand (PrintDefinition n) = (SExpList [SymbolAtom "print-definition", StringAtom n])
|
||||
putIDECommand (ReplCompletions n) = (SExpList [SymbolAtom "repl-completions", StringAtom n])
|
||||
putIDECommand (Directive n) = (SExpList [SymbolAtom "directive", StringAtom n])
|
||||
putIDECommand (EnableSyntax b) = (SExpList [SymbolAtom "enable-syntax", BoolAtom b])
|
||||
putIDECommand GetOptions = (SExpList [SymbolAtom "get-options"])
|
||||
putIDECommand Version = SymbolAtom "version"
|
||||
Cast (FileName, NonEmptyFC) FileContext where
|
||||
cast (filename, _, (startLine, startCol), (endLine, endCol)) =
|
||||
MkFileContext
|
||||
{ file = filename
|
||||
, range = MkBounds {startLine, startCol, endLine, endCol}
|
||||
}
|
||||
|
||||
export
|
||||
getMsg : SExp -> Maybe (IDECommand, Integer)
|
||||
getMsg (SExpList [cmdexp, IntegerAtom num])
|
||||
= do cmd <- getIDECommand cmdexp
|
||||
= do cmd <- fromSExp cmdexp
|
||||
pure (cmd, num)
|
||||
getMsg _ = Nothing
|
||||
|
||||
escape : String -> String
|
||||
escape = pack . concatMap escapeChar . unpack
|
||||
where
|
||||
escapeChar : Char -> List Char
|
||||
escapeChar '\\' = ['\\', '\\']
|
||||
escapeChar '"' = ['\\', '\"']
|
||||
escapeChar c = [c]
|
||||
|
||||
export
|
||||
Show SExp where
|
||||
show (SExpList xs) = assert_total $ "(" ++ showSep " " (map show xs) ++ ")"
|
||||
show (StringAtom str) = "\"" ++ escape str ++ "\""
|
||||
show (BoolAtom b) = ":" ++ show b
|
||||
show (IntegerAtom i) = show i
|
||||
show (SymbolAtom s) = ":" ++ s
|
||||
|
||||
public export
|
||||
interface SExpable a where
|
||||
toSExp : a -> SExp
|
||||
|
||||
export
|
||||
SExpable IDECommand where
|
||||
toSExp = putIDECommand
|
||||
|
||||
export
|
||||
SExpable SExp where
|
||||
toSExp = id
|
||||
|
||||
export
|
||||
SExpable Bool where
|
||||
toSExp = BoolAtom
|
||||
|
||||
export
|
||||
SExpable String where
|
||||
toSExp = StringAtom
|
||||
|
||||
export
|
||||
SExpable Integer where
|
||||
toSExp = IntegerAtom
|
||||
|
||||
export
|
||||
SExpable Int where
|
||||
toSExp = IntegerAtom . cast
|
||||
|
||||
export
|
||||
SExpable Nat where
|
||||
toSExp = IntegerAtom . cast
|
||||
|
||||
export
|
||||
SExpable Name where
|
||||
toSExp = SymbolAtom . show
|
||||
|
||||
export
|
||||
(SExpable a, SExpable b) => SExpable (a, b) where
|
||||
toSExp (x, y)
|
||||
= case toSExp y of
|
||||
SExpList xs => SExpList (toSExp x :: xs)
|
||||
y' => SExpList [toSExp x, y']
|
||||
|
||||
export
|
||||
SExpable a => SExpable (List a) where
|
||||
toSExp xs
|
||||
= SExpList (map toSExp xs)
|
||||
|
||||
export
|
||||
SExpable a => SExpable (Maybe a) where
|
||||
toSExp Nothing = SExpList []
|
||||
toSExp (Just x) = toSExp x
|
||||
|
||||
export
|
||||
version : Int -> Int -> SExp
|
||||
version maj min = toSExp (SymbolAtom "protocol-version", maj, min)
|
||||
|
||||
sendStr : File -> String -> IO ()
|
||||
sendStr f st =
|
||||
map (const ()) (fPutStr f st)
|
||||
|
||||
export
|
||||
send : {auto c : Ref Ctxt Defs} -> SExpable a => File -> a -> Core ()
|
||||
send : {auto c : Ref Ctxt Defs} -> File -> Reply -> Core ()
|
||||
send f resp
|
||||
= do let r = show (toSExp resp) ++ "\n"
|
||||
log "ide-mode.send" 20 r
|
||||
|
@ -16,7 +16,7 @@ import Libraries.Data.String.Extra as L
|
||||
%default covering
|
||||
|
||||
public export
|
||||
record HolePremise where
|
||||
record Premise where
|
||||
constructor MkHolePremise
|
||||
name : Name
|
||||
type : IPTerm
|
||||
@ -25,14 +25,14 @@ record HolePremise where
|
||||
|
||||
|
||||
public export
|
||||
record HoleData where
|
||||
record Data where
|
||||
constructor MkHoleData
|
||||
name : Name
|
||||
type : IPTerm
|
||||
context : List HolePremise
|
||||
context : List Holes.Premise
|
||||
|
||||
export
|
||||
prettyHoles : List HoleData -> Doc IdrisSyntax
|
||||
prettyHoles : List Holes.Data -> Doc IdrisSyntax
|
||||
prettyHoles holes = case holes of
|
||||
[] => "No holes"
|
||||
[x] => "1 hole" <+> colon <++> prettyHole x
|
||||
@ -41,11 +41,10 @@ prettyHoles holes = case holes of
|
||||
|
||||
where
|
||||
|
||||
prettyHole : HoleData -> Doc IdrisSyntax
|
||||
prettyHole : Holes.Data -> Doc IdrisSyntax
|
||||
prettyHole x = pretty x.name <++> colon <++> prettyTerm x.type
|
||||
|
||||
|
||||
|
||||
||| If input is a hole, return number of locals in scope at binding
|
||||
||| point
|
||||
export
|
||||
@ -89,7 +88,7 @@ extractHoleData : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
Defs -> Env Term vars -> Name -> Nat -> Term vars ->
|
||||
Core HoleData
|
||||
Core Holes.Data
|
||||
extractHoleData defs env fn (S args) (Bind fc x (Let _ c val ty) sc)
|
||||
= extractHoleData defs env fn args (subst val sc)
|
||||
extractHoleData defs env fn (S args) (Bind fc x b sc)
|
||||
@ -116,7 +115,7 @@ holeData : {vars : _} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
Defs -> Env Term vars -> Name -> Nat -> Term vars ->
|
||||
Core HoleData
|
||||
Core Holes.Data
|
||||
|
||||
holeData gam env fn args ty
|
||||
= do hdata <- extractHoleData gam env fn args ty
|
||||
@ -125,7 +124,7 @@ holeData gam env fn args ty
|
||||
then hdata
|
||||
else { context $= dropShadows } hdata
|
||||
where
|
||||
dropShadows : List HolePremise -> List HolePremise
|
||||
dropShadows : List Holes.Premise -> List Holes.Premise
|
||||
dropShadows [] = []
|
||||
dropShadows (premise :: rest)
|
||||
= if premise.name `elem` map name rest
|
||||
@ -136,7 +135,7 @@ export
|
||||
getUserHolesData :
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
Core (List HoleData)
|
||||
Core (List Holes.Data)
|
||||
getUserHolesData
|
||||
= do defs <- get Ctxt
|
||||
let ctxt = gamma defs
|
||||
@ -157,12 +156,13 @@ showHole : {vars : _} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
Defs -> Env Term vars -> Name -> Nat -> Term vars ->
|
||||
Core String
|
||||
|
||||
showHole defs env fn args ty
|
||||
= do hdata <- holeData defs env fn args ty
|
||||
case hdata.context of
|
||||
[] => pure $ show (hdata.name) ++ " : " ++ show hdata.type
|
||||
_ => pure $ concat
|
||||
(map (\premise => " " ++ showCount premise.multiplicity ++ " "
|
||||
(map (\ premise : Holes.Premise => " " ++ showCount premise.multiplicity ++ " "
|
||||
++ (impBracket premise.isImplicit $
|
||||
tidy premise.name ++ " : " ++ (show premise.type) ++ "\n" )
|
||||
) hdata.context)
|
||||
@ -192,20 +192,19 @@ prettyHole defs env fn args ty
|
||||
<+> (pretty $ L.replicate 30 '-') <+> hardline
|
||||
<+> pretty (nameRoot $ hdata.name) <++> colon <++> prettyTerm hdata.type
|
||||
|
||||
sexpPremise : HolePremise -> SExp
|
||||
sexpPremise premise =
|
||||
SExpList [StringAtom $ " " ++ showCount premise.multiplicity ++ " "
|
||||
++ (impBracket premise.isImplicit $
|
||||
tidy premise.name)
|
||||
,StringAtom $ show premise.type
|
||||
,SExpList [] -- TODO: metadata
|
||||
]
|
||||
|
||||
premiseIDE : Holes.Premise -> HolePremise
|
||||
premiseIDE premise = IDE.MkHolePremise
|
||||
{ name = " " ++ showCount premise.multiplicity ++ " "
|
||||
++ (impBracket premise.isImplicit $
|
||||
tidy premise.name)
|
||||
, type = show premise.type
|
||||
}
|
||||
|
||||
export
|
||||
sexpHole : HoleData -> SExp
|
||||
sexpHole hole = SExpList
|
||||
[ StringAtom (show hole.name)
|
||||
, SExpList $ map sexpPremise hole.context -- Premises
|
||||
, SExpList [ StringAtom $ show hole.type -- Conclusion
|
||||
, SExpList[]] -- TODO: Highlighting information
|
||||
]
|
||||
holeIDE : Holes.Data -> IDE.HoleData
|
||||
holeIDE hole = IDE.MkHoleData
|
||||
{ name = show hole.name
|
||||
, type = show hole.type
|
||||
, context = map premiseIDE hole.context
|
||||
}
|
||||
|
40
src/Idris/IDEMode/Pretty.idr
Normal file
40
src/Idris/IDEMode/Pretty.idr
Normal file
@ -0,0 +1,40 @@
|
||||
module Idris.IDEMode.Pretty
|
||||
|
||||
import Protocol.IDE
|
||||
import Idris.Pretty
|
||||
import Idris.Doc.Annotations
|
||||
|
||||
------------------------------------------------------------------------
|
||||
-- Text properties supported by the IDE mode
|
||||
------------------------------------------------------------------------
|
||||
|
||||
export
|
||||
syntaxToProperties : IdrisSyntax -> Maybe Properties
|
||||
syntaxToProperties syn = mkDecor <$> syntaxToDecoration syn
|
||||
|
||||
export
|
||||
annToProperties : IdrisAnn -> Maybe Properties
|
||||
annToProperties Warning = Just $ MkProperties
|
||||
{ decor = Just Postulate
|
||||
, format = Just Bold
|
||||
}
|
||||
annToProperties Error = Just $ MkProperties
|
||||
{ decor = Just $ Postulate
|
||||
, format = Just Bold
|
||||
}
|
||||
annToProperties ErrorDesc = Nothing
|
||||
annToProperties FileCtxt = Just $ mkDecor Typ
|
||||
annToProperties Code = Just $ mkDecor Bound
|
||||
annToProperties Meta = Just $ mkDecor Function
|
||||
annToProperties (Syntax syn) = syntaxToProperties syn
|
||||
annToProperties UserDocString = Just $ mkDecor Comment
|
||||
|
||||
export
|
||||
docToProperties : IdrisDocAnn -> Maybe Properties
|
||||
docToProperties Header = pure $ mkFormat Underline
|
||||
docToProperties Deprecation = pure $ mkFormat Bold
|
||||
docToProperties Declarations = Nothing
|
||||
docToProperties (Decl _) = Nothing
|
||||
docToProperties DocStringBody = Nothing
|
||||
docToProperties UserDocString = Nothing
|
||||
docToProperties (Syntax syn) = syntaxToProperties syn
|
@ -19,8 +19,9 @@ import Idris.IDEMode.Commands
|
||||
import Idris.IDEMode.Holes
|
||||
import Idris.IDEMode.Parser
|
||||
import Idris.IDEMode.SyntaxHighlight
|
||||
import Idris.IDEMode.Pretty
|
||||
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
import Libraries.Utils.Path
|
||||
|
||||
import Data.List
|
||||
@ -122,7 +123,7 @@ todoCmd cmdName = iputStrLn $ reflow $ cmdName ++ ": command not yet implemented
|
||||
data IDEResult
|
||||
= REPL REPLResult
|
||||
| NameList (List Name)
|
||||
| FoundHoles (List HoleData)
|
||||
| FoundHoles (List Holes.Data)
|
||||
| Term String -- should be a PTerm + metadata, or SExp.
|
||||
| TTTerm String -- should be a TT Term + metadata, or perhaps SExp
|
||||
| NameLocList (List (Name, FC))
|
||||
@ -169,7 +170,7 @@ process (AddMissing l n)
|
||||
pure $ REPL $ Edited $ DisplayEdit emptyDoc
|
||||
process (ExprSearch l n hs all)
|
||||
= replWrap $ Idris.REPL.process (Editing (ExprSearch False (fromInteger l)
|
||||
(UN $ Basic n) (map (UN . Basic) hs)))
|
||||
(UN $ Basic n) (map (UN . Basic) hs.list)))
|
||||
process ExprSearchNext
|
||||
= replWrap $ Idris.REPL.process (Editing ExprSearchNext)
|
||||
process (GenerateDef l n)
|
||||
@ -249,58 +250,45 @@ processCatch cmd
|
||||
|
||||
idePutStrLn : {auto c : Ref Ctxt Defs} -> File -> Integer -> String -> Core ()
|
||||
idePutStrLn outf i msg
|
||||
= send outf (SExpList [SymbolAtom "write-string",
|
||||
toSExp msg, toSExp i])
|
||||
= send outf $ WriteString msg i
|
||||
|
||||
returnFromIDE : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core ()
|
||||
returnFromIDE outf i msg
|
||||
= do send outf (SExpList [SymbolAtom "return", msg, toSExp i])
|
||||
returnFromIDE : {auto c : Ref Ctxt Defs} -> File -> Integer -> IDE.ReplyPayload -> Core ()
|
||||
returnFromIDE outf i payload
|
||||
= do send outf (Immediate payload i)
|
||||
|
||||
printIDEResult : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core ()
|
||||
printIDEResult outf i msg
|
||||
= returnFromIDE outf i
|
||||
$ SExpList [ SymbolAtom "ok"
|
||||
, toSExp msg
|
||||
]
|
||||
|
||||
export
|
||||
SExpable a => SExpable (Span a) where
|
||||
toSExp (MkSpan start width ann)
|
||||
= SExpList [ IntegerAtom (cast start)
|
||||
, IntegerAtom (cast width)
|
||||
, toSExp ann
|
||||
]
|
||||
printIDEResult : {auto c : Ref Ctxt Defs} -> File -> Integer -> IDE.Result -> Core ()
|
||||
printIDEResult outf i result
|
||||
= returnFromIDE outf i $ OK result []
|
||||
|
||||
printIDEResultWithHighlight :
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
File -> Integer -> (SExp, List (Span Properties)) ->
|
||||
File -> Integer -> (Result, List (Span Properties)) ->
|
||||
Core ()
|
||||
printIDEResultWithHighlight outf i (msg, spans) = do
|
||||
printIDEResultWithHighlight outf i (result, spans) = do
|
||||
-- log "ide-mode.highlight" 10 $ show spans
|
||||
returnFromIDE outf i
|
||||
$ SExpList [ SymbolAtom "ok"
|
||||
, msg
|
||||
, toSExp spans
|
||||
]
|
||||
$ OK result spans
|
||||
|
||||
-- TODO: refactor to construct an error response
|
||||
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 $
|
||||
uncurry IDE.Error !(renderWithDecorations annToProperties msg)
|
||||
|
||||
SExpable REPLEval where
|
||||
toSExp EvalTC = SymbolAtom "typecheck"
|
||||
toSExp NormaliseAll = SymbolAtom "normalise"
|
||||
toSExp Execute = SymbolAtom "execute"
|
||||
toSExp Scheme = SymbolAtom "scheme"
|
||||
Cast REPLEval String where
|
||||
cast EvalTC = "typecheck"
|
||||
cast NormaliseAll = "normalise"
|
||||
cast Execute = "execute"
|
||||
cast Scheme = "scheme"
|
||||
|
||||
SExpable REPLOpt where
|
||||
toSExp (ShowImplicits impl) = SExpList [ SymbolAtom "show-implicits", toSExp impl ]
|
||||
toSExp (ShowNamespace ns) = SExpList [ SymbolAtom "show-namespace", toSExp ns ]
|
||||
toSExp (ShowTypes typs) = SExpList [ SymbolAtom "show-types", toSExp typs ]
|
||||
toSExp (EvalMode mod) = SExpList [ SymbolAtom "eval", toSExp mod ]
|
||||
toSExp (Editor editor) = SExpList [ SymbolAtom "editor", toSExp editor ]
|
||||
toSExp (CG str) = SExpList [ SymbolAtom "cg", toSExp str ]
|
||||
toSExp (Profile p) = SExpList [ SymbolAtom "profile", toSExp p ]
|
||||
toSExp (EvalTiming p) = SExpList [ SymbolAtom "evaltiming", toSExp p ]
|
||||
Cast REPLOpt REPLOption where
|
||||
cast (ShowImplicits impl) = MkOption "show-implicits" BOOL impl
|
||||
cast (ShowNamespace ns) = MkOption "show-namespace" BOOL ns
|
||||
cast (ShowTypes typs) = MkOption "show-types" BOOL typs
|
||||
cast (EvalMode mod) = MkOption "eval" ATOM $ cast mod
|
||||
cast (Editor editor) = MkOption "editor" STRING editor
|
||||
cast (CG str) = MkOption "cg" STRING str
|
||||
cast (Profile p) = MkOption "profile" BOOL p
|
||||
cast (EvalTiming p) = MkOption "evaltiming" BOOL p
|
||||
|
||||
|
||||
displayIDEResult : {auto c : Ref Ctxt Defs} ->
|
||||
@ -312,32 +300,31 @@ displayIDEResult : {auto c : Ref Ctxt Defs} ->
|
||||
displayIDEResult outf i (REPL $ REPLError err)
|
||||
= printIDEError outf i err
|
||||
displayIDEResult outf i (REPL RequestedHelp )
|
||||
= printIDEResult outf i
|
||||
$ StringAtom $ displayHelp
|
||||
= printIDEResult outf i $ AString displayHelp
|
||||
displayIDEResult outf i (REPL $ Evaluated x Nothing)
|
||||
= printIDEResultWithHighlight outf i
|
||||
$ mapFst StringAtom
|
||||
$ mapFst AString
|
||||
!(renderWithDecorations syntaxToProperties $ prettyTerm x)
|
||||
displayIDEResult outf i (REPL $ Evaluated x (Just y))
|
||||
= printIDEResultWithHighlight outf i
|
||||
$ mapFst StringAtom
|
||||
$ mapFst AString
|
||||
!(renderWithDecorations syntaxToProperties
|
||||
$ prettyTerm x <++> ":" <++> prettyTerm y)
|
||||
displayIDEResult outf i (REPL $ Printed xs)
|
||||
= printIDEResultWithHighlight outf i
|
||||
$ mapFst StringAtom
|
||||
$ mapFst AString
|
||||
$ !(renderWithDecorations annToProperties xs)
|
||||
displayIDEResult outf i (REPL (PrintedDoc xs))
|
||||
= printIDEResultWithHighlight outf i
|
||||
$ mapFst StringAtom
|
||||
$ mapFst AString
|
||||
$ !(renderWithDecorations docToProperties xs)
|
||||
displayIDEResult outf i (REPL $ TermChecked x y)
|
||||
= printIDEResultWithHighlight outf i
|
||||
$ mapFst StringAtom
|
||||
$ mapFst AString
|
||||
!(renderWithDecorations syntaxToProperties
|
||||
$ prettyTerm x <++> ":" <++> prettyTerm y)
|
||||
displayIDEResult outf i (REPL $ FileLoaded x)
|
||||
= printIDEResult outf i $ SExpList []
|
||||
= printIDEResult outf i $ AUnit
|
||||
displayIDEResult outf i (REPL $ ErrorLoadingFile x err)
|
||||
= printIDEError outf i $ reflow "Error loading file" <++> pretty x <+> colon <++> pretty (show err)
|
||||
displayIDEResult outf i (REPL $ ErrorsBuildingFile x errs)
|
||||
@ -346,76 +333,61 @@ displayIDEResult outf i (REPL $ ErrorsBuildingFile x errs)
|
||||
displayIDEResult outf i (REPL $ NoFileLoaded)
|
||||
= printIDEError outf i $ reflow "No file can be reloaded"
|
||||
displayIDEResult outf i (REPL $ CurrentDirectory dir)
|
||||
= printIDEResult outf i
|
||||
$ StringAtom $ "Current working directory is \"" ++ dir ++ "\""
|
||||
= printIDEResult outf i $ AString $ "Current working directory is \"\{dir}\""
|
||||
displayIDEResult outf i (REPL CompilationFailed)
|
||||
= printIDEError outf i $ reflow "Compilation failed"
|
||||
displayIDEResult outf i (REPL $ Compiled f)
|
||||
= printIDEResult outf i $ StringAtom
|
||||
$ "File " ++ f ++ " written"
|
||||
= printIDEResult outf i $ AString "File \{f} written"
|
||||
displayIDEResult outf i (REPL $ ProofFound x)
|
||||
= printIDEResult outf i
|
||||
$ StringAtom $ show x
|
||||
= printIDEResult outf i $ AString $ show x
|
||||
displayIDEResult outf i (REPL $ Missed cases)
|
||||
= printIDEResult outf i
|
||||
$ StringAtom $ showSep "\n"
|
||||
$ AString $ showSep "\n"
|
||||
$ map handleMissing' cases
|
||||
displayIDEResult outf i (REPL $ CheckedTotal xs)
|
||||
= printIDEResult outf i
|
||||
$ StringAtom $ showSep "\n"
|
||||
$ AString $ showSep "\n"
|
||||
$ map (\ (fn, tot) => (show fn ++ " is " ++ show tot)) xs
|
||||
displayIDEResult outf i (REPL $ LogLevelSet k)
|
||||
= printIDEResult outf i
|
||||
$ StringAtom $ "Set loglevel to " ++ show k
|
||||
$ AString $ "Set loglevel to " ++ show k
|
||||
displayIDEResult outf i (REPL $ OptionsSet opts)
|
||||
= printIDEResult outf i optionsSexp
|
||||
where
|
||||
optionsSexp : SExp
|
||||
optionsSexp = SExpList $ map toSExp opts
|
||||
= printIDEResult outf i $ AnOptionList $ map cast opts
|
||||
displayIDEResult outf i (REPL $ VersionIs x)
|
||||
= printIDEResult outf i versionSExp
|
||||
where
|
||||
semverSexp : SExp
|
||||
semverSexp = case (semVer x) of
|
||||
(maj, min, patch) => SExpList (map toSExp [maj, min, patch])
|
||||
tagSexp : SExp
|
||||
tagSexp = case versionTag x of
|
||||
Nothing => SExpList [ StringAtom "" ]
|
||||
Just t => SExpList [ StringAtom t ]
|
||||
versionSExp : SExp
|
||||
versionSExp = SExpList [ semverSexp, tagSexp ]
|
||||
|
||||
= let (major, minor, patch) = semVer x
|
||||
in printIDEResult outf i $ AVersion $ MkIdrisVersion
|
||||
{major, minor, patch, tag = versionTag x}
|
||||
displayIDEResult outf i (REPL $ Edited (DisplayEdit xs))
|
||||
= printIDEResult outf i $ StringAtom $ show xs
|
||||
= printIDEResult outf i $ AString $ show xs
|
||||
displayIDEResult outf i (REPL $ Edited (EditError x))
|
||||
= printIDEError outf i x
|
||||
displayIDEResult outf i (REPL $ Edited (MadeLemma lit name pty pappstr))
|
||||
= printIDEResult outf i
|
||||
$ SExpList [ SymbolAtom "metavariable-lemma"
|
||||
, SExpList [ SymbolAtom "replace-metavariable", StringAtom pappstr ]
|
||||
, SExpList [ SymbolAtom "definition-type", StringAtom $ relit lit $ show name ++ " : " ++ show pty ]
|
||||
]
|
||||
= printIDEResult outf i $ AMetaVarLemma $ MkMetaVarLemma
|
||||
{ application = pappstr
|
||||
, lemma = relit lit $ show name ++ " : " ++ show pty
|
||||
}
|
||||
displayIDEResult outf i (REPL $ Edited (MadeWith lit wapp))
|
||||
= printIDEResult outf i
|
||||
$ StringAtom $ showSep "\n" (map (relit lit) wapp)
|
||||
$ AString $ showSep "\n" (map (relit lit) wapp)
|
||||
displayIDEResult outf i (REPL $ (Edited (MadeCase lit cstr)))
|
||||
= printIDEResult outf i
|
||||
$ StringAtom $ showSep "\n" (map (relit lit) cstr)
|
||||
$ AString $ showSep "\n" (map (relit lit) cstr)
|
||||
displayIDEResult outf i (FoundHoles holes)
|
||||
= printIDEResult outf i $ SExpList $ map sexpHole holes
|
||||
= printIDEResult outf i $ AHoleList $ map holeIDE holes
|
||||
displayIDEResult outf i (NameList ns)
|
||||
= printIDEResult outf i $ SExpList $ map toSExp ns
|
||||
= printIDEResult outf i $ ANameList $ map show ns
|
||||
displayIDEResult outf i (Term t)
|
||||
= printIDEResult outf i $ StringAtom t
|
||||
= printIDEResult outf i $ AString t
|
||||
displayIDEResult outf i (TTTerm t)
|
||||
= printIDEResult outf i $ StringAtom t
|
||||
= printIDEResult outf i $ AString t
|
||||
displayIDEResult outf i (REPL $ ConsoleWidthSet mn)
|
||||
= let width = case mn of
|
||||
Just k => show k
|
||||
Nothing => "auto"
|
||||
in printIDEResult outf i $ StringAtom $ "Set consolewidth to " ++ width
|
||||
in printIDEResult outf i $ AString $ "Set consolewidth to " ++ width
|
||||
displayIDEResult outf i (NameLocList dat)
|
||||
= printIDEResult outf i $ SExpList !(traverse (constructSExp . map toNonEmptyFC) dat)
|
||||
= printIDEResult outf i $ ANameLocList $
|
||||
!(traverse (constructFileContext . map toNonEmptyFC) dat)
|
||||
where
|
||||
-- In order to recover the full path to the module referenced by FC,
|
||||
-- which stores a module identifier as opposed to a full path,
|
||||
@ -441,15 +413,15 @@ displayIDEResult outf i (NameLocList dat)
|
||||
sexpOriginDesc (PhysicalPkgSrc fname) = pure fname
|
||||
sexpOriginDesc (Virtual Interactive) = pure "(Interactive)"
|
||||
|
||||
constructSExp : (Name, NonEmptyFC) -> Core SExp
|
||||
constructSExp (name, origin, (startLine, startCol), (endLine, endCol)) = pure $
|
||||
SExpList [ StringAtom !(render $ pretty name)
|
||||
, StringAtom !(sexpOriginDesc origin)
|
||||
, IntegerAtom $ cast $ startLine
|
||||
, IntegerAtom $ cast $ startCol
|
||||
, IntegerAtom $ cast $ endLine
|
||||
, IntegerAtom $ cast $ endCol
|
||||
]
|
||||
constructFileContext : (Name, NonEmptyFC) -> Core (String, FileContext)
|
||||
constructFileContext (name, origin, (startLine, startCol), (endLine, endCol)) = pure $
|
||||
-- TODO: fix the emacs mode to use the more structured SExpr representation
|
||||
(!(render $ pretty name)
|
||||
, MkFileContext
|
||||
{ file = !(sexpOriginDesc origin)
|
||||
, range = MkBounds {startCol, startLine, endCol, endLine}
|
||||
})
|
||||
|
||||
-- do not use a catchall so that we are warned about missing cases when adding a
|
||||
-- new construtor to the enumeration.
|
||||
displayIDEResult _ _ (REPL Done) = pure ()
|
||||
@ -518,5 +490,5 @@ replIDE
|
||||
case res of
|
||||
REPL _ => printError $ reflow "Running idemode but output isn't"
|
||||
IDEMode _ inf outf => do
|
||||
send outf (version 2 0)
|
||||
send outf (ProtocolVersion 2 1) -- TODO: Move this info somewhere more central
|
||||
loop
|
||||
|
@ -17,164 +17,30 @@ import Libraries.Data.PosMap
|
||||
|
||||
%default total
|
||||
|
||||
------------------------------------------------------------------------
|
||||
-- Text properties supported by the IDE mode
|
||||
------------------------------------------------------------------------
|
||||
|
||||
data Formatting : Type where
|
||||
Bold : Formatting
|
||||
Italic : Formatting
|
||||
Underline : Formatting
|
||||
|
||||
-- CAREFUL: this instance is used in SExpable Formatting. If you change
|
||||
-- it then you need to fix the SExpable implementation in order not to
|
||||
-- break the IDE mode.
|
||||
Show Formatting where
|
||||
show Bold = "bold"
|
||||
show Italic = "italic"
|
||||
show Underline = "underline"
|
||||
|
||||
-- At most one decoration & one formatting
|
||||
-- (We could use `These` to guarantee non-emptiness but I am not
|
||||
-- convinced this will stay being just 2 fields e.g. the emacs mode
|
||||
-- has support for tagging things as errors, adding links, etc.)
|
||||
public export
|
||||
record Properties where
|
||||
constructor MkProperties
|
||||
decor : Maybe Decoration
|
||||
format : Maybe Formatting
|
||||
|
||||
export
|
||||
mkDecor : Decoration -> Properties
|
||||
mkDecor dec = MkProperties (Just dec) Nothing
|
||||
|
||||
export
|
||||
mkFormat : Formatting -> Properties
|
||||
mkFormat = MkProperties Nothing . Just
|
||||
|
||||
export
|
||||
syntaxToProperties : IdrisSyntax -> Maybe Properties
|
||||
syntaxToProperties syn = mkDecor <$> syntaxToDecoration syn
|
||||
|
||||
export
|
||||
annToProperties : IdrisAnn -> Maybe Properties
|
||||
annToProperties Warning = Nothing
|
||||
annToProperties Error = Nothing
|
||||
annToProperties ErrorDesc = Nothing
|
||||
annToProperties FileCtxt = Nothing
|
||||
annToProperties Code = Nothing
|
||||
annToProperties Meta = Nothing
|
||||
annToProperties (Syntax syn) = syntaxToProperties syn
|
||||
annToProperties UserDocString = Nothing
|
||||
|
||||
export
|
||||
docToProperties : IdrisDocAnn -> Maybe Properties
|
||||
docToProperties Header = pure $ mkFormat Underline
|
||||
docToProperties Deprecation = pure $ mkFormat Bold
|
||||
docToProperties Declarations = Nothing
|
||||
docToProperties (Decl _) = Nothing
|
||||
docToProperties DocStringBody = Nothing
|
||||
docToProperties UserDocString = Nothing
|
||||
docToProperties (Syntax syn) = syntaxToProperties syn
|
||||
|
||||
SExpable Decoration where
|
||||
toSExp decor = SExpList [ SymbolAtom "decor"
|
||||
, SymbolAtom (show decor)
|
||||
]
|
||||
|
||||
SExpable Formatting where
|
||||
toSExp format = SExpList [ SymbolAtom "text-formatting"
|
||||
, SymbolAtom (show format)
|
||||
]
|
||||
|
||||
export
|
||||
SExpable Properties where
|
||||
toSExp (MkProperties dec form) = SExpList $ catMaybes
|
||||
[ toSExp <$> form
|
||||
, toSExp <$> dec
|
||||
]
|
||||
|
||||
|
||||
record Highlight where
|
||||
constructor MkHighlight
|
||||
location : NonEmptyFC
|
||||
filename : String
|
||||
name : String
|
||||
isImplicit : Bool
|
||||
key : String
|
||||
decor : Decoration
|
||||
docOverview : String
|
||||
typ : String
|
||||
ns : String
|
||||
|
||||
SExpable (FileName, FC) where
|
||||
toSExp (fname, fc) = case isNonEmptyFC fc of
|
||||
Just (origin, (startLine, startCol), (endLine, endCol)) =>
|
||||
SExpList [ SExpList [ SymbolAtom "filename", StringAtom fname ]
|
||||
, SExpList [ SymbolAtom "start"
|
||||
, IntegerAtom (cast startLine + 1)
|
||||
, IntegerAtom (cast startCol + 1)
|
||||
]
|
||||
, SExpList [ SymbolAtom "end"
|
||||
, IntegerAtom (cast endLine + 1)
|
||||
, IntegerAtom (cast endCol)
|
||||
]
|
||||
]
|
||||
Nothing => SExpList []
|
||||
|
||||
SExpable Highlight where
|
||||
toSExp (MkHighlight loc fname nam impl k dec doc t ns)
|
||||
= SExpList [ toSExp $ (fname, justFC loc)
|
||||
, SExpList [ SExpList [ SymbolAtom "name", StringAtom nam ]
|
||||
, SExpList [ SymbolAtom "namespace", StringAtom ns ]
|
||||
, toSExp dec
|
||||
, SExpList [ SymbolAtom "implicit", toSExp impl ]
|
||||
, SExpList [ SymbolAtom "key", StringAtom k ]
|
||||
, SExpList [ SymbolAtom "doc-overview", StringAtom doc ]
|
||||
, SExpList [ SymbolAtom "type", StringAtom t ]
|
||||
]
|
||||
]
|
||||
|
||||
||| Output some data using current dialog index
|
||||
export
|
||||
printOutput : {auto c : Ref Ctxt Defs} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
SExp -> Core ()
|
||||
printOutput msg
|
||||
SourceHighlight -> Core ()
|
||||
printOutput highlight
|
||||
= do opts <- get ROpts
|
||||
case idemode opts of
|
||||
REPL _ => pure ()
|
||||
IDEMode i _ f =>
|
||||
send f (SExpList [SymbolAtom "output",
|
||||
msg, toSExp i])
|
||||
|
||||
send f (Intermediate (HighlightSource [highlight]) i)
|
||||
|
||||
outputHighlight : {auto c : Ref Ctxt Defs} ->
|
||||
{auto opts : Ref ROpts REPLOpts} ->
|
||||
Highlight -> Core ()
|
||||
outputHighlight h =
|
||||
printOutput $ SExpList [ SymbolAtom "ok"
|
||||
, SExpList [ SymbolAtom "highlight-source"
|
||||
, toSExp hlt
|
||||
]
|
||||
]
|
||||
where
|
||||
hlt : List Highlight
|
||||
hlt = [h]
|
||||
outputHighlight hl =
|
||||
printOutput $ Full hl
|
||||
|
||||
lwOutputHighlight :
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto opts : Ref ROpts REPLOpts} ->
|
||||
(FileName, NonEmptyFC, Decoration) -> Core ()
|
||||
lwOutputHighlight (fname, nfc, decor) =
|
||||
printOutput $ SExpList [ SymbolAtom "ok"
|
||||
, SExpList [ SymbolAtom "highlight-source"
|
||||
, toSExp $ the (List _) [
|
||||
SExpList [ toSExp $ (fname, justFC nfc)
|
||||
, SExpList [ toSExp decor]
|
||||
]]]]
|
||||
|
||||
|
||||
printOutput $ Lw $ MkLwHighlight {location = cast (fname, nfc), decor}
|
||||
|
||||
outputNameSyntax : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
@ -188,9 +54,8 @@ outputNameSyntax (fname, nfc, decor, nm) = do
|
||||
let fc = justFC nfc
|
||||
let (mns, name) = displayName nm
|
||||
outputHighlight $ MkHighlight
|
||||
{ location = nfc
|
||||
{ location = cast (fname, nfc)
|
||||
, name
|
||||
, filename = fname
|
||||
, isImplicit = False
|
||||
, key = ""
|
||||
, decor
|
||||
|
@ -14,6 +14,7 @@ import Idris.Doc.String
|
||||
|
||||
import Idris.Error
|
||||
import Idris.IDEMode.Commands
|
||||
import Idris.IDEMode.Pretty
|
||||
import Idris.Pretty
|
||||
import public Idris.REPL.Opts
|
||||
import Idris.Resugar
|
||||
@ -42,9 +43,9 @@ iputStrLn msg
|
||||
-- output silenced
|
||||
REPL _ => pure ()
|
||||
IDEMode i _ f =>
|
||||
send f (SExpList [SymbolAtom "write-string",
|
||||
toSExp !(renderWithoutColor msg), toSExp i])
|
||||
|
||||
send f (WriteString
|
||||
!(renderWithoutColor msg)
|
||||
i)
|
||||
|
||||
||| Sampled against `VerbosityLvl`.
|
||||
public export
|
||||
@ -114,7 +115,7 @@ emitProblem a replDocCreator idemodeDocCreator getFC status
|
||||
-- TODO: Display a better message when the error doesn't contain a location
|
||||
case map toNonEmptyFC (getFC a) of
|
||||
Nothing => iputStrLn msg
|
||||
Just (origin, startPos, endPos) => do
|
||||
Just nfc@(origin, startPos, endPos) => do
|
||||
fname <- case origin of
|
||||
PhysicalIdrSrc ident => do
|
||||
-- recover the file name relative to the working directory.
|
||||
@ -125,17 +126,9 @@ emitProblem a replDocCreator idemodeDocCreator getFC status
|
||||
pure fname
|
||||
Virtual Interactive =>
|
||||
pure "(Interactive)"
|
||||
send f (SExpList [SymbolAtom "warning",
|
||||
SExpList [toSExp {a = String} fname,
|
||||
toSExp (addOne startPos),
|
||||
toSExp (addOne endPos),
|
||||
toSExp !(renderWithoutColor msg),
|
||||
-- highlighting; currently blank
|
||||
SExpList []],
|
||||
toSExp i])
|
||||
where
|
||||
addOne : (Int, Int) -> (Int, Int)
|
||||
addOne (l, c) = (l + 1, c + 1)
|
||||
let (str,spans) = !(renderWithDecorations annToProperties msg)
|
||||
send f (Warning (cast (the String fname, nfc)) str spans
|
||||
i)
|
||||
|
||||
-- Display an error message from checking a source file
|
||||
export
|
||||
|
29
src/Libraries/Data/Span.idr
Normal file
29
src/Libraries/Data/Span.idr
Normal file
@ -0,0 +1,29 @@
|
||||
-- This ought to go into base at some future point
|
||||
module Libraries.Data.Span
|
||||
|
||||
public export
|
||||
record Span (a : Type) where
|
||||
constructor MkSpan
|
||||
start : Nat
|
||||
length : Nat
|
||||
property : a
|
||||
|
||||
export
|
||||
Functor Span where
|
||||
map f = { property $= f }
|
||||
|
||||
export
|
||||
Foldable Span where
|
||||
foldr c n span = c span.property n
|
||||
|
||||
export
|
||||
Traversable Span where
|
||||
traverse f (MkSpan start width prop)
|
||||
= MkSpan start width <$> f prop
|
||||
|
||||
export
|
||||
Show a => Show (Span a) where
|
||||
show (MkSpan start width prop)
|
||||
= concat {t = List} [ "[", show start, "-", show width, "]"
|
||||
, show prop
|
||||
]
|
@ -6,9 +6,13 @@ module Libraries.Text.Bounded
|
||||
public export
|
||||
record Bounds where
|
||||
constructor MkBounds
|
||||
||| 0-based first line
|
||||
startLine : Int
|
||||
||| 0-based first col
|
||||
startCol : Int
|
||||
||| 0-based last line of bound
|
||||
endLine : Int
|
||||
||| 0-based first column after bound
|
||||
endCol : Int
|
||||
|
||||
Show Bounds where
|
||||
|
@ -6,6 +6,7 @@ import Data.Maybe
|
||||
import Data.SnocList
|
||||
import Data.String
|
||||
import public Libraries.Data.String.Extra
|
||||
import public Libraries.Data.Span
|
||||
|
||||
%hide Data.String.lines
|
||||
%hide Data.String.lines'
|
||||
@ -774,33 +775,6 @@ Show (Doc ann) where
|
||||
-- Turn the document into a string, and a list of annotation spans
|
||||
------------------------------------------------------------------------
|
||||
|
||||
public export
|
||||
record Span (a : Type) where
|
||||
constructor MkSpan
|
||||
start : Nat
|
||||
length : Nat
|
||||
property : a
|
||||
|
||||
export
|
||||
Functor Span where
|
||||
map f = { property $= f }
|
||||
|
||||
export
|
||||
Foldable Span where
|
||||
foldr c n span = c span.property n
|
||||
|
||||
export
|
||||
Traversable Span where
|
||||
traverse f (MkSpan start width prop)
|
||||
= MkSpan start width <$> f prop
|
||||
|
||||
export
|
||||
Show a => Show (Span a) where
|
||||
show (MkSpan start width prop)
|
||||
= concat {t = List} [ "[", show start, "-", show width, "]"
|
||||
, show prop
|
||||
]
|
||||
|
||||
export
|
||||
displaySpans : SimpleDocStream a -> (String, List (Span a))
|
||||
displaySpans p =
|
||||
|
@ -12,7 +12,7 @@ import Libraries.Text.Lexer.Tokenizer
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter.Util
|
||||
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
import Libraries.Utils.Octal
|
||||
import Libraries.Utils.String
|
||||
|
||||
|
@ -1,4 +1,4 @@
|
||||
module Libraries.Utils.Hex
|
||||
module Protocol.Hex
|
||||
|
||||
import Data.Bits
|
||||
import Data.List
|
168
src/Protocol/IDE.idr
Normal file
168
src/Protocol/IDE.idr
Normal file
@ -0,0 +1,168 @@
|
||||
||| Messages exchanged during the IDE protocol
|
||||
module Protocol.IDE
|
||||
|
||||
import Protocol.SExp
|
||||
import Data.List
|
||||
import Data.Maybe
|
||||
|
||||
import public Libraries.Data.Span
|
||||
|
||||
import public Protocol.IDE.Command as Protocol.IDE
|
||||
import public Protocol.IDE.Decoration as Protocol.IDE
|
||||
import public Protocol.IDE.Formatting as Protocol.IDE
|
||||
import public Protocol.IDE.FileContext as Protocol.IDE
|
||||
import public Protocol.IDE.Holes as Protocol.IDE
|
||||
import public Protocol.IDE.Result as Protocol.IDE
|
||||
import public Protocol.IDE.Highlight as Protocol.IDE
|
||||
|
||||
%default total
|
||||
------------------------------------------------------------------------
|
||||
|
||||
public export
|
||||
Highlighting : Type
|
||||
Highlighting = List (Span Properties)
|
||||
|
||||
export
|
||||
SExpable a => SExpable (Span a) where
|
||||
toSExp (MkSpan start width ann)
|
||||
= SExpList [ IntegerAtom (cast start)
|
||||
, IntegerAtom (cast width)
|
||||
, toSExp ann
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable a => FromSExpable (Span a) where
|
||||
fromSExp (SExpList [ start
|
||||
, width
|
||||
, ann
|
||||
]) = do pure $ MkSpan { start = !(fromSExp start)
|
||||
, length = !(fromSExp width)
|
||||
, property = !(fromSExp ann)}
|
||||
fromSExp _ = Nothing
|
||||
|
||||
------------------------------------------------------------------------
|
||||
|
||||
public export
|
||||
data ReplyPayload =
|
||||
OK Result Highlighting
|
||||
| HighlightSource (List SourceHighlight)
|
||||
| Error String Highlighting
|
||||
|
||||
export
|
||||
SExpable ReplyPayload where
|
||||
toSExp (OK result hl) = SExpList (SymbolAtom "ok" :: toSExp result ::
|
||||
case hl of
|
||||
[] => []
|
||||
_ => [SExpList (map toSExp hl)])
|
||||
toSExp (HighlightSource hls) = SExpList
|
||||
[ SymbolAtom "ok"
|
||||
, SExpList
|
||||
[ SymbolAtom "highlight-source"
|
||||
, toSExp hls]
|
||||
]
|
||||
toSExp (Error msg hl) = SExpList (SymbolAtom "error" :: toSExp msg ::
|
||||
case hl of
|
||||
[] => []
|
||||
_ => [SExpList (map toSExp hl)])
|
||||
|
||||
-- Again, not the most efficient. Probably better to index by the
|
||||
-- expected return type in the future
|
||||
export
|
||||
FromSExpable ReplyPayload where
|
||||
fromSExp (SExpList [SymbolAtom "ok", result])
|
||||
= do pure $ OK !(fromSExp result) []
|
||||
fromSExp (SExpList [SymbolAtom "ok", result, hl])
|
||||
= do pure $ OK !(fromSExp result) !(fromSExp hl)
|
||||
fromSExp (SExpList
|
||||
[ SymbolAtom "ok"
|
||||
, SExpList
|
||||
[ SymbolAtom "highlight-source"
|
||||
, hls]
|
||||
]) = do pure $ HighlightSource !(fromSExp hls)
|
||||
fromSExp (SExpList [SymbolAtom "error", msg])
|
||||
= do pure $ Error !(fromSExp msg) []
|
||||
fromSExp (SExpList [SymbolAtom "error", msg, hl])
|
||||
= do pure $ Error !(fromSExp msg) !(fromSExp hl)
|
||||
fromSExp _ = Nothing
|
||||
|
||||
public export
|
||||
data Reply =
|
||||
ProtocolVersion Int Int
|
||||
| Immediate ReplyPayload Integer
|
||||
| Intermediate ReplyPayload Integer
|
||||
| WriteString String Integer
|
||||
| SetPrompt String Integer
|
||||
| Warning FileContext String Highlighting Integer
|
||||
|
||||
export
|
||||
SExpable Reply where
|
||||
toSExp (ProtocolVersion maj min) = toSExp (SymbolAtom "protocol-version", maj, min)
|
||||
toSExp ( Immediate payload id) = SExpList [SymbolAtom "return",
|
||||
toSExp payload, toSExp id]
|
||||
toSExp (Intermediate payload id) = SExpList [SymbolAtom "output",
|
||||
toSExp payload, toSExp id]
|
||||
toSExp (WriteString str id) = SExpList [SymbolAtom "write-string", toSExp str, toSExp id]
|
||||
toSExp (SetPrompt str id) = SExpList [SymbolAtom "set-prompt" , toSExp str, toSExp id]
|
||||
toSExp (Warning fc str spans id) = SExpList [SymbolAtom "warning",
|
||||
SExpList $ toSExp fc.file :: toSExp (fc.range.startLine, fc.range.startCol)
|
||||
:: toSExp (fc.range.endLine , fc.range.endCol )
|
||||
:: toSExp str :: case spans of
|
||||
[] => []
|
||||
_ => [SExpList (map toSExp spans)]
|
||||
, toSExp id]
|
||||
|
||||
export
|
||||
FromSExpable Reply where
|
||||
fromSExp (SExpList [SymbolAtom "protocol-version", major, minor]) =
|
||||
do Just $ ProtocolVersion !(fromSExp major) !(fromSExp minor)
|
||||
fromSExp (SExpList [SymbolAtom "return", payload, iden]) =
|
||||
do Just $ Immediate !(fromSExp payload) !(fromSExp iden)
|
||||
fromSExp (SExpList [SymbolAtom "output", payload, iden]) =
|
||||
do Just $ Intermediate !(fromSExp payload) !(fromSExp iden)
|
||||
fromSExp (SExpList [SymbolAtom "write-string", str, iden]) =
|
||||
do Just $ WriteString !(fromSExp str) !(fromSExp iden)
|
||||
fromSExp (SExpList [SymbolAtom "set-prompt", str, iden]) =
|
||||
do Just $ SetPrompt !(fromSExp str) !(fromSExp iden)
|
||||
fromSExp (SExpList [SymbolAtom "warning"
|
||||
, SExpList [filename, SExpList [startLine, startCol]
|
||||
, SExpList [endLine , endCol ]
|
||||
, str]
|
||||
, iden]) = do
|
||||
pure $ Warning (MkFileContext
|
||||
{ file = !(fromSExp filename)
|
||||
, range = MkBounds { startLine = !(fromSExp startLine)
|
||||
, startCol = !(fromSExp startCol)
|
||||
, endLine = !(fromSExp endLine)
|
||||
, endCol = !(fromSExp endCol)}
|
||||
})
|
||||
!(fromSExp str)
|
||||
[]
|
||||
!(fromSExp iden)
|
||||
fromSExp (SExpList [SymbolAtom "warning"
|
||||
, SExpList [filename, SExpList [startLine, startCol]
|
||||
, SExpList [endLine , endCol ]
|
||||
, str, hl]
|
||||
, iden]) = do
|
||||
pure $ Warning (MkFileContext
|
||||
{ file = !(fromSExp filename)
|
||||
, range = MkBounds { startLine = !(fromSExp startLine)
|
||||
, startCol = !(fromSExp startCol)
|
||||
, endLine = !(fromSExp endLine)
|
||||
, endCol = !(fromSExp endCol)}
|
||||
})
|
||||
!(fromSExp str)
|
||||
!(fromSExp hl)
|
||||
!(fromSExp iden)
|
||||
fromSExp _ = Nothing
|
||||
|
||||
public export
|
||||
data Request =
|
||||
Cmd IDECommand
|
||||
|
||||
export
|
||||
SExpable Request where
|
||||
toSExp (Cmd cmd) = toSExp cmd
|
||||
|
||||
export
|
||||
FromSExpable Request where
|
||||
fromSExp cmd = do pure $ Cmd !(fromSExp cmd)
|
190
src/Protocol/IDE/Command.idr
Normal file
190
src/Protocol/IDE/Command.idr
Normal file
@ -0,0 +1,190 @@
|
||||
module Protocol.IDE.Command
|
||||
import Protocol.SExp
|
||||
|
||||
%default total
|
||||
|
||||
-- TODO: introduce SExpable DocMode and refactor below to use it
|
||||
public export
|
||||
data DocMode = Overview | Full
|
||||
|
||||
public export
|
||||
record Hints where
|
||||
constructor MkHints
|
||||
list : List String
|
||||
|
||||
export
|
||||
SExpable Hints where
|
||||
toSExp hs = toSExp hs.list
|
||||
|
||||
export
|
||||
FromSExpable Hints where
|
||||
fromSExp hs = MkHints <$> fromSExp hs
|
||||
|
||||
public export
|
||||
data IDECommand
|
||||
= Interpret String
|
||||
| LoadFile String (Maybe Integer)
|
||||
| TypeOf String (Maybe (Integer, Integer))
|
||||
| NameAt String (Maybe (Integer, Integer))
|
||||
| CaseSplit Integer Integer String
|
||||
| AddClause Integer String
|
||||
-- deprecated: | AddProofClause
|
||||
| AddMissing Integer String
|
||||
| ExprSearch Integer String Hints Bool
|
||||
| ExprSearchNext
|
||||
| GenerateDef Integer String
|
||||
| GenerateDefNext
|
||||
| MakeLemma Integer String
|
||||
| MakeCase Integer String
|
||||
| MakeWith Integer String
|
||||
| DocsFor String (Maybe DocMode)
|
||||
| Directive String
|
||||
| Apropos String
|
||||
| Metavariables Integer
|
||||
| WhoCalls String
|
||||
| CallsWho String
|
||||
| BrowseNamespace String
|
||||
| NormaliseTerm String -- TODO: implement a Binary lib
|
||||
| ShowTermImplicits String -- and implement Binary (Term)
|
||||
| HideTermImplicits String -- for these four defintions,
|
||||
| ElaborateTerm String -- then change String to Term, as in idris1
|
||||
| PrintDefinition String
|
||||
| ReplCompletions String
|
||||
| EnableSyntax Bool
|
||||
| Version
|
||||
| GetOptions
|
||||
|
||||
|
||||
getIDECommand : SExp -> Maybe IDECommand
|
||||
getIDECommand (SExpList [SymbolAtom "interpret", StringAtom cmd])
|
||||
= Just $ Interpret cmd
|
||||
getIDECommand (SExpList [SymbolAtom "load-file", StringAtom fname])
|
||||
= Just $ LoadFile fname Nothing
|
||||
getIDECommand (SExpList [SymbolAtom "load-file", StringAtom fname, IntegerAtom l])
|
||||
= Just $ LoadFile fname (Just l)
|
||||
getIDECommand (SExpList [SymbolAtom "type-of", StringAtom n])
|
||||
= Just $ TypeOf n Nothing
|
||||
getIDECommand (SExpList [SymbolAtom "type-of", StringAtom n,
|
||||
IntegerAtom l, IntegerAtom c])
|
||||
= Just $ TypeOf n (Just (l, c))
|
||||
getIDECommand (SExpList [SymbolAtom "name-at", StringAtom n])
|
||||
= Just $ NameAt n Nothing
|
||||
getIDECommand (SExpList [SymbolAtom "name-at", StringAtom n,
|
||||
IntegerAtom l, IntegerAtom c])
|
||||
= Just $ NameAt n (Just (l, c))
|
||||
getIDECommand (SExpList [SymbolAtom "case-split", IntegerAtom l, IntegerAtom c,
|
||||
StringAtom n])
|
||||
= Just $ CaseSplit l c n
|
||||
getIDECommand (SExpList [SymbolAtom "case-split", IntegerAtom l, StringAtom n])
|
||||
= Just $ CaseSplit l 0 n
|
||||
getIDECommand (SExpList [SymbolAtom "add-clause", IntegerAtom l, StringAtom n])
|
||||
= Just $ AddClause l n
|
||||
getIDECommand (SExpList [SymbolAtom "add-missing", IntegerAtom line, StringAtom n])
|
||||
= Just $ AddMissing line n
|
||||
getIDECommand (SExpList [SymbolAtom "proof-search", IntegerAtom l, StringAtom n])
|
||||
= Just $ ExprSearch l n (MkHints []) False
|
||||
getIDECommand (SExpList [SymbolAtom "proof-search", IntegerAtom l, StringAtom n, hs])
|
||||
= (\hs' => ExprSearch l n hs' False) <$> fromSExp hs
|
||||
getIDECommand (SExpList [SymbolAtom "proof-search", IntegerAtom l, StringAtom n, hs, SymbolAtom mode])
|
||||
= (\hs' => ExprSearch l n hs' (getMode mode)) <$> fromSExp hs
|
||||
where
|
||||
getMode : String -> Bool
|
||||
getMode m = m == "all"
|
||||
getIDECommand (SymbolAtom "proof-search-next") = Just ExprSearchNext
|
||||
getIDECommand (SExpList [SymbolAtom "generate-def", IntegerAtom l, StringAtom n])
|
||||
= Just $ GenerateDef l n
|
||||
getIDECommand (SymbolAtom "generate-def-next") = Just GenerateDefNext
|
||||
getIDECommand (SExpList [SymbolAtom "make-lemma", IntegerAtom l, StringAtom n])
|
||||
= Just $ MakeLemma l n
|
||||
getIDECommand (SExpList [SymbolAtom "make-case", IntegerAtom l, StringAtom n])
|
||||
= Just $ MakeCase l n
|
||||
getIDECommand (SExpList [SymbolAtom "make-with", IntegerAtom l, StringAtom n])
|
||||
= Just $ MakeWith l n
|
||||
getIDECommand (SExpList (SymbolAtom "docs-for" :: StringAtom n :: modeTail))
|
||||
= do -- Maybe monad
|
||||
modeOpt <- case modeTail of
|
||||
[] => Just Nothing
|
||||
[SymbolAtom "overview"] => Just $ Just Overview
|
||||
[SymbolAtom "full" ] => Just $ Just Full
|
||||
_ => Nothing
|
||||
Just $ DocsFor n modeOpt
|
||||
getIDECommand (SExpList [SymbolAtom "apropos", StringAtom n])
|
||||
= Just $ Apropos n
|
||||
getIDECommand (SExpList [SymbolAtom "directive", StringAtom n])
|
||||
= Just $ Directive n
|
||||
getIDECommand (SExpList [SymbolAtom "metavariables", IntegerAtom n])
|
||||
= Just $ Metavariables n
|
||||
getIDECommand (SExpList [SymbolAtom "who-calls", StringAtom n])
|
||||
= Just $ WhoCalls n
|
||||
getIDECommand (SExpList [SymbolAtom "calls-who", StringAtom n])
|
||||
= Just $ CallsWho n
|
||||
getIDECommand (SExpList [SymbolAtom "browse-namespace", StringAtom ns])
|
||||
= Just $ BrowseNamespace ns
|
||||
getIDECommand (SExpList [SymbolAtom "normalise-term", StringAtom tm])
|
||||
= Just $ NormaliseTerm tm
|
||||
getIDECommand (SExpList [SymbolAtom "show-term-implicits", StringAtom tm])
|
||||
= Just $ ShowTermImplicits tm
|
||||
getIDECommand (SExpList [SymbolAtom "hide-term-implicits", StringAtom tm])
|
||||
= Just $ HideTermImplicits tm
|
||||
getIDECommand (SExpList [SymbolAtom "elaborate-term", StringAtom tm])
|
||||
= Just $ ElaborateTerm tm
|
||||
getIDECommand (SExpList [SymbolAtom "print-definition", StringAtom n])
|
||||
= Just $ PrintDefinition n
|
||||
getIDECommand (SExpList [SymbolAtom "repl-completions", StringAtom n])
|
||||
= Just $ ReplCompletions n
|
||||
getIDECommand (SExpList [SymbolAtom "enable-syntax", BoolAtom b])
|
||||
= Just $ EnableSyntax b
|
||||
getIDECommand (SymbolAtom "version") = Just Version
|
||||
getIDECommand (SExpList [SymbolAtom "get-options"]) = Just GetOptions
|
||||
getIDECommand _ = Nothing
|
||||
|
||||
export
|
||||
FromSExpable IDECommand where
|
||||
fromSExp = getIDECommand
|
||||
|
||||
putIDECommand : IDECommand -> SExp
|
||||
putIDECommand (Interpret cmd) = (SExpList [SymbolAtom "interpret", StringAtom cmd])
|
||||
putIDECommand (LoadFile fname Nothing) = (SExpList [SymbolAtom "load-file", StringAtom fname])
|
||||
putIDECommand (LoadFile fname (Just line)) = (SExpList [SymbolAtom "load-file", StringAtom fname, IntegerAtom line])
|
||||
putIDECommand (TypeOf cmd Nothing) = (SExpList [SymbolAtom "type-of", StringAtom cmd])
|
||||
putIDECommand (TypeOf cmd (Just (line, col))) = (SExpList [SymbolAtom "type-of", StringAtom cmd, IntegerAtom line, IntegerAtom col])
|
||||
putIDECommand (NameAt cmd Nothing) = (SExpList [SymbolAtom "name-at", StringAtom cmd])
|
||||
putIDECommand (NameAt cmd (Just (line, col))) = (SExpList [SymbolAtom "name-at", StringAtom cmd, IntegerAtom line, IntegerAtom col])
|
||||
putIDECommand (CaseSplit line col n) = (SExpList [SymbolAtom "case-split", IntegerAtom line, IntegerAtom col, StringAtom n])
|
||||
putIDECommand (AddClause line n) = (SExpList [SymbolAtom "add-clause", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (AddMissing line n) = (SExpList [SymbolAtom "add-missing", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (ExprSearch line n hints mode) = (SExpList [SymbolAtom "proof-search", IntegerAtom line, StringAtom n, toSExp hints, getMode mode])
|
||||
where
|
||||
getMode : Bool -> SExp
|
||||
getMode True = SymbolAtom "all"
|
||||
getMode False = SymbolAtom "other"
|
||||
putIDECommand ExprSearchNext = SymbolAtom "proof-search-next"
|
||||
putIDECommand (GenerateDef line n) = (SExpList [SymbolAtom "generate-def", IntegerAtom line, StringAtom n])
|
||||
putIDECommand GenerateDefNext = SymbolAtom "generate-def-next"
|
||||
putIDECommand (MakeLemma line n) = (SExpList [SymbolAtom "make-lemma", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (MakeCase line n) = (SExpList [SymbolAtom "make-case", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (MakeWith line n) = (SExpList [SymbolAtom "make-with", IntegerAtom line, StringAtom n])
|
||||
putIDECommand (DocsFor n modeOpt) = let modeTail = case modeOpt of
|
||||
Nothing => []
|
||||
Just Overview => [SymbolAtom "overview"]
|
||||
Just Full => [SymbolAtom "full"] in
|
||||
(SExpList (SymbolAtom "docs-for" :: StringAtom n :: modeTail))
|
||||
putIDECommand (Apropos n) = (SExpList [SymbolAtom "apropos", StringAtom n])
|
||||
putIDECommand (Metavariables n) = (SExpList [SymbolAtom "metavariables", IntegerAtom n])
|
||||
putIDECommand (WhoCalls n) = (SExpList [SymbolAtom "who-calls", StringAtom n])
|
||||
putIDECommand (CallsWho n) = (SExpList [SymbolAtom "calls-who", StringAtom n])
|
||||
putIDECommand (BrowseNamespace ns) = (SExpList [SymbolAtom "browse-namespace", StringAtom ns])
|
||||
putIDECommand (NormaliseTerm tm) = (SExpList [SymbolAtom "normalise-term", StringAtom tm])
|
||||
putIDECommand (ShowTermImplicits tm) = (SExpList [SymbolAtom "show-term-implicits", StringAtom tm])
|
||||
putIDECommand (HideTermImplicits tm) = (SExpList [SymbolAtom "hide-term-implicits", StringAtom tm])
|
||||
putIDECommand (ElaborateTerm tm) = (SExpList [SymbolAtom "elaborate-term", StringAtom tm])
|
||||
putIDECommand (PrintDefinition n) = (SExpList [SymbolAtom "print-definition", StringAtom n])
|
||||
putIDECommand (ReplCompletions n) = (SExpList [SymbolAtom "repl-completions", StringAtom n])
|
||||
putIDECommand (Directive n) = (SExpList [SymbolAtom "directive", StringAtom n])
|
||||
putIDECommand (EnableSyntax b) = (SExpList [SymbolAtom "enable-syntax", BoolAtom b])
|
||||
putIDECommand GetOptions = (SExpList [SymbolAtom "get-options"])
|
||||
putIDECommand Version = SymbolAtom "version"
|
||||
|
||||
export
|
||||
SExpable IDECommand where
|
||||
toSExp = putIDECommand
|
79
src/Protocol/IDE/Decoration.idr
Normal file
79
src/Protocol/IDE/Decoration.idr
Normal file
@ -0,0 +1,79 @@
|
||||
module Protocol.IDE.Decoration
|
||||
|
||||
import Protocol.SExp
|
||||
|
||||
%default total
|
||||
|
||||
public export
|
||||
data Decoration : Type where
|
||||
Comment : Decoration
|
||||
Typ : Decoration
|
||||
Function : Decoration
|
||||
Data : Decoration
|
||||
Keyword : Decoration
|
||||
Bound : Decoration
|
||||
Namespace : Decoration
|
||||
Postulate : Decoration
|
||||
Module : Decoration
|
||||
|
||||
public export
|
||||
Eq Decoration where
|
||||
Comment == Comment = True
|
||||
Typ == Typ = True
|
||||
Function == Function = True
|
||||
Data == Data = True
|
||||
Keyword == Keyword = True
|
||||
Bound == Bound = True
|
||||
Namespace == Namespace = True
|
||||
Postulate == Postulate = True
|
||||
Module == Module = True
|
||||
_ == _ = False
|
||||
|
||||
public export
|
||||
Show Decoration where
|
||||
show Comment = "comment"
|
||||
show Typ = "type"
|
||||
show Function = "function"
|
||||
show Data = "data"
|
||||
show Keyword = "keyword"
|
||||
show Bound = "bound"
|
||||
show Namespace = "namespace"
|
||||
show Postulate = "postulate"
|
||||
show Module = "module"
|
||||
|
||||
|
||||
export
|
||||
SExpable Decoration where
|
||||
toSExp decor = SExpList [ SymbolAtom "decor"
|
||||
, SymbolAtom (show decor)
|
||||
]
|
||||
where
|
||||
{- This definition looks like it repeats `Show`, but `display` is part
|
||||
of the IDE protocol, whereas the `Show` instance doesn't have to be.
|
||||
-}
|
||||
display : Decoration -> String
|
||||
display Comment = "comment"
|
||||
display Typ = "type"
|
||||
display Function = "function"
|
||||
display Data = "data"
|
||||
display Keyword = "keyword"
|
||||
display Bound = "bound"
|
||||
display Namespace = "namespace"
|
||||
display Postulate = "postulate"
|
||||
display Module = "module"
|
||||
|
||||
export
|
||||
FromSExpable Decoration where
|
||||
fromSExp (SExpList [SymbolAtom "decor", SymbolAtom decor]) =
|
||||
case decor of
|
||||
"comment" => Just Comment
|
||||
"type" => Just Typ
|
||||
"function" => Just Function
|
||||
"data" => Just Data
|
||||
"keyword" => Just Keyword
|
||||
"bound" => Just Bound
|
||||
"namespace" => Just Namespace
|
||||
"postulate" => Just Postulate
|
||||
"module" => Just Module
|
||||
_ => Nothing
|
||||
fromSExp _ = Nothing
|
49
src/Protocol/IDE/FileContext.idr
Normal file
49
src/Protocol/IDE/FileContext.idr
Normal file
@ -0,0 +1,49 @@
|
||||
module Protocol.IDE.FileContext
|
||||
|
||||
import Protocol.SExp
|
||||
|
||||
import public Libraries.Text.Bounded
|
||||
|
||||
%default total
|
||||
|
||||
public export
|
||||
record FileContext where
|
||||
constructor MkFileContext
|
||||
file : String
|
||||
range : Bounds
|
||||
|
||||
export
|
||||
SExpable FileContext where
|
||||
toSExp fc =
|
||||
SExpList [ SExpList
|
||||
[ SymbolAtom "filename", toSExp fc.file ]
|
||||
, SExpList [ SymbolAtom "start"
|
||||
, IntegerAtom (cast fc.range.startLine)
|
||||
, IntegerAtom (cast fc.range.startCol)
|
||||
]
|
||||
, SExpList [ SymbolAtom "end"
|
||||
, IntegerAtom (cast fc.range.endLine)
|
||||
, IntegerAtom (cast fc.range.endCol)
|
||||
]
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable FileContext where
|
||||
fromSExp (SExpList [ SExpList
|
||||
[ SymbolAtom "filename", filenameSExp ]
|
||||
, SExpList [ SymbolAtom "start"
|
||||
, IntegerAtom startLine
|
||||
, IntegerAtom startCol
|
||||
]
|
||||
, SExpList [ SymbolAtom "end"
|
||||
, IntegerAtom endLine
|
||||
, IntegerAtom endCol
|
||||
]
|
||||
]) = do file <- fromSExp filenameSExp
|
||||
pure $ MkFileContext {file, range = MkBounds
|
||||
{ startLine = cast startLine
|
||||
, startCol = cast startCol
|
||||
, endLine = cast endLine
|
||||
, endCol = cast endCol
|
||||
}}
|
||||
fromSExp _ = Nothing
|
91
src/Protocol/IDE/Formatting.idr
Normal file
91
src/Protocol/IDE/Formatting.idr
Normal file
@ -0,0 +1,91 @@
|
||||
module Protocol.IDE.Formatting
|
||||
|
||||
import Protocol.SExp
|
||||
import Protocol.IDE.Decoration
|
||||
|
||||
import Data.Maybe
|
||||
import Data.List
|
||||
|
||||
%default total
|
||||
|
||||
public export
|
||||
data Formatting : Type where
|
||||
Bold : Formatting
|
||||
Italic : Formatting
|
||||
Underline : Formatting
|
||||
|
||||
export
|
||||
Show Formatting where
|
||||
show Bold = "bold"
|
||||
show Italic = "italic"
|
||||
show Underline = "underline"
|
||||
|
||||
export
|
||||
SExpable Formatting where
|
||||
toSExp format = SExpList [ SymbolAtom "text-formatting"
|
||||
, SymbolAtom (show format)
|
||||
]
|
||||
where
|
||||
{- This definition looks like it repeats `Show`, but `display` is part
|
||||
of the IDE protocol, whereas the `Show` instance doesn't have to be.
|
||||
-}
|
||||
display : Formatting -> String
|
||||
display Bold = "bold"
|
||||
display Italic = "italic"
|
||||
display Underline = "underline"
|
||||
|
||||
|
||||
export
|
||||
FromSExpable Formatting where
|
||||
fromSExp (SExpList [ SymbolAtom "text-formatting"
|
||||
, SymbolAtom format
|
||||
]) =
|
||||
case format of
|
||||
"bold" => Just Bold
|
||||
"italic" => Just Italic
|
||||
"underline" => Just Underline
|
||||
_ => Nothing
|
||||
fromSExp _ = Nothing
|
||||
|
||||
-- At most one decoration & one formatting
|
||||
-- (We could use `These` to guarantee non-emptiness but I am not
|
||||
-- convinced this will stay being just 2 fields e.g. the emacs mode
|
||||
-- has support for tagging things as errors, adding links, etc.)
|
||||
public export
|
||||
record Properties where
|
||||
constructor MkProperties
|
||||
decor : Maybe Decoration
|
||||
format : Maybe Formatting
|
||||
|
||||
export
|
||||
mkDecor : Decoration -> Properties
|
||||
mkDecor dec = MkProperties (Just dec) Nothing
|
||||
|
||||
export
|
||||
mkFormat : Formatting -> Properties
|
||||
mkFormat = MkProperties Nothing . Just
|
||||
|
||||
export
|
||||
SExpable Properties where
|
||||
toSExp (MkProperties dec form) = SExpList $ catMaybes
|
||||
[ toSExp <$> form
|
||||
, toSExp <$> dec
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable Properties where
|
||||
fromSExp (SExpList props) =
|
||||
case props of
|
||||
[] => Just $ MkProperties {decor = Nothing, format = Nothing}
|
||||
[prop] => do
|
||||
let Nothing = fromSExp prop
|
||||
| Just decor => Just $ mkDecor decor
|
||||
format <- fromSExp prop
|
||||
pure $ mkFormat format
|
||||
[prop1, prop2] => -- assume the same order as in toSExp
|
||||
-- not part of the protocol though
|
||||
do let format = Just !(fromSExp prop1)
|
||||
let decor = Just !(fromSExp prop2)
|
||||
pure $ MkProperties {format, decor}
|
||||
_ => Nothing
|
||||
fromSExp _ = Nothing
|
97
src/Protocol/IDE/Highlight.idr
Normal file
97
src/Protocol/IDE/Highlight.idr
Normal file
@ -0,0 +1,97 @@
|
||||
module Protocol.IDE.Highlight
|
||||
|
||||
import Protocol.SExp
|
||||
|
||||
import Protocol.IDE.FileContext
|
||||
import Protocol.IDE.Decoration
|
||||
|
||||
%default total
|
||||
|
||||
public export
|
||||
record Highlight where
|
||||
constructor MkHighlight
|
||||
location : FileContext
|
||||
name : String
|
||||
isImplicit : Bool
|
||||
key : String
|
||||
decor : Decoration
|
||||
docOverview : String
|
||||
typ : String
|
||||
ns : String
|
||||
|
||||
public export
|
||||
record LwHighlight where
|
||||
constructor MkLwHighlight
|
||||
location : FileContext
|
||||
decor : Decoration
|
||||
|
||||
export
|
||||
SExpable Highlight where
|
||||
toSExp (MkHighlight fc nam impl k dec doc t ns)
|
||||
= SExpList [ toSExp fc
|
||||
, SExpList [ SExpList [ SymbolAtom "name", StringAtom nam ]
|
||||
, SExpList [ SymbolAtom "namespace", StringAtom ns ]
|
||||
, toSExp dec
|
||||
, SExpList [ SymbolAtom "implicit", toSExp impl ]
|
||||
, SExpList [ SymbolAtom "key", StringAtom k ]
|
||||
, SExpList [ SymbolAtom "doc-overview", StringAtom doc ]
|
||||
, SExpList [ SymbolAtom "type", StringAtom t ]
|
||||
]
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable Highlight where
|
||||
fromSExp (SExpList [ fc
|
||||
, SExpList [ SExpList [ SymbolAtom "name", StringAtom nam ]
|
||||
, SExpList [ SymbolAtom "namespace", StringAtom ns ]
|
||||
, dec
|
||||
, SExpList [ SymbolAtom "implicit", impl ]
|
||||
, SExpList [ SymbolAtom "key", StringAtom key ]
|
||||
, SExpList [ SymbolAtom "doc-overview", StringAtom doc ]
|
||||
, SExpList [ SymbolAtom "type", StringAtom typ ]
|
||||
]
|
||||
]) = do
|
||||
pure $ MkHighlight
|
||||
{ location = !(fromSExp fc)
|
||||
, name = nam
|
||||
, ns
|
||||
, isImplicit = !(fromSExp impl)
|
||||
, decor = !(fromSExp dec)
|
||||
, docOverview = doc
|
||||
, key, typ}
|
||||
fromSExp _ = Nothing
|
||||
|
||||
export
|
||||
SExpable LwHighlight where
|
||||
toSExp lwhl = SExpList
|
||||
[ toSExp lwhl.location
|
||||
, SExpList [ toSExp lwhl.decor]
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable LwHighlight where
|
||||
fromSExp (SExpList
|
||||
[ location
|
||||
, SExpList [ decor ]
|
||||
]) = do pure $ MkLwHighlight
|
||||
{ location = !(fromSExp location)
|
||||
, decor = !(fromSExp decor)}
|
||||
fromSExp _ = Nothing
|
||||
|
||||
public export
|
||||
data SourceHighlight =
|
||||
Full Highlight
|
||||
| Lw LwHighlight
|
||||
|
||||
export
|
||||
SExpable SourceHighlight where
|
||||
toSExp (Full hl) = toSExp hl
|
||||
toSExp (Lw hl) = toSExp hl
|
||||
|
||||
export
|
||||
FromSExpable SourceHighlight where
|
||||
fromSExp shl = do
|
||||
let Nothing = fromSExp shl
|
||||
| Just hl => Just (Full hl)
|
||||
hl <- fromSExp shl
|
||||
pure $ Lw hl
|
55
src/Protocol/IDE/Holes.idr
Normal file
55
src/Protocol/IDE/Holes.idr
Normal file
@ -0,0 +1,55 @@
|
||||
module Protocol.IDE.Holes
|
||||
|
||||
import Protocol.SExp
|
||||
|
||||
%default total
|
||||
|
||||
public export
|
||||
record HolePremise where
|
||||
constructor MkHolePremise
|
||||
name : String -- Future: send more structured representation of:
|
||||
-- im/explicity (+ default value?) + quantity
|
||||
type : String -- Future: String + highlighting info
|
||||
|
||||
export
|
||||
SExpable HolePremise where
|
||||
toSExp premise =
|
||||
SExpList [ StringAtom premise.name
|
||||
, StringAtom premise.type
|
||||
, SExpList [] -- TODO: metadata
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable HolePremise where
|
||||
fromSExp (SExpList [ StringAtom name
|
||||
, StringAtom type
|
||||
, SExpList [] -- TODO: metadata
|
||||
]) = do pure $ MkHolePremise
|
||||
{name, type}
|
||||
fromSExp _ = Nothing
|
||||
|
||||
public export
|
||||
record HoleData where
|
||||
constructor MkHoleData
|
||||
name : String
|
||||
type : String -- Future : String + highlighting info
|
||||
context : List HolePremise
|
||||
|
||||
export
|
||||
SExpable HoleData where
|
||||
toSExp hole = SExpList
|
||||
[ StringAtom (show hole.name)
|
||||
, toSExp hole.context
|
||||
, SExpList [ toSExp hole.type -- Conclusion
|
||||
, SExpList[]] -- TODO: Highlighting information
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable HoleData where
|
||||
fromSExp (SExpList
|
||||
[ StringAtom name
|
||||
, context
|
||||
, SExpList [ conclusion
|
||||
, SExpList[]] -- TODO: Highlighting information
|
||||
]) = do pure $ MkHoleData {name, type = !(fromSExp conclusion), context = !(fromSExp context)}
|
||||
fromSExp _ = Nothing
|
146
src/Protocol/IDE/Result.idr
Normal file
146
src/Protocol/IDE/Result.idr
Normal file
@ -0,0 +1,146 @@
|
||||
module Protocol.IDE.Result
|
||||
|
||||
import Protocol.SExp
|
||||
|
||||
import Protocol.IDE.Holes
|
||||
import Protocol.IDE.FileContext
|
||||
|
||||
import Data.Maybe
|
||||
|
||||
%default total
|
||||
|
||||
public export
|
||||
data OptionType = BOOL | STRING | ATOM
|
||||
|
||||
public export
|
||||
(.sem) : OptionType -> Type
|
||||
BOOL .sem = Bool
|
||||
STRING .sem = String
|
||||
ATOM .sem = String
|
||||
|
||||
%unbound_implicits off
|
||||
public export
|
||||
record REPLOption where
|
||||
constructor MkOption
|
||||
name : String
|
||||
type : OptionType
|
||||
val : type.sem
|
||||
%unbound_implicits on
|
||||
|
||||
sexpOptionVal : {type : OptionType} -> type.sem -> SExp
|
||||
sexpOptionVal {type = BOOL } = toSExp
|
||||
sexpOptionVal {type = STRING} = toSExp
|
||||
sexpOptionVal {type = ATOM } = toSExp
|
||||
|
||||
export
|
||||
SExpable REPLOption where
|
||||
toSExp opt@(MkOption {}) = SExpList
|
||||
[ SymbolAtom opt.name
|
||||
, sexpOptionVal opt.val
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable REPLOption where
|
||||
fromSExp (SExpList
|
||||
[ SymbolAtom name
|
||||
, val
|
||||
]) = do
|
||||
let Nothing = fromSExp val
|
||||
| Just val => Just $ MkOption {name, type = BOOL, val}
|
||||
let Nothing = fromSExp val
|
||||
| Just val => Just $ MkOption {name, type = STRING, val}
|
||||
val <- fromSExp val
|
||||
Just $ MkOption {name, type = ATOM, val}
|
||||
fromSExp _ = Nothing
|
||||
|
||||
public export
|
||||
record MetaVarLemma where
|
||||
constructor MkMetaVarLemma
|
||||
application, lemma : String
|
||||
|
||||
export
|
||||
SExpable MetaVarLemma where
|
||||
toSExp mvl = SExpList [ SymbolAtom "metavariable-lemma"
|
||||
, SExpList [ SymbolAtom "replace-metavariable", StringAtom mvl.application ]
|
||||
, SExpList [ SymbolAtom "definition-type", StringAtom mvl.lemma ]
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable MetaVarLemma where
|
||||
fromSExp (SExpList [ SymbolAtom "metavariable-lemma"
|
||||
, SExpList [ SymbolAtom "replace-metavariable", StringAtom application ]
|
||||
, SExpList [ SymbolAtom "definition-type", StringAtom lemma ]
|
||||
]) = Just $ MkMetaVarLemma {application, lemma}
|
||||
fromSExp _ = Nothing
|
||||
|
||||
public export
|
||||
record IdrisVersion where
|
||||
constructor MkIdrisVersion
|
||||
major, minor, patch : Nat
|
||||
tag : Maybe String
|
||||
|
||||
export
|
||||
SExpable IdrisVersion where
|
||||
toSExp version = SExpList
|
||||
[ SExpList (map toSExp [version.major, version.minor, version.patch])
|
||||
, SExpList [StringAtom $ fromMaybe "" version.tag]
|
||||
]
|
||||
|
||||
export
|
||||
FromSExpable IdrisVersion where
|
||||
fromSExp (SExpList
|
||||
[ SExpList [majorSExp, minorSExp, patchSExp]
|
||||
, SExpList [StringAtom tagSExp]
|
||||
]) = do pure $ MkIdrisVersion
|
||||
{ major = !(fromSExp majorSExp)
|
||||
, minor = !(fromSExp minorSExp)
|
||||
, patch = !(fromSExp patchSExp)
|
||||
, tag = case tagSExp of
|
||||
"" => Nothing
|
||||
str => Just str}
|
||||
fromSExp _ = Nothing
|
||||
|
||||
public export
|
||||
data Result =
|
||||
AString String
|
||||
| AUnit
|
||||
| AVersion IdrisVersion
|
||||
| AMetaVarLemma MetaVarLemma
|
||||
| ANameLocList (List (String, FileContext))
|
||||
| AHoleList (List HoleData)
|
||||
| ANameList (List String)
|
||||
| AnOptionList (List REPLOption)
|
||||
|
||||
export
|
||||
SExpable Result where
|
||||
toSExp (AString s) = toSExp s
|
||||
toSExp (AUnit ) = toSExp (the (List Int) [])
|
||||
toSExp (AVersion version) = toSExp version
|
||||
toSExp (AMetaVarLemma mvl) = toSExp mvl
|
||||
toSExp (ANameLocList fcs) = toSExp fcs
|
||||
toSExp (AHoleList holes) = toSExp holes
|
||||
toSExp (ANameList names) = SExpList (map SymbolAtom names)
|
||||
toSExp (AnOptionList opts) = toSExp opts
|
||||
|
||||
|
||||
-- This code is not efficient. Usually the client knows what kind of
|
||||
-- result to expect based on the request it issued.
|
||||
export
|
||||
FromSExpable Result where
|
||||
fromSExp (SExpList []) = Just AUnit -- resolve ambiguity somewhat arbitrarily...
|
||||
fromSExp sexp = do
|
||||
let Nothing = fromSExp sexp
|
||||
| Just str => pure $ AString str
|
||||
let Nothing = fromSExp sexp
|
||||
| Just version => pure $ AVersion version
|
||||
let Nothing = fromSExp sexp
|
||||
| Just mvl => pure $ AMetaVarLemma mvl
|
||||
let Nothing = fromSExp sexp
|
||||
| Just nll => pure $ ANameLocList nll
|
||||
let Nothing = fromSExp sexp
|
||||
| Just hl => pure $ AHoleList hl
|
||||
let Nothing = fromSExp sexp
|
||||
| Just nl => pure $ ANameList nl
|
||||
let Nothing = fromSExp sexp
|
||||
| Just optl => pure $ AnOptionList optl
|
||||
Nothing
|
118
src/Protocol/SExp.idr
Normal file
118
src/Protocol/SExp.idr
Normal file
@ -0,0 +1,118 @@
|
||||
module Protocol.SExp
|
||||
|
||||
import Data.List
|
||||
|
||||
%default total
|
||||
|
||||
-- should be in base somewhere!
|
||||
join : String -> List String -> String
|
||||
join sep xs = concat $ intersperse sep xs
|
||||
|
||||
public export
|
||||
data SExp = SExpList (List SExp)
|
||||
| StringAtom String
|
||||
| BoolAtom Bool
|
||||
| IntegerAtom Integer
|
||||
| SymbolAtom String
|
||||
|
||||
escape : String -> String
|
||||
escape = pack . concatMap escapeChar . unpack
|
||||
where
|
||||
escapeChar : Char -> List Char
|
||||
escapeChar '\\' = ['\\', '\\']
|
||||
escapeChar '"' = ['\\', '\"']
|
||||
escapeChar c = [c]
|
||||
|
||||
export
|
||||
Show SExp where
|
||||
show (SExpList xs) = assert_total $ "(" ++ join " " (map show xs) ++ ")"
|
||||
show (StringAtom str) = "\"" ++ escape str ++ "\""
|
||||
show (BoolAtom b) = ":" ++ show b
|
||||
show (IntegerAtom i) = show i
|
||||
show (SymbolAtom s) = ":" ++ s
|
||||
|
||||
public export
|
||||
interface SExpable a where
|
||||
toSExp : a -> SExp
|
||||
|
||||
-- TODO: Merge these into 1 interface later
|
||||
public export
|
||||
interface FromSExpable a where
|
||||
fromSExp : SExp -> Maybe a
|
||||
|
||||
export
|
||||
SExpable SExp where
|
||||
toSExp = id
|
||||
|
||||
export
|
||||
SExpable Bool where
|
||||
toSExp = BoolAtom
|
||||
|
||||
export
|
||||
FromSExpable Bool where
|
||||
fromSExp (BoolAtom b) = Just b
|
||||
fromSExp _ = Nothing
|
||||
|
||||
export
|
||||
SExpable String where
|
||||
toSExp = StringAtom
|
||||
|
||||
export
|
||||
FromSExpable String where
|
||||
fromSExp (StringAtom s) = Just s
|
||||
fromSExp _ = Nothing
|
||||
|
||||
export
|
||||
SExpable Integer where
|
||||
toSExp = IntegerAtom
|
||||
|
||||
export
|
||||
FromSExpable Integer where
|
||||
fromSExp (IntegerAtom a) = Just a
|
||||
fromSExp _ = Nothing
|
||||
|
||||
export
|
||||
SExpable Int where
|
||||
toSExp = IntegerAtom . cast
|
||||
|
||||
export
|
||||
FromSExpable Int where
|
||||
fromSExp a = do Just $ cast {from = Integer }$ !(fromSExp a)
|
||||
|
||||
export
|
||||
SExpable Nat where
|
||||
toSExp = IntegerAtom . cast
|
||||
|
||||
export
|
||||
FromSExpable Nat where
|
||||
fromSExp a = do Just $ cast {from = Integer }$ !(fromSExp a)
|
||||
|
||||
export
|
||||
(SExpable a, SExpable b) => SExpable (a, b) where
|
||||
toSExp (x, y)
|
||||
= case toSExp y of
|
||||
SExpList xs => SExpList (toSExp x :: xs)
|
||||
y' => SExpList [toSExp x, y']
|
||||
|
||||
export
|
||||
(FromSExpable a, FromSExpable b) => FromSExpable (a, b) where
|
||||
fromSExp (SExpList xs) = case xs of
|
||||
[x,y] => do pure $ (!(fromSExp x), !(fromSExp y))
|
||||
(x :: xs) => do pure $ (!(fromSExp x), !(fromSExp $ SExpList xs))
|
||||
_ => Nothing
|
||||
fromSExp _ = Nothing
|
||||
|
||||
export
|
||||
SExpable a => SExpable (List a) where
|
||||
toSExp xs
|
||||
= SExpList (map toSExp xs)
|
||||
|
||||
export
|
||||
FromSExpable a => FromSExpable (List a) where
|
||||
fromSExp (SExpList sexps) = traverse fromSExp sexps
|
||||
fromSExp _ = Nothing
|
||||
|
||||
export
|
||||
SExpable a => SExpable (Maybe a) where
|
||||
toSExp Nothing = SExpList []
|
||||
toSExp (Just x) = toSExp x
|
@ -17,7 +17,7 @@ import TTImp.TTImp
|
||||
import TTImp.TTImp.Functor
|
||||
import TTImp.Unelab
|
||||
|
||||
import Libraries.Utils.Hex
|
||||
import Protocol.Hex
|
||||
|
||||
import Data.List
|
||||
import Libraries.Data.NameMap
|
||||
|
@ -1,69 +1,69 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
000038(:write-string "1/1: Building LocType (LocType.idr)" 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 1) (:end 1 4)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 6) (:end 1 9)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 13) (:end 1 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 17) (:end 1 18)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 20) (:end 1 23)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 25) (:end 1 26)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 28) (:end 1 31)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 33) (:end 1 37)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 6) (:end 2 8)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 10)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 12) (:end 2 15)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 17) (:end 2 17)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 19) (:end 2 19)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 6) (:end 3 9)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 11) (:end 3 11)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 13) (:end 3 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 18) (:end 3 21)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 23) (:end 3 23)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 25) (:end 3 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 27) (:end 3 28)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 30) (:end 3 33)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 35) (:end 3 35)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 36) (:end 3 36)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 38) (:end 3 38)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 39) (:end 3 39)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 41) (:end 3 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 5 6)) ((:decor :function)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 8) (:end 5 8)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 10) (:end 5 13)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 15) (:end 5 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 17) (:end 5 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 19) (:end 5 20)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 22) (:end 5 25)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 27) (:end 5 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 29) (:end 5 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 31) (:end 5 32)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 34) (:end 5 37)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 39) (:end 5 39)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 42) (:end 5 42)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 44)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 45) (:end 5 45)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 5 47)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 1) (:end 6 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 8) (:end 6 9)) ((:name "Nil") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 12)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 14) (:end 6 14)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 6 17)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 1) (:end 7 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 8) (:end 7 8)) ((:decor :keyword)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 11) (:end 7 12)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 15)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 16) (:end 7 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 19)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 21) (:end 7 21)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 23)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 25) (:end 7 26)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 28) (:end 7 33)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 36)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 7 39)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 0) (:end 0 4)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 5) (:end 0 9)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 10) (:end 0 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 12) (:end 0 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 16) (:end 0 18)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 19) (:end 0 23)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 24) (:end 0 26)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 27) (:end 0 31)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 32) (:end 0 37)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 5) (:end 1 8)) ((:decor :data)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 9) (:end 1 10)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 15)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 16) (:end 1 17)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 18) (:end 1 19)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 5) (:end 2 9)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 11)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 12) (:end 2 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 14) (:end 2 16)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 17) (:end 2 21)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 22) (:end 2 23)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 24) (:end 2 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 26) (:end 2 28)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 29) (:end 2 33)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 34) (:end 2 35)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 35) (:end 2 36)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 37) (:end 2 38)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 38) (:end 2 39)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 40) (:end 2 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 0) (:end 4 6)) ((:decor :function)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 9) (:end 4 13)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 14) (:end 4 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 16) (:end 4 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 18) (:end 4 20)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 21) (:end 4 25)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 26) (:end 4 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 28) (:end 4 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 30) (:end 4 32)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 33) (:end 4 37)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 38) (:end 4 39)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 39) (:end 4 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 41) (:end 4 42)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 43) (:end 4 44)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 44) (:end 4 45)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 46) (:end 4 47)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 0) (:end 5 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 7) (:end 5 9)) ((:name "Nil") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 10) (:end 5 12)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 13) (:end 5 14)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 15) (:end 5 17)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 0) (:end 6 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 7) (:end 6 8)) ((:decor :keyword)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 8) (:end 6 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 10) (:end 6 12)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 13) (:end 6 15)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 15) (:end 6 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 17) (:end 6 19)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 20) (:end 6 21)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 22) (:end 6 23)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 24) (:end 6 26)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 27) (:end 6 33)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 34) (:end 6 36)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 37) (:end 6 39)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
000040(:return (:ok "Syntax highlight option changed to False" ()) 2)
|
||||
00003d(:return (:ok "Syntax highlight option changed to False") 2)
|
||||
000015(:return (:ok ()) 3)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,3 +1,3 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
__EXPECTED_LINE__
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,68 +1,68 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
000038(:write-string "1/1: Building LocType (LocType.idr)" 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 1) (:end 1 4)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 6) (:end 1 9)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 13) (:end 1 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 17) (:end 1 18)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 20) (:end 1 23)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 25) (:end 1 26)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 28) (:end 1 31)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 33) (:end 1 37)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 6) (:end 2 8)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 10)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 12) (:end 2 15)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 17) (:end 2 17)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 19) (:end 2 19)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 6) (:end 3 9)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 11) (:end 3 11)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 13) (:end 3 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 18) (:end 3 21)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 23) (:end 3 23)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 25) (:end 3 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 27) (:end 3 28)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 30) (:end 3 33)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 35) (:end 3 35)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 36) (:end 3 36)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 38) (:end 3 38)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 39) (:end 3 39)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 3 41) (:end 3 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 1) (:end 5 6)) ((:decor :function)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 8) (:end 5 8)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 10) (:end 5 13)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 15) (:end 5 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 17) (:end 5 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 19) (:end 5 20)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 22) (:end 5 25)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 27) (:end 5 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 29) (:end 5 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 31) (:end 5 32)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 34) (:end 5 37)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 39) (:end 5 39)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 40) (:end 5 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 42) (:end 5 42)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 44) (:end 5 44)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 45) (:end 5 45)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 47) (:end 5 47)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 1) (:end 6 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 8) (:end 6 9)) ((:name "Nil") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 11) (:end 6 12)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 14) (:end 6 14)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 16) (:end 6 17)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 1) (:end 7 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 8) (:end 7 8)) ((:decor :keyword)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 9) (:end 7 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 11) (:end 7 12)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 14) (:end 7 15)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 16) (:end 7 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 18) (:end 7 19)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 21) (:end 7 21)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 23) (:end 7 23)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 25) (:end 7 26)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 28) (:end 7 33)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 35) (:end 7 36)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 7 38) (:end 7 39)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 0) (:end 0 4)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 5) (:end 0 9)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 10) (:end 0 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 12) (:end 0 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 16) (:end 0 18)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 19) (:end 0 23)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 24) (:end 0 26)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 27) (:end 0 31)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 0 32) (:end 0 37)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 5) (:end 1 8)) ((:decor :data)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 9) (:end 1 10)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 11) (:end 1 15)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 16) (:end 1 17)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 1 18) (:end 1 19)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 5) (:end 2 9)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 10) (:end 2 11)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 12) (:end 2 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 14) (:end 2 16)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 17) (:end 2 21)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 22) (:end 2 23)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 24) (:end 2 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 26) (:end 2 28)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 29) (:end 2 33)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 34) (:end 2 35)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 35) (:end 2 36)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 37) (:end 2 38)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 38) (:end 2 39)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 2 40) (:end 2 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 0) (:end 4 6)) ((:decor :function)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 9) (:end 4 13)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 14) (:end 4 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 16) (:end 4 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 18) (:end 4 20)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 21) (:end 4 25)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 26) (:end 4 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 28) (:end 4 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 30) (:end 4 32)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 33) (:end 4 37)) ((:name "Vect") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 38) (:end 4 39)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 39) (:end 4 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 41) (:end 4 42)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 43) (:end 4 44)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 44) (:end 4 45)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 4 46) (:end 4 47)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 0) (:end 5 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 7) (:end 5 9)) ((:name "Nil") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 10) (:end 5 12)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 13) (:end 5 14)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 5 15) (:end 5 17)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 0) (:end 6 6)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 7) (:end 6 8)) ((:decor :keyword)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 8) (:end 6 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 10) (:end 6 12)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 13) (:end 6 15)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 15) (:end 6 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 17) (:end 6 19)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 20) (:end 6 21)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 22) (:end 6 23)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 24) (:end 6 26)) ((:name "::") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 27) (:end 6 33)) ((:name "append") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 34) (:end 6 36)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "LocType.idr") (:start 6 37) (:end 6 39)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
000095(:return (:ok "Main.Vect : Nat -> Type -> Type" ((0 9 ((:decor :type))) (12 3 ((:decor :type))) (19 4 ((:decor :type))) (27 4 ((:decor :type))))) 2)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,3 +1,3 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
000037(:return (:ok "\"Test\"" ((0 6 ((:decor :data))))) 32)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,307 +1,307 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
000036(:write-string "1/1: Building Syntax (Syntax.idr)" 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 29) (:end 42 29)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 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 1 8) (:end 1 13)) ((:decor :module)))))) 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 5 1) (:end 5 9)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 11) (:end 5 11)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 13) (:end 5 13)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 14) (:end 5 14)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 16) (:end 5 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 18) (:end 5 18)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 20) (:end 5 23)) ((:decor :type)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 24) (:end 5 24)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 28) (:end 41 29)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 0 7) (:end 0 13)) ((:decor :module)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 2 0) (:end 2 8)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 2 9) (:end 2 14)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 0) (:end 4 9)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 10) (:end 4 11)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 12) (:end 4 13)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 13) (:end 4 14)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 15) (:end 4 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 17) (:end 4 18)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 19) (:end 4 23)) ((:decor :type)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 23) (:end 4 24)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 25) (:end 4 27)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 28) (:end 4 29)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 29) (:end 4 39)) ((:name "assumption") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 40) (:end 4 41)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 42) (:end 4 46)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 47) (:end 4 48)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 48) (:end 4 49)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 50) (:end 4 52)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 53) (:end 4 58)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 59) (:end 4 60)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 61) (:end 4 63)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 4 64) (:end 4 70)) ((:decor :type)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 0) (:end 5 9)) ((:name "showMaybe") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 10) (:end 5 11)) ((:decor :bound)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 11) (:end 5 12)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 12) (:end 5 14)) ((:name "ma") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 15) (:end 5 16)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 17) (:end 5 21)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 22) (:end 5 25)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 26) (:end 5 27)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 29) (:end 5 29)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 30) (:end 5 39)) ((:name "assumption") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 41) (:end 5 41)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 43) (:end 5 46)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 48) (:end 5 48)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 49) (:end 5 49)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 51) (:end 5 52)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 54) (:end 5 58)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 60) (:end 5 60)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 62) (:end 5 63)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 65) (:end 5 70)) ((:decor :type)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 1) (:end 6 9)) ((:name "showMaybe") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 11) (:end 6 11)) ((:decor :bound)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 12) (:end 6 12)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 13) (:end 6 14)) ((:name "ma") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 16) (:end 6 16)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 18) (:end 6 21)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 23) (:end 6 25)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 27) (:end 6 27)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 28) (:end 6 29)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 31) (:end 6 31)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 33) (:end 6 34)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 35) (:end 6 35)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 37) (:end 6 38)) ((:name "ma") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 40) (:end 6 41)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 5) (:end 7 11)) ((:name "Nothing") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 13) (:end 7 14)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 16) (:end 7 24)) ((:decor :data)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 8 5) (:end 8 8)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 8 10) (:end 8 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 8 12) (:end 8 13)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 8 15) (:end 8 21)) ((:decor :data)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 8 23) (:end 8 24)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 8 26) (:end 8 29)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 8 31) (:end 8 31)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 1) (:end 10 4)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 6) (:end 10 6)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 8) (:end 10 11)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 13) (:end 10 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 1) (:end 11 4)) ((:name "nats") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 6) (:end 11 6)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 3) (:end 12 5)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 7) (:end 12 7)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 9) (:end 12 9)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 11) (:end 12 11)) ((:decor :data)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 7) (:end 13 7)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 9) (:end 13 9)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 11) (:end 13 11)) ((:decor :data)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 7) (:end 14 8)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 10) (:end 14 10)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 12) (:end 14 12)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 13) (:end 14 14)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 18) (:end 14 19)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 21) (:end 14 21)) ((:decor :data)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 22) (:end 14 22)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 7) (:end 15 8)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 10) (:end 15 10)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 12) (:end 15 12)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 13) (:end 15 13)) ((:decor :data)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 14) (:end 15 14)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 16) (:end 15 16)) ((:decor :data)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 17) (:end 15 17)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 19) (:end 15 19)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 20) (:end 15 20)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 22) (:end 15 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 23) (:end 15 23)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 25) (:end 15 25)) ((:decor :data)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 26) (:end 15 26)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 28) (:end 15 29)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 31) (:end 15 32)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 3) (:end 16 4)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 6) (:end 16 6)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 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)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 8) (:end 16 8)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 9) (:end 16 10)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 14) (:end 16 15)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 17) (:end 16 17)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 18) (:end 16 18)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 20) (:end 16 21)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 23) (:end 16 23)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 24) (:end 16 24)) ((:decor :data)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 25) (:end 16 25)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 27) (:end 16 27)) ((:decor :data)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 28) (:end 16 28)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 30) (:end 16 30)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 31) (:end 16 31)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 33) (:end 16 33)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 34) (:end 16 34)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 36) (:end 16 36)) ((:decor :data)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 37) (:end 16 37)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 39) (:end 16 40)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 16 42) (:end 16 43)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 1) (:end 19 6)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 8) (:end 19 11)) ((:decor :type)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 13) (:end 19 17)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 15) (:end 20 20)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 21 3) (:end 21 6)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 21 8) (:end 21 8)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 21 10) (:end 21 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 21 10) (:end 21 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 21 10) (:end 21 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 23 1) (:end 23 7)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 23 9) (:end 23 9)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 23 11) (:end 23 15)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 23 17) (:end 23 20)) ((:name "ANat") (:namespace "Syntax") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 24 1) (:end 24 7)) ((:name "doBlock") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 3) (:end 25 3)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 8) (:end 25 10)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 12) (:end 25 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 14) (:end 25 14)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 16) (:end 25 16)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 8) (:end 26 10)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 12) (:end 26 12)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 14) (:end 26 14)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 16) (:end 26 16)) ((:decor :data)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 8) (:end 27 8)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 9) (:end 27 9)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 10) (:end 27 10)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 12) (:end 27 12)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 13) (:end 27 13)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 15) (:end 27 16)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 18) (:end 27 21)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 23) (:end 27 23)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 24) (:end 27 24)) ((:decor :data)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 25) (:end 27 25)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 27) (:end 27 27)) ((:decor :data)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 28) (:end 27 28)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 8) (:end 28 10)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 12) (:end 28 12)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 13) (:end 28 13)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 14) (:end 28 14)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 16) (:end 28 16)) ((:name "e") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 17) (:end 28 17)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 19) (:end 28 19)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 21) (:end 28 21)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 22) (:end 28 22)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 23) (:end 28 23)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 25) (:end 28 25)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 26) (:end 28 26)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 8) (:end 29 8)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 10) (:end 29 11)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 13) (:end 29 14)) ((:decor :keyword)))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 16) (:end 29 22)) ((:name "Nothing") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 24) (:end 29 24)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 26) (:end 29 29)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 31) (:end 29 31)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 33) (:end 29 34)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 8) (:end 30 11)) ((:name "pure") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 15) (:end 30 20)) ((:name "MkANat") (:namespace "Syntax") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 24) (:end 30 26)) ((:name "sum") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 28) (:end 30 28)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 29) (:end 30 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 30) (:end 30 30)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 31) (:end 30 31)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 32) (:end 30 32)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 33) (:end 30 33)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 34) (:end 30 34)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 35) (:end 30 35)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 36) (:end 30 36)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 37) (:end 30 37)) ((:name "e") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 38) (:end 30 38)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 39) (:end 30 39)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 30 40) (:end 30 40)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 1) (:end 32 10)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 12) (:end 32 12)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 13) (:end 32 13)) ((:name "x") (: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 "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 17) (:end 32 17)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 19) (:end 32 19)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 21) (:end 32 21)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 23) (:end 32 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 23) (:end 32 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 23) (:end 32 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 32 26) (:end 32 26)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 3) (:end 34 6)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 8) (:end 34 8)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 10) (:end 34 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 35 3) (:end 35 6)) ((:name "add3") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 35 8) (:end 35 8)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 35 10) (:end 35 10)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 35 12) (:end 35 12)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 35 14) (:end 35 14)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 35 16) (:end 35 16)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 35 18) (:end 35 18)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 1) (:end 37 10)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 12) (:end 37 12)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 13) (:end 37 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 14) (:end 37 14)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 16) (:end 37 16)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 18) (:end 37 18)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 20) (:end 37 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 20) (:end 37 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 23) (:end 37 23)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 25) (:end 37 25)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 26) (:end 37 26)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 27) (:end 37 27)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 29) (:end 37 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 31) (:end 37 31)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 33) (:end 37 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 33) (:end 37 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 37 36) (:end 37 36)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 3) (:end 39 6)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 8) (:end 39 8)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 10) (:end 39 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 3) (:end 40 6)) ((:name "add4") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 8) (:end 40 8)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 10) (:end 40 10)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 12) (:end 40 12)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 14) (:end 40 14)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 16) (:end 40 16)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 18) (:end 40 18)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 20) (:end 40 20)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 40 22) (:end 40 22)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 1) (:end 42 7)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 9) (:end 42 9)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 11) (:end 42 15)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 17) (:end 42 17)) ((:decor :type)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 18) (:end 42 18)) ((:decor :bound)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 20) (:end 42 20)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 22) (:end 42 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 22) (:end 42 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 26) (:end 42 27)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 31) (:end 42 32)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 34) (:end 42 34)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 36) (:end 42 38)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 40) (:end 42 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 41) (:end 42 41)) ((:decor :type)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 1) (:end 43 7)) ((:name "anonLam") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 9) (:end 43 9)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 11) (:end 43 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 15) (:end 43 15)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 16) (:end 43 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 17) (:end 43 17)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 19) (:end 43 20)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 23) (:end 43 23)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 25) (:end 43 26)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 28) (:end 43 28)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 30) (:end 43 31)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 33) (:end 43 36)) ((:name "Refl") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 38) (:end 43 38)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 11) (:end 44 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 15) (:end 44 15)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 16) (:end 44 22)) ((:name "uncurry") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 26) (:end 44 26)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 28) (:end 44 28)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 29) (:end 44 29)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 31) (:end 44 31)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 27) (:end 5 29)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 30) (:end 5 31)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 32) (:end 5 34)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 34) (:end 5 35)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 36) (:end 5 38)) ((:name "ma") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 5 39) (:end 5 41)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 4) (:end 6 11)) ((:name "Nothing") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 12) (:end 6 14)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 6 15) (:end 6 24)) ((:decor :data)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 4) (:end 7 8)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 9) (:end 7 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 11) (:end 7 13)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 14) (:end 7 21)) ((:decor :data)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 22) (:end 7 24)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 25) (:end 7 29)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 7 30) (:end 7 31)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 0) (:end 9 4)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 5) (:end 9 6)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 7) (:end 9 11)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 9 12) (:end 9 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 0) (:end 10 4)) ((:name "nats") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 10 5) (:end 10 6)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 2) (:end 11 5)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 6) (:end 11 7)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 8) (:end 11 9)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 11 10) (:end 11 11)) ((:decor :data)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 6) (:end 12 7)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 8) (:end 12 9)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 12 10) (:end 12 11)) ((:decor :data)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 6) (:end 13 8)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 9) (:end 13 10)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 11) (:end 13 12)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 12) (:end 13 14)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 17) (:end 13 19)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 20) (:end 13 21)) ((:decor :data)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 13 21) (:end 13 22)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 6) (:end 14 8)) ((:name "ys") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 9) (:end 14 10)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 11) (:end 14 12)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 12) (:end 14 13)) ((:decor :data)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 13) (:end 14 14)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 15) (:end 14 16)) ((:decor :data)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 16) (:end 14 17)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 18) (:end 14 19)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 19) (:end 14 20)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 21) (:end 14 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 22) (:end 14 23)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 24) (:end 14 25)) ((:decor :data)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 25) (:end 14 26)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 27) (:end 14 29)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 14 30) (:end 14 32)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 2) (:end 15 4)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 5) (:end 15 6)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 6) (:end 15 7)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 7) (:end 15 8)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 8) (:end 15 10)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 13) (:end 15 15)) ((:name "id") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 16) (:end 15 17)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 17) (:end 15 18)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 19) (:end 15 21)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 22) (:end 15 23)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 23) (:end 15 24)) ((:decor :data)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 24) (:end 15 25)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 26) (:end 15 27)) ((:decor :data)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 27) (:end 15 28)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 29) (:end 15 30)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 30) (:end 15 31)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 32) (:end 15 33)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 33) (:end 15 34)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 35) (:end 15 36)) ((:decor :data)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 36) (:end 15 37)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 38) (:end 15 40)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 15 41) (:end 15 43)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 0) (:end 18 6)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 7) (:end 18 11)) ((:decor :type)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 18 12) (:end 18 17)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 19 14) (:end 19 20)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 2) (:end 20 6)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 7) (:end 20 8)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 9) (:end 20 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 9) (:end 20 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 20 9) (:end 20 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 0) (:end 22 7)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 8) (:end 22 9)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 10) (:end 22 15)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 22 16) (:end 22 20)) ((:name "ANat") (:namespace "Syntax") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 23 0) (:end 23 7)) ((:name "doBlock") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 24 2) (:end 24 3)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 24 7) (:end 24 10)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 24 11) (:end 24 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 24 13) (:end 24 14)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 24 15) (:end 24 16)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 7) (:end 25 10)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 11) (:end 25 12)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 13) (:end 25 14)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 25 15) (:end 25 16)) ((:decor :data)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 7) (:end 26 8)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 8) (:end 26 9)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 9) (:end 26 10)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 11) (:end 26 12)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 12) (:end 26 13)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 14) (:end 26 16)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 17) (:end 26 21)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 22) (:end 26 23)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 23) (:end 26 24)) ((:decor :data)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 24) (:end 26 25)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 26) (:end 26 27)) ((:decor :data)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 26 27) (:end 26 28)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 7) (:end 27 10)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 11) (:end 27 12)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 12) (:end 27 13)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 13) (:end 27 14)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 15) (:end 27 16)) ((:name "e") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 16) (:end 27 17)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 18) (:end 27 19)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 20) (:end 27 21)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 21) (:end 27 22)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 22) (:end 27 23)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 24) (:end 27 25)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 27 25) (:end 27 26)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 7) (:end 28 8)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 9) (:end 28 11)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 12) (:end 28 14)) ((:decor :keyword)))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 15) (:end 28 22)) ((:name "Nothing") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 23) (:end 28 24)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 25) (:end 28 29)) ((:name "Just") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 30) (:end 28 31)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 28 32) (:end 28 34)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 7) (:end 29 11)) ((:name "pure") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 14) (:end 29 20)) ((:name "MkANat") (:namespace "Syntax") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 23) (:end 29 26)) ((:name "sum") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 27) (:end 29 28)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 28) (:end 29 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 29) (:end 29 30)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 30) (:end 29 31)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 31) (:end 29 32)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 32) (:end 29 33)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 33) (:end 29 34)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 34) (:end 29 35)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 35) (:end 29 36)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 36) (:end 29 37)) ((:name "e") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 37) (:end 29 38)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 38) (:end 29 39)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 29 39) (:end 29 40)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 0) (:end 31 10)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 11) (:end 31 12)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 12) (:end 31 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 13) (:end 31 14)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 15) (:end 31 16)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 16) (:end 31 17)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 18) (:end 31 19)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 20) (:end 31 21)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 22) (:end 31 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 22) (:end 31 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 22) (:end 31 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 31 25) (:end 31 26)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 2) (:end 33 6)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 7) (:end 33 8)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 33 9) (:end 33 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 2) (:end 34 6)) ((:name "add3") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 7) (:end 34 8)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 9) (:end 34 10)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 11) (:end 34 12)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 13) (:end 34 14)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 15) (:end 34 16)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 34 17) (:end 34 18)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 0) (:end 36 10)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 11) (:end 36 12)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 12) (: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 13) (:end 36 14)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 15) (: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 18)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 19) (:end 36 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 19) (:end 36 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 22) (:end 36 23)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 24) (:end 36 25)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 25) (:end 36 26)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 26) (:end 36 27)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 28) (:end 36 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 30) (:end 36 31)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 32) (:end 36 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 32) (:end 36 35)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 36 35) (:end 36 36)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 2) (:end 38 6)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 7) (:end 38 8)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 38 9) (:end 38 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 2) (:end 39 6)) ((:name "add4") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 7) (:end 39 8)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 9) (: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 11) (: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 13) (: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 15) (: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 17) (:end 39 18)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 19) (:end 39 20)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 39 21) (:end 39 22)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 0) (:end 41 7)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 8) (:end 41 9)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 10) (:end 41 15)) ((:name "Maybe") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 16) (:end 41 17)) ((:decor :type)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 17) (:end 41 18)) ((:decor :bound)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 19) (:end 41 20)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 21) (:end 41 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 21) (:end 41 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 25) (:end 41 27)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 30) (:end 41 32)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 33) (:end 41 34)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 35) (:end 41 38)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 39) (:end 41 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 41 40) (:end 41 41)) ((:decor :type)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 0) (:end 42 7)) ((:name "anonLam") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 8) (:end 42 9)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 10) (:end 42 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 14) (:end 42 15)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 15) (:end 42 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 16) (:end 42 17)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 18) (:end 42 20)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 22) (:end 42 23)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 24) (:end 42 26)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 27) (:end 42 28)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 29) (:end 42 31)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 32) (:end 42 36)) ((:name "Refl") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 42 37) (:end 42 38)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 10) (:end 43 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 14) (:end 43 15)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 15) (:end 43 22)) ((:name "uncurry") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 25) (:end 43 26)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 27) (:end 43 28)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 28) (:end 43 29)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 30) (:end 43 31)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 32) (:end 43 34)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 35) (:end 43 36)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 37) (:end 43 38)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 39) (:end 43 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 43 40) (:end 43 41)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 10) (:end 44 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 14) (:end 44 15)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 15) (:end 44 16)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 17) (:end 44 18)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 18) (:end 44 19)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 19) (:end 44 20)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 21) (:end 44 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 22) (:end 44 23)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 24) (:end 44 26)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 27) (:end 44 28)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 28) (:end 44 29)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 29) (:end 44 30)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 31) (:end 44 32)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 32) (:end 44 33)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 33) (:end 44 34)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 36) (:end 44 36)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 38) (:end 44 38)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 40) (:end 44 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 44 41) (:end 44 41)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 11) (:end 45 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 15) (:end 45 15)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 16) (:end 45 16)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 18) (:end 45 18)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 19) (:end 45 19)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 20) (:end 45 20)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 22) (:end 45 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 23) (:end 45 23)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 25) (:end 45 26)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 28) (:end 45 28)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 29) (:end 45 29)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 30) (:end 45 30)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 32) (:end 45 32)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 33) (:end 45 33)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 34) (:end 45 34)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 11) (:end 46 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 15) (:end 46 15)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 16) (:end 46 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 18) (:end 46 18)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 20) (:end 46 21)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 23) (:end 46 23)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 24) (:end 46 24)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 25) (:end 46 25)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 27) (:end 46 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 28) (:end 46 28)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 29) (:end 46 29)) ((: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 18) (:end 47 18)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 20) (:end 47 21)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 23) (:end 47 23)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 25) (:end 47 25)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 26) (:end 47 30)) ((:name "aNat") (:namespace "Syntax.ANat") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 31) (:end 47 31)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 48 9) (:end 48 15)) ((:name "doBlock") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 10) (:end 45 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 14) (:end 45 15)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 15) (:end 45 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 17) (:end 45 18)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 19) (:end 45 21)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 22) (:end 45 23)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 23) (:end 45 24)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 24) (:end 45 25)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 26) (:end 45 27)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 27) (:end 45 28)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 45 28) (:end 45 29)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 10) (:end 46 13)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 14) (:end 46 15)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 15) (:end 46 16)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 17) (:end 46 18)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 19) (:end 46 21)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 22) (:end 46 23)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 24) (:end 46 25)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 25) (:end 46 30)) ((:name "aNat") (:namespace "Syntax.ANat") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 46 30) (:end 46 31)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Syntax.idr") (:start 47 8) (:end 47 15)) ((:name "doBlock") (:namespace "Syntax") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,78 +1,78 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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 1 8) (:end 1 16)) ((:decor :module)))))) 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)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 13) (:end 12 16)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 13) (:end 12 16)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 13) (:end 12 16)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 23) (:end 12 26)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 23) (:end 12 26)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 23) (:end 12 26)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 11) (:end 13 12)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 16) (:end 13 17)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 14) (:end 14 15)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 14 32) (:end 14 33)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 0 7) (:end 0 16)) ((:decor :module)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 0) (:end 2 9)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 10) (:end 2 14)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 10) (:end 2 14)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 15) (:end 2 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 15) (:end 2 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 17) (:end 2 19)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 20) (:end 2 26)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 27) (:end 2 28)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 28) (:end 2 29)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 30) (:end 2 31)) ((:decor :bound)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 32) (:end 2 33)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 34) (:end 2 38)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 38) (:end 2 39)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 2 40) (:end 2 45)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 3 14) (:end 3 22)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 2) (:end 5 3)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 4) (:end 5 7)) ((:name "Doc") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 4) (:end 5 7)) ((:decor :function)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 8) (:end 5 9)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 5 10) (:end 5 14)) ((:decor :type)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 2) (:end 6 7)) ((:name "toDoc") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 2) (:end 6 7)) ((:decor :function)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 8) (:end 6 9)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 10) (:end 6 16)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 17) (:end 6 19)) ((:decor :keyword)))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 6 20) (:end 6 23)) ((:name "Doc") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 2) (:end 8 8)) ((:name "pretty") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 2) (:end 8 8)) ((:decor :function)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 9) (:end 8 10)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 11) (:end 8 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 11) (:end 8 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 11) (:end 8 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 13) (:end 8 15)) ((:decor :keyword)))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 8 16) (:end 8 19)) ((:name "Doc") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 2) (:end 9 8)) ((:name "pretty") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 9) (:end 9 10)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 11) (:end 9 12)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 13) (:end 9 18)) ((:name "toDoc") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 19) (:end 9 20)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 20) (:end 9 24)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 25) (:end 9 26)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 9 26) (:end 9 27)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 2) (:end 11 9)) ((:name "prettys") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 2) (:end 11 9)) ((:decor :function)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 10) (:end 11 11)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 12) (:end 11 16)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 12) (:end 11 16)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 12) (:end 11 16)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 17) (:end 11 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 17) (:end 11 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 17) (:end 11 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 19) (:end 11 21)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 22) (:end 11 26)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 22) (:end 11 26)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 22) (:end 11 26)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 11 27) (:end 11 30)) ((:name "Doc") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 2) (:end 12 9)) ((:name "prettys") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 10) (:end 12 12)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 13) (:end 12 14)) ((:decor :keyword)))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 12 15) (:end 12 17)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 2) (:end 13 9)) ((:name "prettys") (:namespace "") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 10) (:end 13 11)) ((:decor :keyword)))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 11) (:end 13 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 13) (:end 13 15)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 16) (:end 13 18)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 18) (:end 13 19)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 20) (:end 13 21)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 22) (:end 13 28)) ((:name "pretty") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 29) (:end 13 30)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 31) (:end 13 33)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 34) (:end 13 41)) ((:name "prettys") (:namespace "Interface") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "Interface.idr") (:start 13 42) (:end 13 44)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,64 +1,64 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
00003e(:write-string "1/1: Building SimpleData (SimpleData.idr)" 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 1 8) (:end 1 17)) ((:decor :module)))))) 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)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 0 7) (:end 0 17)) ((:decor :module)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 0) (:end 2 4)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 5) (:end 2 8)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 5) (:end 2 8)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 5) (:end 2 8)) ((:decor :type)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 9) (:end 2 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 9) (:end 2 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 9) (:end 2 10)) ((:decor :bound)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 11) (:end 2 12)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 13) (:end 2 17)) ((:decor :data)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 18) (:end 2 19)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 20) (:end 2 24)) ((:decor :data)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 25) (:end 2 26)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 27) (:end 2 28)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 28) (:end 2 31)) ((:name "Fwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 32) (:end 2 33)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 2 33) (:end 2 34)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 0) (:end 3 4)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 5) (:end 3 8)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 5) (:end 3 8)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 5) (:end 3 8)) ((:decor :type)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 9) (:end 3 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 9) (:end 3 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 9) (:end 3 10)) ((:decor :bound)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 11) (:end 3 12)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 13) (:end 3 17)) ((:decor :data)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 18) (:end 3 19)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 20) (:end 3 24)) ((:decor :data)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 25) (:end 3 26)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 26) (:end 3 29)) ((:name "Bwd") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 30) (:end 3 31)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 31) (:end 3 32)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 3 33) (:end 3 34)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 0) (:end 5 4)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 5) (:end 5 9)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 5) (:end 5 9)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 5) (:end 5 9)) ((:decor :type)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 10) (:end 5 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 10) (:end 5 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 10) (:end 5 11)) ((:decor :bound)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 12) (:end 5 13)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 12) (:end 5 13)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 5 12) (:end 5 13)) ((:decor :bound)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 2) (:end 6 3)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 4) (:end 6 8)) ((:decor :data)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 6 9) (:end 6 10)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 2) (:end 7 3)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 4) (:end 7 8)) ((:decor :data)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 9) (:end 7 10)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 10) (:end 7 14)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 15) (:end 7 16)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 17) (:end 7 18)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 18) (:end 7 19)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 20) (:end 7 21)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 22) (:end 7 23)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 23) (:end 7 27)) ((:name "Tree") (:namespace "SimpleData") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 28) (:end 7 29)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 30) (:end 7 31)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SimpleData.idr") (:start 7 31) (:end 7 32)) ((:decor :keyword)))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1 +1 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
|
@ -1,93 +1,93 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
000038(:write-string "1/1: Building Rainbow (Rainbow.idr)" 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 1 11) (:end 1 12)) ((:decor :namespace)))))) 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)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 6 11) (:end 6 12)) ((:decor :namespace)))))) 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)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 16 11) (:end 16 12)) ((:decor :namespace)))))) 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)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 20 11) (:end 20 12)) ((:decor :namespace)))))) 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)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 27 5) (:end 27 14)) ((:name "believe_me") (:namespace "Builtin") (:decor :postulate) (: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)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 0 10) (:end 0 12)) ((:decor :namespace)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 1 2) (:end 1 8)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 1 9) (:end 1 15)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 2) (:end 2 6)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 7) (:end 2 12)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 13) (:end 2 14)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 2 15) (:end 2 19)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 5 10) (:end 5 12)) ((:decor :namespace)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 6 2) (:end 6 8)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 6 9) (:end 6 15)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 2) (:end 7 6)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 7) (:end 7 12)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 13) (:end 7 14)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 15) (:end 7 19)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 7 20) (:end 7 25)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 4) (:end 8 7)) ((:decor :data)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 8) (:end 8 9)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 8 10) (:end 8 15)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 4) (:end 9 9)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 12) (:end 9 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 16) (:end 9 18)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 19) (:end 9 24)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 25) (:end 9 27)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 9 28) (:end 9 33)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 11 2) (:end 11 8)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 11 9) (:end 11 15)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 2) (:end 12 6)) ((:decor :function)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 7) (:end 12 8)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 9) (:end 12 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 13) (:end 12 15)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 16) (:end 12 21)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 22) (:end 12 24)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 12 25) (:end 12 30)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 2) (:end 13 6)) ((:name "::") (:namespace "Main.L1") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 7) (:end 13 8)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 13 9) (:end 13 14)) ((:name "Cons1") (:namespace "Main.L1") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 15 10) (:end 15 12)) ((:decor :namespace)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 16 2) (:end 16 8)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 16 9) (:end 16 15)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 2) (:end 17 6)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 7) (:end 17 11)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 12) (:end 17 13)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 14) (:end 17 17)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 18) (:end 17 20)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 21) (:end 17 26)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 27) (:end 17 29)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 30) (:end 17 34)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 17 35) (:end 17 40)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 19 10) (:end 19 12)) ((:decor :namespace)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 20 2) (:end 20 8)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 20 9) (:end 20 15)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 2) (:end 21 6)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 7) (:end 21 12)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 13) (:end 21 14)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 15) (:end 21 19)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 21 20) (:end 21 25)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 4) (:end 22 7)) ((:decor :data)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 8) (:end 22 9)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 22 10) (:end 22 15)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 4) (:end 23 8)) ((:decor :data)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 9) (:end 23 10)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 11) (:end 23 14)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 15) (:end 23 17)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 17) (:end 23 21)) ((:decor :type)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 22) (:end 23 24)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 23 25) (:end 23 30)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 25 0) (:end 25 1)) ((:decor :function)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 25 2) (:end 25 3)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 25 4) (:end 25 7)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 0) (:end 26 1)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 2) (:end 26 3)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 4) (:end 26 14)) ((:name "believe_me") (:namespace "Builtin") (:decor :postulate) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 26 15) (:end 26 23)) ((:decor :data)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 0) (:end 28 7)) ((:decor :function)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 8) (:end 28 9)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 10) (:end 28 13)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 14) (:end 28 16)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 28 17) (:end 28 21)) ((:decor :type)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 0) (:end 29 7)) ((:name "Rainbow") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 8) (:end 29 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 10) (:end 29 11)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 13) (:end 29 14)) ((:name "::") (:namespace "Main.L2") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 15) (:end 29 16)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 17) (:end 29 18)) ((:name "::") (:namespace "Main.L0") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 19) (:end 29 20)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 21) (:end 29 22)) ((:name "::") (:namespace "Main.L2") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 23) (:end 29 24)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Rainbow.idr") (:start 29 25) (:end 29 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
|
||||
|
@ -1,217 +1,217 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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 1 8) (:end 1 21)) ((:decor :module)))))) 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 3 8) (:end 3 18)) ((:decor :module)))))) 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)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 28) (:end 18 28)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 34) (:end 18 34)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 42) (:end 18 42)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 50) (:end 18 50)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 32) (:end 19 32)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 40) (:end 19 40)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 48) (:end 19 48)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 56) (:end 19 56)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 19 64) (:end 19 64)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 0 7) (:end 0 21)) ((:decor :module)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 2 0) (:end 2 6)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 2 7) (:end 2 18)) ((:decor :module)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 0) (:end 4 4)) ((:decor :keyword)))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 5) (:end 4 16)) ((:decor :type)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 4 17) (:end 4 18)) ((:decor :bound)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 2) (:end 5 3)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 4) (:end 5 7)) ((:decor :data)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 5 8) (:end 5 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 2) (:end 6 3)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 4) (:end 6 7)) ((:decor :data)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 8) (:end 6 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 6 10) (:end 6 11)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 2) (:end 7 3)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 4) (:end 7 9)) ((:decor :data)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 10) (:end 7 11)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 12) (:end 7 13)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 7 14) (:end 7 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 0) (:end 9 1)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 1) (:end 9 4)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 4) (:end 9 5)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 6) (:end 9 13)) ((:decor :type)))))) 1)
|
||||
0000e7(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 14) (:end 9 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e7(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 14) (:end 9 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e7(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 14) (:end 9 25)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 9 26) (:end 9 31)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 2) (:end 10 5)) ((:name "map") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 6) (:end 10 7)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 8) (:end 10 9)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 10) (:end 10 11)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 10 11) (:end 10 15)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 4) (:end 11 7)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 8) (:end 11 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 10) (:end 11 12)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 13) (:end 11 16)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 17) (:end 11 18)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 18) (:end 11 19)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 20) (:end 11 21)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 11 21) (:end 11 22)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 4) (:end 12 7)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 8) (:end 12 9)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 10) (:end 12 11)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 12) (:end 12 14)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 15) (:end 12 18)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 19) (:end 12 20)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 20) (:end 12 21)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 22) (:end 12 23)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 23) (:end 12 24)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 25) (:end 12 26)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 26) (:end 12 27)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 28) (:end 12 29)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 12 29) (:end 12 30)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 4) (:end 13 9)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 10) (:end 13 11)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 12) (:end 13 13)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 14) (:end 13 15)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 16) (:end 13 18)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 19) (:end 13 24)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 25) (:end 13 26)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 26) (: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 28) (:end 13 29)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 29) (:end 13 30)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 31) (:end 13 32)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 32) (:end 13 33)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 34) (:end 13 35)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 35) (:end 13 36)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 37) (:end 13 38)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 38) (:end 13 39)) ((:name "f") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 40) (:end 13 41)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 13 41) (:end 13 42)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 0) (:end 15 4)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 0) (:end 15 4)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 0) (:end 15 4)) ((:name "Show") (:namespace "Prelude.Show") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 5) (:end 15 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 5) (:end 15 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 5) (:end 15 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 7) (:end 15 9)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 10) (:end 15 14)) ((:decor :type)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 15) (:end 15 16)) ((:decor :keyword)))))) 1)
|
||||
0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 16) (:end 15 27)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 16) (:end 15 27)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 16) (:end 15 27)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 28) (:end 15 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 28) (:end 15 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 28) (:end 15 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 29) (:end 15 30)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 15 31) (:end 15 36)) ((:decor :keyword)))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 2) (:end 16 6)) ((:name "show") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 7) (:end 16 8)) ((:decor :keyword)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 8) (:end 16 11)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 12) (:end 16 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 13) (:end 16 14)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 15) (:end 16 16)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 17) (:end 16 23)) ((:decor :data)))))) 1)
|
||||
0000ea(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 24) (:end 16 26)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 27) (:end 16 31)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 16 32) (:end 16 33)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 2) (: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 7) (:end 17 8)) ((:decor :keyword)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 8) (:end 17 11)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 12) (:end 17 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 14) (:end 17 15)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 15) (:end 17 16)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 17) (:end 17 18)) ((:decor :keyword)))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 19) (:end 17 26)) ((:name "unwords") (:namespace "Data.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 27) (:end 17 28)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 28) (:end 17 33)) ((:decor :data)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 33) (:end 17 34)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 35) (:end 17 39)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 40) (:end 17 41)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 41) (:end 17 42)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 43) (:end 17 47)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 48) (:end 17 49)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 17 49) (:end 17 50)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 2) (: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 7) (:end 18 8)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 8) (:end 18 13)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 14) (:end 18 15)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 16) (:end 18 17)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 18) (:end 18 19)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 19) (:end 18 20)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 21) (:end 18 22)) ((:decor :keyword)))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 23) (:end 18 30)) ((:name "unwords") (:namespace "Data.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 31) (:end 18 32)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 32) (:end 18 39)) ((:decor :data)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 39) (:end 18 40)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 41) (:end 18 45)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 46) (:end 18 47)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 47) (:end 18 48)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 49) (:end 18 53)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 54) (:end 18 55)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 55) (:end 18 56)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 57) (:end 18 61)) ((:name "show") (:namespace "Prelude.Show") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 62) (:end 18 63)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 18 63) (:end 18 64)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 0) (:end 20 2)) ((:name "Eq") (:namespace "Prelude.EqOrd") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 0) (:end 20 2)) ((:name "Eq") (:namespace "Prelude.EqOrd") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 0) (:end 20 2)) ((:name "Eq") (:namespace "Prelude.EqOrd") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 3) (:end 20 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 3) (:end 20 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 3) (:end 20 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 5) (:end 20 7)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 8) (:end 20 10)) ((:decor :type)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 11) (:end 20 12)) ((:decor :keyword)))))) 1)
|
||||
0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 12) (:end 20 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 12) (:end 20 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 12) (:end 20 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 12) (:end 20 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e9(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 12) (:end 20 23)) ((:name "OneTwoThree") (:namespace "Implementation") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 24) (:end 20 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 24) (:end 20 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 24) (:end 20 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 24) (:end 20 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 24) (:end 20 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 25) (:end 20 26)) ((:decor :keyword)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 20 27) (:end 20 32)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 2) (:end 21 5)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 6) (:end 21 7)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 8) (:end 21 10)) ((:name "==") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 11) (:end 21 14)) ((:name "One") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 15) (:end 21 16)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 17) (:end 21 18)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 19) (:end 21 20)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 21) (:end 21 23)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 21 24) (:end 21 25)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 2) (:end 22 5)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 6) (:end 22 7)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 8) (:end 22 9)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 10) (:end 22 12)) ((:name "==") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 13) (:end 22 16)) ((:name "Two") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 17) (:end 22 18)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 19) (:end 22 20)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 21) (:end 22 22)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 23) (:end 22 24)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 25) (:end 22 27)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 28) (:end 22 29)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 30) (:end 22 32)) ((:name "&&") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 33) (:end 22 34)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 35) (:end 22 37)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 22 38) (:end 22 39)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 2) (:end 23 7)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 8) (:end 23 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 10) (:end 23 11)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 12) (:end 23 13)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 14) (:end 23 16)) ((:name "==") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 17) (:end 23 22)) ((:name "Three") (:namespace "Implementation") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 23) (:end 23 24)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 25) (:end 23 26)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 27) (:end 23 28)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 29) (:end 23 30)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 31) (:end 23 32)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 33) (:end 23 35)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 36) (:end 23 37)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 38) (:end 23 40)) ((:name "&&") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 41) (:end 23 42)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 43) (:end 23 45)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 46) (:end 23 47)) ((:name "y") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 48) (:end 23 50)) ((:name "&&") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 51) (:end 23 52)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 53) (:end 23 55)) ((:name "==") (:namespace "Prelude.EqOrd") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 23 56) (:end 23 57)) ((:name "z") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 2) (:end 24 3)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 4) (:end 24 6)) ((:name "==") (:namespace "Implementation") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 7) (:end 24 8)) ((:decor :keyword)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 9) (:end 24 10)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "Implementation.idr") (:start 24 11) (:end 24 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
|
||||
|
@ -1,109 +1,109 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 1 8) (:end 1 11)) ((:decor :module)))))) 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)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 10) (:end 3 13)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 20) (:end 3 23)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 3) (:end 5 4)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 9) (:end 5 10)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 6) (:end 6 7)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 6 18) (:end 6 19)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 11) (:end 8 14)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 21) (:end 8 24)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 3) (:end 10 4)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 9) (:end 10 10)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 6) (:end 11 7)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 30) (:end 11 30)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 11 32) (:end 11 32)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 37) (:end 13 40)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 47) (:end 13 50)) ((:name "List") (:namespace "Prelude.Basics") (: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 16 8) (:end 16 17)) ((:decor :comment)))))) 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)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 18 21) (:end 18 22)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 19 8) (:end 19 16)) ((:decor :comment)))))) 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)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 21 21) (:end 21 22)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 0 7) (:end 0 11)) ((:decor :module)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 0) (:end 2 6)) ((:decor :function)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 7) (:end 2 8)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 9) (:end 2 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 14) (:end 2 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 16) (:end 2 18)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 19) (:end 2 23)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 2 24) (:end 2 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 0) (:end 3 6)) ((:name "listId") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 7) (:end 3 9)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 10) (:end 3 11)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 12) (:end 3 16)) ((:decor :keyword)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 17) (:end 3 19)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 3 20) (:end 3 22)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 2) (:end 4 4)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 5) (:end 4 7)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 4 8) (:end 4 10)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 2) (:end 5 3)) ((:decor :keyword)))))) 1)
|
||||
0000c4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 3) (:end 5 4)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 5) (:end 5 7)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 8) (:end 5 10)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 10) (:end 5 11)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 12) (:end 5 14)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 15) (:end 5 16)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 17) (:end 5 19)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 20) (:end 5 26)) ((:name "listId") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 5 27) (:end 5 29)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 0) (:end 7 7)) ((:decor :function)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 8) (:end 7 9)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 10) (:end 7 14)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 15) (:end 7 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 17) (:end 7 19)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 20) (:end 7 24)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 7 25) (:end 7 26)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 0) (:end 8 7)) ((:name "listRev") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 8) (:end 8 9)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 10) (:end 8 11)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 8 11) (:end 8 15)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 9 2) (:end 9 4)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 9 5) (:end 9 7)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 9 8) (:end 9 10)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 2) (:end 10 3)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 3) (:end 10 4)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 5) (:end 10 7)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 8) (:end 10 10)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 10) (:end 10 11)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 12) (:end 10 14)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 15) (:end 10 22)) ((:name "listRev") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 23) (:end 10 25)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 26) (:end 10 28)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 29) (:end 10 30)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 30) (:end 10 31)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 10 31) (:end 10 32)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 0) (:end 12 11)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 12) (:end 12 13)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 14) (:end 12 15)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 15) (:end 12 16)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 16) (:end 12 17)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 18) (:end 12 19)) ((:name "q") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 20) (:end 12 21)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 22) (:end 12 23)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 22) (:end 12 23)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 24) (:end 12 26)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 27) (:end 12 31)) ((:name "Bool") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 27) (:end 12 31)) ((:name "Bool") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 31) (:end 12 32)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 33) (:end 12 35)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 36) (:end 12 40)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 41) (:end 12 42)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 43) (:end 12 45)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 46) (:end 12 50)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 12 51) (:end 12 52)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 0) (:end 13 11)) ((:name "listFilter2") (:namespace "Case") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 12) (:end 13 13)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 14) (:end 13 15)) ((:name "q") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 13 16) (:end 13 18)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 14 2) (:end 14 3)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 14 7) (:end 14 8)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 14 9) (:end 14 11)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 14 12) (:end 14 14)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 15 7) (:end 15 17)) ((:decor :comment)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 16 7) (:end 16 10)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 16 11) (:end 16 15)) ((:name "True") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 16 16) (:end 16 17)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 16 18) (:end 16 19)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 16 20) (:end 16 21)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 17 9) (:end 17 10)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 17 11) (:end 17 16)) ((:name "False") (: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 19)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 17 20) (:end 17 22)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 18 7) (:end 18 16)) ((:decor :comment)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 19 7) (:end 19 11)) ((:name "True") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 19 12) (:end 19 14)) ((:decor :keyword)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 19 15) (:end 19 19)) ((:name "pure") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 19 20) (:end 19 21)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 19 21) (:end 19 22)) ((:name "q") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 19 23) (:end 19 24)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 19 24) (:end 19 25)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 9) (:end 20 10)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 11) (:end 20 16)) ((:name "False") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 17) (:end 20 19)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 20 20) (:end 20 22)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 21 7) (:end 21 11)) ((:name "pure") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Case.idr") (:start 21 12) (:end 21 13)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,81 +1,81 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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 1 8) (:end 1 19)) ((:decor :module)))))) 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 11)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 13) (:end 8 16)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 18) (:end 8 19)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 21) (:end 8 21)) ((:decor :keyword)))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 22) (:end 8 23)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 25) (:end 8 30)) ((:decor :data)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 31) (:end 8 31)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 11) (:end 9 11)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 13) (:end 9 16)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 18) (:end 9 19)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 21) (:end 9 22)) ((:decor :data)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 10 11) (:end 10 11)) ((: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)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 0 7) (:end 0 19)) ((:decor :module)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 2 0) (:end 2 6)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 2 7) (:end 2 17)) ((:decor :type)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 2 18) (:end 2 23)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 2) (:end 3 6)) ((:decor :function)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 7) (:end 3 8)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 3 9) (:end 3 15)) ((:decor :type)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 2) (:end 4 6)) ((:decor :function)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 9) (:end 4 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 9) (:end 4 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 4 9) (:end 4 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 0) (:end 6 7)) ((:decor :function)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 8) (:end 6 9)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 10) (:end 6 20)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 21) (:end 6 23)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 6 24) (:end 6 34)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 0) (:end 7 7)) ((:name "bigMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 8) (:end 7 9)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 10) (:end 7 11)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 12) (:end 7 16)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 17) (:end 7 19)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 20) (:end 7 21)) ((:decor :keyword)))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 21) (:end 7 23)) ((:name "++") (:namespace "Prelude.Types.String") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 24) (:end 7 30)) ((:decor :data)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 7 30) (:end 7 31)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 10) (:end 8 11)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 12) (:end 8 16)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 17) (:end 8 19)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 8 20) (:end 8 22)) ((:decor :data)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 0) (:end 11 9)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 10) (:end 11 11)) ((:decor :keyword)))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 12) (:end 11 22)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 23) (:end 11 25)) ((:decor :keyword)))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 11 26) (:end 11 36)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 0) (:end 12 9)) ((:name "smallMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 10) (:end 12 11)) ((:decor :keyword)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 12) (:end 12 13)) ((:decor :keyword)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 14) (:end 12 18)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 19) (:end 12 21)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 22) (:end 12 23)) ((:decor :data)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 24) (:end 12 25)) ((:decor :keyword)))))) 1)
|
||||
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)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 15 1) (:end 15 23)) ((:decor :comment)))))) 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)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 26) (:end 12 27)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 12 28) (:end 12 35)) ((:name "bigMono") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 14 0) (:end 14 23)) ((:decor :comment)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 15 0) (:end 15 6)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 15 7) (:end 15 11)) ((:decor :type)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 15 12) (:end 15 17)) ((:decor :keyword)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 2) (:end 16 12)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 13) (:end 16 14)) ((:decor :keyword)))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 15) (:end 16 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 15) (:end 16 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 16 15) (:end 16 25)) ((:name "Attributes") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 2) (:end 17 9)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 10) (:end 17 11)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 17 12) (:end 17 18)) ((:decor :type)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 0) (:end 19 10)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 11) (:end 19 12)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 13) (:end 19 17)) ((:name "Text") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 18) (:end 19 20)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 19 21) (:end 19 25)) ((:name "Text") (:namespace "RecordUpdate") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e7(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 0) (:end 20 10)) ((:name "setArial10") (:namespace "RecordUpdate") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 11) (:end 20 12)) ((:decor :keyword)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 13) (:end 20 14)) ((:decor :keyword)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 15) (:end 20 25)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 25) (:end 20 27)) ((:decor :keyword)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 27) (:end 20 31)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 32) (:end 20 34)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 20 35) (:end 20 42)) ((:decor :data)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 13) (:end 21 14)) ((:decor :keyword)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 15) (:end 21 25)) ((:decor :function)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 25) (:end 21 30)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 31) (:end 21 33)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 21 34) (:end 21 36)) ((:decor :data)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordUpdate.idr") (:start 22 13) (:end 22 14)) ((:decor :keyword)))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,328 +1,328 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 11 8) (:end 11 9)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 1 8) (:end 1 11)) ((:decor :module)))))) 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)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 2) (:end 4 3)) ((:name "f") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 4) (:end 4 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 2) (: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 4) (:end 5 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 2) (:end 10 3)) ((:name "g") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 5) (:end 10 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 7) (:end 10 9)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 10) (:end 10 12)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 2) (:end 14 8)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 9) (:end 14 10)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 4) (:end 15 10)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 11) (:end 15 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 11) (:end 15 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 15) (:end 15 16)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 4) (: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 11) (: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 11) (: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 15) (: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 6) (:end 17 12)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 13) (:end 17 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 13) (:end 17 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 13) (:end 17 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 17) (:end 17 18)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 21) (:end 17 22)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 23) (:end 17 24)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 6) (: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 13) (: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 13) (: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 13) (: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 17) (: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 21) (: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 23) (: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 8) (:end 19 14)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 15) (:end 19 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 15) (:end 19 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 15) (:end 19 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 15) (:end 19 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 19) (:end 19 20)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 23) (:end 19 24)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 25) (:end 19 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 25) (:end 19 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 29) (:end 19 30)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 31) (:end 19 32)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 8) (: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 15) (: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 15) (: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 15) (: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 15) (: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 19) (: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 23) (: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 25) (: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 25) (: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 29) (: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 31) (:end 20 32)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 2) (:end 21 8)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 9) (:end 21 10)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 2) (:end 28 10)) ((:name "someNats") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 11) (:end 28 12)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 4) (:end 29 12)) ((:name "someNats") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 13) (:end 29 14)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 13) (:end 29 14)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 20) (:end 29 26)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 27) (:end 29 28)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 0 7) (:end 0 11)) ((:decor :module)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 0) (:end 2 1)) ((:decor :function)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 2) (:end 2 3)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 4) (:end 2 5)) ((:decor :keyword)))))) 1)
|
||||
0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 5) (:end 2 6)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 7) (:end 2 8)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 9) (:end 2 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 12) (:end 2 13)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 14) (:end 2 16)) ((:decor :keyword)))))) 1)
|
||||
00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 17) (:end 2 18)) ((:decor :type)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 18) (:end 2 19)) ((:decor :bound)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 20) (:end 2 21)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 22) (:end 2 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 22) (:end 2 25)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 26) (:end 2 28)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 29) (:end 2 30)) ((:decor :bound)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 31) (:end 2 32)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 33) (:end 2 36)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 33) (:end 2 36)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 37) (:end 2 39)) ((:name "DPair") (:namespace "Builtin.DPair") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 40) (:end 2 41)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 42) (:end 2 43)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 44) (:end 2 45)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 46) (:end 2 47)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 48) (:end 2 49)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 2 49) (:end 2 50)) ((:decor :type)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 0) (:end 3 1)) ((:name "f") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 2) (:end 3 3)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 4) (:end 3 8)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 9) (:end 3 10)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 9) (:end 3 10)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 10) (:end 3 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 10) (:end 3 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 10) (:end 3 11)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 12) (:end 3 13)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 12) (:end 3 13)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 12) (:end 3 13)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 14) (:end 3 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 14) (:end 3 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 14) (:end 3 15)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 15) (:end 3 16)) ((:decor :keyword)))))) 1)
|
||||
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)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 5) (:end 8 8)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 3) (:end 9 4)) ((:name "Nil") (:namespace "Prelude.Basics") (: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)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 6) (:end 10 7)) ((:name "::") (:namespace "Prelude.Basics") (: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)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 17) (:end 3 22)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 3 23) (:end 3 25)) ((:decor :bound)))))) 1)
|
||||
0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 4) (:end 4 5)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 6) (:end 4 7)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 8) (:end 4 9)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 14) (:end 4 15)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 17) (:end 4 18)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 19) (:end 4 21)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 22) (:end 4 23)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 24) (:end 4 26)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 27) (:end 4 30)) ((:name "sym") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 4 31) (:end 4 33)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 4) (: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 6) (:end 5 7)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 8) (:end 5 9)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 9) (:end 5 10)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 11) (:end 5 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 12) (:end 5 13)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 14) (:end 5 15)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 17) (:end 5 18)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 19) (:end 5 20)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 21) (:end 5 23)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 24) (:end 5 25)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 26) (:end 5 28)) ((:name "MkDPair") (:namespace "Builtin.DPair") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 29) (:end 5 32)) ((:name "sym") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 5 33) (:end 5 35)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 0) (:end 7 1)) ((:decor :function)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 2) (:end 7 3)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 4) (:end 7 8)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c5(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 9) (:end 7 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 11) (:end 7 13)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 7 14) (:end 7 17)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 0) (:end 8 1)) ((:name "g") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 2) (:end 8 4)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 5) (:end 8 6)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 8 7) (:end 8 8)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 0) (:end 9 1)) ((:name "g") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 2) (:end 9 3)) ((:decor :keyword)))))) 1)
|
||||
0000c4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 3) (:end 9 4)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 5) (:end 9 7)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 8) (:end 9 10)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 12) (:end 9 16)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 17) (:end 9 18)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 17) (:end 9 18)) ((:decor :keyword)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 18) (:end 9 20)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 18) (:end 9 20)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 21) (:end 9 23)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 21) (:end 9 23)) ((:name "++") (:namespace "Prelude.Types.List") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 24) (:end 9 26)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 24) (:end 9 26)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 9 26) (:end 9 27)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 4) (:end 10 5)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 5) (:end 10 6)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 10) (:end 10 12)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 12) (:end 10 13)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 14) (:end 10 15)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 16) (:end 10 20)) ((:name "asas") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 21) (:end 10 22)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 10 23) (:end 10 24)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 0) (:end 12 6)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 7) (:end 12 8)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 9) (:end 12 12)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 13) (:end 12 15)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 12 16) (:end 12 19)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 0) (:end 13 6)) ((:name "nested") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 7) (:end 13 8)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 9) (:end 13 13)) ((:decor :keyword)))))) 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 13 14) (:end 13 15)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 15) (:end 13 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 15) (:end 13 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 13 16) (:end 13 17)) ((:decor :keyword)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 9) (:end 14 10)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 11) (:end 14 12)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 13) (:end 14 14)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 15) (:end 14 19)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 20) (:end 14 21)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 20) (:end 14 21)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 21) (:end 14 22)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 21) (:end 14 22)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 23) (:end 14 24)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 23) (:end 14 24)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 25) (:end 14 26)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 25) (:end 14 26)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 14 26) (:end 14 27)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 11) (:end 15 12)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 13) (:end 15 14)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 17) (:end 15 18)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 19) (:end 15 20)) ((:decor :data)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 21) (:end 15 22)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 15 23) (:end 15 24)) ((:decor :data)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 11) (: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 13) (:end 16 14)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 17) (:end 16 18)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 19) (:end 16 20)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 21) (:end 16 22)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 23) (:end 16 27)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 28) (:end 16 29)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 28) (:end 16 29)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 29) (:end 16 30)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 29) (:end 16 30)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 31) (:end 16 32)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 31) (:end 16 32)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 33) (:end 16 34)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 33) (:end 16 34)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 16 34) (:end 16 35)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 13) (:end 17 14)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 15) (:end 17 16)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 19) (:end 17 20)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 23) (:end 17 24)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 25) (:end 17 26)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 27) (:end 17 28)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 29) (:end 17 30)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 17 31) (:end 17 32)) ((:decor :data)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 13) (: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 15) (:end 18 16)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 19) (:end 18 20)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 23) (: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 25) (:end 18 26)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 27) (:end 18 28)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 29) (:end 18 30)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 31) (:end 18 35)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 36) (:end 18 37)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 36) (:end 18 37)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 37) (:end 18 38)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 37) (:end 18 38)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 39) (:end 18 40)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 39) (:end 18 40)) ((:name "+") (:namespace "Prelude.Num") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 41) (:end 18 42)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 41) (:end 18 42)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 18 42) (:end 18 43)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 15) (:end 19 16)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 17) (:end 19 18)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 21) (:end 19 22)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 25) (:end 19 26)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 27) (:end 19 28)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 31) (:end 19 32)) ((:name "l") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 33) (:end 19 34)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 35) (:end 19 36)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 37) (:end 19 38)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 19 39) (:end 19 40)) ((:decor :data)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 15) (: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 17) (:end 20 18)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 21) (:end 20 22)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 25) (: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 27) (:end 20 28)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 31) (: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 33) (:end 20 34)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 35) (:end 20 36)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 37) (:end 20 38)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 39) (:end 20 40)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 20 41) (:end 20 42)) ((:decor :data)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 9) (:end 21 10)) ((:name "m") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 11) (:end 21 12)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 13) (:end 21 14)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 15) (:end 21 16)) ((:name "k") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 17) (:end 21 18)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 21 19) (:end 21 20)) ((:decor :data)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 0) (:end 23 4)) ((:decor :keyword)))))) 1)
|
||||
00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 5) (:end 23 9)) ((:decor :type)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 10) (:end 23 11)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 12) (:end 23 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 16) (:end 23 18)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 19) (:end 23 23)) ((:decor :type)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 23 24) (:end 23 29)) ((:decor :keyword)))))) 1)
|
||||
00006e(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 2) (:end 24 8)) ((:decor :data)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 9) (:end 24 10)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 11) (:end 24 12)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 12) (:end 24 13)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 14) (:end 24 15)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 16) (:end 24 19)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 19) (:end 24 20)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 21) (:end 24 23)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 24) (:end 24 28)) ((:name "ANat") (:namespace "With") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 24 29) (:end 24 30)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 0) (:end 26 8)) ((:decor :function)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 9) (:end 26 10)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 11) (:end 26 14)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 15) (:end 26 17)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 26 18) (:end 26 21)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 0) (:end 27 8)) ((:name "someNats") (:namespace "With") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 9) (:end 27 10)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 11) (:end 27 15)) ((:decor :keyword)))))) 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)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 16) (:end 27 17)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 17) (:end 27 23)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 17) (:end 27 23)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 24) (:end 27 25)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 24) (:end 27 25)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 27 25) (:end 27 26)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 11) (:end 28 12)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 13) (:end 28 14)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 15) (:end 28 16)) ((:decor :bound)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 16) (:end 28 17)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 17) (:end 28 18)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 18) (:end 28 24)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 25) (:end 28 26)) ((: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 27)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 28) (:end 28 32)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 33) (:end 28 34)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 33) (:end 28 34)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 34) (:end 28 40)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 34) (:end 28 40)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 41) (:end 28 42)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 41) (:end 28 42)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 28 42) (:end 28 43)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 15) (:end 29 16)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 17) (:end 29 18)) ((:name "p") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 17) (:end 29 18)) ((:decor :bound)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 18) (:end 29 19)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 19) (:end 29 20)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 27) (:end 29 28)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 28) (:end 29 29)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 30) (:end 29 31)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 32) (:end 29 38)) ((:name "MkANat") (:namespace "With") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 39) (:end 29 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 41) (:end 29 42)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "With.idr") (:start 29 43) (:end 29 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
|
||||
|
@ -1,27 +1,27 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 9) (:end 1 12)) ((:name "List") (:namespace "Prelude.Basics") (: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)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 0) (:end 0 5)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 6) (:end 0 7)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 8) (:end 0 12)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 0 13) (:end 0 16)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 0) (:end 1 5)) ((:name "hours") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 6) (:end 1 7)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 8) (:end 1 9)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 9) (:end 1 10)) ((:decor :data)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 10) (:end 1 12)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 12) (:end 1 14)) ((:decor :data)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 1 14) (:end 1 15)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 0) (:end 3 4)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 5) (:end 3 6)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 7) (:end 3 13)) ((:name "Stream") (:namespace "Prelude.Types.Stream") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 3 14) (:end 3 17)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 0) (:end 4 4)) ((:name "nats") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 5) (:end 4 6)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1)
|
||||
00006e(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 8) (:end 4 9)) ((:decor :data)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 9) (:end 4 10)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 10) (:end 4 11)) ((:decor :data)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 11) (:end 4 13)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Ranges.idr") (:start 4 13) (:end 4 14)) ((:decor :keyword)))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,56 +1,56 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 1 8) (:end 1 21)) ((:decor :module)))))) 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)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 0 7) (:end 0 21)) ((:decor :module)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 0) (:end 2 5)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 6) (:end 2 7)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 8) (:end 2 18)) ((:name "FromString") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 19) (:end 2 20)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 21) (:end 2 23)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 2 24) (:end 2 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 0) (:end 3 5)) ((:name "hello") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 6) (:end 3 7)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 3 8) (:end 3 15)) ((:decor :data)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 0) (:end 5 9)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 10) (:end 5 11)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 12) (:end 5 18)) ((:decor :type)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 19) (:end 5 21)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 5 22) (:end 5 28)) ((:decor :type)))))) 1)
|
||||
0000e7(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 0) (:end 6 9)) ((:name "helloName") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 10) (:end 6 14)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 15) (:end 6 16)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 17) (:end 6 48)) ((:decor :data)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 20) (:end 6 25)) ((:name "hello") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 26) (:end 6 27)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 27) (:end 6 28)) ((:decor :bound)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 29) (:end 6 30)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 31) (:end 6 37)) ((:decor :type)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 37) (:end 6 38)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 6 42) (:end 6 46)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 0) (:end 8 11)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 12) (:end 8 13)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 14) (:end 8 20)) ((:decor :type)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 21) (:end 8 23)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 8 24) (:end 8 30)) ((:decor :type)))))) 1)
|
||||
0000ea(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 0) (:end 9 11)) ((:name "welcomeName") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 12) (:end 9 16)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 17) (:end 9 18)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 9 19) (:end 12 5)) ((:decor :data)))))) 1)
|
||||
0000ea(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 10 4) (:end 10 13)) ((:name "helloName") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 10 14) (:end 10 18)) ((:name "name") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 14 0) (:end 14 11)) ((:decor :function)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 14 12) (:end 14 13)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 14 14) (:end 14 20)) ((:decor :type)))))) 1)
|
||||
0000ec(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 0) (:end 15 11)) ((:name "scareQuotes") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 12) (:end 15 13)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 15 14) (:end 15 25)) ((:decor :data)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 0) (:end 17 4)) ((:decor :function)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 5) (:end 17 6)) ((:decor :keyword)))))) 1)
|
||||
0000ec(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 7) (:end 17 33)) ((:name "scareQuotes") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 34) (:end 17 35)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 17 36) (:end 17 47)) ((:decor :data)))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 18 0) (:end 18 4)) ((:name "test") (:namespace "StringLiterals") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 18 5) (:end 18 6)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "StringLiterals.idr") (:start 18 7) (:end 18 11)) ((:name "Refl") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,81 +1,81 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 2) (:end 10 4)) ((:name "id") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 2) (:end 11 4)) ((:name "id") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 15) (:end 11 17)) ((:name "id") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 0) (:end 0 4)) ((:decor :keyword)))))) 1)
|
||||
000070(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 5) (:end 0 10)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 11) (:end 0 12)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 13) (:end 0 16)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 17) (:end 0 19)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 20) (:end 0 24)) ((:decor :type)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 0 25) (:end 0 30)) ((:decor :keyword)))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 2) (:end 1 3)) ((:decor :data)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 4) (:end 1 5)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 6) (:end 1 11)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 1 12) (:end 1 13)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00006f(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 2) (:end 2 3)) ((:decor :data)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 4) (:end 2 5)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 6) (:end 2 11)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 12) (:end 2 13)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 14) (:end 2 16)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 17) (:end 2 22)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 23) (:end 2 24)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 24) (:end 2 25)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 26) (:end 2 27)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 2 27) (:end 2 28)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 0) (:end 4 4)) ((:decor :function)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 5) (:end 4 6)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 7) (:end 4 8)) ((:decor :keyword)))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 8) (:end 4 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 9) (:end 4 10)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 11) (:end 4 14)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 14) (:end 4 15)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 16) (:end 4 18)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 19) (:end 4 24)) ((:name "Natty") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 4 25) (:end 4 26)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 0) (:end 5 4)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 5) (:end 5 6)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 7) (:end 5 8)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 5 9) (:end 5 10)) ((:name "Z") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 0) (:end 6 4)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 5) (:end 6 6)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 6) (:end 6 7)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 8) (:end 6 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 9) (:end 6 10)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 11) (:end 6 12)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 13) (:end 6 14)) ((:name "S") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 15) (:end 6 16)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 16) (:end 6 20)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 21) (:end 6 22)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 6 22) (:end 6 23)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 0) (:end 8 2)) ((:decor :function)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 3) (:end 8 4)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 5) (:end 8 8)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 9) (:end 8 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 8 12) (:end 8 15)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 0) (:end 9 2)) ((:name "id") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 3) (:end 9 4)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 5) (:end 9 9)) ((:decor :keyword)))))) 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)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 11) (:end 9 15)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 11) (:end 9 15)) ((:name "view") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 16) (:end 9 17)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 16) (:end 9 17)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 9 17) (:end 9 18)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 5) (:end 10 6)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 7) (:end 10 8)) ((:decor :keyword)))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 9) (:end 10 10)) ((:name "Z") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 11) (:end 10 12)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 10 13) (:end 10 14)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 5) (:end 11 6)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 7) (:end 11 8)) ((:decor :keyword)))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 9) (:end 11 10)) ((:name "S") (:namespace "Main") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 11) (:end 11 12)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 13) (:end 11 14)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 18) (:end 11 19)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 20) (:end 11 21)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "WithApp.idr") (:start 11 22) (:end 11 23)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,70 +1,70 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
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)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 0) (:end 0 9)) ((:decor :function)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 10) (:end 0 11)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 12) (:end 0 13)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 13) (:end 0 14)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 15) (:end 0 17)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 18) (:end 0 19)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 20) (:end 0 21)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 22) (:end 0 25)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 26) (:end 0 27)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 27) (:end 0 28)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 29) (:end 0 31)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 32) (:end 0 33)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 34) (:end 0 36)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 0 37) (:end 0 38)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 0) (:end 1 9)) ((:name "transport") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 10) (:end 1 12)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 13) (:end 1 14)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 15) (:end 1 16)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 17) (:end 1 24)) ((:decor :keyword)))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 25) (:end 1 28)) ((:name "sym") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 29) (:end 1 31)) ((:name "eq") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 32) (:end 1 34)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 1 35) (:end 1 36)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 0) (:end 3 6)) ((:decor :function)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 7) (:end 3 8)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 9) (:end 3 10)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 10) (:end 3 11)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 12) (:end 3 13)) ((:name "_") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 14) (:end 3 15)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 16) (:end 3 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 18) (:end 3 21)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 22) (:end 3 23)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 23) (:end 3 24)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 25) (:end 3 27)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 28) (:end 3 29)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 29) (:end 3 30)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 31) (:end 3 32)) ((:name "_") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 33) (:end 3 34)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 35) (:end 3 36)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 37) (:end 3 40)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 41) (:end 3 42)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 42) (:end 3 43)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 44) (:end 3 46)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 47) (:end 3 48)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 49) (:end 3 52)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 3 53) (:end 3 54)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 0) (:end 4 6)) ((:name "nested") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 7) (:end 4 10)) ((:name "eq1") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 11) (:end 4 14)) ((:name "eq2") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 4 15) (:end 4 16)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 5 2) (:end 5 9)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 5 10) (:end 5 13)) ((:name "eq1") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 5 14) (:end 5 16)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 6 4) (:end 6 11)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 6 12) (:end 6 15)) ((:name "eq2") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 6 16) (:end 6 18)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 6) (:end 7 9)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 10) (:end 7 13)) ((:name "prf") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 14) (:end 7 15)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 16) (:end 7 17)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 18) (:end 7 21)) ((:name "===") (:namespace "Builtin") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 22) (:end 7 23)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 24) (:end 7 26)) ((:decor :keyword)))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 27) (:end 7 31)) ((:name "Refl") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 7 32) (:end 7 34)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Rewrite.idr") (:start 8 6) (:end 8 9)) ((:name "prf") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,97 +1,97 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
00004c(:write-string "1/1: Building RecordProjections (RecordProjections.idr)" 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 8) (:end 1 11)) ((:decor :type)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 13) (:end 1 13)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 14) (:end 1 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 16) (:end 1 16)) ((:decor :keyword)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 18) (:end 1 21)) ((:decor :type)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 22) (:end 1 22)) ((:decor :keyword)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 24) (:end 1 28)) ((:decor :keyword)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 2 15) (:end 2 20)) ((:decor :data)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 3) (:end 3 8)) ((:decor :function)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 10) (:end 3 10)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 12) (:end 3 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 4 3) (:end 4 5)) ((:decor :function)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 4 7) (:end 4 7)) ((:decor :keyword)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 4 9) (:end 4 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 4 9) (:end 4 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 4 9) (:end 4 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 1) (:end 6 4)) ((:decor :function)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 6) (:end 6 6)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 8) (:end 6 11)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 13) (:end 6 13)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 14) (:end 6 17)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 19) (:end 6 22)) ((:name "Bool") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 23) (:end 6 23)) ((:decor :keyword)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 25) (:end 6 26)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 28) (:end 6 31)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 33) (:end 6 33)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 34) (:end 6 37)) ((:name "Bool") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 38) (:end 6 38)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 40) (:end 6 42)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 43) (:end 6 43)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 1) (:end 7 4)) ((:name "proj") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 6) (:end 7 7)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 16) (:end 7 16)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 18) (:end 7 19)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 1) (:end 8 4)) ((:name "proj") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 6) (:end 8 6)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 7) (:end 8 7)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 9) (:end 8 10)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 12) (:end 8 13)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 14) (:end 8 14)) ((:decor :keyword)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 16) (:end 8 16)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 18) (:end 8 23)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 25) (:end 8 25)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 26) (:end 8 32)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 34) (:end 8 34)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 35) (:end 8 37)) ((:name "nat") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 39) (:end 8 39)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 40) (:end 8 40)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 42) (:end 8 43)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 45) (:end 8 48)) ((:name "proj") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 8 50) (:end 8 51)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007f(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 1) (:end 10 8)) ((:decor :function)))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 10) (:end 10 10)) ((:decor :keyword)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 12) (:end 10 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 17) (:end 10 17)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 18) (:end 10 21)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 23) (:end 10 23)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 24) (:end 10 27)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 29) (:end 10 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 30) (:end 10 30)) ((:decor :keyword)))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 31) (:end 10 31)) ((:decor :keyword)))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 33) (:end 10 34)) ((:decor :keyword)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 36) (:end 10 39)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 41) (:end 10 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 11 1) (:end 11 8)) ((:name "ununpack") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 11 10) (:end 11 10)) ((:decor :keyword)))))) 1)
|
||||
0000ec(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 11 12) (:end 11 14)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 11 16) (:end 11 16)) ((:decor :keyword)))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 11 17) (:end 11 23)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 11 24) (:end 11 30)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 11 31) (:end 11 31)) ((:decor :keyword)))))) 1)
|
||||
00007f(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 1) (:end 13 8)) ((:decor :function)))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 10) (:end 13 10)) ((:decor :keyword)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 12) (:end 13 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 17) (:end 13 17)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 18) (:end 13 21)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 23) (:end 13 23)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 24) (:end 13 27)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 29) (:end 13 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 30) (:end 13 30)) ((:decor :keyword)))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 31) (:end 13 31)) ((:decor :keyword)))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 33) (:end 13 34)) ((:decor :keyword)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 36) (:end 13 39)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 41) (:end 13 43)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 14 1) (:end 14 8)) ((:name "deepNats") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 14 10) (:end 14 10)) ((:decor :keyword)))))) 1)
|
||||
0000ec(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 14 12) (:end 14 14)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 14 16) (:end 14 16)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 14 17) (:end 14 22)) ((:name "nat") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 14 24) (:end 14 24)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 14 26) (:end 14 34)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 14 35) (:end 14 35)) ((:decor :keyword)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 7) (:end 0 11)) ((:decor :type)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 12) (:end 0 13)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 13) (:end 0 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 15) (:end 0 16)) ((:decor :keyword)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 17) (:end 0 21)) ((:decor :type)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 21) (:end 0 22)) ((:decor :keyword)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 0 23) (:end 0 28)) ((:decor :keyword)))))) 1)
|
||||
00007b(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 1 14) (:end 1 20)) ((:decor :data)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 2 2) (:end 2 8)) ((:decor :function)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 2 9) (:end 2 10)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 2 11) (:end 2 12)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 2) (:end 3 5)) ((:decor :function)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 6) (:end 3 7)) ((:decor :keyword)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 8) (:end 3 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 8) (:end 3 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 3 8) (:end 3 11)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 0) (:end 5 4)) ((:decor :function)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 5) (:end 5 6)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 7) (:end 5 11)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 12) (:end 5 13)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 13) (:end 5 17)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 18) (:end 5 22)) ((:name "Bool") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 22) (:end 5 23)) ((:decor :keyword)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 24) (:end 5 26)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 27) (:end 5 31)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 32) (:end 5 33)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 33) (:end 5 37)) ((:name "Bool") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 37) (:end 5 38)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 39) (:end 5 42)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 5 42) (:end 5 43)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 0) (:end 6 4)) ((:name "proj") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 5) (:end 6 7)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 15) (:end 6 16)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 6 17) (:end 6 19)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 0) (:end 7 4)) ((:name "proj") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 5) (:end 7 6)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 6) (:end 7 7)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 8) (:end 7 10)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 11) (:end 7 13)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 13) (:end 7 14)) ((:decor :keyword)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 15) (:end 7 16)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 17) (:end 7 23)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 24) (:end 7 25)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e4(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 25) (:end 7 32)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 33) (:end 7 34)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 34) (:end 7 37)) ((:name "nat") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 38) (:end 7 39)) ((:name "x") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 39) (:end 7 40)) ((:decor :keyword)))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 41) (:end 7 43)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 44) (:end 7 48)) ((:name "proj") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 7 49) (:end 7 51)) ((:name "xs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 0) (:end 9 8)) ((:decor :function)))))) 1)
|
||||
00007d(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 9) (:end 9 10)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 11) (:end 9 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 16) (:end 9 17)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 17) (:end 9 21)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 22) (:end 9 23)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 23) (:end 9 27)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 28) (:end 9 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 29) (:end 9 30)) ((:decor :keyword)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 30) (:end 9 31)) ((:decor :keyword)))))) 1)
|
||||
00007e(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 32) (:end 9 34)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 35) (:end 9 39)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 9 40) (:end 9 41)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 0) (:end 10 8)) ((:name "ununpack") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007f(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 9) (:end 10 10)) ((:decor :keyword)))))) 1)
|
||||
0000ec(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 11) (:end 10 14)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 15) (:end 10 16)) ((:decor :keyword)))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 16) (:end 10 23)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 23) (:end 10 30)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 10 30) (:end 10 31)) ((:decor :keyword)))))) 1)
|
||||
00007f(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 0) (:end 12 8)) ((:decor :function)))))) 1)
|
||||
00007f(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 9) (:end 12 10)) ((:decor :keyword)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 11) (:end 12 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 16) (:end 12 17)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 17) (:end 12 21)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 22) (:end 12 23)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 23) (:end 12 27)) ((:name "Pack") (:namespace "Main") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 28) (:end 12 29)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 29) (:end 12 30)) ((:decor :keyword)))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 30) (:end 12 31)) ((:decor :keyword)))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 32) (:end 12 34)) ((:decor :keyword)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 35) (:end 12 39)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 12 40) (:end 12 43)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e1(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 0) (:end 13 8)) ((:name "deepNats") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007f(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 9) (:end 13 10)) ((:decor :keyword)))))) 1)
|
||||
0000ec(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 11) (:end 13 14)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 15) (:end 13 16)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 16) (:end 13 22)) ((:name "nat") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 23) (:end 13 24)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e6(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 25) (:end 13 34)) ((:name "nested") (:namespace "Main.Pack") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000080(:output (:ok (:highlight-source ((((:filename "RecordProjections.idr") (:start 13 34) (:end 13 35)) ((:decor :keyword)))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,185 +1,185 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
000036(:write-string "1/1: Building Tuples (Tuples.idr)" 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 1) (:end 1 4)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 6) (:end 1 6)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 8) (:end 1 8)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 10) (:end 1 11)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 13) (:end 1 13)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 0) (:end 0 4)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 5) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 7) (:end 0 8)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 9) (:end 0 11)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 12) (:end 0 13)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 14) (:end 0 16)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 17) (:end 0 21)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 22) (:end 0 23)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 24) (:end 0 26)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 27) (:end 0 31)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 32) (:end 0 33)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 33) (:end 0 34)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 34) (:end 0 35)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 36) (:end 0 37)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 37) (:end 0 38)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 39) (:end 0 40)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 0 40) (:end 0 41)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 0) (:end 1 4)) ((:name "init") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 5) (:end 1 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 7) (:end 1 8)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 9) (:end 1 10)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 11) (:end 1 14)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 15) (:end 1 16)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 18) (:end 1 21)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 23) (:end 1 23)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 25) (:end 1 26)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 28) (:end 1 31)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 33) (:end 1 33)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 34) (:end 1 34)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 35) (:end 1 35)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 37) (:end 1 37)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 38) (:end 1 38)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 40) (:end 1 40)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 41) (:end 1 41)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 1) (:end 2 4)) ((:name "init") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 6) (:end 2 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 8) (:end 2 8)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 10) (:end 2 10)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 12) (:end 2 14)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 16) (:end 2 16)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 17) (:end 2 23)) ((:name "uncurry") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 25) (:end 2 25)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 26) (:end 2 26)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 27) (:end 2 27)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 28) (:end 2 28)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 30) (:end 2 30)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 32) (:end 2 34)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 36) (:end 2 36)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 37) (:end 2 37)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 38) (:end 2 38)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 39) (:end 2 39)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 41) (:end 2 41)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 43) (:end 2 45)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 47) (:end 2 47)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 48) (:end 2 48)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 49) (:end 2 49)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 2 50) (:end 2 50)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 1) (:end 4 5)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 7) (:end 4 7)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 9) (:end 4 12)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 14) (:end 4 14)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 15) (:end 4 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 16) (:end 4 16)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 18) (:end 4 18)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 19) (:end 4 19)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 21) (:end 4 22)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 24) (:end 4 24)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 25) (:end 4 28)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 30) (:end 4 30)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 31) (:end 4 31)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 33) (:end 4 36)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 38) (:end 4 38)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 39) (:end 4 39)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 1) (:end 5 5)) ((:name "unzip") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 7) (:end 5 8)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 10) (:end 5 10)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 12) (:end 5 12)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 13) (:end 5 14)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 15) (:end 5 15)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 17) (:end 5 18)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 19) (:end 5 19)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 1) (:end 6 5)) ((:name "unzip") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 7) (:end 6 7)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 8) (:end 6 8)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 9) (:end 6 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 10) (:end 6 10)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 12) (:end 6 12)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 13) (:end 6 13)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 15) (:end 6 16)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 18) (:end 6 20)) ((:name "abs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 21) (:end 6 21)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 3) (:end 7 3)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 5) (:end 7 7)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 9) (:end 7 9)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 10) (:end 7 11)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 12) (:end 7 12)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 14) (:end 7 15)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 16) (:end 7 16)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 18) (:end 7 18)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 20) (:end 7 24)) ((:name "unzip") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 26) (:end 7 28)) ((:name "abs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 30) (:end 7 31)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 5) (:end 8 5)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 6) (:end 8 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 8) (:end 8 9)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 11) (:end 8 12)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 13) (:end 8 13)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 15) (:end 8 15)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 17) (:end 8 18)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 20) (:end 8 21)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 8 22) (:end 8 22)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 1) (:end 10 6)) ((:decor :function)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 8) (:end 10 8)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 10) (:end 10 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 15) (:end 10 15)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 16) (:end 10 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 17) (:end 10 17)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 19) (:end 10 19)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 20) (:end 10 20)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 22) (:end 10 22)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 23) (:end 10 23)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 25) (:end 10 25)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 26) (:end 10 26)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 28) (:end 10 29)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 31) (:end 10 31)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 32) (:end 10 35)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 37) (:end 10 37)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 38) (:end 10 38)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 40) (:end 10 43)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 45) (:end 10 45)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 46) (:end 10 46)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 48) (:end 10 51)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 53) (:end 10 53)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 54) (:end 10 54)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 56) (:end 10 59)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 61) (:end 10 61)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 62) (:end 10 62)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 1) (:end 11 6)) ((:name "unzip4") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 8) (:end 11 9)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 11) (:end 11 11)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 13) (:end 11 13)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 14) (:end 11 15)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 16) (:end 11 16)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 18) (:end 11 19)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 20) (:end 11 20)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 22) (:end 11 23)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 24) (:end 11 24)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 26) (:end 11 27)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 28) (:end 11 28)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 1) (:end 12 6)) ((:name "unzip4") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 8) (:end 12 8)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 9) (:end 12 12)) ((:name "abcd") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 14) (:end 12 15)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 17) (:end 12 21)) ((:name "abcds") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 22) (:end 12 22)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 3) (:end 13 3)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 5) (:end 13 7)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 9) (:end 13 9)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 10) (:end 13 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 11) (:end 13 11)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 13) (:end 13 13)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 14) (:end 13 14)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 16) (:end 13 16)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 17) (:end 13 17)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 19) (:end 13 19)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 20) (:end 13 20)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 22) (:end 13 22)) ((:decor :keyword)))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 24) (:end 13 27)) ((:name "abcd") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 9) (:end 14 9)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 10) (:end 14 11)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 12) (:end 14 12)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 14) (:end 14 15)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 16) (:end 14 16)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 18) (:end 14 19)) ((:name "cs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 20) (:end 14 20)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 22) (:end 14 23)) ((:name "ds") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 24) (:end 14 24)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 26) (:end 14 26)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 28) (:end 14 33)) ((:name "unzip4") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 35) (:end 14 39)) ((:name "abcds") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 5) (:end 15 6)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 8) (:end 15 8)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 9) (:end 15 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 11) (:end 15 12)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 14) (:end 15 15)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 16) (:end 15 16)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 18) (:end 15 18)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 20) (:end 15 21)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 23) (:end 15 24)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 25) (:end 15 25)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 27) (:end 15 27)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 29) (:end 15 30)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 32) (:end 15 33)) ((:name "cs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 34) (:end 15 34)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 36) (:end 15 36)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 38) (:end 15 39)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 41) (:end 15 42)) ((:name "ds") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 15 43) (:end 15 43)) ((:decor :keyword)))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 16) (:end 1 23)) ((:name "uncurry") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 24) (:end 1 25)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 25) (:end 1 26)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 26) (:end 1 27)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 27) (:end 1 28)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 29) (:end 1 30)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 31) (:end 1 34)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 35) (:end 1 36)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 36) (:end 1 37)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 37) (:end 1 38)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 38) (:end 1 39)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 40) (:end 1 41)) ((:name ".") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000df(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 42) (:end 1 45)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 46) (:end 1 47)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 47) (:end 1 48)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 48) (:end 1 49)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 1 49) (:end 1 50)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 0) (:end 3 5)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 6) (:end 3 7)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 8) (:end 3 12)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 13) (:end 3 14)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 14) (:end 3 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 15) (:end 3 16)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 17) (:end 3 18)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 18) (:end 3 19)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 20) (:end 3 22)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 23) (:end 3 24)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 24) (:end 3 28)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 29) (:end 3 30)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 30) (:end 3 31)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 32) (:end 3 36)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 37) (:end 3 38)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 3 38) (:end 3 39)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 0) (:end 4 5)) ((:name "unzip") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 6) (:end 4 8)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 9) (:end 4 10)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 11) (:end 4 12)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 12) (:end 4 14)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 14) (:end 4 15)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 16) (:end 4 18)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 4 18) (:end 4 19)) ((:decor :keyword)))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 0) (:end 5 5)) ((:name "unzip") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 6) (:end 5 7)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 7) (:end 5 8)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 8) (:end 5 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d2(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 9) (:end 5 10)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 11) (:end 5 12)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 12) (:end 5 13)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 14) (:end 5 16)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 17) (:end 5 20)) ((:name "abs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 5 20) (:end 5 21)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 2) (:end 6 3)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 4) (:end 6 7)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 8) (:end 6 9)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 9) (:end 6 11)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 11) (:end 6 12)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 13) (:end 6 15)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 15) (:end 6 16)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 17) (:end 6 18)) ((:decor :keyword)))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 19) (:end 6 24)) ((:name "unzip") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 25) (:end 6 28)) ((:name "abs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 6 29) (:end 6 31)) ((:decor :keyword)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 4) (:end 7 5)) ((:decor :keyword)))))) 1)
|
||||
0000c6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 5) (:end 7 6)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 7) (:end 7 9)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 10) (:end 7 12)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 12) (:end 7 13)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 14) (:end 7 15)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 16) (:end 7 18)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 19) (:end 7 21)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 7 21) (:end 7 22)) ((:decor :keyword)))))) 1)
|
||||
000072(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 0) (:end 9 6)) ((:decor :function)))))) 1)
|
||||
000071(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 7) (:end 9 8)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 9) (:end 9 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 14) (:end 9 15)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 15) (:end 9 16)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 16) (:end 9 17)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 18) (:end 9 19)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 19) (:end 9 20)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 21) (:end 9 22)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 22) (:end 9 23)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 24) (:end 9 25)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 25) (:end 9 26)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 27) (:end 9 29)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 30) (:end 9 31)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 31) (:end 9 35)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 36) (:end 9 37)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 37) (:end 9 38)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 39) (:end 9 43)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 44) (:end 9 45)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 45) (:end 9 46)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 47) (:end 9 51)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 52) (:end 9 53)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 53) (:end 9 54)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 55) (:end 9 59)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 60) (:end 9 61)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 9 61) (:end 9 62)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 0) (:end 10 6)) ((:name "unzip4") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 7) (:end 10 9)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 10) (:end 10 11)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 12) (:end 10 13)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 13) (:end 10 15)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 15) (:end 10 16)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 17) (:end 10 19)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 19) (:end 10 20)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 21) (:end 10 23)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 23) (:end 10 24)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 25) (:end 10 27)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 10 27) (:end 10 28)) ((:decor :keyword)))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 0) (:end 11 6)) ((:name "unzip4") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 7) (:end 11 8)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 8) (:end 11 12)) ((:name "abcd") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 13) (:end 11 15)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 16) (:end 11 21)) ((:name "abcds") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 11 21) (:end 11 22)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 2) (:end 12 3)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 4) (:end 12 7)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 8) (:end 12 9)) ((:decor :keyword)))))) 1)
|
||||
0000c9(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 9) (:end 12 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 10) (:end 12 11)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 12) (:end 12 13)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 13) (:end 12 14)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 15) (:end 12 16)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 16) (:end 12 17)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 18) (:end 12 19)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 19) (:end 12 20)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 21) (:end 12 22)) ((:decor :keyword)))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 12 23) (:end 12 27)) ((:name "abcd") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 8) (:end 13 9)) ((:decor :keyword)))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 9) (:end 13 11)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 11) (:end 13 12)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 13) (:end 13 15)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 15) (:end 13 16)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 17) (:end 13 19)) ((:name "cs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 19) (:end 13 20)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 21) (:end 13 23)) ((:name "ds") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 23) (:end 13 24)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 25) (:end 13 26)) ((:decor :keyword)))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 27) (:end 13 33)) ((:name "unzip4") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 13 34) (:end 13 39)) ((:name "abcds") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 4) (:end 14 6)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 7) (:end 14 8)) ((:decor :keyword)))))) 1)
|
||||
0000c8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 8) (:end 14 9)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 10) (:end 14 12)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 13) (:end 14 15)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 15) (:end 14 16)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 17) (:end 14 18)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 19) (:end 14 21)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 22) (:end 14 24)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 24) (:end 14 25)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 26) (:end 14 27)) ((:name "c") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 28) (:end 14 30)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 31) (:end 14 33)) ((:name "cs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d5(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 33) (:end 14 34)) ((:name "MkPair") (:namespace "Builtin") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ca(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 35) (:end 14 36)) ((:name "d") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 37) (:end 14 39)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 40) (:end 14 42)) ((:name "ds") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "Tuples.idr") (:start 14 42) (:end 14 43)) ((:decor :keyword)))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,121 +1,121 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
00003e(:write-string "1/1: Building LetBinders (LetBinders.idr)" 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 1 1) (:end 1 6)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 1 8) (:end 1 17)) ((:decor :module)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 1) (:end 3 5)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 7) (:end 3 7)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 1) (:end 4 6)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 8) (:end 4 12)) ((:decor :type)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 14) (:end 4 14)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 15) (:end 4 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 17) (:end 4 17)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 19) (:end 4 22)) ((:decor :type)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 23) (:end 4 23)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 25) (:end 4 29)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 5 15) (:end 5 19)) ((:decor :data)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 3) (:end 6 6)) ((:decor :function)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 8) (:end 6 8)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 10) (:end 6 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 7 3) (:end 7 6)) ((:decor :function)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 7 8) (:end 7 8)) ((:decor :keyword)))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 7 10) (:end 7 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 7 10) (:end 7 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 7 10) (:end 7 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 7 15) (:end 7 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 1) (:end 9 4)) ((:decor :function)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 6) (:end 9 6)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 8) (:end 9 12)) ((:name "List1") (:namespace "LetBinders") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 14) (:end 9 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 16) (:end 9 17)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 19) (:end 9 23)) ((:name "List1") (:namespace "LetBinders") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 25) (:end 9 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 1) (:end 10 4)) ((:name "swap") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 6) (:end 10 8)) ((:name "aas") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 10) (:end 10 10)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 3) (:end 11 5)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 7) (:end 11 7)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 8) (:end 11 8)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 10) (:end 11 12)) ((:name ":::") (:namespace "LetBinders") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 14) (:end 11 15)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 16) (:end 11 16)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 18) (:end 11 19)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 21) (:end 11 23)) ((:name "aas") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 25) (:end 11 26)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 3) (:end 12 5)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 7) (:end 12 7)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 8) (:end 12 8)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 10) (:end 12 11)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 13) (:end 12 14)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 15) (:end 12 15)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 17) (:end 12 17)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 19) (:end 12 20)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 9) (:end 13 9)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 11) (:end 13 12)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 14) (:end 13 15)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 17) (:end 13 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 19) (:end 13 21)) ((:name ":::") (:namespace "LetBinders") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 23) (:end 13 24)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 7) (:end 14 10)) ((:name "rest") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 12) (:end 14 12)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 14) (:end 14 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 16) (:end 14 17)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 19) (:end 14 20)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 15 3) (:end 15 4)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 15 6) (:end 15 6)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 15 8) (:end 15 10)) ((:name ":::") (:namespace "LetBinders") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 15 12) (:end 15 15)) ((:name "rest") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 1) (:end 17 8)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 10) (:end 17 10)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 12) (:end 17 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 17) (:end 17 17)) ((:decor :keyword)))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 18) (:end 17 20)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 21) (:end 17 21)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 23) (:end 17 23)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 24) (:end 17 24)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 26) (:end 17 27)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 29) (:end 17 32)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 34) (:end 17 34)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 35) (:end 17 38)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 40) (:end 17 40)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 41) (:end 17 41)) ((:decor :keyword)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 18 1) (:end 18 8)) ((:name "identity") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 18 10) (:end 18 10)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 19 3) (:end 19 5)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 5) (:end 21 13)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 15) (:end 21 15)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 17) (:end 21 17)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 18) (:end 21 18)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 20) (:end 21 20)) ((:decor :keyword)))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 22) (:end 21 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 25) (:end 21 25)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 27) (:end 21 28)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 30) (:end 21 30)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 32) (:end 21 33)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 35) (:end 21 38)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 40) (:end 21 40)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 5) (:end 22 13)) ((:name "replicate") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 15) (:end 22 15)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 17) (:end 22 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 19) (:end 22 19)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 21) (:end 22 22)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 5) (:end 23 13)) ((:name "replicate") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 15) (:end 23 15)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 16) (:end 23 16)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 18) (:end 23 18)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 19) (:end 23 19)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 21) (:end 23 21)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 23) (:end 23 23)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 25) (:end 23 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 27) (:end 23 28)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 30) (:end 23 38)) ((:name "replicate") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 40) (:end 23 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 23 42) (:end 23 42)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 25 5) (:end 25 11)) ((:name "closure") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 25 13) (:end 25 14)) ((:decor :keyword)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 25 16) (:end 25 22)) ((:name "uncurry") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 25 24) (:end 25 32)) ((:name "replicate") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 27 3) (:end 27 4)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 27 6) (:end 27 8)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 27 10) (:end 27 16)) ((:name "closure") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 0 7) (:end 0 17)) ((:decor :module)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 0) (:end 2 5)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 6) (:end 2 7)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 0) (:end 3 6)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 7) (:end 3 12)) ((:decor :type)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 13) (:end 3 14)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 14) (:end 3 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 16) (:end 3 17)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 18) (:end 3 22)) ((:decor :type)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 22) (:end 3 23)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 24) (:end 3 29)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 4 14) (:end 4 19)) ((:decor :data)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 5 2) (:end 5 6)) ((:decor :function)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 5 7) (:end 5 8)) ((:decor :keyword)))))) 1)
|
||||
0000cb(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 5 9) (:end 5 10)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 2) (:end 6 6)) ((:decor :function)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 7) (:end 6 8)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 9) (:end 6 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 9) (:end 6 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 9) (:end 6 13)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 6 14) (:end 6 15)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 0) (:end 8 4)) ((:decor :function)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 5) (:end 8 6)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 7) (:end 8 12)) ((:name "List1") (:namespace "LetBinders") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 13) (:end 8 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 15) (:end 8 17)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 18) (:end 8 23)) ((:name "List1") (:namespace "LetBinders") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 8 24) (:end 8 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 0) (:end 9 4)) ((:name "swap") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 5) (:end 9 8)) ((:name "aas") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 9 9) (:end 9 10)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 2) (:end 10 5)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 6) (:end 10 7)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 7) (:end 10 8)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 9) (:end 10 12)) ((:name ":::") (:namespace "LetBinders") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 13) (:end 10 15)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 15) (:end 10 16)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 17) (:end 10 19)) ((:decor :keyword)))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 20) (:end 10 23)) ((:name "aas") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 10 24) (:end 10 26)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 2) (:end 11 5)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 6) (:end 11 7)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 7) (:end 11 8)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 9) (:end 11 11)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 12) (:end 11 14)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 14) (:end 11 15)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 16) (:end 11 17)) ((:decor :keyword)))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 11 18) (:end 11 20)) ((:name "as") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 8) (:end 12 9)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 10) (:end 12 12)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 13) (:end 12 15)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 16) (:end 12 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 18) (:end 12 21)) ((:name ":::") (:namespace "LetBinders") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 12 22) (:end 12 24)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 6) (:end 13 10)) ((:name "rest") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 11) (:end 13 12)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 13) (:end 13 14)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 15) (:end 13 17)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 13 18) (:end 13 20)) ((:name "bs") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 2) (:end 14 4)) ((:decor :keyword)))))) 1)
|
||||
0000cc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 5) (:end 14 6)) ((:name "b") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 7) (:end 14 10)) ((:name ":::") (:namespace "LetBinders") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d1(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 14 11) (:end 14 15)) ((:name "rest") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 0) (:end 16 8)) ((:decor :function)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 9) (:end 16 10)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 11) (:end 16 15)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 16) (:end 16 17)) ((:decor :keyword)))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 17) (:end 16 20)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 20) (:end 16 21)) ((:name "Pair") (:namespace "Builtin") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 22) (:end 16 23)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 23) (:end 16 24)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 25) (:end 16 27)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 28) (:end 16 32)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 33) (:end 16 34)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 34) (:end 16 38)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 39) (:end 16 40)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 16 40) (:end 16 41)) ((:decor :keyword)))))) 1)
|
||||
0000e0(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 0) (:end 17 8)) ((:name "identity") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 17 9) (:end 17 10)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 18 2) (:end 18 5)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 4) (:end 20 13)) ((:decor :function)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 14) (:end 20 15)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 16) (:end 20 17)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 17) (:end 20 18)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 19) (:end 20 20)) ((:decor :keyword)))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 21) (:end 20 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 24) (:end 20 25)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 26) (:end 20 28)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 29) (:end 20 30)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 31) (:end 20 33)) ((:decor :keyword)))))) 1)
|
||||
0000de(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 34) (:end 20 38)) ((:name "List") (:namespace "Prelude.Basics") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 20 39) (:end 20 40)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 4) (:end 21 13)) ((:name "replicate") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 14) (:end 21 15)) ((:name "Z") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 16) (:end 21 17)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 18) (:end 21 19)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 21 20) (:end 21 22)) ((:name "Nil") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 4) (:end 22 13)) ((:name "replicate") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 14) (:end 22 15)) ((:decor :keyword)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 15) (:end 22 16)) ((:name "S") (:namespace "Prelude.Types") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 17) (:end 22 18)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 18) (:end 22 19)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 20) (:end 22 21)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 22) (:end 22 23)) ((:decor :keyword)))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 24) (:end 22 25)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000dc(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 26) (:end 22 28)) ((:name "::") (:namespace "Prelude.Basics") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 29) (:end 22 38)) ((:name "replicate") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 39) (:end 22 40)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000ce(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 22 41) (:end 22 42)) ((:name "a") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 24 4) (:end 24 11)) ((:name "closure") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 24 12) (:end 24 14)) ((:decor :keyword)))))) 1)
|
||||
0000e5(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 24 15) (:end 24 22)) ((:name "uncurry") (:namespace "Prelude.Basics") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 24 23) (:end 24 32)) ((:name "replicate") (:namespace "LetBinders") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 26 2) (:end 26 4)) ((:decor :keyword)))))) 1)
|
||||
0000e3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 26 5) (:end 26 8)) ((:name "map") (:namespace "Prelude.Interfaces") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d3(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 26 9) (:end 26 16)) ((:name "closure") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1,93 +1,93 @@
|
||||
000018(:protocol-version 2 0)
|
||||
000018(:protocol-version 2 1)
|
||||
000040(:write-string "1/1: Building SnocRainbow (SnocRainbow.idr)" 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 1 11) (:end 1 12)) ((:decor :namespace)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 3) (:end 2 8)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 10) (:end 2 15)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 3 3) (:end 3 6)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 3 8) (:end 3 12)) ((:decor :type)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 3 14) (:end 3 14)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 3 16) (:end 3 19)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 6 11) (:end 6 12)) ((:decor :namespace)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 3) (:end 7 8)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 10) (:end 7 15)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 3) (:end 8 6)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 8) (:end 8 12)) ((:decor :type)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 14) (:end 8 14)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 16) (:end 8 19)) ((:decor :type)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 21) (:end 8 25)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 5) (:end 9 7)) ((:decor :data)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 9) (:end 9 9)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 11) (:end 9 15)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 10 5) (:end 10 9)) ((:decor :data)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 10 11) (:end 10 11)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 10 13) (:end 10 17)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 10 19) (:end 10 20)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 10 22) (:end 10 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 10 26) (:end 10 27)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 10 29) (:end 10 33)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 3) (:end 12 8)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 10) (:end 12 15)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 3) (:end 13 6)) ((:decor :function)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 8) (:end 13 8)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 10) (:end 13 14)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 16) (:end 13 17)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 19) (:end 13 21)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 23) (:end 13 24)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 26) (:end 13 30)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 14 3) (:end 14 6)) ((:name ":<") (:namespace "Main.L1") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 14 8) (:end 14 8)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 14 10) (:end 14 14)) ((:name "Cons1") (:namespace "Main.L1") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 16 11) (:end 16 12)) ((:decor :namespace)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 3) (:end 17 8)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 10) (:end 17 15)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 3) (:end 18 6)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 8) (:end 18 11)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 13) (:end 18 13)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 15) (:end 18 19)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 21) (:end 18 22)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 24) (:end 18 26)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 28) (:end 18 29)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 31) (:end 18 34)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 18 36) (:end 18 40)) ((:decor :keyword)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 20 11) (:end 20 12)) ((:decor :namespace)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 3) (:end 21 8)) ((:decor :keyword)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 10) (:end 21 15)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 3) (:end 22 6)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 8) (:end 22 12)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 14) (:end 22 14)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 16) (:end 22 19)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 21) (:end 22 25)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 5) (:end 23 7)) ((:decor :data)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 9) (:end 23 9)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 11) (:end 23 15)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 24 5) (:end 24 8)) ((:decor :data)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 24 10) (:end 24 10)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 24 12) (:end 24 15)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 24 17) (:end 24 18)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 24 20) (:end 24 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 24 24) (:end 24 25)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 24 27) (:end 24 31)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 1) (:end 26 1)) ((:decor :function)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 3) (:end 26 3)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 5) (:end 26 7)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 27 1) (:end 27 1)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 27 3) (:end 27 3)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 27 5) (:end 27 14)) ((:name "believe_me") (:namespace "Builtin") (:decor :postulate) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 27 16) (:end 27 23)) ((:decor :data)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 1) (:end 29 7)) ((:decor :function)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 9) (:end 29 9)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 11) (:end 29 13)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 15) (:end 29 16)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 18) (:end 29 21)) ((:decor :type)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 1) (:end 30 7)) ((:name "Rainbow") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 9) (:end 30 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 11) (:end 30 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 14) (:end 30 15)) ((:name "Lin") (:namespace "Main.L0") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 17) (:end 30 17)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 19) (:end 30 19)) ((:name ":<") (:namespace "Main.L2") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 21) (:end 30 21)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 23) (:end 30 23)) ((:name ":<") (:namespace "Main.L0") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 25) (:end 30 25)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 30 27) (:end 30 27)) ((:name ":<") (:namespace "Main.L2") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 0 10) (:end 0 12)) ((:decor :namespace)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 1 2) (:end 1 8)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 1 9) (:end 1 15)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 2) (:end 2 6)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 7) (:end 2 12)) ((:decor :type)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 13) (:end 2 14)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 2 15) (:end 2 19)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 5 10) (:end 5 12)) ((:decor :namespace)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 6 2) (:end 6 8)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 6 9) (:end 6 15)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 2) (:end 7 6)) ((:decor :keyword)))))) 1)
|
||||
000074(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 7) (:end 7 12)) ((:decor :type)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 13) (:end 7 14)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 15) (:end 7 19)) ((:decor :type)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 7 20) (:end 7 25)) ((:decor :keyword)))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 4) (:end 8 7)) ((:decor :data)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 8) (:end 8 9)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 8 10) (:end 8 15)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000073(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 4) (:end 9 9)) ((:decor :data)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 10) (:end 9 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 12) (:end 9 17)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 18) (:end 9 20)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 21) (:end 9 24)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 25) (:end 9 27)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 9 28) (:end 9 33)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 11 2) (:end 11 8)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 11 9) (:end 11 15)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 2) (:end 12 6)) ((:decor :function)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 7) (:end 12 8)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 9) (:end 12 14)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 15) (:end 12 17)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 18) (:end 12 21)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 22) (:end 12 24)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 12 25) (:end 12 30)) ((:name "List1") (:namespace "Main.L1") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 2) (:end 13 6)) ((:name ":<") (:namespace "Main.L1") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 7) (:end 13 8)) ((:decor :keyword)))))) 1)
|
||||
0000d8(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 13 9) (:end 13 14)) ((:name "Cons1") (:namespace "Main.L1") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 15 10) (:end 15 12)) ((:decor :namespace)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 16 2) (:end 16 8)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 16 9) (:end 16 15)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 2) (:end 17 6)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 7) (:end 17 11)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 12) (:end 17 13)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 14) (:end 17 19)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 20) (:end 17 22)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 23) (:end 17 26)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 27) (:end 17 29)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 30) (:end 17 34)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 17 35) (:end 17 40)) ((:decor :keyword)))))) 1)
|
||||
00007c(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 19 10) (:end 19 12)) ((:decor :namespace)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 20 2) (:end 20 8)) ((:decor :keyword)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 20 9) (:end 20 15)) ((:decor :keyword)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 2) (:end 21 6)) ((:decor :keyword)))))) 1)
|
||||
000076(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 7) (:end 21 12)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 13) (:end 21 14)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 15) (:end 21 19)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 21 20) (:end 21 25)) ((:decor :keyword)))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 4) (:end 22 7)) ((:decor :data)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 8) (:end 22 9)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 22 10) (:end 22 15)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000075(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 4) (:end 23 8)) ((:decor :data)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 9) (:end 23 10)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 11) (:end 23 15)) ((:decor :type)))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 16) (:end 23 18)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 19) (:end 23 22)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 23) (:end 23 25)) ((:decor :keyword)))))) 1)
|
||||
0000d9(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 23 26) (:end 23 31)) ((:name "List0") (:namespace "Main.L0") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 25 0) (:end 25 1)) ((:decor :function)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 25 2) (:end 25 3)) ((:decor :keyword)))))) 1)
|
||||
0000db(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 25 4) (:end 25 7)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d4(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 0) (:end 26 1)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 2) (:end 26 3)) ((:decor :keyword)))))) 1)
|
||||
0000e2(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 4) (:end 26 14)) ((:name "believe_me") (:namespace "Builtin") (:decor :postulate) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 26 15) (:end 26 23)) ((:decor :data)))))) 1)
|
||||
000079(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 0) (:end 28 7)) ((:decor :function)))))) 1)
|
||||
000078(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 8) (:end 28 9)) ((:decor :keyword)))))) 1)
|
||||
0000dd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 10) (:end 28 13)) ((:name "Nat") (:namespace "Prelude.Types") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 14) (:end 28 16)) ((:decor :keyword)))))) 1)
|
||||
000077(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 28 17) (:end 28 21)) ((:decor :type)))))) 1)
|
||||
0000da(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 0) (:end 29 7)) ((:name "Rainbow") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cd(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 8) (:end 29 9)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
00007a(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 10) (:end 29 11)) ((:decor :keyword)))))) 1)
|
||||
0000d7(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 13) (:end 29 15)) ((:name "Lin") (:namespace "Main.L0") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 16) (:end 29 17)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 18) (:end 29 19)) ((:name ":<") (:namespace "Main.L2") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 20) (:end 29 21)) ((:name "m") (:namespace "Main") (:decor :function) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 22) (:end 29 23)) ((:name ":<") (:namespace "Main.L0") (:decor :data) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000cf(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 24) (:end 29 25)) ((:name "n") (:namespace "") (:decor :bound) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
0000d6(:output (:ok (:highlight-source ((((:filename "SnocRainbow.idr") (:start 29 26) (:end 29 27)) ((:name ":<") (:namespace "Main.L2") (:decor :type) (:implicit :False) (:key "") (:doc-overview "") (:type "")))))) 1)
|
||||
000015(:return (:ok ()) 1)
|
||||
Alas the file is done, aborting
|
||||
|
Loading…
Reference in New Issue
Block a user