mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-21 02:31:50 +03:00
Add an option that dumps package details to JSON (#3293)
* Initial stab at package json dump * expose ipkg json dump as new option * make dependency output easier to ingest by another tool * Add a test for ipkg json dump * cleanup * maybe just don't collide with existing equally good fixity * make new operator private * Add new module to api ipkg file * Add note to CHANGELOG_NEXT * correct the docs for the dump-ipkg-json command
This commit is contained in:
parent
0174618724
commit
88a5328ec1
@ -24,6 +24,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
|
|||||||
* The Nix flake now exposes the Idris2 API package as `idris2Api` and Idris2's
|
* The Nix flake now exposes the Idris2 API package as `idris2Api` and Idris2's
|
||||||
C support library as `support`.
|
C support library as `support`.
|
||||||
|
|
||||||
|
* A new `idris2 --dump-ipkg-json` option (requires either `--find-ipkg` or
|
||||||
|
specifying the `.ipkg` file) dumps JSON information about an Idris package.
|
||||||
|
|
||||||
### Language changes
|
### Language changes
|
||||||
|
|
||||||
* Autobind and Typebind modifier on operators allow the user to
|
* Autobind and Typebind modifier on operators allow the user to
|
||||||
|
@ -128,6 +128,7 @@ modules =
|
|||||||
|
|
||||||
Idris.Package,
|
Idris.Package,
|
||||||
Idris.Package.Init,
|
Idris.Package.Init,
|
||||||
|
Idris.Package.ToJson,
|
||||||
Idris.Package.Types,
|
Idris.Package.Types,
|
||||||
|
|
||||||
Idris.Parser,
|
Idris.Parser,
|
||||||
|
@ -24,6 +24,7 @@ data PkgCommand
|
|||||||
| Clean
|
| Clean
|
||||||
| REPL
|
| REPL
|
||||||
| Init
|
| Init
|
||||||
|
| DumpJson
|
||||||
|
|
||||||
export
|
export
|
||||||
Show PkgCommand where
|
Show PkgCommand where
|
||||||
@ -35,6 +36,7 @@ Show PkgCommand where
|
|||||||
show Clean = "--clean"
|
show Clean = "--clean"
|
||||||
show REPL = "--repl"
|
show REPL = "--repl"
|
||||||
show Init = "--init"
|
show Init = "--init"
|
||||||
|
show DumpJson = "--dump-ipkg-json"
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data DirCommand
|
data DirCommand
|
||||||
@ -279,6 +281,9 @@ options = [MkOpt ["--check", "-c"] [] [CheckOnly]
|
|||||||
MkOpt ["--init"] [Optional "package file"]
|
MkOpt ["--init"] [Optional "package file"]
|
||||||
(\ f => [Package Init f])
|
(\ f => [Package Init f])
|
||||||
(Just "Interactively initialise a new project"),
|
(Just "Interactively initialise a new project"),
|
||||||
|
MkOpt ["--dump-ipkg-json"] [Optional "package file"]
|
||||||
|
(\ f => [Package DumpJson f])
|
||||||
|
(Just "Dump an Idris2 package file in the JSON format"),
|
||||||
MkOpt ["--build"] [Optional "package file"]
|
MkOpt ["--build"] [Optional "package file"]
|
||||||
(\f => [Package Build f])
|
(\f => [Package Build f])
|
||||||
(Just "Build modules/executable for the given package"),
|
(Just "Build modules/executable for the given package"),
|
||||||
|
@ -47,6 +47,7 @@ import Idris.Version
|
|||||||
|
|
||||||
import public Idris.Package.Types
|
import public Idris.Package.Types
|
||||||
import Idris.Package.Init
|
import Idris.Package.Init
|
||||||
|
import Idris.Package.ToJson
|
||||||
|
|
||||||
%default covering
|
%default covering
|
||||||
|
|
||||||
@ -943,6 +944,7 @@ processPackage opts (cmd, mfile)
|
|||||||
| errs => coreLift (exitWith (ExitFailure 1))
|
| errs => coreLift (exitWith (ExitFailure 1))
|
||||||
runRepl (map snd $ mainmod pkg)
|
runRepl (map snd $ mainmod pkg)
|
||||||
Init => pure () -- already handled earlier
|
Init => pure () -- already handled earlier
|
||||||
|
DumpJson => coreLift . putStrLn $ toJson pkg
|
||||||
|
|
||||||
record PackageOpts where
|
record PackageOpts where
|
||||||
constructor MkPFR
|
constructor MkPFR
|
||||||
|
116
src/Idris/Package/ToJson.idr
Normal file
116
src/Idris/Package/ToJson.idr
Normal file
@ -0,0 +1,116 @@
|
|||||||
|
module Idris.Package.ToJson
|
||||||
|
|
||||||
|
import Idris.Package.Types
|
||||||
|
import Libraries.Data.String.Extra
|
||||||
|
import Data.List
|
||||||
|
import Core.FC
|
||||||
|
import Core.Name.Namespace
|
||||||
|
import Data.Maybe
|
||||||
|
|
||||||
|
%default total
|
||||||
|
|
||||||
|
-- We don't bother with a robust JSON implementation
|
||||||
|
-- for this one-off JSON serialization.
|
||||||
|
|
||||||
|
private infixl 0 ~~=
|
||||||
|
|
||||||
|
interface ToJson v where
|
||||||
|
toJson : v -> String
|
||||||
|
|
||||||
|
ToJson String where
|
||||||
|
toJson str = "\"\{str}\""
|
||||||
|
|
||||||
|
ToJson Bool where
|
||||||
|
toJson False = "false"
|
||||||
|
toJson True = "true"
|
||||||
|
|
||||||
|
ToJson a => ToJson (List a) where
|
||||||
|
toJson xs = "[\{join "," $ map toJson xs}]"
|
||||||
|
|
||||||
|
(~=) : ToJson v => String -> v -> String
|
||||||
|
(~=) field value = "\{toJson field}: \{toJson value}"
|
||||||
|
|
||||||
|
(~~=) : ToJson v => String -> Maybe v -> Maybe String
|
||||||
|
(~~=) field = map (field ~=)
|
||||||
|
|
||||||
|
ToJson PkgVersion where
|
||||||
|
toJson v = "\"\{show v}\""
|
||||||
|
|
||||||
|
ToJson PkgVersionBounds where
|
||||||
|
toJson (MkPkgVersionBounds lowerBound lowerInclusive upperBound upperInclusive) =
|
||||||
|
let fields =
|
||||||
|
[ "lowerInclusive" ~= lowerInclusive
|
||||||
|
, "lowerBound" ~= maybe "*" show lowerBound
|
||||||
|
, "upperInclusive" ~= upperInclusive
|
||||||
|
, "upperBound" ~= maybe "*" show upperBound
|
||||||
|
]
|
||||||
|
in
|
||||||
|
"{\{join "," fields}}"
|
||||||
|
|
||||||
|
ToJson Depends where
|
||||||
|
toJson (MkDepends pkgname pkgbounds) = "{\{pkgname ~= pkgbounds}}"
|
||||||
|
|
||||||
|
ToJson ModuleIdent where
|
||||||
|
toJson ident = "\"\{show ident}\""
|
||||||
|
|
||||||
|
namespace Package
|
||||||
|
export
|
||||||
|
toJson : PkgDesc -> String
|
||||||
|
toJson (MkPkgDesc
|
||||||
|
name
|
||||||
|
version
|
||||||
|
langversion
|
||||||
|
authors
|
||||||
|
maintainers
|
||||||
|
license
|
||||||
|
brief
|
||||||
|
readme
|
||||||
|
homepage
|
||||||
|
sourceloc
|
||||||
|
bugtracker
|
||||||
|
depends
|
||||||
|
modules
|
||||||
|
mainmod
|
||||||
|
executable
|
||||||
|
options
|
||||||
|
sourcedir
|
||||||
|
builddir
|
||||||
|
outputdir
|
||||||
|
prebuild
|
||||||
|
postbuild
|
||||||
|
preinstall
|
||||||
|
postinstall
|
||||||
|
preclean
|
||||||
|
postclean) =
|
||||||
|
let optionalFields = catMaybes $
|
||||||
|
[ "version" ~~= version
|
||||||
|
, "langversion" ~~= langversion
|
||||||
|
, "authors" ~~= authors
|
||||||
|
, "maintainers" ~~= maintainers
|
||||||
|
, "license" ~~= license
|
||||||
|
, "brief" ~~= brief
|
||||||
|
, "readme" ~~= readme
|
||||||
|
, "homepage" ~~= homepage
|
||||||
|
, "sourceloc" ~~= sourceloc
|
||||||
|
, "bugtracker" ~~= bugtracker
|
||||||
|
, "main" ~~= fst <$> mainmod
|
||||||
|
, "executable" ~~= executable
|
||||||
|
, "opts" ~~= snd <$> options
|
||||||
|
, "sourcedir" ~~= sourcedir
|
||||||
|
, "builddir" ~~= builddir
|
||||||
|
, "outputdir" ~~= outputdir
|
||||||
|
, "prebuild" ~~= snd <$> prebuild
|
||||||
|
, "postbuild" ~~= snd <$> postbuild
|
||||||
|
, "preinstall" ~~= snd <$> preinstall
|
||||||
|
, "postinstall" ~~= snd <$> postinstall
|
||||||
|
, "preclean" ~~= snd <$> preclean
|
||||||
|
, "postclean" ~~= snd <$> postclean
|
||||||
|
]
|
||||||
|
fields =
|
||||||
|
[ "name" ~= name
|
||||||
|
, "depends" ~= depends
|
||||||
|
, "modules" ~= fst <$> modules
|
||||||
|
] ++ optionalFields
|
||||||
|
in
|
||||||
|
"{\{join "," fields}}"
|
||||||
|
|
@ -233,6 +233,7 @@ opts "--" "--output-dir" = pure []
|
|||||||
|
|
||||||
-- with package files
|
-- with package files
|
||||||
opts x "--build" = prefixOnlyIfNonEmpty x <$> findIpkg
|
opts x "--build" = prefixOnlyIfNonEmpty x <$> findIpkg
|
||||||
|
opts x "--dump-ipkg-json" = prefixOnlyIfNonEmpty x <$> findIpkg
|
||||||
opts x "--install" = prefixOnlyIfNonEmpty x <$> findIpkg
|
opts x "--install" = prefixOnlyIfNonEmpty x <$> findIpkg
|
||||||
opts x "--mkdoc" = prefixOnlyIfNonEmpty x <$> findIpkg
|
opts x "--mkdoc" = prefixOnlyIfNonEmpty x <$> findIpkg
|
||||||
opts x "--typecheck" = prefixOnlyIfNonEmpty x <$> findIpkg
|
opts x "--typecheck" = prefixOnlyIfNonEmpty x <$> findIpkg
|
||||||
|
1
tests/idris2/pkg/pkg019/expected
Normal file
1
tests/idris2/pkg/pkg019/expected
Normal file
@ -0,0 +1 @@
|
|||||||
|
{"name": "test","depends": [{"contrib": {"lowerInclusive": true,"lowerBound": "*","upperInclusive": true,"upperBound": "*"}},{"idris2": {"lowerInclusive": true,"lowerBound": "0.7.0","upperInclusive": true,"upperBound": "*"}}],"modules": ["And.A.Proof","Yet.Another.Path"],"version": "0.1.0","langversion": {"lowerInclusive": true,"lowerBound": "0.7.0","upperInclusive": true,"upperBound": "*"},"authors": "Average Joe","maintainers": "Average Joe","license": "MIT","brief": "A basic test","readme": "https://github.com/idris-lang/Idris2/README.md","homepage": "https://github.com/idris-lang/Idris2","sourceloc": "https://github.com/idris-lang/Idris2","bugtracker": "https://github.com/idris-lang/Idris2/issues","main": "Main","executable": "test-exec","opts": "--codegen node","sourcedir": "src","builddir": "build","outputdir": "out","prebuild": "make pre","postbuild": "make post","preinstall": "make pre-install","postinstall": "make post-install","preclean": "make pre-clean","postclean": "make post-clean"}
|
4
tests/idris2/pkg/pkg019/run
Executable file
4
tests/idris2/pkg/pkg019/run
Executable file
@ -0,0 +1,4 @@
|
|||||||
|
. ../../../testutils.sh
|
||||||
|
|
||||||
|
idris2 --dump-ipkg-json test.ipkg
|
||||||
|
|
8
tests/idris2/pkg/pkg019/src/And/A/Proof.idr
Normal file
8
tests/idris2/pkg/pkg019/src/And/A/Proof.idr
Normal file
@ -0,0 +1,8 @@
|
|||||||
|
module And.A.Proof
|
||||||
|
|
||||||
|
import Yet.Another.Path as Val
|
||||||
|
|
||||||
|
%default total
|
||||||
|
|
||||||
|
equality : Val.val === 2+3
|
||||||
|
equality = Refl
|
4
tests/idris2/pkg/pkg019/src/Main.idr
Normal file
4
tests/idris2/pkg/pkg019/src/Main.idr
Normal file
@ -0,0 +1,4 @@
|
|||||||
|
module Main
|
||||||
|
|
||||||
|
main : IO ()
|
||||||
|
main = pure ()
|
7
tests/idris2/pkg/pkg019/src/Yet/Another/Path.idr
Normal file
7
tests/idris2/pkg/pkg019/src/Yet/Another/Path.idr
Normal file
@ -0,0 +1,7 @@
|
|||||||
|
module Yet.Another.Path
|
||||||
|
|
||||||
|
%default total
|
||||||
|
|
||||||
|
public export
|
||||||
|
val : Nat
|
||||||
|
val = 5
|
49
tests/idris2/pkg/pkg019/test.ipkg
Normal file
49
tests/idris2/pkg/pkg019/test.ipkg
Normal file
@ -0,0 +1,49 @@
|
|||||||
|
package test
|
||||||
|
version = 0.1.0
|
||||||
|
authors = "Average Joe"
|
||||||
|
maintainers = "Average Joe"
|
||||||
|
license = "MIT"
|
||||||
|
brief = "A basic test"
|
||||||
|
readme = "https://github.com/idris-lang/Idris2/README.md"
|
||||||
|
homepage = "https://github.com/idris-lang/Idris2"
|
||||||
|
sourceloc = "https://github.com/idris-lang/Idris2"
|
||||||
|
bugtracker = "https://github.com/idris-lang/Idris2/issues"
|
||||||
|
|
||||||
|
-- the Idris2 version required (e.g. langversion >= 0.5.1)
|
||||||
|
langversion >= 0.7.0
|
||||||
|
|
||||||
|
-- packages to add to search path
|
||||||
|
depends = contrib,
|
||||||
|
idris2 >= 0.7.0
|
||||||
|
|
||||||
|
-- modules to install
|
||||||
|
modules = And.A.Proof,
|
||||||
|
Yet.Another.Path
|
||||||
|
|
||||||
|
-- main file (i.e. file to load at REPL)
|
||||||
|
main = Main
|
||||||
|
|
||||||
|
-- name of executable
|
||||||
|
executable = test-exec
|
||||||
|
opts = "--codegen node"
|
||||||
|
sourcedir = "src"
|
||||||
|
builddir = "build"
|
||||||
|
outputdir = "out"
|
||||||
|
|
||||||
|
-- script to run before building
|
||||||
|
prebuild = "make pre"
|
||||||
|
|
||||||
|
-- script to run after building
|
||||||
|
postbuild = "make post"
|
||||||
|
|
||||||
|
-- script to run after building, before installing
|
||||||
|
preinstall = "make pre-install"
|
||||||
|
|
||||||
|
-- script to run after installing
|
||||||
|
postinstall = "make post-install"
|
||||||
|
|
||||||
|
-- script to run before cleaning
|
||||||
|
preclean = "make pre-clean"
|
||||||
|
|
||||||
|
-- script to run after cleaning
|
||||||
|
postclean = "make post-clean"
|
Loading…
Reference in New Issue
Block a user