Version number constraints in 'depends' field

Now reporting an error if we can't find a package that satisfies the
constraints. The version number field can still be a string (as it used
to be) but will give a deprecation warning - and the old style version
string wasn't used anyway.

Version constraints can have an upper and/or lower bound, which can be
inclusive or not.
This commit is contained in:
Edwin Brady 2021-02-27 17:58:52 +00:00
parent 1052c41a3c
commit 1d8166b1b2
21 changed files with 226 additions and 19 deletions

View File

@ -9,8 +9,10 @@ REPL/IDE mode changes:
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.
a directory which includes its version number, and dependencies can have
version number ranges using `<=`, `<`, `>=`, `>`, `==` to express version
constraints. Version numbers must be in the form of integers, separated by
dots (e.g. `1.0`, `0.3.0`, `3.1.4.1.5` etc)
* Idris now looks in the current working directory, under a subdirectory
`depends` for local installations of packages before looking globally.

View File

@ -1,6 +1,6 @@
package idris2app
depends = idris2, network
depends = network
sourcedir = "src"

View File

@ -1,4 +1,5 @@
package idris2
version = 0.3.0
modules =
Algebra,

View File

@ -155,6 +155,7 @@ public export
data Warning : Type where
UnreachableClause : {vars : _} ->
FC -> Env Term vars -> Term vars -> Warning
Deprecated : String -> Warning
export
Show TTCErrorMsg where
@ -398,6 +399,7 @@ getErrorLoc (InRHS _ _ err) = getErrorLoc err
export
getWarningLoc : Warning -> Maybe FC
getWarningLoc (UnreachableClause fc _ _) = Just fc
getWarningLoc (Deprecated _) = Nothing
-- Core is a wrapper around IO that is specialised for efficiency.
export

View File

@ -63,23 +63,39 @@ public export
record PkgVersionBounds where
constructor MkPkgVersionBounds
lowerBound : Maybe PkgVersion
lowerInclusive : Bool -- >= if true
upperBound : Maybe PkgVersion
upperInclusive : Bool -- <= if true
export
Show PkgVersionBounds where
show p = if noBounds p then "any"
else maybe "" (\v => (if p.lowerInclusive then ">= " else "> ")
++ show v ++ " ") p.lowerBound ++
maybe "" (\v => (if p.upperInclusive then "<= " else "< ") ++ show v) p.upperBound
where
noBounds : PkgVersionBounds -> Bool
noBounds p = isNothing p.lowerBound && isNothing p.upperBound
export
anyBounds : PkgVersionBounds
anyBounds = MkPkgVersionBounds Nothing Nothing
anyBounds = MkPkgVersionBounds Nothing True Nothing True
export
current : PkgVersionBounds
current = let (maj,min,patch) = semVer version
version = Just (MkPkgVersion [maj, min, patch]) in
MkPkgVersionBounds version version
MkPkgVersionBounds version True version True
export
inBounds : PkgVersion -> PkgVersionBounds -> Bool
inBounds v bounds
= maybe True (\v' => v >= v') bounds.lowerBound &&
maybe True (\v' => v <= v') bounds.upperBound
= maybe True (\v' => if bounds.lowerInclusive
then v >= v'
else v > v') bounds.lowerBound &&
maybe True (\v' => if bounds.upperInclusive
then v <= v'
else v < v') bounds.upperBound
||| CLOpt - possible command line options
public export

View File

@ -455,6 +455,8 @@ pwarning (UnreachableClause fc env tm)
= pure $ errorDesc (reflow "Unreachable clause:"
<++> code !(pshow env tm))
<+> line <+> !(ploc fc)
pwarning (Deprecated s)
= pure $ pretty "Deprecation warning:" <++> pretty s
prettyMaybeLoc : Maybe FC -> Doc IdrisAnn
prettyMaybeLoc Nothing = emptyDoc

View File

@ -32,6 +32,7 @@ import Idris.CommandLine
import Idris.ModTree
import Idris.ProcessIdr
import Idris.REPL
import Idris.REPLCommon
import Idris.REPLOpts
import Idris.SetOptions
import Idris.Syntax
@ -40,6 +41,16 @@ import IdrisPaths
%default covering
public export
record Depends where
constructor MkDepends
pkgname : String
pkgbounds : PkgVersionBounds
export
Show Depends where
show p = p.pkgname ++ " " ++ show p.pkgbounds
public export
record PkgDesc where
constructor MkPkgDesc
@ -53,7 +64,7 @@ record PkgDesc where
homepage : Maybe String
sourceloc : Maybe String
bugtracker : Maybe String
depends : List String -- packages to add to search path
depends : List Depends -- packages to add to search path
modules : List (ModuleIdent, String) -- modules to install (namespace, filename)
mainmod : Maybe (ModuleIdent, String) -- main file (i.e. file to load at REPL)
executable : Maybe String -- name of executable
@ -108,6 +119,7 @@ initPkgDesc pname
data DescField : Type where
PVersion : FC -> PkgVersion -> DescField
PVersionDep : FC -> String -> DescField
PAuthors : FC -> String -> DescField
PMaintainers : FC -> String -> DescField
PLicense : FC -> String -> DescField
@ -116,7 +128,7 @@ data DescField : Type where
PHomePage : FC -> String -> DescField
PSourceLoc : FC -> String -> DescField
PBugTracker : FC -> String -> DescField
PDepends : List String -> DescField
PDepends : List Depends -> DescField
PModules : List (FC, ModuleIdent) -> DescField
PMainMod : FC -> ModuleIdent -> DescField
PExec : String -> DescField
@ -159,9 +171,15 @@ field fname
end <- location
pure (PVersion (MkFC fname start end)
(MkPkgVersion (map fromInteger vs)))
<|> do start <- location
ignore $ exactProperty "version"
equals
v <- stringLit
end <- location
pure (PVersionDep (MkFC fname start end) v)
<|> do ignore $ exactProperty "depends"
equals
ds <- sep packageName
ds <- sep depends
pure (PDepends ds)
<|> do ignore $ exactProperty "modules"
equals
@ -181,6 +199,47 @@ field fname
e <- (stringLit <|> packageName)
pure (PExec e)
where
data Bound = LT PkgVersion Bool | GT PkgVersion Bool
bound : Rule (List Bound)
bound
= do lte
vs <- sepBy1 dot' integerLit
pure [LT (MkPkgVersion (map fromInteger vs)) True]
<|> do gte
vs <- sepBy1 dot' integerLit
pure [GT (MkPkgVersion (map fromInteger vs)) True]
<|> do lt
vs <- sepBy1 dot' integerLit
pure [LT (MkPkgVersion (map fromInteger vs)) False]
<|> do gt
vs <- sepBy1 dot' integerLit
pure [GT (MkPkgVersion (map fromInteger vs)) False]
<|> do eqop
vs <- sepBy1 dot' integerLit
pure [LT (MkPkgVersion (map fromInteger vs)) True,
GT (MkPkgVersion (map fromInteger vs)) True]
mkBound : List Bound -> PkgVersionBounds -> PackageEmptyRule PkgVersionBounds
mkBound (LT b i :: bs) pkgbs
= maybe (mkBound bs (record { upperBound = Just b,
upperInclusive = i } pkgbs))
(\_ => fail "Dependency already has an upper bound")
pkgbs.upperBound
mkBound (GT b i :: bs) pkgbs
= maybe (mkBound bs (record { lowerBound = Just b,
lowerInclusive = i } pkgbs))
(\_ => fail "Dependency already has a lower bound")
pkgbs.lowerBound
mkBound [] pkgbs = pure pkgbs
depends : Rule Depends
depends
= do name <- packageName
bs <- sepBy andop bound
pure (MkDepends name !(mkBound (concat bs) anyBounds))
strField : (FC -> String -> DescField) -> String -> Rule DescField
strField fieldConstructor fieldName
= do start <- location
@ -202,10 +261,15 @@ data ParsedMods : Type where
data MainMod : Type where
addField : {auto c : Ref Ctxt Defs} ->
{auto s : Ref Syn SyntaxInfo} ->
{auto o : Ref ROpts REPLOpts} ->
{auto p : Ref ParsedMods (List (FC, ModuleIdent))} ->
{auto m : Ref MainMod (Maybe (FC, ModuleIdent))} ->
DescField -> PkgDesc -> Core PkgDesc
addField (PVersion fc n) pkg = pure $ record { version = n } pkg
addField (PVersionDep fc n) pkg
= do emitWarning (Deprecated "version numbers must now be of the form x.y.z")
pure pkg
addField (PAuthors fc a) pkg = pure $ record { authors = a } pkg
addField (PMaintainers fc a) pkg = pure $ record { maintainers = Just a } pkg
addField (PLicense fc a) pkg = pure $ record { license = Just a } pkg
@ -234,6 +298,8 @@ addField (PPreclean fc e) pkg = pure $ record { preclean = Just (fc, e) } pkg
addField (PPostclean fc e) pkg = pure $ record { postclean = Just (fc, e) } pkg
addFields : {auto c : Ref Ctxt Defs} ->
{auto s : Ref Syn SyntaxInfo} ->
{auto o : Ref ROpts REPLOpts} ->
List DescField -> PkgDesc -> Core PkgDesc
addFields xs desc = do p <- newRef ParsedMods []
m <- newRef MainMod Nothing
@ -264,7 +330,7 @@ addDeps : {auto c : Ref Ctxt Defs} ->
PkgDesc -> Core ()
addDeps pkg
= do defs <- get Ctxt
traverse_ (\p => addPkgDir p anyBounds) (depends pkg)
traverse_ (\p => addPkgDir p.pkgname p.pkgbounds) (depends pkg)
processOptions : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} ->
@ -505,6 +571,8 @@ runRepl fname = do
export
parsePkgFile : {auto c : Ref Ctxt Defs} ->
{auto s : Ref Syn SyntaxInfo} ->
{auto o : Ref ROpts REPLOpts} ->
String -> Core PkgDesc
parsePkgFile file = do
Right (pname, fs) <- coreLift $ parseFile file
@ -621,6 +689,7 @@ processPackageOpts opts
export
findIpkg : {auto c : Ref Ctxt Defs} ->
{auto r : Ref ROpts REPLOpts} ->
{auto s : Ref Syn SyntaxInfo} ->
Maybe String -> Core (Maybe String)
findIpkg fname
= do Just (dir, ipkgn, up) <- coreLift findIpkgFile
@ -650,5 +719,6 @@ findIpkg fname
dropHead str [] = []
dropHead str (x :: xs)
= if x == str then xs else x :: xs
loadDependencies : List String -> Core ()
loadDependencies = traverse_ (\p => addPkgDir p anyBounds)
loadDependencies : List Depends -> Core ()
loadDependencies = traverse_ (\p => addPkgDir p.pkgname p.pkgbounds)

View File

@ -65,7 +65,6 @@ candidateDirs dname pkg bounds
then getFiles d (((dname </> str), ver) :: acc)
else getFiles d acc
-- TODO: Version numbers on dependencies
export
addPkgDir : {auto c : Ref Ctxt Defs} ->
String -> PkgVersionBounds -> Core ()
@ -94,7 +93,7 @@ addPkgDir p bounds
-- 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)
[] => throw (CantFindPackage (p ++ " (" ++ show bounds ++ ")"))
((p, _) :: ps) => addExtraDir p
dirOption : Dirs -> DirCommand -> Core ()

View File

@ -24,6 +24,12 @@ data Token
| DotSepIdent (Maybe Namespace) String
| Separator
| Dot
| LTE
| GTE
| LT
| GT
| EqOp
| AndOp
| Space
| StringLit String
| IntegerLit Integer
@ -36,6 +42,12 @@ Show Token where
show (DotSepIdent ns n) = "DotSepIdentifier: " ++ show ns ++ "." ++ show n
show Separator = "Separator"
show Dot = "Dot"
show LTE = "LTE"
show GTE = "GTE"
show LT = "LT"
show GT = "GT"
show EqOp = "EqOp"
show AndOp = "AndOp"
show Space = "Space"
show (StringLit s) = "StringLit: " ++ s
show (IntegerLit i) = "IntegerLit: " ++ show i
@ -48,6 +60,12 @@ Pretty Token where
pretty (DotSepIdent ns n) = "DotSepIdentifier:" <++> pretty ns <+> dot <+> pretty n
pretty Separator = "Separator"
pretty Dot = "Dot"
pretty LTE = "LTE"
pretty GTE = "GTE"
pretty LT = "LT"
pretty GT = "GT"
pretty EqOp = "EqOp"
pretty AndOp = "AndOp"
pretty Space = "Space"
pretty (StringLit s) = "StringLit:" <++> pretty s
pretty (IntegerLit i) = "IntegerLit:" <++> pretty i
@ -61,14 +79,38 @@ separator = is ','
dot : Lexer
dot = is '.'
lte : Lexer
lte = is '<' <+> is '='
gte : Lexer
gte = is '>' <+> is '='
lt : Lexer
lt = is '<'
gt : Lexer
gt = is '>'
eqop : Lexer
eqop = is '=' <+> is '='
andop : Lexer
andop = is '&' <+> is '&'
rawTokens : TokenMap Token
rawTokens =
[ (equals, const Equals)
, (comment, Comment . drop 2)
[ (comment, Comment . drop 2)
, (namespacedIdent, uncurry DotSepIdent . mkNamespacedIdent)
, (identAllowDashes, DotSepIdent Nothing)
, (separator, const Separator)
, (dot, const Dot)
, (lte, const LTE)
, (gte, const GTE)
, (lt, const LT)
, (gt, const GT)
, (eqop, const EqOp)
, (andop, const AndOp)
, (equals, const Equals)
, (spacesOrNewlines, const Space)
, (stringLit, \s => StringLit (stripQuotes s))
, (intLit, \i => IntegerLit (cast i))

View File

@ -24,6 +24,48 @@ equals = terminal "Expected equals"
Equals => Just ()
_ => Nothing)
export
lte : Rule ()
lte = terminal "Expected <="
(\x => case x.val of
LTE => Just ()
_ => Nothing)
export
gte : Rule ()
gte = terminal "Expected >="
(\x => case x.val of
GTE => Just ()
_ => Nothing)
export
lt : Rule ()
lt = terminal "Expected <="
(\x => case x.val of
LT => Just ()
_ => Nothing)
export
gt : Rule ()
gt = terminal "Expected >="
(\x => case x.val of
GT => Just ()
_ => Nothing)
export
eqop : Rule ()
eqop = terminal "Expected =="
(\x => case x.val of
EqOp => Just ()
_ => Nothing)
export
andop : Rule ()
andop = terminal "Expected &&"
(\x => case x.val of
AndOp => Just ()
_ => Nothing)
export
eoi : Rule ()
eoi = terminal "Expected end of input"

View File

@ -26,7 +26,7 @@ fromLexError : String -> (StopReason, Int, Int, String) -> Error
fromLexError fname (ComposeNotClosing begin end, _, _, _)
= LexFail (MkFC fname begin end) "Bracket is not properly closed."
fromLexError fname (_, l, c, _)
= LexFail (MkFC fname (l, c) (l, c + 1)) "Can't recognoise token."
= LexFail (MkFC fname (l, c) (l, c + 1)) "Can't recognise token."
export
fromParsingError : (Show token, Pretty token) => String -> ParsingError token -> Error

View File

@ -137,7 +137,7 @@ idrisTests = MkTestPool []
-- Parameters blocks
"params001",
-- Packages and ipkg files
"pkg001", "pkg002", "pkg003", "pkg004", "pkg005",
"pkg001", "pkg002", "pkg003", "pkg004", "pkg005", "pkg006",
-- Positivity checking
"positivity001", "positivity002", "positivity003",
-- Larger programs arising from real usage. Typically things with

View File

@ -0,0 +1,2 @@
package bar
version = 1.0.1

View File

@ -0,0 +1,2 @@
package foo
version = 0.5

View File

@ -0,0 +1,3 @@
Uncaught error: Can't find package foo (>= 0.4 < 0.5)
Uncaught error: Can't find package baz (any)
Warning: Deprecation warning: version numbers must now be of the form x.y.z

7
tests/idris2/pkg006/run Executable file
View File

@ -0,0 +1,7 @@
$1 --build test1.ipkg
$1 --build test2.ipkg
$1 --build test3.ipkg
$1 --build test4.ipkg
$1 --build test5.ipkg
rm -rf build

View File

@ -0,0 +1,3 @@
package test1
depends = foo, bar

View File

@ -0,0 +1,4 @@
package test2
depends = foo >= 0.5 && < 0.6,
bar >= 1

View File

@ -0,0 +1,4 @@
package test3
depends = foo >= 0.4 && < 0.5,
bar >= 1

View File

@ -0,0 +1,4 @@
package test4
depends = baz,
bar >= 1

View File

@ -0,0 +1,2 @@
package test5
version = "this one"