mirror of
https://github.com/edwinb/Idris2-boot.git
synced 2024-11-28 14:06:26 +03:00
Merge branch 'master' into configurable-ide-mode-socket
This commit is contained in:
commit
3dcd48e08b
@ -21,3 +21,4 @@ This is a placeholder, to get set up with readthedocs.
|
||||
:maxdepth: 1
|
||||
|
||||
faq/faq
|
||||
reference/index
|
||||
|
19
docs/reference/index.rst
Normal file
19
docs/reference/index.rst
Normal file
@ -0,0 +1,19 @@
|
||||
**********************
|
||||
Idris2 Reference Guide
|
||||
**********************
|
||||
|
||||
.. note::
|
||||
|
||||
The documentation for Idris 2 has been published under the Creative
|
||||
Commons CC0 License. As such to the extent possible under law, *The
|
||||
Idris Community* has waived all copyright and related or neighboring
|
||||
rights to Documentation for Idris.
|
||||
|
||||
More information concerning the CC0 can be found online at: http://creativecommons.org/publicdomain/zero/1.0/
|
||||
|
||||
This is a placeholder, to get set up with readthedocs.
|
||||
|
||||
.. toctree::
|
||||
:maxdepth: 1
|
||||
|
||||
packages
|
102
docs/reference/packages.rst
Normal file
102
docs/reference/packages.rst
Normal file
@ -0,0 +1,102 @@
|
||||
.. _ref-sect-packages:
|
||||
|
||||
********
|
||||
Packages
|
||||
********
|
||||
|
||||
Idris includes a simple system for building packages from a package
|
||||
description file. These files can be used with the Idris compiler to
|
||||
manage the development process of your Idris programmes and packages.
|
||||
|
||||
Package Descriptions
|
||||
====================
|
||||
|
||||
A package description includes the following:
|
||||
|
||||
+ A header, consisting of the keyword package followed by the package
|
||||
name. Package names can be any valid Idris identifier. The iPKG
|
||||
format also takes a quoted version that accepts any valid filename.
|
||||
+ Fields describing package contents, ``<field> = <value>``
|
||||
|
||||
At least one field must be the modules field, where the value is a
|
||||
comma separated list of modules. For example, a library test which
|
||||
has two modules ``Foo.idr`` and ``Bar.idr`` as source files would be
|
||||
written as follows::
|
||||
|
||||
package test
|
||||
|
||||
modules = Foo, Bar
|
||||
|
||||
Other examples of package files can be found in the ``libs`` directory
|
||||
of the main Idris repository, and in `third-party libraries <https://github.com/idris-lang/Idris-dev/wiki/Libraries>`_.
|
||||
|
||||
Metadata
|
||||
--------
|
||||
|
||||
The `iPKG` format supports additional metadata associated with the package.
|
||||
The added fields are:
|
||||
|
||||
+ ``brief = "<text>"``, a string literal containing a brief description
|
||||
of the package.
|
||||
|
||||
+ ``version = "<text>""``, a version string to associate with the package.
|
||||
|
||||
+ ``readme = "<file>""``, location of the README file.
|
||||
|
||||
+ ``license = "<text>"``, a string description of the licensing
|
||||
information.
|
||||
|
||||
+ ``authors = "<text>"``, the author information.
|
||||
|
||||
+ ``maintainers = "<text>"``, Maintainer information.
|
||||
|
||||
+ ``homepage = "<url>"``, the website associated with the package.
|
||||
|
||||
+ ``sourceloc = "<url>"``, the location of the DVCS where the source
|
||||
can be found.
|
||||
|
||||
+ ``bugtracker = "<url>"``, the location of the project's bug tracker.
|
||||
|
||||
|
||||
Common Fields
|
||||
-------------
|
||||
|
||||
Other common fields which may be present in an ``ipkg`` file are:
|
||||
|
||||
+ ``executable = <output>``, which takes the name of the executable
|
||||
file to generate. Executable names can be any valid Idris
|
||||
identifier. the iPKG format also takes a quoted version that accepts
|
||||
any valid filename.
|
||||
|
||||
+ ``main = <module>``, which takes the name of the main module, and
|
||||
must be present if the executable field is present.
|
||||
|
||||
+ ``opts = "<idris options>"``, which allows options to be passed to
|
||||
Idris.
|
||||
|
||||
+ ``depends = <pkg name> (',' <pkg name>)+``, a comma separated list of
|
||||
package names that the Idris package requires.
|
||||
|
||||
|
||||
Comments
|
||||
---------
|
||||
|
||||
Package files support comments using the standard Idris singleline ``--`` and multiline ``{- -}`` format.
|
||||
|
||||
Using Package files
|
||||
===================
|
||||
|
||||
Given an Idris package file ``test.ipkg`` it can be used with the Idris compiler as follows:
|
||||
|
||||
+ ``idris2 --build test.ipkg`` will build all modules in the package
|
||||
|
||||
+ ``idris2 --install test.ipkg`` will install the package, making it
|
||||
accessible by other Idris libraries and programs.
|
||||
|
||||
+ ``idris2 --clean test.ipkg`` will clean the intermediate build files.
|
||||
|
||||
Once the test package has been installed, the command line option
|
||||
``--package test`` makes it accessible (abbreviated to ``-p test``).
|
||||
For example::
|
||||
|
||||
idris -p test Main.idr
|
@ -29,6 +29,7 @@ modules =
|
||||
Core.Name,
|
||||
Core.Normalise,
|
||||
Core.Options,
|
||||
Core.Reflect,
|
||||
Core.Termination,
|
||||
Core.TT,
|
||||
Core.TTC,
|
||||
|
77
libs/base/Language/Reflection.idr
Normal file
77
libs/base/Language/Reflection.idr
Normal file
@ -0,0 +1,77 @@
|
||||
module Language.Reflection
|
||||
|
||||
public export
|
||||
FilePos : Type
|
||||
FilePos = (Int, Int)
|
||||
|
||||
public export
|
||||
data FC : Type where
|
||||
MkFC : String -> FilePos -> FilePos -> FC
|
||||
EmptyFC : FC
|
||||
|
||||
public export
|
||||
emptyFC : FC
|
||||
emptyFC = MkFC "(empty)" (0, 0) (0, 0)
|
||||
|
||||
public export
|
||||
data NameType : Type where
|
||||
Bound : NameType
|
||||
Func : NameType
|
||||
DataCon : (tag : Int) -> (arity : Nat) -> NameType
|
||||
TyCon : (tag : Int) -> (arity : Nat) -> NameType
|
||||
|
||||
public export
|
||||
data Constant
|
||||
= I Int
|
||||
| BI Integer
|
||||
| Str String
|
||||
| Ch Char
|
||||
| Db Double
|
||||
| WorldVal
|
||||
|
||||
| IntType
|
||||
| IntegerType
|
||||
| StringType
|
||||
| CharType
|
||||
| DoubleType
|
||||
| WorldType
|
||||
|
||||
public export
|
||||
data Name = UN String
|
||||
| MN String Int
|
||||
| NS (List String) Name
|
||||
|
||||
public export
|
||||
data Count = M0 | M1 | MW
|
||||
|
||||
public export
|
||||
data PiInfo = ImplicitArg | ExplicitArg | AutoImplicit
|
||||
|
||||
public export
|
||||
data IsVar : Name -> Nat -> List Name -> Type where
|
||||
First : IsVar n Z (n :: ns)
|
||||
Later : IsVar n i ns -> IsVar n (S i) (m :: ns)
|
||||
|
||||
public export
|
||||
data LazyReason = LInf | LLazy | LUnknown
|
||||
|
||||
-- Type checked terms in the core TT
|
||||
public export
|
||||
data TT : List Name -> Type where
|
||||
Local : FC -> (idx : Nat) -> (n : Name) ->
|
||||
(0 prf : IsVar name idx vars) -> TT vars
|
||||
Ref : FC -> NameType -> Name -> TT vars
|
||||
Pi : FC -> Count -> PiInfo ->
|
||||
(x : Name) -> (argTy : TT vars) -> (retTy : TT (x :: vars)) ->
|
||||
TT vars
|
||||
Lam : FC -> Count -> PiInfo ->
|
||||
(x : Name) -> (argTy : TT vars) -> (scope : TT (x :: vars)) ->
|
||||
TT vars
|
||||
App : FC -> TT vars -> TT vars -> TT vars
|
||||
TDelayed : FC -> LazyReason -> TT vars -> TT vars
|
||||
TDelay : FC -> LazyReason -> (ty : TT vars) -> (arg : TT vars) -> TT vars
|
||||
TForce : FC -> TT vars -> TT vars
|
||||
PrimVal : FC -> Constant -> TT vars
|
||||
Erased : FC -> TT vars
|
||||
TType : FC -> TT vars
|
||||
|
@ -24,6 +24,8 @@ modules = Control.Monad.Identity,
|
||||
|
||||
Decidable.Equality,
|
||||
|
||||
Language.Reflection,
|
||||
|
||||
System,
|
||||
System.Concurrency.Raw,
|
||||
System.File,
|
||||
|
8
libs/network/.gitignore
vendored
Normal file
8
libs/network/.gitignore
vendored
Normal file
@ -0,0 +1,8 @@
|
||||
|
||||
network-tests
|
||||
|
||||
*.o
|
||||
*.so
|
||||
*.dylib
|
||||
*.dll
|
||||
|
@ -33,7 +33,13 @@ runServer = do
|
||||
putStrLn ("Received: " ++ str)
|
||||
Right n <- send s ("echo: " ++ str)
|
||||
| Left err => putStrLn ("Server failed to send data with error: " ++ show err)
|
||||
putStrLn ("Server sent " ++ show n ++ " bytes")
|
||||
-- This might be printed either before or after the client prints
|
||||
-- what it's received, and I think there's enough to check it's
|
||||
-- working without this message so I've removed it. If you disagree,
|
||||
-- please put it back, but also please make sure it's synchronised
|
||||
-- such that the messages are always printed in the same order. - EB
|
||||
-- putStrLn ("Server sent " ++ show n ++ " bytes")
|
||||
pure ()
|
||||
|
||||
runClient : Port -> IO ()
|
||||
runClient serverPort = do
|
||||
|
@ -1,4 +1,3 @@
|
||||
Client sent 12 bytes
|
||||
Received: hello world!
|
||||
Server sent 18 bytes
|
||||
Received: echo: hello world!
|
||||
|
@ -3,7 +3,7 @@
|
||||
${IDRIS2} --exec main Echo.idr > build/output
|
||||
|
||||
if ! [ -z "$(diff build/output expected)" ]; then
|
||||
echo "TEST FAILURE: unexpected build/output"
|
||||
echo "TEST FAILURE: unexpected build/output: "
|
||||
cat build/output
|
||||
exit 2
|
||||
else
|
||||
|
@ -564,7 +564,7 @@ data Nat = Z | S Nat
|
||||
public export
|
||||
integerToNat : Integer -> Nat
|
||||
integerToNat x
|
||||
= if intToBool (prim__eq_Integer x 0)
|
||||
= if intToBool (prim__lte_Integer x 0)
|
||||
then Z
|
||||
else S (assert_total (integerToNat (prim__sub_Integer x 1)))
|
||||
|
||||
|
465
src/Core/Reflect.idr
Normal file
465
src/Core/Reflect.idr
Normal file
@ -0,0 +1,465 @@
|
||||
module Core.Reflect
|
||||
|
||||
import Core.Context
|
||||
import Core.Env
|
||||
import Core.Normalise
|
||||
import Core.Value
|
||||
import Core.TT
|
||||
|
||||
%default covering
|
||||
|
||||
public export
|
||||
interface Reify a where
|
||||
reify : Defs -> NF vars -> Core a
|
||||
|
||||
public export
|
||||
interface Reflect a where
|
||||
reflect : FC -> Defs -> Env Term vars -> a -> Core (Term vars)
|
||||
|
||||
export
|
||||
getCon : FC -> Defs -> Name -> Core (Term vars)
|
||||
getCon fc defs n
|
||||
= case !(lookupDefExact n (gamma defs)) of
|
||||
Just (DCon t a) => resolved (gamma defs) (Ref fc (DataCon t a) n)
|
||||
Just (TCon t a _ _ _ _) => resolved (gamma defs) (Ref fc (TyCon t a) n)
|
||||
Just _ => resolved (gamma defs) (Ref fc Func n)
|
||||
_ => throw (UndefinedName fc n)
|
||||
|
||||
export
|
||||
appCon : FC -> Defs -> Name -> List (Term vars) -> Core (Term vars)
|
||||
appCon fc defs n args
|
||||
= do fn <- getCon fc defs n
|
||||
resolved (gamma defs) (apply fc fn args)
|
||||
|
||||
export
|
||||
prelude : String -> Name
|
||||
prelude n = NS ["Prelude"] (UN n)
|
||||
|
||||
export
|
||||
builtin : String -> Name
|
||||
builtin n = NS ["Builtin"] (UN n)
|
||||
|
||||
export
|
||||
primio : String -> Name
|
||||
primio n = NS ["PrimIO"] (UN n)
|
||||
|
||||
export
|
||||
reflection : String -> Name
|
||||
reflection n = NS ["Reflection", "Language"] (UN n)
|
||||
|
||||
export
|
||||
cantReify : NF vars -> String -> Core a
|
||||
cantReify val ty
|
||||
= throw (GenericMsg (getLoc val) ("Can't reify as " ++ ty))
|
||||
|
||||
export
|
||||
cantReflect : FC -> String -> Core a
|
||||
cantReflect fc ty
|
||||
= throw (GenericMsg fc ("Can't reflect as " ++ ty))
|
||||
|
||||
export
|
||||
Reify () where
|
||||
reify defs _ = pure ()
|
||||
|
||||
export
|
||||
Reflect () where
|
||||
reflect fc defs env _ = getCon fc defs (builtin "MkUnit")
|
||||
|
||||
export
|
||||
Reify String where
|
||||
reify defs (NPrimVal _ (Str str)) = pure str
|
||||
reify defs val = cantReify val "String"
|
||||
|
||||
export
|
||||
Reflect String where
|
||||
reflect fc defs env x = pure (PrimVal fc (Str x))
|
||||
|
||||
export
|
||||
Reify Int where
|
||||
reify defs (NPrimVal _ (I v)) = pure v
|
||||
reify defs val = cantReify val "Int"
|
||||
|
||||
export
|
||||
Reflect Int where
|
||||
reflect fc defs env x = pure (PrimVal fc (I x))
|
||||
|
||||
export
|
||||
Reify Integer where
|
||||
reify defs (NPrimVal _ (BI v)) = pure v
|
||||
reify defs val = cantReify val "Integer"
|
||||
|
||||
export
|
||||
Reflect Integer where
|
||||
reflect fc defs env x = pure (PrimVal fc (BI x))
|
||||
|
||||
export
|
||||
Reify Char where
|
||||
reify defs (NPrimVal _ (Ch v)) = pure v
|
||||
reify defs val = cantReify val "Char"
|
||||
|
||||
export
|
||||
Reflect Char where
|
||||
reflect fc defs env x = pure (PrimVal fc (Ch x))
|
||||
|
||||
export
|
||||
Reify Double where
|
||||
reify defs (NPrimVal _ (Db v)) = pure v
|
||||
reify defs val = cantReify val "Double"
|
||||
|
||||
export
|
||||
Reflect Double where
|
||||
reflect fc defs env x = pure (PrimVal fc (Db x))
|
||||
|
||||
export
|
||||
Reify Bool where
|
||||
reify defs (NDCon _ (NS _ (UN "True")) _ _ _) = pure True
|
||||
reify defs (NDCon _ (NS _ (UN "False")) _ _ _) = pure False
|
||||
reify defs val = cantReify val "Bool"
|
||||
|
||||
export
|
||||
Reflect Bool where
|
||||
reflect fc defs env True = getCon fc defs (prelude "True")
|
||||
reflect fc defs env False = getCon fc defs (prelude "False")
|
||||
|
||||
export
|
||||
Reify Nat where
|
||||
reify defs (NDCon _ (NS _ (UN "Z")) _ _ _)
|
||||
= pure Z
|
||||
reify defs (NDCon _ (NS _ (UN "S")) _ _ [k])
|
||||
= do k' <- reify defs !(evalClosure defs k)
|
||||
pure (S k')
|
||||
reify defs val = cantReify val "Nat"
|
||||
|
||||
export
|
||||
Reflect Nat where
|
||||
reflect fc defs env Z = getCon fc defs (prelude "Z")
|
||||
reflect fc defs env (S k)
|
||||
= do k' <- reflect fc defs env k
|
||||
appCon fc defs (prelude "S") [k']
|
||||
|
||||
export
|
||||
Reify a => Reify (List a) where
|
||||
reify defs (NDCon _ (NS _ (UN "Nil")) _ _ _)
|
||||
= pure []
|
||||
reify defs (NDCon _ (NS _ (UN "::")) _ _ [_, x, xs])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
xs' <- reify defs !(evalClosure defs xs)
|
||||
pure (x' :: xs')
|
||||
reify defs val = cantReify val "List"
|
||||
|
||||
export
|
||||
Reflect a => Reflect (List a) where
|
||||
reflect fc defs env [] = appCon fc defs (prelude "Nil") [Erased fc]
|
||||
reflect fc defs env (x :: xs)
|
||||
= do x' <- reflect fc defs env x
|
||||
xs' <- reflect fc defs env xs
|
||||
appCon fc defs (prelude "::") [Erased fc, x', xs']
|
||||
|
||||
export
|
||||
(Reify a, Reify b) => Reify (a, b) where
|
||||
reify defs (NDCon _ (NS _ (UN "MkPair")) _ _ [_, _, x, y])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
y' <- reify defs !(evalClosure defs y)
|
||||
pure (x', y')
|
||||
reify defs val = cantReify val "Pair"
|
||||
|
||||
export
|
||||
(Reflect a, Reflect b) => Reflect (a, b) where
|
||||
reflect fc defs env (x, y)
|
||||
= do x' <- reflect fc defs env x
|
||||
y' <- reflect fc defs env y
|
||||
appCon fc defs (builtin "MkPair") [Erased fc, Erased fc, x', y']
|
||||
|
||||
export
|
||||
Reify Name where
|
||||
reify defs (NDCon _ (NS _ (UN "UN")) _ _ [str])
|
||||
= do str' <- reify defs !(evalClosure defs str)
|
||||
pure (UN str')
|
||||
reify defs (NDCon _ (NS _ (UN "MN")) _ _ [str, i])
|
||||
= do str' <- reify defs !(evalClosure defs str)
|
||||
i' <- reify defs !(evalClosure defs i)
|
||||
pure (MN str' i')
|
||||
reify defs (NDCon _ (NS _ (UN "NS")) _ _ [ns, n])
|
||||
= do ns' <- reify defs !(evalClosure defs ns)
|
||||
n' <- reify defs !(evalClosure defs n)
|
||||
pure (NS ns' n')
|
||||
reify defs val = cantReify val "Name"
|
||||
|
||||
export
|
||||
Reflect Name where
|
||||
reflect fc defs env (UN x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "UN") [x']
|
||||
reflect fc defs env (MN x i)
|
||||
= do x' <- reflect fc defs env x
|
||||
i' <- reflect fc defs env i
|
||||
appCon fc defs (reflection "MN") [x', i']
|
||||
reflect fc defs env (NS ns n)
|
||||
= do ns' <- reflect fc defs env ns
|
||||
n' <- reflect fc defs env n
|
||||
appCon fc defs (reflection "NS") [ns', n']
|
||||
reflect fc defs env val = cantReflect fc "Name"
|
||||
|
||||
export
|
||||
Reify NameType where
|
||||
reify defs (NDCon _ (NS _ (UN "Bound")) _ _ _)
|
||||
= pure Bound
|
||||
reify defs (NDCon _ (NS _ (UN "Func")) _ _ _)
|
||||
= pure Func
|
||||
reify defs (NDCon _ (NS _ (UN "DataCon")) _ _ [t,i])
|
||||
= do t' <- reify defs !(evalClosure defs t)
|
||||
i' <- reify defs !(evalClosure defs i)
|
||||
pure (DataCon t' i')
|
||||
reify defs (NDCon _ (NS _ (UN "TyCon")) _ _ [t,i])
|
||||
= do t' <- reify defs !(evalClosure defs t)
|
||||
i' <- reify defs !(evalClosure defs i)
|
||||
pure (TyCon t' i')
|
||||
reify defs val = cantReify val "NameType"
|
||||
|
||||
export
|
||||
Reflect NameType where
|
||||
reflect fc defs env Bound = getCon fc defs (reflection "Bound")
|
||||
reflect fc defs env Func = getCon fc defs (reflection "Func")
|
||||
reflect fc defs env (DataCon t i)
|
||||
= do t' <- reflect fc defs env t
|
||||
i' <- reflect fc defs env i
|
||||
appCon fc defs (reflection "DataCon") [t', i']
|
||||
reflect fc defs env (TyCon t i)
|
||||
= do t' <- reflect fc defs env t
|
||||
i' <- reflect fc defs env i
|
||||
appCon fc defs (reflection "TyCon") [t', i']
|
||||
|
||||
export
|
||||
Reify Constant where
|
||||
reify defs (NDCon _ (NS _ (UN "I")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (I x')
|
||||
reify defs (NDCon _ (NS _ (UN "BI")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (BI x')
|
||||
reify defs (NDCon _ (NS _ (UN "Str")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (Str x')
|
||||
reify defs (NDCon _ (NS _ (UN "Ch")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (Ch x')
|
||||
reify defs (NDCon _ (NS _ (UN "Db")) _ _ [x])
|
||||
= do x' <- reify defs !(evalClosure defs x)
|
||||
pure (Db x')
|
||||
reify defs (NDCon _ (NS _ (UN "WorldVal")) _ _ [])
|
||||
= pure WorldVal
|
||||
reify defs (NDCon _ (NS _ (UN "IntType")) _ _ [])
|
||||
= pure IntType
|
||||
reify defs (NDCon _ (NS _ (UN "IntegerType")) _ _ [])
|
||||
= pure IntegerType
|
||||
reify defs (NDCon _ (NS _ (UN "StringType")) _ _ [])
|
||||
= pure StringType
|
||||
reify defs (NDCon _ (NS _ (UN "CharType")) _ _ [])
|
||||
= pure CharType
|
||||
reify defs (NDCon _ (NS _ (UN "DoubleType")) _ _ [])
|
||||
= pure DoubleType
|
||||
reify defs (NDCon _ (NS _ (UN "WorldType")) _ _ [])
|
||||
= pure WorldType
|
||||
reify defs val = cantReify val "Constant"
|
||||
|
||||
export
|
||||
Reflect Constant where
|
||||
reflect fc defs env (I x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "I") [x']
|
||||
reflect fc defs env (BI x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "BI") [x']
|
||||
reflect fc defs env (Str x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "Str") [x']
|
||||
reflect fc defs env (Ch x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "Ch") [x']
|
||||
reflect fc defs env (Db x)
|
||||
= do x' <- reflect fc defs env x
|
||||
appCon fc defs (reflection "Db") [x']
|
||||
reflect fc defs env WorldVal
|
||||
= getCon fc defs (reflection "WorldVal")
|
||||
reflect fc defs env IntType
|
||||
= getCon fc defs (reflection "IntTyoe")
|
||||
reflect fc defs env IntegerType
|
||||
= getCon fc defs (reflection "IntegerType")
|
||||
reflect fc defs env StringType
|
||||
= getCon fc defs (reflection "StringType")
|
||||
reflect fc defs env CharType
|
||||
= getCon fc defs (reflection "CharType")
|
||||
reflect fc defs env DoubleType
|
||||
= getCon fc defs (reflection "DoubleTyoe")
|
||||
reflect fc defs env WorldType
|
||||
= getCon fc defs (reflection "WorldType")
|
||||
|
||||
export
|
||||
Reify Visibility where
|
||||
reify defs (NDCon _ (NS _ (UN "Private")) _ _ [])
|
||||
= pure Private
|
||||
reify defs (NDCon _ (NS _ (UN "Export")) _ _ [])
|
||||
= pure Export
|
||||
reify defs (NDCon _ (NS _ (UN "Public")) _ _ [])
|
||||
= pure Public
|
||||
reify defs val = cantReify val "Visibility"
|
||||
|
||||
export
|
||||
Reflect Visibility where
|
||||
reflect fc defs env Private = getCon fc defs (reflection "Private")
|
||||
reflect fc defs env Export = getCon fc defs (reflection "Export")
|
||||
reflect fc defs env Public = getCon fc defs (reflection "Public")
|
||||
|
||||
export
|
||||
Reify RigCount where
|
||||
reify defs (NDCon _ (NS _ (UN "M0")) _ _ [])
|
||||
= pure Rig0
|
||||
reify defs (NDCon _ (NS _ (UN "M1")) _ _ [])
|
||||
= pure Rig1
|
||||
reify defs (NDCon _ (NS _ (UN "MW")) _ _ [])
|
||||
= pure RigW
|
||||
reify defs val = cantReify val "Count"
|
||||
|
||||
export
|
||||
Reflect RigCount where
|
||||
reflect fc defs env Rig0 = getCon fc defs (reflection "M0")
|
||||
reflect fc defs env Rig1 = getCon fc defs (reflection "M1")
|
||||
reflect fc defs env RigW = getCon fc defs (reflection "MW")
|
||||
|
||||
export
|
||||
Reify PiInfo where
|
||||
reify defs (NDCon _ (NS _ (UN "ImplicitArg")) _ _ [])
|
||||
= pure Implicit
|
||||
reify defs (NDCon _ (NS _ (UN "ExplicitArg")) _ _ [])
|
||||
= pure Explicit
|
||||
reify defs (NDCon _ (NS _ (UN "AutoImplicit")) _ _ [])
|
||||
= pure AutoImplicit
|
||||
reify defs val = cantReify val "PiInfo"
|
||||
|
||||
export
|
||||
Reflect PiInfo where
|
||||
reflect fc defs env Implicit = getCon fc defs (reflection "ImplicitArg")
|
||||
reflect fc defs env Explicit = getCon fc defs (reflection "ExplicitArg")
|
||||
reflect fc defs env AutoImplicit = getCon fc defs (reflection "AutoImplicit")
|
||||
|
||||
export
|
||||
Reify LazyReason where
|
||||
reify defs (NDCon _ (NS _ (UN "LInf")) _ _ [])
|
||||
= pure LInf
|
||||
reify defs (NDCon _ (NS _ (UN "LLazy")) _ _ [])
|
||||
= pure LLazy
|
||||
reify defs (NDCon _ (NS _ (UN "LUnknown")) _ _ [])
|
||||
= pure LUnknown
|
||||
reify defs val = cantReify val "LazyReason"
|
||||
|
||||
export
|
||||
Reflect LazyReason where
|
||||
reflect fc defs env LInf = getCon fc defs (reflection "LInf")
|
||||
reflect fc defs env LLazy = getCon fc defs (reflection "LLazy")
|
||||
reflect fc defs env LUnknown = getCon fc defs (reflection "LUnknown")
|
||||
|
||||
export
|
||||
Reify FC where
|
||||
reify defs (NDCon _ (NS _ (UN "MkFC")) _ _ [fn, start, end])
|
||||
= do fn' <- reify defs !(evalClosure defs fn)
|
||||
start' <- reify defs !(evalClosure defs start)
|
||||
end' <- reify defs !(evalClosure defs start)
|
||||
pure (MkFC fn' start' end')
|
||||
reify defs (NDCon _ (NS _ (UN "EmptyFC")) _ _ [])
|
||||
= pure EmptyFC
|
||||
reify defs val = cantReify val "FC"
|
||||
|
||||
export
|
||||
Reflect FC where
|
||||
reflect fc defs env (MkFC fn start end)
|
||||
= do fn' <- reflect fc defs env fn
|
||||
start' <- reflect fc defs env start
|
||||
end' <- reflect fc defs env end
|
||||
appCon fc defs (reflection "MkFC") [fn', start', end']
|
||||
reflect fc defs env EmptyFC = getCon fc defs (reflection "EmptyFC")
|
||||
|
||||
-- Reflection of well typed terms: We don't reify terms because that involves
|
||||
-- type checking, but we can reflect them
|
||||
|
||||
export
|
||||
Reflect (IsVar name idx vs) where
|
||||
reflect fc defs env First
|
||||
= appCon fc defs (reflection "First") [Erased fc, Erased fc]
|
||||
reflect fc defs env (Later p)
|
||||
= do p' <- reflect fc defs env p
|
||||
appCon fc defs (reflection "Later")
|
||||
[Erased fc, Erased fc, Erased fc, Erased fc, p']
|
||||
|
||||
-- Assume terms are normalised so there's not Let bindings in particular
|
||||
export
|
||||
Reflect (Term vs) where
|
||||
reflect fc defs env (Local {name} lfc _ idx prf)
|
||||
= do name' <- reflect fc defs env name
|
||||
lfc' <- reflect fc defs env lfc
|
||||
idx' <- reflect fc defs env idx
|
||||
prf' <- reflect fc defs env prf
|
||||
appCon fc defs (reflection "Local")
|
||||
[Erased fc, Erased fc, lfc', idx', prf']
|
||||
reflect fc defs env (Ref rfc nt n)
|
||||
= do rfc' <- reflect fc defs env rfc
|
||||
nt' <- reflect fc defs env nt
|
||||
n' <- reflect fc defs env n
|
||||
appCon fc defs (reflection "Ref")
|
||||
[Erased fc, rfc', nt', n']
|
||||
reflect fc defs env (Bind bfc x (Pi c p ty) sc)
|
||||
= do bfc' <- reflect fc defs env bfc
|
||||
x' <- reflect fc defs env x
|
||||
c' <- reflect fc defs env c
|
||||
p' <- reflect fc defs env p
|
||||
ty' <- reflect fc defs env ty
|
||||
sc' <- reflect fc defs env sc
|
||||
appCon fc defs (reflection "Pi")
|
||||
[Erased fc, bfc', c', p', x', ty', sc']
|
||||
reflect fc defs env (Bind bfc x (Lam c p ty) sc)
|
||||
= do bfc' <- reflect fc defs env bfc
|
||||
x' <- reflect fc defs env x
|
||||
c' <- reflect fc defs env c
|
||||
p' <- reflect fc defs env p
|
||||
ty' <- reflect fc defs env ty
|
||||
sc' <- reflect fc defs env sc
|
||||
appCon fc defs (reflection "Lam")
|
||||
[Erased fc, bfc', c', p', x', ty', sc']
|
||||
reflect fc defs env (App afc fn arg)
|
||||
= do afc' <- reflect fc defs env afc
|
||||
fn' <- reflect fc defs env fn
|
||||
arg' <- reflect fc defs env arg
|
||||
appCon fc defs (reflection "App")
|
||||
[Erased fc, afc', fn', arg']
|
||||
reflect fc defs env (TDelayed dfc r tm)
|
||||
= do dfc' <- reflect fc defs env dfc
|
||||
r' <- reflect fc defs env r
|
||||
tm' <- reflect fc defs env tm
|
||||
appCon fc defs (reflection "TDelayed")
|
||||
[Erased fc, dfc', r', tm']
|
||||
reflect fc defs env (TDelay dfc r ty tm)
|
||||
= do dfc' <- reflect fc defs env dfc
|
||||
r' <- reflect fc defs env r
|
||||
ty' <- reflect fc defs env ty
|
||||
tm' <- reflect fc defs env tm
|
||||
appCon fc defs (reflection "TDelay")
|
||||
[Erased fc, dfc', r', ty', tm']
|
||||
reflect fc defs env (TForce dfc tm)
|
||||
= do dfc' <- reflect fc defs env dfc
|
||||
tm' <- reflect fc defs env tm
|
||||
appCon fc defs (reflection "TForce")
|
||||
[Erased fc, dfc', tm']
|
||||
reflect fc defs env (PrimVal pfc c)
|
||||
= do pfc' <- reflect fc defs env pfc
|
||||
c' <- reflect fc defs env c
|
||||
appCon fc defs (reflection "PrimVal")
|
||||
[Erased fc, pfc', c']
|
||||
reflect fc defs env (Erased efc)
|
||||
= do efc' <- reflect fc defs env efc
|
||||
appCon fc defs (reflection "Erased")
|
||||
[Erased fc, efc']
|
||||
reflect fc defs env (TType tfc)
|
||||
= do tfc' <- reflect fc defs env tfc
|
||||
appCon fc defs (reflection "TType")
|
||||
[Erased fc, tfc']
|
||||
reflect fc defs env val = cantReflect fc "Term"
|
||||
|
@ -25,6 +25,8 @@ data CLOpt
|
||||
=
|
||||
||| Only typecheck the given file
|
||||
CheckOnly |
|
||||
||| The output file from the code generator
|
||||
OutputFile String |
|
||||
||| Execute a given function after checking the source file
|
||||
ExecFn String |
|
||||
||| Use a specific code generator (default chez)
|
||||
@ -69,6 +71,8 @@ record OptDesc where
|
||||
options : List OptDesc
|
||||
options = [MkOpt ["--check", "-c"] [] [CheckOnly]
|
||||
(Just "Exit after checking source file"),
|
||||
MkOpt ["--output", "-o"] ["file"] (\f => [OutputFile f, Quiet])
|
||||
(Just "Specify output file"),
|
||||
MkOpt ["--exec", "-x"] ["name"] (\f => [ExecFn f, Quiet])
|
||||
(Just "Execute function after checking source file"),
|
||||
MkOpt ["--no-prelude"] [] [NoPrelude]
|
||||
|
@ -129,7 +129,8 @@ fnameModified fname
|
||||
= do Right f <- coreLift $ openFile fname Read
|
||||
| Left err => throw (FileErr fname err)
|
||||
Right t <- coreLift $ fileModifiedTime f
|
||||
| Left err => throw (FileErr fname err)
|
||||
| Left err => do coreLift $ closeFile f
|
||||
throw (FileErr fname err)
|
||||
coreLift $ closeFile f
|
||||
pure t
|
||||
|
||||
|
@ -30,6 +30,13 @@ record PkgDesc where
|
||||
name : String
|
||||
version : String
|
||||
authors : String
|
||||
maintainers : Maybe String
|
||||
license : Maybe String
|
||||
brief : Maybe String
|
||||
readme : Maybe String
|
||||
homepage : Maybe String
|
||||
sourceloc : Maybe String
|
||||
bugtracker : Maybe String
|
||||
depends : List String -- packages to add to search path
|
||||
modules : List (List String, String) -- modules to install (namespace, filename)
|
||||
mainmod : Maybe (List String, String) -- main file (i.e. file to load at REPL)
|
||||
@ -44,6 +51,13 @@ Show PkgDesc where
|
||||
show pkg = "Package: " ++ name pkg ++ "\n" ++
|
||||
"Version: " ++ version pkg ++ "\n" ++
|
||||
"Authors: " ++ authors pkg ++ "\n" ++
|
||||
maybe "" (\m => "Maintainers: " ++ m ++ "\n") (maintainers pkg) ++
|
||||
maybe "" (\m => "License: " ++ m ++ "\n") (license pkg) ++
|
||||
maybe "" (\m => "Brief: " ++ m ++ "\n") (brief pkg) ++
|
||||
maybe "" (\m => "ReadMe: " ++ m ++ "\n") (readme pkg) ++
|
||||
maybe "" (\m => "HomePage: " ++ m ++ "\n") (homepage pkg) ++
|
||||
maybe "" (\m => "SourceLoc: " ++ m ++ "\n") (sourceloc pkg) ++
|
||||
maybe "" (\m => "BugTracker: " ++ m ++ "\n") (bugtracker pkg) ++
|
||||
"Depends: " ++ show (depends pkg) ++ "\n" ++
|
||||
"Modules: " ++ show (map snd (modules pkg)) ++ "\n" ++
|
||||
maybe "" (\m => "Main: " ++ snd m ++ "\n") (mainmod pkg) ++
|
||||
@ -56,26 +70,42 @@ Show PkgDesc where
|
||||
|
||||
initPkgDesc : String -> PkgDesc
|
||||
initPkgDesc pname
|
||||
= MkPkgDesc pname "0" "Anonymous" [] []
|
||||
= MkPkgDesc pname "0" "Anonymous" Nothing Nothing
|
||||
Nothing Nothing Nothing Nothing Nothing
|
||||
[] []
|
||||
Nothing Nothing Nothing Nothing Nothing Nothing Nothing
|
||||
|
||||
data DescField : Type where
|
||||
PVersion : FC -> String -> DescField
|
||||
PAuthors : FC -> String -> DescField
|
||||
PDepends : List String -> DescField
|
||||
PModules : List (FC, List String) -> DescField
|
||||
PMainMod : FC -> List String -> DescField
|
||||
PExec : String -> DescField
|
||||
POpts : FC -> String -> DescField
|
||||
PPrebuild : FC -> String -> DescField
|
||||
PPostbuild : FC -> String -> DescField
|
||||
PPreinstall : FC -> String -> DescField
|
||||
PVersion : FC -> String -> DescField
|
||||
PAuthors : FC -> String -> DescField
|
||||
PMaintainers : FC -> String -> DescField
|
||||
PLicense : FC -> String -> DescField
|
||||
PBrief : FC -> String -> DescField
|
||||
PReadMe : FC -> String -> DescField
|
||||
PHomePage : FC -> String -> DescField
|
||||
PSourceLoc : FC -> String -> DescField
|
||||
PBugTracker : FC -> String -> DescField
|
||||
PDepends : List String -> DescField
|
||||
PModules : List (FC, List String) -> DescField
|
||||
PMainMod : FC -> List String -> DescField
|
||||
PExec : String -> DescField
|
||||
POpts : FC -> String -> DescField
|
||||
PPrebuild : FC -> String -> DescField
|
||||
PPostbuild : FC -> String -> DescField
|
||||
PPreinstall : FC -> String -> DescField
|
||||
PPostinstall : FC -> String -> DescField
|
||||
|
||||
field : String -> Rule DescField
|
||||
field fname
|
||||
= strField PVersion "version"
|
||||
<|> strField PAuthors "authors"
|
||||
<|> strField PMaintainers "maintainers"
|
||||
<|> strField PLicense "license"
|
||||
<|> strField PBrief "brief"
|
||||
<|> strField PReadMe "readme"
|
||||
<|> strField PHomePage "homepage"
|
||||
<|> strField PSourceLoc "sourceloc"
|
||||
<|> strField PBugTracker "bugtracker"
|
||||
<|> strField POpts "options"
|
||||
<|> strField POpts "opts"
|
||||
<|> strField PPrebuild "prebuild"
|
||||
@ -86,7 +116,7 @@ field fname
|
||||
ds <- sepBy1 (symbol ",") unqualifiedName
|
||||
pure (PDepends ds)
|
||||
<|> do exactIdent "modules"; symbol "="
|
||||
ms <- sepBy1 (symbol ",")
|
||||
ms <- sepBy1 (symbol ",")
|
||||
(do start <- location
|
||||
ns <- namespace_
|
||||
end <- location
|
||||
@ -126,8 +156,16 @@ addField : {auto c : Ref Ctxt Defs} ->
|
||||
DescField -> PkgDesc -> Core PkgDesc
|
||||
addField (PVersion fc n) pkg = pure (record { version = n } 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)
|
||||
addField (PBrief fc a) pkg = pure (record { brief = Just a } pkg)
|
||||
addField (PReadMe fc a) pkg = pure (record { readme = Just a } pkg)
|
||||
addField (PHomePage fc a) pkg = pure (record { homepage = Just a } pkg)
|
||||
addField (PSourceLoc fc a) pkg = pure (record { sourceloc = Just a } pkg)
|
||||
addField (PBugTracker fc a) pkg = pure (record { bugtracker = Just a } pkg)
|
||||
|
||||
addField (PDepends ds) pkg = pure (record { depends = ds } pkg)
|
||||
addField (PModules ms) pkg
|
||||
addField (PModules ms) pkg
|
||||
= pure (record { modules = !(traverse toSource ms) } pkg)
|
||||
where
|
||||
toSource : (FC, List String) -> Core (List String, String)
|
||||
@ -148,7 +186,7 @@ addFields (x :: xs) desc = addFields xs !(addField x desc)
|
||||
|
||||
runScript : Maybe (FC, String) -> Core ()
|
||||
runScript Nothing = pure ()
|
||||
runScript (Just (fc, s))
|
||||
runScript (Just (fc, s))
|
||||
= do res <- coreLift $ system s
|
||||
when (res /= 0) $
|
||||
throw (GenericMsg fc "Script failed")
|
||||
@ -167,7 +205,7 @@ processOptions Nothing = pure ()
|
||||
processOptions (Just (fc, opts))
|
||||
= do let Right clopts = getOpts (words opts)
|
||||
| Left err => throw (GenericMsg fc err)
|
||||
preOptions clopts
|
||||
preOptions clopts
|
||||
|
||||
build : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
@ -200,7 +238,7 @@ installFrom pname builddir destdir ns@(m :: dns)
|
||||
let ttcPath = builddir ++ dirSep ++ ttcfile ++ ".ttc"
|
||||
let destPath = destdir ++ dirSep ++ showSep "/" (reverse dns)
|
||||
let destFile = destdir ++ dirSep ++ ttcfile ++ ".ttc"
|
||||
Right _ <- coreLift $ mkdirs (reverse dns)
|
||||
Right _ <- coreLift $ mkdirs (reverse dns)
|
||||
| Left err => throw (FileErr pname err)
|
||||
coreLift $ putStrLn $ "Installing " ++ ttcPath ++ " to " ++ destPath
|
||||
Right _ <- coreLift $ copyFile ttcPath destFile
|
||||
@ -210,10 +248,10 @@ installFrom pname builddir destdir ns@(m :: dns)
|
||||
-- Install all the built modules in prefix/package/
|
||||
-- We've already built and checked for success, so if any don't exist that's
|
||||
-- an internal error.
|
||||
install : {auto c : Ref Ctxt Defs} ->
|
||||
install : {auto c : Ref Ctxt Defs} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
PkgDesc -> Core ()
|
||||
install pkg
|
||||
install pkg
|
||||
= do defs <- get Ctxt
|
||||
let build = build_dir (dirs (options defs))
|
||||
runScript (preinstall pkg)
|
||||
@ -230,12 +268,12 @@ install pkg
|
||||
True <- coreLift $ changeDir (name pkg)
|
||||
| False => throw (FileErr (name pkg) FileReadError)
|
||||
|
||||
-- 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
|
||||
traverse (installFrom (name pkg)
|
||||
(srcdir ++ dirSep ++ build)
|
||||
(srcdir ++ dirSep ++ build)
|
||||
(installPrefix ++ dirSep ++ name pkg)) toInstall
|
||||
coreLift $ changeDir srcdir
|
||||
coreLift $ changeDir srcdir
|
||||
runScript (postinstall pkg)
|
||||
|
||||
-- Data.These.bitraverse hand specialised for Core
|
||||
@ -257,10 +295,10 @@ foldWithKeysC {a} {b} fk fv = go []
|
||||
map bifold $ bitraverseC
|
||||
(fv as)
|
||||
(\sm => foldlC
|
||||
(\x, (k, vs) =>
|
||||
(\x, (k, vs) =>
|
||||
do let as' = as ++ [k]
|
||||
y <- assert_total $ go as' vs
|
||||
z <- fk as'
|
||||
z <- fk as'
|
||||
pure $ x <+> y <+> z)
|
||||
neutral
|
||||
(toList sm))
|
||||
@ -272,10 +310,10 @@ Semigroup () where
|
||||
Monoid () where
|
||||
neutral = ()
|
||||
|
||||
clean : {auto c : Ref Ctxt Defs} ->
|
||||
clean : {auto c : Ref Ctxt Defs} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
PkgDesc -> Core ()
|
||||
clean pkg
|
||||
clean pkg
|
||||
= do defs <- get Ctxt
|
||||
let build = build_dir (dirs (options defs))
|
||||
let toClean = mapMaybe (\ks => [| MkPair (tail' ks) (head' ks) |]) $
|
||||
@ -286,12 +324,12 @@ clean pkg
|
||||
let builddir = srcdir ++ dirSep ++ build
|
||||
-- the usual pair syntax breaks with `No such variable a` here for some reason
|
||||
let pkgTrie = the (StringTrie (List String)) $
|
||||
foldl (\trie, ksv =>
|
||||
foldl (\trie, ksv =>
|
||||
let ks = fst ksv
|
||||
v = snd ksv
|
||||
in
|
||||
v = snd ksv
|
||||
in
|
||||
insertWith ks (maybe [v] (v::)) trie) empty toClean
|
||||
foldWithKeysC (deleteFolder builddir)
|
||||
foldWithKeysC (deleteFolder builddir)
|
||||
(\ks => map concat . traverse (deleteBin builddir ks))
|
||||
pkgTrie
|
||||
deleteFolder builddir []
|
||||
@ -311,11 +349,11 @@ clean pkg
|
||||
|
||||
-- Just load the 'Main' module, if it exists, which will involve building
|
||||
-- it if necessary
|
||||
runRepl : {auto c : Ref Ctxt Defs} ->
|
||||
runRepl : {auto c : Ref Ctxt Defs} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
PkgDesc -> Core ()
|
||||
runRepl pkg
|
||||
= do addDeps pkg
|
||||
runRepl pkg
|
||||
= do addDeps pkg
|
||||
processOptions (options pkg)
|
||||
throw (InternalError "Not implemented")
|
||||
|
||||
@ -323,8 +361,8 @@ processPackage : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
PkgCommand -> String -> Core ()
|
||||
processPackage cmd file
|
||||
= do Right (pname, fs) <- coreLift $ parseFile file
|
||||
processPackage cmd file
|
||||
= do Right (pname, fs) <- coreLift $ parseFile file
|
||||
(do desc <- parsePkgDesc file
|
||||
eoi
|
||||
pure desc)
|
||||
@ -354,7 +392,7 @@ processPackageOpts : {auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
List CLOpt -> Core Bool
|
||||
processPackageOpts [Package cmd f]
|
||||
processPackageOpts [Package cmd f]
|
||||
= do processPackage cmd f
|
||||
pure True
|
||||
processPackageOpts opts = rejectPackageOpts opts
|
||||
|
@ -1395,9 +1395,8 @@ editCmd
|
||||
pure (MakeWith (fromInteger line) n)
|
||||
<|> fatalError "Unrecognised command"
|
||||
|
||||
export
|
||||
command : Rule REPLCmd
|
||||
command
|
||||
nonEmptyCommand : Rule REPLCmd
|
||||
nonEmptyCommand
|
||||
= do symbol ":"; replCmd ["t", "type"]
|
||||
tm <- expr pdef "(interactive)" init
|
||||
pure (Check tm)
|
||||
@ -1445,3 +1444,9 @@ command
|
||||
<|> do tm <- expr pdef "(interactive)" init
|
||||
pure (Eval tm)
|
||||
|
||||
export
|
||||
command : EmptyRule REPLCmd
|
||||
command
|
||||
= do eoi
|
||||
pure NOP
|
||||
<|> nonEmptyCommand
|
||||
|
@ -157,7 +157,8 @@ modTime fname
|
||||
= do Right f <- coreLift $ openFile fname Read
|
||||
| Left err => pure 0 -- Beginning of Time :)
|
||||
Right t <- coreLift $ fileModifiedTime f
|
||||
| Left err => pure 0
|
||||
| Left err => do coreLift $ closeFile f
|
||||
pure 0
|
||||
coreLift $ closeFile f
|
||||
pure t
|
||||
|
||||
|
@ -218,6 +218,24 @@ findCG
|
||||
Chicken => pure codegenChicken
|
||||
Racket => pure codegenRacket
|
||||
|
||||
export
|
||||
compileExp : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
PTerm -> String -> Core ()
|
||||
compileExp ctm outfile
|
||||
= do inidx <- resolveName (UN "[input]")
|
||||
ttimp <- desugar AnyExpr [] (PApp replFC (PRef replFC (UN "unsafePerformIO")) ctm)
|
||||
(tm, gty) <- elabTerm inidx InExpr [] (MkNested [])
|
||||
[] ttimp Nothing
|
||||
tm_erased <- linearCheck replFC Rig1 True [] tm
|
||||
ok <- compile !findCG tm_erased outfile
|
||||
maybe (pure ())
|
||||
(\fname => iputStrLn (outfile ++ " written"))
|
||||
ok
|
||||
|
||||
export
|
||||
execExp : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
@ -467,15 +485,7 @@ process Edit
|
||||
loadMainFile f
|
||||
pure True
|
||||
process (Compile ctm outfile)
|
||||
= do inidx <- resolveName (UN "[input]")
|
||||
ttimp <- desugar AnyExpr [] (PApp replFC (PRef replFC (UN "unsafePerformIO")) ctm)
|
||||
(tm, gty) <- elabTerm inidx InExpr [] (MkNested [])
|
||||
[] ttimp Nothing
|
||||
tm_erased <- linearCheck replFC Rig1 True [] tm
|
||||
ok <- compile !findCG tm_erased outfile
|
||||
maybe (pure ())
|
||||
(\fname => iputStrLn (outfile ++ " written"))
|
||||
ok
|
||||
= do compileExp ctm outfile
|
||||
pure True
|
||||
process (Exec ctm)
|
||||
= do execExp ctm
|
||||
@ -551,6 +561,8 @@ process (Editing cmd)
|
||||
process Quit
|
||||
= do iputStrLn "Bye for now!"
|
||||
pure False
|
||||
process NOP
|
||||
= pure True
|
||||
|
||||
processCatch : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
@ -619,7 +631,10 @@ repl
|
||||
inp <- coreLift getLine
|
||||
end <- coreLift $ fEOF stdin
|
||||
if end
|
||||
then iputStrLn "Bye for now!"
|
||||
then do
|
||||
-- start a new line in REPL mode (not relevant in IDE mode)
|
||||
coreLift $ putStrLn ""
|
||||
iputStrLn "Bye for now!"
|
||||
else if !(interpret inp)
|
||||
then repl
|
||||
else pure ()
|
||||
|
@ -56,8 +56,13 @@ postOptions : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
{auto m : Ref MD Metadata} ->
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
List CLOpt -> Core Bool
|
||||
postOptions [] = pure True
|
||||
postOptions (OutputFile outfile :: rest)
|
||||
= do compileExp (PRef (MkFC "(script)" (0, 0) (0, 0)) (UN "main")) outfile
|
||||
postOptions rest
|
||||
pure False
|
||||
postOptions (ExecFn str :: rest)
|
||||
= do execExp (PRef (MkFC "(script)" (0, 0) (0, 0)) (UN str))
|
||||
postOptions rest
|
||||
|
@ -280,6 +280,7 @@ data REPLCmd : Type where
|
||||
Metavars : REPLCmd
|
||||
Editing : EditCmd -> REPLCmd
|
||||
Quit : REPLCmd
|
||||
NOP : REPLCmd
|
||||
|
||||
public export
|
||||
record Import where
|
||||
|
@ -79,10 +79,13 @@ readFromFile fname
|
||||
= do Right h <- openFile fname Read
|
||||
| Left err => pure (Left err)
|
||||
Right fsize <- fileSize h
|
||||
| Left err => pure (Left err)
|
||||
| Left err => do closeFile h
|
||||
pure (Left err)
|
||||
Just b <- newBuffer fsize
|
||||
| Nothing => pure (Left (GenericFileError 0)) --- um, not really
|
||||
| Nothing => do closeFile h
|
||||
pure (Left (GenericFileError 0)) --- um, not really
|
||||
b <- readBufferFromFile h b fsize
|
||||
closeFile h
|
||||
pure (Right (MkBin b 0 fsize fsize))
|
||||
|
||||
public export
|
||||
|
@ -46,6 +46,7 @@ idrisTests
|
||||
"perf001", "perf002",
|
||||
"perror001", "perror002", "perror003", "perror004", "perror005",
|
||||
"perror006",
|
||||
"pkg001",
|
||||
"record001", "record002",
|
||||
"reg001",
|
||||
"total001", "total002", "total003", "total004", "total005",
|
||||
@ -88,9 +89,15 @@ runTest prog testPath
|
||||
Right exp <- readFile "expected"
|
||||
| Left err => do print err
|
||||
pure False
|
||||
|
||||
if (out == exp)
|
||||
then putStrLn "success"
|
||||
else putStrLn $ "FAILURE: found " ++ out
|
||||
else do
|
||||
putStrLn "FAILURE"
|
||||
putStrLn "Expected:"
|
||||
printLn exp
|
||||
putStrLn "Given:"
|
||||
printLn out
|
||||
chdir "../.."
|
||||
pure (out == exp)
|
||||
|
||||
|
69
tests/idris2/pkg001/Dummy.idr
Normal file
69
tests/idris2/pkg001/Dummy.idr
Normal file
@ -0,0 +1,69 @@
|
||||
module Dummy
|
||||
|
||||
import Data.Vect
|
||||
|
||||
namespace DList
|
||||
|
||||
||| A list construct for dependent types.
|
||||
|||
|
||||
||| @aTy The type of the value contained within the list element type.
|
||||
||| @elemTy The type of the elements within the list
|
||||
||| @as The List used to contain the different values within the type.
|
||||
public export
|
||||
data DList : (aTy : Type)
|
||||
-> (elemTy : aTy -> Type)
|
||||
-> (as : List aTy)
|
||||
-> Type where
|
||||
||| Create an empty List
|
||||
Nil : DList aTy elemTy Nil
|
||||
||| Cons
|
||||
|||
|
||||
||| @elem The element to add
|
||||
||| @rest The list for `elem` to be added to.
|
||||
(::) : (elem : elemTy x)
|
||||
-> (rest : DList aTy elemTy xs)
|
||||
-> DList aTy elemTy (x::xs)
|
||||
|
||||
namespace DVect
|
||||
||| A list construct for dependent types.
|
||||
|||
|
||||
||| @aTy The type of the value contained within the list element type.
|
||||
||| @elemTy The type of the elements within the list
|
||||
||| @len The length of the list.
|
||||
||| @as The List used to contain the different values within the type.
|
||||
public export
|
||||
data DVect : (aTy : Type)
|
||||
-> (elemTy : aTy -> Type)
|
||||
-> (len : Nat)
|
||||
-> (as : Vect len aTy)
|
||||
-> Type where
|
||||
||| Create an empty List
|
||||
Nil : DVect aTy elemTy Z Nil
|
||||
||| Cons
|
||||
|||
|
||||
||| @ex The element to add
|
||||
||| @rest The list for `elem` to be added to.
|
||||
(::) : (ex : elemTy x)
|
||||
-> (rest : DVect aTy elemTy n xs)
|
||||
-> DVect aTy elemTy (S n) (x::xs)
|
||||
|
||||
namespace PList
|
||||
public export
|
||||
data PList : (aTy : Type)
|
||||
-> (elemTy : aTy -> Type)
|
||||
-> (predTy : aTy -> Type)
|
||||
-> (as : List aTy)
|
||||
-> (prf : DList aTy predTy as)
|
||||
-> Type
|
||||
where
|
||||
||| Create an empty List
|
||||
Nil : PList aTy elemTy predTy Nil Nil
|
||||
|
||||
||| Cons
|
||||
|||
|
||||
||| @elem The element to add and proof that the element's type satisfies a certain predicate.
|
||||
||| @rest The list for `elem` to be added to.
|
||||
(::) : (elem : elemTy x)
|
||||
-> {prf : predTy x}
|
||||
-> (rest : PList aTy elemTy predTy xs prfs)
|
||||
-> PList aTy elemTy predTy (x :: xs) (prf :: prfs)
|
9
tests/idris2/pkg001/dummy.ipkg
Normal file
9
tests/idris2/pkg001/dummy.ipkg
Normal file
@ -0,0 +1,9 @@
|
||||
package dummy
|
||||
|
||||
authors = "Joe Bloggs"
|
||||
maintainers = "Joe Bloggs"
|
||||
license = "BSD3 but see LICENSE for more information"
|
||||
brief = "This is a dummy package."
|
||||
readme = "README.md"
|
||||
|
||||
modules = Dummy
|
1
tests/idris2/pkg001/expected
Normal file
1
tests/idris2/pkg001/expected
Normal file
@ -0,0 +1 @@
|
||||
1/1: Building Dummy (Dummy.idr)
|
3
tests/idris2/pkg001/run
Executable file
3
tests/idris2/pkg001/run
Executable file
@ -0,0 +1,3 @@
|
||||
$1 --build dummy.ipkg
|
||||
|
||||
rm -rf build
|
Loading…
Reference in New Issue
Block a user