Use version number field in ipkg

Packages are now installed in a directory with their version number.
On adding a package directory, we now look in a local 'depends'
directory first (to allow packages to be installed locally to another
project) before the global install directory.
Dependencies can have version bounds (details yet to be implemented) and
we pick the package with the highest version number that matches.
This commit is contained in:
Edwin Brady 2021-02-27 14:15:19 +00:00
parent f334a050b7
commit 1052c41a3c
16 changed files with 195 additions and 26 deletions

View File

@ -6,6 +6,14 @@ REPL/IDE mode changes:
* Added `:search` command, which searches for functions by type * Added `:search` command, which searches for functions by type
* `:load`/`:l` and `:cd` commands now only accept paths surrounded by double quotes * `:load`/`:l` and `:cd` commands now only accept paths surrounded by double quotes
Other changes:
* The `version` field in `.ipkg` files is now used. Packages are installed into
a directory which includes its version number, and dependencies can give
version number ranges.
* Idris now looks in the current working directory, under a subdirectory
`depends` for local installations of packages before looking globally.
Changes since Idris 2 v0.2.1 Changes since Idris 2 v0.2.1
---------------------------- ----------------------------

View File

@ -1,5 +1,6 @@
[ ] Change version number (MAJOR, MINOR, PATCH) in Makefile [ ] Change version number (MAJOR, MINOR, PATCH) in Makefile
[ ] Change version numbers in doc listings [ ] Change version numbers in doc listings
[ ] Change version numbers in prelude, base, contrib, network ipkgs
[ ] Update bootstrap chez and racket [ ] Update bootstrap chez and racket
[ ] Tag on github with version number (in the form vX.Y.Z) [ ] Tag on github with version number (in the form vX.Y.Z)
[ ] Run release script [ ] Run release script

View File

@ -1,4 +1,5 @@
package base package base
version = 0.3.0
modules = Control.App, modules = Control.App,
Control.App.Console, Control.App.Console,

View File

@ -1,4 +1,5 @@
package contrib package contrib
version = 0.3.0
modules = Control.ANSI, modules = Control.ANSI,
Control.ANSI.SGR, Control.ANSI.SGR,

View File

@ -1,4 +1,5 @@
package network package network
version = 0.3.0
opts = "-p contrib" opts = "-p contrib"

View File

@ -1,4 +1,5 @@
package prelude package prelude
version = 0.3.0
opts = "--no-prelude" opts = "--no-prelude"

View File

@ -2035,6 +2035,12 @@ setBuildDir dir
= do defs <- get Ctxt = do defs <- get Ctxt
put Ctxt (record { options->dirs->build_dir = dir } defs) put Ctxt (record { options->dirs->build_dir = dir } defs)
export
setDependsDir : {auto c : Ref Ctxt Defs} -> String -> Core ()
setDependsDir dir
= do defs <- get Ctxt
put Ctxt (record { options->dirs->depends_dir = dir } defs)
export export
setOutputDir : {auto c : Ref Ctxt Defs} -> Maybe String -> Core () setOutputDir : {auto c : Ref Ctxt Defs} -> Maybe String -> Core ()
setOutputDir dir setOutputDir dir

View File

@ -134,6 +134,7 @@ data Error : Type where
GenericMsg : FC -> String -> Error GenericMsg : FC -> String -> Error
TTCError : TTCErrorMsg -> Error TTCError : TTCErrorMsg -> Error
FileErr : String -> FileError -> Error FileErr : String -> FileError -> Error
CantFindPackage : String -> Error
LitFail : FC -> Error LitFail : FC -> Error
LexFail : FC -> String -> Error LexFail : FC -> String -> Error
ParseFail : (Show token, Pretty token) => ParseFail : (Show token, Pretty token) =>
@ -298,6 +299,7 @@ Show Error where
show (GenericMsg fc str) = show fc ++ ":" ++ str show (GenericMsg fc str) = show fc ++ ":" ++ str
show (TTCError msg) = "Error in TTC file: " ++ show msg show (TTCError msg) = "Error in TTC file: " ++ show msg
show (FileErr fname err) = "File error (" ++ fname ++ "): " ++ show err show (FileErr fname err) = "File error (" ++ fname ++ "): " ++ show err
show (CantFindPackage fname) = "Can't find package " ++ fname
show (LitFail fc) = show fc ++ ":Can't parse literate" show (LitFail fc) = show fc ++ ":Can't parse literate"
show (LexFail fc err) = show fc ++ ":Lexer error (" ++ show err ++ ")" show (LexFail fc err) = show fc ++ ":Lexer error (" ++ show err ++ ")"
show (ParseFail fc err toks) = "Parse error (" ++ show err ++ "): " ++ show toks show (ParseFail fc err toks) = "Parse error (" ++ show err ++ "): " ++ show toks
@ -378,6 +380,7 @@ getErrorLoc (BadRunElab loc _ _) = Just loc
getErrorLoc (GenericMsg loc _) = Just loc getErrorLoc (GenericMsg loc _) = Just loc
getErrorLoc (TTCError _) = Nothing getErrorLoc (TTCError _) = Nothing
getErrorLoc (FileErr _ _) = Nothing getErrorLoc (FileErr _ _) = Nothing
getErrorLoc (CantFindPackage _) = Nothing
getErrorLoc (LitFail loc) = Just loc getErrorLoc (LitFail loc) = Just loc
getErrorLoc (LexFail loc _) = Just loc getErrorLoc (LexFail loc _) = Just loc
getErrorLoc (ParseFail loc _ _) = Just loc getErrorLoc (ParseFail loc _ _) = Just loc

View File

@ -21,6 +21,7 @@ record Dirs where
working_dir : String working_dir : String
source_dir : Maybe String -- source directory, relative to working directory source_dir : Maybe String -- source directory, relative to working directory
build_dir : String -- build directory, relative to working directory build_dir : String -- build directory, relative to working directory
depends_dir : String -- local dependencies directory, relative to working directory
output_dir : Maybe String -- output directory, relative to working directory output_dir : Maybe String -- output directory, relative to working directory
prefix_dir : String -- installation prefix, for finding data files (e.g. run time support) prefix_dir : String -- installation prefix, for finding data files (e.g. run time support)
extra_dirs : List String -- places to look for import files extra_dirs : List String -- places to look for import files
@ -37,10 +38,11 @@ outputDirWithDefault d = fromMaybe (build_dir d </> "exec") (output_dir d)
public export public export
toString : Dirs -> String toString : Dirs -> String
toString d@(MkDirs wdir sdir bdir odir dfix edirs ldirs ddirs) = toString d@(MkDirs wdir sdir bdir ldir odir dfix edirs ldirs ddirs) =
unlines [ "+ Working Directory :: " ++ show wdir unlines [ "+ Working Directory :: " ++ show wdir
, "+ Source Directory :: " ++ show sdir , "+ Source Directory :: " ++ show sdir
, "+ Build Directory :: " ++ show bdir , "+ Build Directory :: " ++ show bdir
, "+ Local Depend Directory :: " ++ show ldir
, "+ Output Directory :: " ++ show (outputDirWithDefault d) , "+ Output Directory :: " ++ show (outputDirWithDefault d)
, "+ Installation Prefix :: " ++ show dfix , "+ Installation Prefix :: " ++ show dfix
, "+ Extra Directories :: " ++ show edirs , "+ Extra Directories :: " ++ show edirs
@ -175,7 +177,7 @@ getCG : Options -> String -> Maybe CG
getCG o cg = lookup (toLower cg) (availableCGs o) getCG o cg = lookup (toLower cg) (availableCGs o)
defaultDirs : Dirs defaultDirs : Dirs
defaultDirs = MkDirs "." Nothing "build" Nothing defaultDirs = MkDirs "." Nothing "build" "depends" Nothing
"/usr/local" ["."] [] [] "/usr/local" ["."] [] []
defaultPPrint : PPrinter defaultPPrint : PPrinter

View File

@ -5,6 +5,7 @@ import IdrisPaths
import Idris.Env import Idris.Env
import Idris.Version import Idris.Version
import Core.Name.Namespace
import Core.Options import Core.Options
import Data.List import Data.List
@ -40,6 +41,46 @@ export
Show DirCommand where Show DirCommand where
show LibDir = "--libdir" show LibDir = "--libdir"
public export
data PkgVersion = MkPkgVersion (List Nat)
export
Show PkgVersion where
show (MkPkgVersion vs) = showSep "." (map show vs)
export
Eq PkgVersion where
MkPkgVersion p == MkPkgVersion q = p == q
export
Ord PkgVersion where
-- list ordering gives us what we want
compare (MkPkgVersion p) (MkPkgVersion q) = compare p q
-- version must be >= lowerBound and <= upperBound
-- Do we want < and > as well?
public export
record PkgVersionBounds where
constructor MkPkgVersionBounds
lowerBound : Maybe PkgVersion
upperBound : Maybe PkgVersion
export
anyBounds : PkgVersionBounds
anyBounds = MkPkgVersionBounds Nothing Nothing
export
current : PkgVersionBounds
current = let (maj,min,patch) = semVer version
version = Just (MkPkgVersion [maj, min, patch]) in
MkPkgVersionBounds version version
export
inBounds : PkgVersion -> PkgVersionBounds -> Bool
inBounds v bounds
= maybe True (\v' => v >= v') bounds.lowerBound &&
maybe True (\v' => v <= v') bounds.upperBound
||| CLOpt - possible command line options ||| CLOpt - possible command line options
public export public export
data CLOpt data CLOpt

View File

@ -74,8 +74,8 @@ updateEnv
-- for the tests means they test the local version not the installed -- for the tests means they test the local version not the installed
-- version -- version
defs <- get Ctxt defs <- get Ctxt
addPkgDir "prelude" addPkgDir "prelude" anyBounds
addPkgDir "base" addPkgDir "base" anyBounds
addDataDir (prefix_dir (dirs (options defs)) </> addDataDir (prefix_dir (dirs (options defs)) </>
("idris2-" ++ showVersion False version) </> "support") ("idris2-" ++ showVersion False version) </> "support")
addLibDir (prefix_dir (dirs (options defs)) </> addLibDir (prefix_dir (dirs (options defs)) </>

View File

@ -406,6 +406,8 @@ perror (TTCError msg)
<++> parens (pretty "the most likely case is that the ./build directory in your current project contains files from a previous build of idris2 or the idris2 executable is from a different build than the installed .ttc files") <++> parens (pretty "the most likely case is that the ./build directory in your current project contains files from a previous build of idris2 or the idris2 executable is from a different build than the installed .ttc files")
perror (FileErr fname err) perror (FileErr fname err)
= pure $ errorDesc (reflow "File error in" <++> pretty fname <++> colon) <++> pretty (show err) = pure $ errorDesc (reflow "File error in" <++> pretty fname <++> colon) <++> pretty (show err)
perror (CantFindPackage fname)
= pure $ errorDesc (reflow "Can't find package " <++> pretty fname)
perror (LitFail fc) perror (LitFail fc)
= pure $ errorDesc (reflow "Can't parse literate.") <+> line <+> !(ploc fc) = pure $ errorDesc (reflow "Can't parse literate.") <+> line <+> !(ploc fc)
perror (LexFail fc msg) perror (LexFail fc msg)

View File

@ -44,7 +44,7 @@ public export
record PkgDesc where record PkgDesc where
constructor MkPkgDesc constructor MkPkgDesc
name : String name : String
version : String version : PkgVersion
authors : String authors : String
maintainers : Maybe String maintainers : Maybe String
license : Maybe String license : Maybe String
@ -68,10 +68,13 @@ record PkgDesc where
preclean : Maybe (FC, String) -- Script to run before cleaning preclean : Maybe (FC, String) -- Script to run before cleaning
postclean : Maybe (FC, String) -- Script to run after cleaning postclean : Maybe (FC, String) -- Script to run after cleaning
installDir : PkgDesc -> String
installDir p = name p ++ "-" ++ show (version p)
export export
Show PkgDesc where Show PkgDesc where
show pkg = "Package: " ++ name pkg ++ "\n" ++ show pkg = "Package: " ++ name pkg ++ "\n" ++
"Version: " ++ version pkg ++ "\n" ++ "Version: " ++ show (version pkg) ++ "\n" ++
"Authors: " ++ authors pkg ++ "\n" ++ "Authors: " ++ authors pkg ++ "\n" ++
maybe "" (\m => "Maintainers: " ++ m ++ "\n") (maintainers pkg) ++ maybe "" (\m => "Maintainers: " ++ m ++ "\n") (maintainers pkg) ++
maybe "" (\m => "License: " ++ m ++ "\n") (license pkg) ++ maybe "" (\m => "License: " ++ m ++ "\n") (license pkg) ++
@ -97,14 +100,14 @@ Show PkgDesc where
initPkgDesc : String -> PkgDesc initPkgDesc : String -> PkgDesc
initPkgDesc pname initPkgDesc pname
= MkPkgDesc pname "0" "Anonymous" Nothing Nothing = MkPkgDesc pname (MkPkgVersion [0,0]) "Anonymous" Nothing Nothing
Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing
[] [] [] []
Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing
Nothing Nothing Nothing Nothing Nothing Nothing Nothing Nothing
data DescField : Type where data DescField : Type where
PVersion : FC -> String -> DescField PVersion : FC -> PkgVersion -> DescField
PAuthors : FC -> String -> DescField PAuthors : FC -> String -> DescField
PMaintainers : FC -> String -> DescField PMaintainers : FC -> String -> DescField
PLicense : FC -> String -> DescField PLicense : FC -> String -> DescField
@ -130,8 +133,7 @@ data DescField : Type where
field : String -> Rule DescField field : String -> Rule DescField
field fname field fname
= strField PVersion "version" = strField PAuthors "authors"
<|> strField PAuthors "authors"
<|> strField PMaintainers "maintainers" <|> strField PMaintainers "maintainers"
<|> strField PLicense "license" <|> strField PLicense "license"
<|> strField PBrief "brief" <|> strField PBrief "brief"
@ -150,6 +152,13 @@ field fname
<|> strField PPostinstall "postinstall" <|> strField PPostinstall "postinstall"
<|> strField PPreclean "preclean" <|> strField PPreclean "preclean"
<|> strField PPostclean "postclean" <|> strField PPostclean "postclean"
<|> do start <- location
ignore $ exactProperty "version"
equals
vs <- sepBy1 dot' integerLit
end <- location
pure (PVersion (MkFC fname start end)
(MkPkgVersion (map fromInteger vs)))
<|> do ignore $ exactProperty "depends" <|> do ignore $ exactProperty "depends"
equals equals
ds <- sep packageName ds <- sep packageName
@ -255,7 +264,7 @@ addDeps : {auto c : Ref Ctxt Defs} ->
PkgDesc -> Core () PkgDesc -> Core ()
addDeps pkg addDeps pkg
= do defs <- get Ctxt = do defs <- get Ctxt
traverse_ addPkgDir (depends pkg) traverse_ (\p => addPkgDir p anyBounds) (depends pkg)
processOptions : {auto c : Ref Ctxt Defs} -> processOptions : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} -> {auto o : Ref ROpts REPLOpts} ->
@ -325,8 +334,8 @@ copyFile src dest
writeToFile dest buf writeToFile dest buf
installFrom : {auto c : Ref Ctxt Defs} -> installFrom : {auto c : Ref Ctxt Defs} ->
String -> String -> String -> ModuleIdent -> Core () String -> String -> ModuleIdent -> Core ()
installFrom pname builddir destdir ns installFrom builddir destdir ns
= do let ttcfile = joinPath (reverse $ unsafeUnfoldModuleIdent ns) = do let ttcfile = joinPath (reverse $ unsafeUnfoldModuleIdent ns)
let ttcPath = builddir </> "ttc" </> ttcfile <.> "ttc" let ttcPath = builddir </> "ttc" </> ttcfile <.> "ttc"
@ -368,18 +377,17 @@ install pkg opts -- not used but might be in the future
"idris2-" ++ showVersion False version "idris2-" ++ showVersion False version
True <- coreLift $ changeDir installPrefix True <- coreLift $ changeDir installPrefix
| False => throw $ InternalError $ "Can't change directory to " ++ installPrefix | False => throw $ InternalError $ "Can't change directory to " ++ installPrefix
Right _ <- coreLift $ mkdirAll (name pkg) Right _ <- coreLift $ mkdirAll (installDir pkg)
| Left err => throw $ InternalError $ unlines | Left err => throw $ InternalError $ unlines
[ "Can't make directory " ++ name pkg [ "Can't make directory " ++ installDir pkg
, show err ] , show err ]
True <- coreLift $ changeDir (name pkg) True <- coreLift $ changeDir (installDir pkg)
| False => throw $ InternalError $ "Can't change directory to " ++ name pkg | False => throw $ InternalError $ "Can't change directory to " ++ installDir pkg
-- We're in that directory now, so copy the files from -- We're in that directory now, so copy the files from
-- srcdir/build into it -- srcdir/build into it
traverse_ (installFrom (name pkg) traverse_ (installFrom (srcdir </> build)
(srcdir </> build) (installPrefix </> installDir pkg)) toInstall
(installPrefix </> name pkg)) toInstall
coreLift_ $ changeDir srcdir coreLift_ $ changeDir srcdir
runScript (postinstall pkg) runScript (postinstall pkg)
@ -643,4 +651,4 @@ findIpkg fname
dropHead str (x :: xs) dropHead str (x :: xs)
= if x == str then xs else x :: xs = if x == str then xs else x :: xs
loadDependencies : List String -> Core () loadDependencies : List String -> Core ()
loadDependencies = traverse_ addPkgDir loadDependencies = traverse_ (\p => addPkgDir p anyBounds)

View File

@ -14,19 +14,88 @@ import Idris.Version
import IdrisPaths import IdrisPaths
import Data.List
import Data.So import Data.So
import Data.Strings
import System import System
import System.Directory
%default covering %default covering
-- Get a list of all the candidate directories that match a package spec
-- in a given path. Return an empty list on file error (e.g. path not existing)
export
candidateDirs : String -> String -> PkgVersionBounds ->
IO (List (String, PkgVersion))
candidateDirs dname pkg bounds
= do Right d <- openDir dname
| Left err => pure []
getFiles d []
where
toVersionNum : String -> List Nat
toVersionNum str
= case span (/= '.') str of
(num, rest) =>
case strM rest of
StrNil => [stringToNatOrZ num]
StrCons _ rest =>
(stringToNatOrZ num :: toVersionNum rest)
getVersion : String -> (String, PkgVersion)
getVersion str
= case span (/= '-') str of
(name, ver) => case strM ver of
StrNil => (name, MkPkgVersion [0])
StrCons _ ver =>
(name, MkPkgVersion (toVersionNum ver))
_ => (str, MkPkgVersion [0])
-- Return a list of paths that match the version spec
-- (full name, version string)
-- We'll order by version string that the highest version number is the
-- one we use
getFiles : Directory -> List (String, PkgVersion) ->
IO (List (String, PkgVersion))
getFiles d acc
= do Right str <- dirEntry d
| Left err => pure (reverse acc)
let (pkgdir, ver) = getVersion str
if pkgdir == pkg && inBounds ver bounds
then getFiles d (((dname </> str), ver) :: acc)
else getFiles d acc
-- TODO: Version numbers on dependencies -- TODO: Version numbers on dependencies
export export
addPkgDir : {auto c : Ref Ctxt Defs} -> addPkgDir : {auto c : Ref Ctxt Defs} ->
String -> Core () String -> PkgVersionBounds -> Core ()
addPkgDir p addPkgDir p bounds
= do defs <- get Ctxt = do defs <- get Ctxt
addExtraDir (prefix_dir (dirs (options defs)) </> let globaldir = prefix_dir (dirs (options defs)) </>
"idris2-" ++ showVersion False version </> p) "idris2-" ++ showVersion False version
let depends = depends_dir (dirs (options defs))
Just srcdir <- coreLift currentDir
| Nothing => throw (InternalError "Can't get current directory")
let localdir = srcdir </> depends
-- Get candidate directories from the global install location,
-- and the local package directory
locFiles <- coreLift $ candidateDirs localdir p bounds
globFiles <- coreLift $ candidateDirs globaldir p bounds
-- If there's anything locally, use that and ignore the global ones
let allFiles = if isNil locFiles
then globFiles
else locFiles
-- Sort in reverse order of version number
let sorted = sortBy (\x, y => compare (snd y) (snd x)) allFiles
-- From what remains, pick the one with the highest version number.
-- If there's none, report it
-- (TODO: Can't do this quite yet due to idris2 build system...)
case sorted of
[] => pure () -- throw (CantFindPackage p)
((p, _) :: ps) => addExtraDir p
dirOption : Dirs -> DirCommand -> Core () dirOption : Dirs -> DirCommand -> Core ()
dirOption dirs LibDir dirOption dirs LibDir
@ -78,7 +147,7 @@ preOptions (Directive d :: opts)
= do setSession (record { directives $= (d::) } !getSession) = do setSession (record { directives $= (d::) } !getSession)
preOptions opts preOptions opts
preOptions (PkgPath p :: opts) preOptions (PkgPath p :: opts)
= do addPkgDir p = do addPkgDir p anyBounds
preOptions opts preOptions opts
preOptions (SourceDir d :: opts) preOptions (SourceDir d :: opts)
= do setSourceDir (Just d) = do setSourceDir (Just d)

View File

@ -23,8 +23,10 @@ data Token
| Equals | Equals
| DotSepIdent (Maybe Namespace) String | DotSepIdent (Maybe Namespace) String
| Separator | Separator
| Dot
| Space | Space
| StringLit String | StringLit String
| IntegerLit Integer
public export public export
Show Token where Show Token where
@ -33,8 +35,10 @@ Show Token where
show Equals = "Equals" show Equals = "Equals"
show (DotSepIdent ns n) = "DotSepIdentifier: " ++ show ns ++ "." ++ show n show (DotSepIdent ns n) = "DotSepIdentifier: " ++ show ns ++ "." ++ show n
show Separator = "Separator" show Separator = "Separator"
show Dot = "Dot"
show Space = "Space" show Space = "Space"
show (StringLit s) = "StringLit: " ++ s show (StringLit s) = "StringLit: " ++ s
show (IntegerLit i) = "IntegerLit: " ++ show i
public export public export
Pretty Token where Pretty Token where
@ -43,8 +47,10 @@ Pretty Token where
pretty Equals = "Equals" pretty Equals = "Equals"
pretty (DotSepIdent ns n) = "DotSepIdentifier:" <++> pretty ns <+> dot <+> pretty n pretty (DotSepIdent ns n) = "DotSepIdentifier:" <++> pretty ns <+> dot <+> pretty n
pretty Separator = "Separator" pretty Separator = "Separator"
pretty Dot = "Dot"
pretty Space = "Space" pretty Space = "Space"
pretty (StringLit s) = "StringLit:" <++> pretty s pretty (StringLit s) = "StringLit:" <++> pretty s
pretty (IntegerLit i) = "IntegerLit:" <++> pretty i
equals : Lexer equals : Lexer
equals = is '=' equals = is '='
@ -52,6 +58,9 @@ equals = is '='
separator : Lexer separator : Lexer
separator = is ',' separator = is ','
dot : Lexer
dot = is '.'
rawTokens : TokenMap Token rawTokens : TokenMap Token
rawTokens = rawTokens =
[ (equals, const Equals) [ (equals, const Equals)
@ -59,8 +68,10 @@ rawTokens =
, (namespacedIdent, uncurry DotSepIdent . mkNamespacedIdent) , (namespacedIdent, uncurry DotSepIdent . mkNamespacedIdent)
, (identAllowDashes, DotSepIdent Nothing) , (identAllowDashes, DotSepIdent Nothing)
, (separator, const Separator) , (separator, const Separator)
, (dot, const Dot)
, (spacesOrNewlines, const Space) , (spacesOrNewlines, const Space)
, (stringLit, \s => StringLit (stripQuotes s)) , (stringLit, \s => StringLit (stripQuotes s))
, (intLit, \i => IntegerLit (cast i))
] ]
export export

View File

@ -47,6 +47,13 @@ stringLit = terminal "Expected string"
StringLit str => Just str StringLit str => Just str
_ => Nothing) _ => Nothing)
export
integerLit : Rule Integer
integerLit = terminal "Expected string"
(\x => case x.val of
IntegerLit i => Just i
_ => Nothing)
export export
namespacedIdent : Rule (Maybe Namespace, String) namespacedIdent : Rule (Maybe Namespace, String)
namespacedIdent = terminal "Expected namespaced identifier" namespacedIdent = terminal "Expected namespaced identifier"
@ -70,6 +77,13 @@ packageName = terminal "Expected package name"
else Nothing else Nothing
_ => Nothing) _ => Nothing)
export
dot' : Rule ()
dot' = terminal "Expected dot"
(\x => case x.val of
Dot => Just ()
_ => Nothing)
sep' : Rule () sep' : Rule ()
sep' = terminal "Expected separator" sep' = terminal "Expected separator"
(\x => case x.val of (\x => case x.val of