Add libraries

This commit is contained in:
Edwin Brady 2020-05-18 13:59:07 +01:00
parent a5793756b7
commit dec7dff622
88 changed files with 12091 additions and 6 deletions

90
Makefile Normal file
View File

@ -0,0 +1,90 @@
include config.mk
# Idris 2 executable used to bootstrap
IDRIS2_BOOT ?= idris2
# Idris 2 executable we're building
NAME = idris2sh
TARGETDIR = build/exec
TARGET = ${TARGETDIR}/${NAME}
MAJOR=0
MINOR=2
PATCH=0
GIT_SHA1=
VER_TAG=
ifeq ($(shell git status >/dev/null 2>&1; echo $$?), 0)
# inside a git repo
ifneq ($(shell git describe --exact-match --tags >/dev/null 2>&1; echo $$?), 0)
# not tagged as a released version, so add sha1 of this build in between releases
GIT_SHA1 := $(shell git rev-parse --short=9 HEAD)
VER_TAG := -${GIT_SHA1}
endif
endif
export IDRIS2_VERSION := ${MAJOR}.${MINOR}.${PATCH}
IDRIS2_VERSION_TAG := ${IDRIS2_VERSION}${VER_TAG}
export IDRIS2_CURDIR = $(CURDIR)
export IDRIS2_BOOT_PATH = ${IDRIS2_CURDIR}/libs/prelude/build/ttc:${IDRIS2_CURDIR}/libs/base/build/ttc:${IDRIS2_CURDIR}/libs/network/build/ttc
.PHONY: all support clean support-clean ${TARGET}
all: support ${TARGET} libs
${TARGET}:
${IDRIS2_BOOT} --build idris2.ipkg
prelude:
${MAKE} -C libs/prelude IDRIS2=../../${TARGET} IDRIS2_PATH=${IDRIS2_BOOT_PATH}
base: prelude
${MAKE} -C libs/base IDRIS2=../../${TARGET} IDRIS2_PATH=${IDRIS2_BOOT_PATH}
network: prelude
${MAKE} -C libs/network IDRIS2=../../${TARGET} IDRIS2_PATH=${IDRIS2_BOOT_PATH}
${MAKE} -C libs/network test IDRIS2=../../${TARGET} IDRIS2_PATH=${IDRIS2_BOOT_PATH}
contrib: prelude
${MAKE} -C libs/contrib IDRIS2=../../${TARGET} IDRIS2_PATH=${IDRIS2_BOOT_PATH}
libs : prelude base network contrib
support:
@${MAKE} -C support/c
support-clean:
@${MAKE} -C support/c clean
clean-libs:
${MAKE} -C libs/prelude clean
${MAKE} -C libs/base clean
${MAKE} -C libs/network clean
${MAKE} -C libs/contrib clean
clean: clean-libs support-clean
${IDRIS2_BOOT} --clean idris2.ipkg
rm -rf ${TARGETDIR}
install: install-idris2 install-support install-libs
install-idris2:
mkdir -p ${PREFIX}/bin/${NAME}_app
install ${TARGET} ${PREFIX}/bin
install ${TARGETDIR}/${NAME}_app/* ${PREFIX}/bin/${NAME}_app
install-support: support
mkdir -p ${PREFIX}/idris2-${IDRIS2_VERSION}/support/chez
mkdir -p ${PREFIX}/idris2-${IDRIS2_VERSION}/support/racket
mkdir -p ${PREFIX}/idris2-${IDRIS2_VERSION}/support/gambit
install support/chez/* ${PREFIX}/idris2-${IDRIS2_VERSION}/support/chez
install support/racket/* ${PREFIX}/idris2-${IDRIS2_VERSION}/support/racket
install support/gambit/* ${PREFIX}/idris2-${IDRIS2_VERSION}/support/gambit
@${MAKE} -C support/c install
install-libs: libs
${MAKE} -C libs/prelude install IDRIS2=../../${TARGET}
${MAKE} -C libs/base install IDRIS2=../../${TARGET}
${MAKE} -C libs/network install IDRIS2=../../${TARGET} IDRIS2_VERSION=${IDRIS2_VERSION}
${MAKE} -C libs/contrib install IDRIS2=../../${TARGET}

44
config.mk Normal file
View File

@ -0,0 +1,44 @@
##### Options which a user might set before building go here #####
PREFIX ?= $(HOME)/.idris2sh
# clang compiles the output much faster than gcc!
CC := clang
##################################################################
RANLIB ?= ranlib
AR ?= ar
CFLAGS := -Wall $(CFLAGS) $(OPT)
LDFLAGS := $(LDFLAGS)
MACHINE := $(shell $(CC) -dumpmachine)
ifneq (,$(findstring cygwin, $(MACHINE)))
OS := windows
SHLIB_SUFFIX := .dll
else ifneq (,$(findstring mingw, $(MACHINE)))
OS := windows
SHLIB_SUFFIX := .dll
else ifneq (,$(findstring windows, $(MACHINE)))
OS := windows
SHLIB_SUFFIX := .dll
else ifneq (,$(findstring darwin, $(MACHINE)))
OS := darwin
SHLIB_SUFFIX := .dylib
CFLAGS += -fPIC
else ifneq (, $(findstring bsd, $(MACHINE)))
OS := bsd
SHLIB_SUFFIX := .so
CFLAGS += -fPIC
else
OS := linux
SHLIB_SUFFIX := .so
CFLAGS += -fPIC
endif
ifeq ($(OS),bsd)
MAKE := gmake
else
MAKE := make
endif

383
libs/base/Control/App.idr Normal file
View File

@ -0,0 +1,383 @@
module Control.App
import Data.IORef
public export
Error : Type
Error = Type
public export
data HasErr : Error -> List Error -> Type where
Here : HasErr e (e :: es)
There : HasErr e es -> HasErr e (e' :: es)
public export
data Path = MayThrow | NoThrow
public export
data Usage = One | Any
public export
0 Has : List (a -> Type) -> a -> Type
Has [] es = ()
Has [e] es = e es
Has (e :: es') es = (e es, Has es' es)
data OneOf : List Type -> Path -> Type where
First : e -> OneOf (e :: es) MayThrow
Later : OneOf es MayThrow -> OneOf (e :: es) MayThrow
public export
data SafeBind : Path -> (l' : Path) -> Type where
[search l']
SafeSame : SafeBind l l
SafeToThrow : SafeBind NoThrow MayThrow
updateP : SafeBind p p' => OneOf es p -> OneOf es p'
updateP @{SafeSame} x = x
updateP @{SafeToThrow} x impossible
Uninhabited (OneOf [] p) where
uninhabited (First x) impossible
uninhabited (Later x) impossible
Uninhabited (OneOf es NoThrow) where
uninhabited (First x) impossible
uninhabited (Later x) impossible
0 execTy : Path -> List Error -> Type -> Type
execTy p es ty = Either (OneOf es p) ty
data AppRes : Type -> Type where
MkAppRes : (result : a) -> (1 x : %World) -> AppRes a
data App1Res : Usage -> Type -> Type where
MkApp1Res1 : (1 result : a) -> (1 x : %World) -> App1Res One a
MkApp1ResW : (result : a) -> (1 x : %World) -> App1Res Any a
PrimApp : Type -> Type
PrimApp a = (1 x : %World) -> AppRes a
export
prim_app_pure : a -> PrimApp a
prim_app_pure x = \w => MkAppRes x w
export
prim_app_bind : (1 act : PrimApp a) -> (1 k : a -> PrimApp b) -> PrimApp b
prim_app_bind fn k w
= let MkAppRes x' w' = fn w in k x' w'
toPrimApp : (1 act : IO a) -> PrimApp a
toPrimApp x
= \w => case toPrim x w of
MkIORes r w => MkAppRes r w
PrimApp1 : Usage -> Type -> Type
PrimApp1 u a = (1 x : %World) -> App1Res u a
toPrimApp1 : {u : _} -> (1 act : IO a) -> PrimApp1 u a
toPrimApp1 x
= \w => case toPrim x w of
MkIORes r w =>
case u of
One => MkApp1Res1 r w
Any => MkApp1ResW r w
prim_app1_bind : (1 act : PrimApp1 Any a) ->
(1 k : a -> PrimApp1 u b) -> PrimApp1 u b
prim_app1_bind fn k w
= let MkApp1ResW x' w' = fn w in k x' w'
export
data App : {default MayThrow l : Path} ->
(es : List Error) -> Type -> Type where
MkApp : (1 prog : (1 w : %World) -> AppRes (execTy l e t)) -> App {l} e t
export
data App1 : {default One u : Usage} ->
(es : List Error) -> Type -> Type where
MkApp1 : (1 prog : (1 w : %World) -> App1Res u t) -> App1 {u} e t
bindApp : SafeBind l l' =>
App {l} e a -> (a -> App {l=l'} e b) -> App {l=l'} e b
bindApp (MkApp prog) next
= MkApp $ \world =>
let MkAppRes x' world' = prog world in
case x' of
Right res =>
let MkApp r = next res in
r world'
Left err => MkAppRes (Left (updateP err)) world'
public export
Cont1Type : Usage -> Type -> Usage -> List Error -> Type -> Type
Cont1Type One a u e b = (1 x : a) -> App1 {u} e b
Cont1Type Any a u e b = (x : a) -> App1 {u} e b
export
bindApp1 : {u : _} -> (1 act : App1 {u} e a) ->
(1 k : Cont1Type u a u' e b) -> App1 {u=u'} e b
bindApp1 {u=One} (MkApp1 fn)
= \k => MkApp1 (\w => let MkApp1Res1 x' w' = fn w
MkApp1 res = k x' in
res w')
bindApp1 {u=Any} (MkApp1 fn)
= \k => MkApp1 (\w => let MkApp1ResW x' w' = fn w
MkApp1 res = k x' in
res w')
absurdWith1 : (1 w : b) -> OneOf e NoThrow -> any
absurdWith1 w (First p) impossible
absurdWithVoid : (1 w : b) -> OneOf [Void] t -> any
absurdWithVoid w (First p) impossible
absurdWith2 : (1 x : a) -> (1 w : b) -> OneOf e NoThrow -> any
absurdWith2 x w (First p) impossible
export
bindL : App {l=NoThrow} e a -> (1 k : a -> App {l} e b) -> App {l} e b
bindL (MkApp prog) next
= MkApp $ \world =>
let MkAppRes x' world' = prog world in
case x' of
Right res =>
let MkApp r = next res in
r world'
Left err => absurdWith2 next world' err
export
app : (1 p : App {l=NoThrow} e a) -> App1 {u=Any} e a
app (MkApp prog)
= MkApp1 $ \world =>
let MkAppRes x' world' = prog world in
case x' of
Left err => absurdWith1 world' err
Right res => MkApp1ResW res world'
export
app1 : (1 p : App1 {u=Any} e a) -> App {l} e a
app1 (MkApp1 prog)
= MkApp $ \world =>
let MkApp1ResW x' world' = prog world in
MkAppRes (Right x') world'
pureApp : a -> App {l} e a
pureApp x = MkApp $ \w => MkAppRes (Right x) w
export
Functor (App {l} es) where
map f ap = bindApp ap $ \ap' => pureApp (f ap')
export
Applicative (App {l} es) where
pure = pureApp
(<*>) f a = bindApp f $ \f' =>
bindApp a $ \a' => pure (f' a')
export
Monad (App es) where
(>>=) = bindApp -- won't get used, but handy to have the instance
export
(>>=) : SafeBind l l' =>
App {l} e a -> (k : a -> App {l=l'} e b) -> App {l=l'} e b
(>>=) = bindApp
namespace App1
export
(>>=) : {u : _} -> (1 act : App1 {u} e a) ->
(1 k : Cont1Type u a u' e b) -> App1 {u=u'} e b
(>>=) = bindApp1
export
pure : (x : a) -> App1 {u=Any} e a
pure x = MkApp1 $ \w => MkApp1ResW x w
export
pure1 : (1 x : a) -> App1 e a
pure1 x = MkApp1 $ \w => MkApp1Res1 x w
export
data State : (tag : a) -> Type -> List Error -> Type where
[search tag]
MkState : IORef t -> State tag t e
%hint export
mapState : State tag t e -> State tag t (eff :: e)
mapState (MkState s) = MkState s
export
get : (0 tag : _) -> State tag t e => App {l} e t
get tag @{MkState r}
= MkApp $
prim_app_bind (toPrimApp $ readIORef r) $ \val =>
MkAppRes (Right val)
export
put : (0 tag : _) -> State tag t e => (1 val : t) -> App {l} e ()
put tag @{MkState r} val
= MkApp $
prim_app_bind (toPrimApp $ writeIORef r val) $ \val =>
MkAppRes (Right ())
export
modify : (0 tag : _) -> State tag t e => (1 p : t -> t) -> App {l} e ()
modify tag f
= let (>>=) = bindL in
do x <- get tag
put tag (f x)
export
new : t -> (1 p : State tag t e => App {l} e a) -> App {l} e a
new val prog
= MkApp $
prim_app_bind (toPrimApp $ newIORef val) $ \ref =>
let st = MkState ref
MkApp res = prog @{st} in
res
public export
interface Exception err e where
throw : err -> App e a
catch : App e a -> (err -> App e a) -> App e a
findException : HasErr e es -> e -> OneOf es MayThrow
findException Here err = First err
findException (There p) err = Later $ findException p err
findError : HasErr e es -> OneOf es MayThrow -> Maybe e
findError Here (First err) = Just err
findError (There p) (Later q) = findError p q
findError _ _ = Nothing
export
HasErr e es => Exception e es where
throw err = MkApp $ MkAppRes (Left (findException %search err))
catch (MkApp prog) handler
= MkApp $
prim_app_bind prog $ \res =>
case res of
Left err =>
case findError %search err of
Nothing => MkAppRes (Left err)
Just err' =>
let MkApp e' = handler err' in
e'
Right ok => MkAppRes (Right ok)
export
handle : App (err :: e) a ->
(onok : a -> App e b) ->
(onerr : err -> App e b) -> App e b
handle (MkApp prog) onok onerr
= MkApp $
prim_app_bind prog $ \res =>
case res of
Left (First err) =>
let MkApp err' = onerr err in
err'
Left (Later p) =>
-- different exception, so rethrow
MkAppRes (Left p)
Right ok =>
let MkApp ok' = onok ok in
ok'
export
lift : App e a -> App (err :: e) a
lift (MkApp prog)
= MkApp $
prim_app_bind prog $ \res =>
case res of
Left err => MkAppRes (Left (Later err))
Right ok => MkAppRes (Right ok)
public export
Init : List Error
Init = [Void]
export
run : App {l} Init a -> IO a
run (MkApp prog)
= primIO $ \w =>
case (prim_app_bind prog $ \r =>
case r of
Right res => MkAppRes res
Left (First err) => absurd err) w of
MkAppRes r w => MkIORes r w
export
noThrow : App Init a -> App Init {l=NoThrow} a
noThrow (MkApp prog)
= MkApp $ \w =>
case prog w of
MkAppRes (Left err) w => absurdWithVoid w err
MkAppRes (Right res) w => MkAppRes (Right res) w
public export
interface PrimIO e where
primIO : IO a -> App {l} e a
primIO1 : (1 act : IO a) -> App1 e a
-- fork starts a new environment, so that any existing state can't get
-- passed to it (since it'll be tagged with the wrong environment)
fork : (forall e' . PrimIO e' => App {l} e' ()) -> App e ()
export
HasErr Void e => PrimIO e where
primIO op =
MkApp $ \w =>
let MkAppRes r w = toPrimApp op w in
MkAppRes (Right r) w
primIO1 op = MkApp1 $ \w => toPrimApp1 op w
fork thread
= MkApp $
prim_app_bind
(toPrimApp $ PrimIO.fork $
do run thread
pure ())
$ \_ =>
MkAppRes (Right ())
infix 5 @@
export
new1 : t -> (1 p : State tag t e => App1 {u} e a) -> App1 {u} e a
new1 val prog
= MkApp1 $ prim_app1_bind (toPrimApp1 $ newIORef val) $ \ref =>
let st = MkState ref
MkApp1 res = prog @{st} in
res
public export
data Res : (a : Type) -> (a -> Type) -> Type where
(@@) : (val : a) -> (1 r : t val) -> Res a t
public export
data FileEx = GenericFileEx Int -- errno
| FileReadError
| FileWriteError
| FileNotFound
| PermissionDenied
| FileExists
export
Show FileEx where
show (GenericFileEx errno) = "File error: " ++ show errno
show FileReadError = "File Read Error"
show FileWriteError = "File Write Error"
show FileNotFound = "File Not Found"
show PermissionDenied = "Permission Denied"
show FileExists = "File Exists"
public export
data IOError
= GenericErr String
| FileErr FileEx
export
Show IOError where
show (GenericErr str) = "IO error: " ++ str
show (FileErr f) = "File error: " ++ show f

View File

@ -0,0 +1,17 @@
module Control.App.Console
import public Control.App
public export
interface Console e where
putStr : String -> App {l} e ()
getStr : App {l} e String
export
PrimIO e => Console e where
putStr str = primIO $ putStr str
getStr = primIO $ getLine
export
putStrLn : Console e => String -> App {l} e ()
putStrLn str = putStr (str ++ "\n")

View File

@ -0,0 +1,69 @@
module Control.App.FileIO
import Control.App
import Data.List
import Data.Strings
import System.File
toFileEx : FileError -> FileEx
toFileEx (GenericFileError i) = GenericFileEx i
toFileEx FileReadError = FileReadError
toFileEx FileWriteError = FileWriteError
toFileEx FileNotFound = FileNotFound
toFileEx PermissionDenied = PermissionDenied
toFileEx FileExists = FileExists
public export
interface Has [Exception IOError] e => FileIO e where
withFile : String -> Mode ->
(onError : IOError -> App e a) ->
(onOpen : File -> App e a) ->
App e a
fGetStr : File -> App e String
fPutStr : File -> String -> App e ()
fEOF : File -> App e Bool
-- TODO : Add Binary File IO with buffers
export
readFile : FileIO e => String -> App e String
readFile f
= withFile f Read throw $ \h =>
do content <- read [] h
pure (fastAppend content)
where
read : List String -> File -> App e (List String)
read acc h
= do eof <- FileIO.fEOF h
if eof
then pure (reverse acc)
else do str <- fGetStr h
read (str :: acc) h
export
Has [PrimIO, Exception IOError] e => FileIO e where
withFile fname m onError proc
= do Right h <- primIO $ openFile fname m
| Left err => onError (FileErr (toFileEx err))
res <- catch (proc h) onError
pure res
fGetStr f
= do Right str <- primIO $ fGetLine f
| Left err => throw (FileErr (toFileEx err))
pure str
fPutStr f str
= do Right () <- primIO $ File.fPutStr f str
| Left err => throw (FileErr (toFileEx err))
pure ()
fEOF f = primIO $ fEOF f
export
withFileIO : Has [PrimIO] e =>
App (IOError :: e) a ->
(ok : a -> App e b) ->
(err : IOError -> App e b) -> App e b
withFileIO prog ok err = handle prog ok err

View File

@ -0,0 +1,28 @@
module Control.App.Network
import Control.App
import Control.App.Network.Data
data SocketState = Ready | Bound | Listening | Open | Closed
data Socket : SocketState -> Type where
MkSocket : Int -> Socket s
data NetworkError = Foo
interface Has [Exception NetworkError] e => Network e where
newSocket : (1 p : (1 s : Socket Ready) -> App e a) -> App e a
done : (1 s : Socket Closed) -> App e ()
bind : (1 sock : Socket Ready) ->
(addr : Maybe SocketAddress) -> (port : Port) ->
Either NetworkError (Socket Bound)
{-
connect : (sock : Socket) -> (addr : SocketAddress) -> (port : Port) -> IO ResultCode
listen : (sock : Socket) -> IO Int
accept : (sock : Socket) -> IO (Either SocketError (Socket, SocketAddress))
send : (sock : Socket) -> (msg : String) -> IO (Either SocketError ResultCode)
recv : (sock : Socket) -> (len : ByteLength) -> IO (Either SocketError (String, ResultCode))
close : Socket -> IO ()
-}

View File

@ -0,0 +1,179 @@
module Control.App.PrimApp
import Data.List
import Data.IORef
public export
data PEffect : Type where
St : Type -> PEffect
Exc : Type -> PEffect
Sys : PEffect
public export
data HasPEff : PEffect -> List PEffect -> Type where
Here : HasPEff e (e :: es)
There : HasPEff e es -> HasPEff e (e' :: es)
public export
data NotException : PEffect -> Type where
StNotE : NotException (St t)
SysNotE : NotException Sys
public export
data NotState : PEffect -> Type where
ExcNotS : NotState (Exc t)
SysNotS : NotState Sys
public export
data Linear : List PEffect -> Type where
EmptyLin : Linear []
RestLin : NotException e => Linear es -> Linear (e :: es)
public export
data Stateless : List PEffect -> Type where
EmptySt : Stateless []
RestSt : NotState e => Stateless es -> Stateless (e :: es)
export
data Env : List PEffect -> Type where
None : Env []
Ref : IORef t -> Env es -> Env (St t :: es)
SkipE : Env es -> Env (Exc t :: es)
SkipP : Env es -> Env (Sys :: es)
getState : Env es -> (p : HasPEff (St t) es) -> IORef t
getState (Ref r env) Here = r
getState (Ref r env) (There p) = getState env p
getState (SkipE env) (There p) = getState env p
getState (SkipP env) (There p) = getState env p
public export
0 execTy : List PEffect -> Type -> Type
execTy [] ty = ty
execTy (St t :: es) ty = execTy es ty
execTy (Exc e :: es) ty = Either e (execTy es ty)
execTy (Sys :: es) ty = execTy es ty
public export
data PApp : List PEffect -> Type -> Type where
Pure : (x : a) -> PApp es a
Bind : PApp es a -> (a -> PApp es b) -> PApp es b
BindL : Linear es =>
(1 act : PApp es a) ->
(1 k : a -> PApp es b) -> PApp es b
New : a -> PApp (St a :: es) t -> PApp es t
Get : HasPEff (St t) es => PApp es t
Put : HasPEff (St t) es => t -> PApp es ()
Throw : HasPEff (Exc e) es => e -> PApp es a
Catch : HasPEff (Exc e) es =>
PApp es a ->
(err : e -> PApp es a) -> PApp es a
Handle : PApp (Exc e :: es) a ->
(ok : a -> PApp es b) ->
(err : e -> PApp es b) -> PApp es b
GetEnv : PApp es (Env es)
Fork : HasPEff Sys es => PApp es () -> PApp es ()
Prim : HasPEff Sys es => PrimIO a -> PApp es a
export
Functor (PApp es) where
map f ap = Bind ap $ \ap' => Pure (f ap')
export
Applicative (PApp es) where
pure = Pure
(<*>) f a
= Bind f $ \f' =>
Bind a $ \a' => pure (f' a')
export
Monad (PApp es) where
(>>=) = Bind
throwIn : Env es -> HasPEff (Exc e) es -> e ->
IO (execTy (es ++ rest) a)
throwIn (SkipE es) Here e = pure (Left e)
throwIn (Ref r es) (There p) e = throwIn es p e
throwIn (SkipE es) (There p) e
= do res <- throwIn es p e
pure (Right res)
throwIn (SkipP es) (There p) e = throwIn es p e
findRes : Env es -> HasPEff (Exc e) es -> execTy es a -> Maybe e
findRes (SkipE env) (There p) (Left _) = Nothing -- wrong exception
findRes (SkipE env) (There p) (Right r) = findRes env p r
findRes (Ref r env) (There p) t = findRes env p t
findRes (SkipP env) (There p) t = findRes env p t
findRes None _ _ = Nothing
findRes _ Here (Left ex) = Just ex
findRes _ Here _ = Nothing
copyEnv : Env es -> IO (Env es)
copyEnv None = pure None
copyEnv (Ref t env)
= do val <- readIORef t
t' <- newIORef val
env' <- copyEnv env
pure (Ref t' env')
copyEnv (SkipE env)
= do env' <- copyEnv env
pure (SkipE env')
copyEnv (SkipP env)
= do env' <- copyEnv env
pure (SkipP env')
clearEnv : Env es -> IO (execTy es ())
clearEnv None = pure ()
clearEnv (Ref t env) = clearEnv env
clearEnv (SkipE env)
= do e' <- clearEnv env
pure (Right e')
clearEnv (SkipP env) = clearEnv env
exec : Env es -> PApp es t -> (t -> IO (execTy es a)) -> IO (execTy es a)
exec env (Pure val) k = k val
exec env (Bind act next) k
= exec env act (\res => exec env (next res) k)
exec env (BindL act next) k
= exec env act (\res => exec env (next res) k)
exec env (New val prog) k
= do r <- newIORef val
exec (Ref r env) prog k
exec env (Get @{p}) k
= do let ref = getState env p
val <- readIORef ref
k val
exec env (Put @{p} val) k
= do let ref = getState env p
writeIORef ref val
k ()
exec env (Throw @{p} e) k
= rewrite sym (appendNilRightNeutral es) in
throwIn env p e {rest = []}
exec env (Handle prog ok err) k
= do res <- exec (SkipE env) prog
(\res => do res' <- exec env (ok res) k
pure (Right res'))
case res of
Left ex => exec env (err ex) k
Right ok => pure ok
exec env (Catch @{p} prog err) k
= do res <- exec env prog k
case findRes env p res of
Just ex => exec env (err ex) k
Nothing => pure res
exec env GetEnv k = k env
exec env (Fork proc) k
= do fork (do env' <- copyEnv env
exec env proc (\u => clearEnv env)
pure ())
k ()
exec env (Prim io) k
= do op <- primIO io
k op
export
run : PApp [Sys] a -> IO a
run prog = exec (SkipP None) prog pure

View File

@ -0,0 +1,265 @@
module Control.App
import Data.List
import Data.IORef
public export
data Effect : Type where
St : Type -> Effect
Exc : Type -> Effect
Sys : Effect
public export
data HasEff : Effect -> List Effect -> Type where
Here : HasEff e (e :: es)
There : HasEff e es -> HasEff e (e' :: es)
public export
0 Has : List (a -> Type) -> a -> Type
Has [] es = ()
Has [e] es = e es
Has (e :: es') es = (e es, Has es' es)
export
data Env : List Effect -> Type where
None : Env []
Ref : IORef t -> Env es -> Env (St t :: es)
SkipE : Env es -> Env (Exc t :: es)
SkipP : Env es -> Env (Sys :: es)
getState : Env es -> (p : HasEff (St t) es) -> IORef t
getState (Ref r env) Here = r
getState (Ref r env) (There p) = getState env p
getState (SkipE env) (There p) = getState env p
getState (SkipP env) (There p) = getState env p
public export
data Path : Type where
[noHints]
MayThrow : Path
NoThrow : Path
%hint public export
dpath : Path
dpath = MayThrow
data OneOf : List Type -> Path -> Type where
First : e -> OneOf (e :: es) MayThrow
Later : OneOf es MayThrow -> OneOf (e :: es) MayThrow
updateP : OneOf es p -> OneOf es p'
updateP {p=MayThrow} {p'=MayThrow} x = x
Uninhabited (OneOf [] p) where
uninhabited (First x) impossible
uninhabited (Later x) impossible
Uninhabited (OneOf es NoThrow) where
uninhabited (First x) impossible
uninhabited (Later x) impossible
public export
0 excTy : List Effect -> List Type
excTy [] = []
excTy (St t :: es) = excTy es
excTy (Exc e :: es) = e :: excTy es
excTy (Sys :: es) = excTy es
0 execTy : Path -> List Effect -> Type -> Type
execTy p es ty = Either (OneOf (excTy es) p) ty
export
data App : (l : Path) => (es : List Effect) -> Type -> Type where
MkApp : (1 prog : Env e -> IO (execTy l e t)) -> App {l} e t
pureApp : a -> App {l} e a
pureApp x = MkApp (\env => pure (Right x))
public export
data SafeBind : Path -> (l' : Path) -> Type where
[search l']
SafeSame : SafeBind l l
SafeToThrow : SafeBind NoThrow MayThrow
bindApp : SafeBind l l' =>
App {l} e a -> (a -> App {l=l'} e b) -> App {l=l'} e b
bindApp (MkApp prog) k
= MkApp $ \env =>
do Right res <- prog env
| Left err => pure (Left (updateP err))
let MkApp ka = k res
ka env
absurdWith : (1 x : a) -> OneOf e NoThrow -> any
absurdWith x (First p) impossible
export
bindL : App {l=NoThrow} e a -> (1 k : a -> App {l} e b) -> App {l} e b
bindL (MkApp prog) k
= MkApp $ \env =>
io_bind (prog env) $ \r =>
case r of
Left err => absurdWith k err
Right res =>
let MkApp ka = k res in ka env
export
lift : App e t -> App (eff :: e) t
lift (MkApp p)
= MkApp (\env =>
case env of
Ref r env' => p env'
SkipP env' => p env'
SkipE env' =>
do res <- p env'
case res of
Left err => pure (Left (Later err))
Right ok => pure (Right ok))
export
Functor (App {l} es) where
map f ap = bindApp ap $ \ap' => pureApp (f ap')
export
Applicative (App {l} es) where
pure = pureApp
(<*>) f a = bindApp f $ \f' =>
bindApp a $ \a' => pure (f' a')
export
Monad (App es) where
(>>=) = bindApp -- won't get used, but handy to have the instance
export
(>>=) : SafeBind l l' =>
App {l} e a -> (k : a -> App {l=l'} e b) -> App {l=l'} e b
(>>=) = bindApp
export
new : a -> App {l} (St a :: es) t -> App {l} es t
new val (MkApp prog)
= MkApp $ \env =>
do ref <- newIORef val
prog (Ref ref env)
public export
interface State t es where
get : App {l} es t
put : t -> App {l} es ()
export
HasEff (St t) es => State t es where
get
= MkApp $ \env =>
do let ref = getState env %search
val <- readIORef ref
pure (Right val)
put val
= MkApp $ \env =>
do let ref = getState env %search
writeIORef ref val
pure (Right ())
public export
interface Exception e es where
throw : e -> App es a
catch : App es a -> (err : e -> App es a) -> App es a
findException : Env es -> HasEff (Exc e) es -> e -> OneOf (excTy es) MayThrow
findException (SkipE env) Here err = First err
findException (Ref r env) (There p) err = findException env p err
findException (SkipE env) (There p) err = Later $ findException env p err
findException (SkipP env) (There p) err = findException env p err
findError : Env es -> HasEff (Exc e) es -> OneOf (excTy es) MayThrow -> Maybe e
findError (SkipE env) Here (First err) = Just err
findError (SkipE env) Here (Later p) = Nothing -- wrong exception
findError (SkipE env) (There p) (First err) = Nothing -- wrong exception
findError (SkipE env) (There p) (Later q) = findError env p q
findError (Ref r env) (There p) err = findError env p err
findError (SkipP env) (There p) err = findError env p err
export
HasEff (Exc e) es => Exception e es where
throw err
= MkApp $ \env =>
pure (Left (findException env %search err))
catch (MkApp prog) handler
= MkApp $ \env =>
do res <- prog env
case res of
Left err =>
case findError env %search err of
Nothing => pure (Left err)
Just err' =>
let MkApp e' = handler err' in
e' env
Right ok => pure (Right ok)
export
handle : App (Exc err :: e) a ->
(onok : a -> App e b) ->
(onerr : err -> App e b) -> App e b
handle (MkApp prog) onok onerr
= MkApp $ \env =>
do res <- prog (SkipE env)
case res of
Left (First err) =>
let MkApp err' = onerr err in
err' env
Left (Later p) =>
-- different exception, so rethrow
pure (Left p)
Right ok =>
let MkApp ok' = onok ok in
ok' env
public export
interface PrimIO es where
primIO : IO a -> App {l} es a
-- Copies the environment, to make sure states are local to threads
fork : App es () -> App {l} es ()
copyEnv : Env es -> IO (Env es)
copyEnv None = pure None
copyEnv (Ref t env)
= do val <- readIORef t
t' <- newIORef val
env' <- copyEnv env
pure (Ref t' env')
copyEnv (SkipE env)
= do env' <- copyEnv env
pure (SkipE env')
copyEnv (SkipP env)
= do env' <- copyEnv env
pure (SkipP env')
export
HasEff Sys es => PrimIO es where
primIO io
= MkApp $ \env =>
do res <- io
pure (Right res)
fork (MkApp p)
= MkApp $ \env =>
do fork (do env' <- copyEnv env
p env'
pure ())
pure (Right ())
public export
Init : List Effect
Init = [Sys]
export
run : App Init a -> IO a
run (MkApp prog)
= do Right res <- prog (SkipP None)
| Left err => absurd err
pure res
infix 5 @@
public export
data Res : (a : Type) -> (a -> Type) -> Type where
(@@) : (val : a) -> (1 r : t val) -> Res a t

View File

@ -0,0 +1,45 @@
module Control.Monad.Identity
public export
record Identity (a : Type) where
constructor Id
runIdentity : a
public export
Functor Identity where
map fn (Id a) = Id (fn a)
public export
Applicative Identity where
pure x = Id x
(Id f) <*> (Id g) = Id (f g)
public export
Monad Identity where
(Id x) >>= k = k x
public export
Show a => Show (Identity a) where
showPrec d (Id x) = showCon d "Id" $ showArg x
public export
Eq a => Eq (Identity a) where
(Id x) == (Id y) = x == y
public export
Ord a => Ord (Identity a) where
compare (Id x) (Id y) = compare x y
-- public export
-- Enum a => Enum (Identity a) where
-- toNat (Id x) = toNat x
-- fromNat n = Id $ fromNat n
-- pred (Id n) = Id $ pred n
public export
(Semigroup a) => Semigroup (Identity a) where
(<+>) x y = Id (runIdentity x <+> runIdentity y)
public export
(Monoid a) => Monoid (Identity a) where
neutral = Id (neutral)

View File

@ -0,0 +1,93 @@
module Control.Monad.State
import public Control.Monad.Identity
import public Control.Monad.Trans
||| A computation which runs in a context and produces an output
public export
interface Monad m => MonadState stateType (m : Type -> Type) | m where
||| Get the context
get : m stateType
||| Write a new context/output
put : stateType -> m ()
||| The transformer on which the State monad is based
public export
record StateT (stateType : Type) (m : Type -> Type) (a : Type) where
constructor ST
runStateT : stateType -> m (a, stateType)
public export
implementation Functor f => Functor (StateT stateType f) where
map f (ST g) = ST (\st => map (mapFst f) (g st)) where
mapFst : (a -> x) -> (a, s) -> (x, s)
mapFst fn (a, b) = (fn a, b)
public export
implementation Monad f => Applicative (StateT stateType f) where
pure x = ST (\st => pure (x, st))
(ST f) <*> (ST a)
= ST (\st =>
do (g, r) <- f st
(b, t) <- a r
pure (g b, t))
public export
implementation Monad m => Monad (StateT stateType m) where
(ST f) >>= k
= ST (\st =>
do (v, st') <- f st
let ST kv = k v
kv st')
public export
implementation Monad m => MonadState stateType (StateT stateType m) where
get = ST (\x => pure (x, x))
put x = ST (\y => pure ((), x))
public export
implementation MonadTrans (StateT stateType) where
lift x
= ST (\st =>
do r <- x
pure (r, st))
public export
implementation (Monad f, Alternative f) => Alternative (StateT st f) where
empty = lift empty
(ST f) <|> (ST g) = ST (\st => f st <|> g st)
||| Apply a function to modify the context of this computation
public export
modify : MonadState stateType m => (stateType -> stateType) -> m ()
modify f
= do s <- get
put (f s)
||| Evaluate a function in the context held by this computation
public export
gets : MonadState stateType m => (stateType -> a) -> m a
gets f
= do s <- get
pure (f s)
||| The State monad. See the MonadState interface
public export
State : (stateType : Type) -> (ty : Type) -> Type
State = \s, a => StateT s Identity a
||| Unwrap a State monad computation.
public export
runState : StateT stateType Identity a -> stateType -> (a, stateType)
runState act = runIdentity . runStateT act
||| Unwrap a State monad computation, but discard the final state.
public export
evalState : State stateType a -> stateType -> a
evalState m = fst . runState m
||| Unwrap a State monad computation, but discard the resulting value.
public export
execState : State stateType a -> stateType -> stateType
execState m = snd . runState m

View File

@ -0,0 +1,6 @@
module Control.Monad.Trans
public export
interface MonadTrans (t : (Type -> Type) -> Type -> Type) where
lift : Monad m => m a -> t m a

View File

@ -0,0 +1,61 @@
import Control.App2
namespace Console
public export
interface Console e where
putStr : String -> App e ()
getStr : App e String
export
PrimIO e => Console e where
putStr str = primIO $ putStr str
getStr = primIO $ getLine
export
putStrLn : Console e => String -> App e ()
putStrLn str = putStr (str ++ "\n")
namespace StateEx
public export
interface Console e => StateEx e where
inc : App e Int
testRes : String -> App e Bool
export
Has [Console, State Int, Exception String] e => StateEx e where
inc
= do count <- get
put (count + the Int 1)
pure count
testRes str
= do count <- get
case str of
"DONE" => pure True
"BAD" => throw "Nope"
_ => do putStrLn $ "Hello " ++ str ++ " " ++ show count
pure False
test : Has [StateEx] e => App e ()
test
= do putStr "Name: "
inc
x <- getStr
done <- testRes x
if done
then pure ()
else test
addState : Has [Console, StateEx] e => App e ()
addState
= do new "foo" $
do putStrLn "Here we go!"
str <- get
putStrLn str
runTest : IO ()
runTest
= run $ do new (the Int 0) $
handle (do test; addState) pure
(\err : String =>
putStrLn $ "Error: " ++ err)

View File

@ -0,0 +1,84 @@
module Control.WellFounded
import Data.Nat
import Data.List
public export
data Accessible : (rel : a -> a -> Type) -> (x : a) -> Type where
Access : (rec : (y : a) -> rel y x -> Accessible rel y) ->
Accessible rel x
public export
interface WellFounded (rel : a -> a -> Type) where
wellFounded : (x : a) -> Accessible rel x
export
accRec : {0 rel : (arg1 : a) -> (arg2 : a) -> Type} ->
(step : (x : a) -> ((y : a) -> rel y x -> b) -> b) ->
(z : a) -> (0 acc : Accessible rel z) -> b
accRec step z (Access f) =
step z $ \yarg, lt => accRec step yarg (f yarg lt)
export
accInd : {0 rel : a -> a -> Type} -> {0 P : a -> Type} ->
(step : (x : a) -> ((y : a) -> rel y x -> P y) -> P x) ->
(z : a) -> (0 acc : Accessible rel z) -> P z
accInd step z (Access f) =
step z $ \y, lt => accInd step y (f y lt)
export
wfRec : WellFounded rel =>
(step : (x : a) -> ((y : a) -> rel y x -> b) -> b) ->
a -> b
wfRec step x = accRec step x (wellFounded {rel} x)
export
wfInd : WellFounded rel => {0 P : a -> Type} ->
(step : (x : a) -> ((y : a) -> rel y x -> P y) -> P x) ->
(myz : a) -> P myz
wfInd step myz = accInd step myz (wellFounded {rel} myz)
public export
interface Sized a where
size : a -> Nat
public export
Smaller : Sized a => a -> a -> Type
Smaller x y = size x `LT` size y
public export
SizeAccessible : Sized a => a -> Type
SizeAccessible = Accessible Smaller
export
sizeAccessible : Sized a => (x : a) -> SizeAccessible x
sizeAccessible x = Access (acc $ size x)
where
acc : (sizeX : Nat) -> (y : a) -> (size y `LT` sizeX) -> SizeAccessible y
acc (S x') y (LTESucc yLEx')
= Access (\z, zLTy => acc x' z (lteTransitive zLTy yLEx'))
export
sizeInd : Sized a => {0 P : a -> Type} ->
(step : (x : a) -> ((y : a) -> Smaller y x -> P y) -> P x) ->
(z : a) ->
P z
sizeInd step z = accInd step z (sizeAccessible z)
export
sizeRec : Sized a =>
(step : (x : a) -> ((y : a) -> Smaller y x -> b) -> b) ->
(z : a) -> b
sizeRec step z = accRec step z (sizeAccessible z)
export
implementation Sized Nat where
size = \x => x
export
implementation Sized (List a) where
size = length
export
implementation (Sized a, Sized b) => Sized (Pair a b) where
size (x,y) = size x + size y

175
libs/base/Data/Buffer.idr Normal file
View File

@ -0,0 +1,175 @@
module Data.Buffer
import System.File
export
data Buffer : Type where [external]
%foreign "scheme:blodwen-buffer-size"
prim__bufferSize : Buffer -> Int
export
rawSize : Buffer -> IO Int
rawSize buf = pure (prim__bufferSize buf)
%foreign "scheme:blodwen-new-buffer"
prim__newBuffer : Int -> PrimIO Buffer
export
newBuffer : Int -> IO (Maybe Buffer)
newBuffer size
= do buf <- primIO (prim__newBuffer size)
pure $ Just buf
-- if prim__nullAnyPtr buf /= 0
-- then pure Nothing
-- else pure $ Just $ MkBuffer buf size 0
-- might be needed if we do this in C...
export
freeBuffer : Buffer -> IO ()
freeBuffer buf = pure ()
%foreign "scheme:blodwen-buffer-setbyte"
prim__setByte : Buffer -> Int -> Int -> PrimIO ()
-- Assumes val is in the range 0-255
export
setByte : Buffer -> (loc : Int) -> (val : Int) -> IO ()
setByte buf loc val
= primIO (prim__setByte buf loc val)
%foreign "scheme:blodwen-buffer-getbyte"
prim__getByte : Buffer -> Int -> PrimIO Int
export
getByte : Buffer -> (loc : Int) -> IO Int
getByte buf loc
= primIO (prim__getByte buf loc)
%foreign "scheme:blodwen-buffer-setint"
prim__setInt : Buffer -> Int -> Int -> PrimIO ()
export
setInt : Buffer -> (loc : Int) -> (val : Int) -> IO ()
setInt buf loc val
= primIO (prim__setInt buf loc val)
%foreign "scheme:blodwen-buffer-getint"
prim__getInt : Buffer -> Int -> PrimIO Int
export
getInt : Buffer -> (loc : Int) -> IO Int
getInt buf loc
= primIO (prim__getInt buf loc)
%foreign "scheme:blodwen-buffer-setdouble"
prim__setDouble : Buffer -> Int -> Double -> PrimIO ()
export
setDouble : Buffer -> (loc : Int) -> (val : Double) -> IO ()
setDouble buf loc val
= primIO (prim__setDouble buf loc val)
%foreign "scheme:blodwen-buffer-getdouble"
prim__getDouble : Buffer -> Int -> PrimIO Double
export
getDouble : Buffer -> (loc : Int) -> IO Double
getDouble buf loc
= primIO (prim__getDouble buf loc)
-- Get the length of a string in bytes, rather than characters
export
%foreign "C:strlen,libc 6"
stringByteLength : String -> Int
%foreign "scheme:blodwen-buffer-setstring"
prim__setString : Buffer -> Int -> String -> PrimIO ()
export
setString : Buffer -> (loc : Int) -> (val : String) -> IO ()
setString buf loc val
= primIO (prim__setString buf loc val)
%foreign "scheme:blodwen-buffer-getstring"
prim__getString : Buffer -> Int -> Int -> PrimIO String
export
getString : Buffer -> (loc : Int) -> (len : Int) -> IO String
getString buf loc len
= primIO (prim__getString buf loc len)
export
bufferData : Buffer -> IO (List Int)
bufferData buf
= do len <- rawSize buf
unpackTo [] len
where
unpackTo : List Int -> Int -> IO (List Int)
unpackTo acc 0 = pure acc
unpackTo acc loc
= do val <- getByte buf (loc - 1)
unpackTo (val :: acc) (loc - 1)
%foreign "scheme:blodwen-buffer-copydata"
prim__copyData : Buffer -> Int -> Int -> Buffer -> Int -> PrimIO ()
export
copyData : (src : Buffer) -> (start, len : Int) ->
(dest : Buffer) -> (loc : Int) -> IO ()
copyData src start len dest loc
= primIO (prim__copyData src start len dest loc)
-- %foreign "scheme:blodwen-readbuffer-bytes"
-- prim__readBufferBytes : FilePtr -> AnyPtr -> Int -> Int -> PrimIO Int
--
-- export
-- readBufferFromFile : BinaryFile -> Buffer -> (maxbytes : Int) ->
-- IO (Either FileError Buffer)
-- readBufferFromFile (FHandle h) (MkBuffer buf size loc) max
-- = do read <- primIO (prim__readBufferBytes h buf loc max)
-- if read >= 0
-- then pure (Right (MkBuffer buf size (loc + read)))
-- else pure (Left FileReadError)
%foreign "scheme:blodwen-read-bytevec"
prim__readBufferFromFile : String -> PrimIO Buffer
%foreign "scheme:blodwen-isbytevec"
prim__isBuffer : Buffer -> Int
-- Create a new buffer by reading all the contents from the given file
-- Fails if no bytes can be read or buffer can't be created
export
createBufferFromFile : String -> IO (Either FileError Buffer)
createBufferFromFile fn
= do buf <- primIO (prim__readBufferFromFile fn)
if prim__isBuffer buf /= 0
then pure (Left FileReadError)
else do let sz = prim__bufferSize buf
pure (Right buf)
%foreign "scheme:blodwen-write-bytevec"
prim__writeBuffer : String -> Buffer -> Int -> PrimIO Int
export
writeBufferToFile : String -> Buffer -> (maxbytes : Int) ->
IO (Either FileError ())
writeBufferToFile fn buf max
= do res <- primIO (prim__writeBuffer fn buf max)
if res /= 0
then pure (Left FileWriteError)
else pure (Right ())
export
resizeBuffer : Buffer -> Int -> IO (Maybe Buffer)
resizeBuffer old newsize
= do Just buf <- newBuffer newsize
| Nothing => pure Nothing
-- If the new buffer is smaller than the old one, just copy what
-- fits
oldsize <- rawSize old
let len = if newsize < oldsize then newsize else oldsize
copyData old 0 len buf 0
freeBuffer old
pure (Just buf)

84
libs/base/Data/Either.idr Normal file
View File

@ -0,0 +1,84 @@
module Data.Either
||| True if the argument is Left, False otherwise
public export
isLeft : Either a b -> Bool
isLeft (Left l) = True
isLeft (Right r) = False
||| True if the argument is Right, False otherwise
public export
isRight : Either a b -> Bool
isRight (Left l) = False
isRight (Right r) = True
||| Keep the payloads of all Left constructors in a list of Eithers
public export
lefts : List (Either a b) -> List a
lefts [] = []
lefts (x::xs) =
case x of
Left l => l :: lefts xs
Right r => lefts xs
||| Keep the payloads of all Right constructors in a list of Eithers
public export
rights : List (Either a b) -> List b
rights [] = []
rights (x::xs) =
case x of
Left l => rights xs
Right r => r :: rights xs
||| Split a list of Eithers into a list of the left elements and a list of the right elements
public export
partitionEithers : List (Either a b) -> (List a, List b)
partitionEithers l = (lefts l, rights l)
||| Remove a "useless" Either by collapsing the case distinction
public export
fromEither : Either a a -> a
fromEither (Left l) = l
fromEither (Right r) = r
||| Right becomes left and left becomes right
public export
mirror : Either a b -> Either b a
mirror (Left x) = Right x
mirror (Right x) = Left x
--------------------------------------------------------------------------------
-- Conversions
--------------------------------------------------------------------------------
||| Convert a Maybe to an Either by using a default value in case of Nothing
||| @ e the default value
public export
maybeToEither : (def : Lazy e) -> Maybe a -> Either e a
maybeToEither def (Just j) = Right j
maybeToEither def Nothing = Left def
||| Convert an Either to a Maybe from Right injection
public export
eitherToMaybe : Either e a -> Maybe a
eitherToMaybe (Left _) = Nothing
eitherToMaybe (Right x) = Just x
-- Injectivity of constructors
||| Left is injective
total
leftInjective : {b : Type}
-> {x : a}
-> {y : a}
-> (Left {b = b} x = Left {b = b} y) -> (x = y)
leftInjective Refl = Refl
||| Right is injective
total
rightInjective : {a : Type}
-> {x : b}
-> {y : b}
-> (Right {a = a} x = Right {a = a} y) -> (x = y)
rightInjective Refl = Refl

173
libs/base/Data/Fin.idr Normal file
View File

@ -0,0 +1,173 @@
module Data.Fin
import Data.Maybe
import Data.Nat
import Decidable.Equality
||| Numbers strictly less than some bound. The name comes from "finite sets".
|||
||| It's probably not a good idea to use `Fin` for arithmetic, and they will be
||| exceedingly inefficient at run time.
||| @ n the upper bound
public export
data Fin : (n : Nat) -> Type where
FZ : Fin (S k)
FS : Fin k -> Fin (S k)
export
implementation Uninhabited (Fin Z) where
uninhabited FZ impossible
uninhabited (FS f) impossible
export
FSInjective : (m : Fin k) -> (n : Fin k) -> FS m = FS n -> m = n
FSInjective left _ Refl = Refl
export
implementation Eq (Fin n) where
(==) FZ FZ = True
(==) (FS k) (FS k') = k == k'
(==) _ _ = False
||| There are no elements of `Fin Z`
export
FinZAbsurd : Fin Z -> Void
FinZAbsurd FZ impossible
export
FinZElim : Fin Z -> a
FinZElim x = void (FinZAbsurd x)
||| Convert a Fin to a Nat
public export
finToNat : Fin n -> Nat
finToNat FZ = Z
finToNat (FS k) = S (finToNat k)
||| `finToNat` is injective
export
finToNatInjective : (fm : Fin k) -> (fn : Fin k) -> (finToNat fm) = (finToNat fn) -> fm = fn
finToNatInjective (FS m) FZ Refl impossible
finToNatInjective FZ (FS n) Refl impossible
finToNatInjective (FS m) (FS n) prf =
cong FS (finToNatInjective m n (succInjective (finToNat m) (finToNat n) prf))
finToNatInjective FZ FZ Refl = Refl
export
implementation Cast (Fin n) Nat where
cast x = finToNat x
||| Convert a Fin to an Integer
public export
finToInteger : Fin n -> Integer
finToInteger FZ = 0
finToInteger (FS k) = 1 + finToInteger k
export
implementation Cast (Fin n) Integer where
cast x = finToInteger x
||| Weaken the bound on a Fin by 1
public export
weaken : Fin n -> Fin (S n)
weaken FZ = FZ
weaken (FS k) = FS (weaken k)
||| Weaken the bound on a Fin by some amount
public export
weakenN : (n : Nat) -> Fin m -> Fin (m + n)
weakenN n FZ = FZ
weakenN n (FS f) = FS (weakenN n f)
||| Attempt to tighten the bound on a Fin.
||| Return `Left` if the bound could not be tightened, or `Right` if it could.
export
strengthen : {n : _} -> Fin (S n) -> Either (Fin (S n)) (Fin n)
strengthen {n = S k} FZ = Right FZ
strengthen {n = S k} (FS i) with (strengthen i)
strengthen (FS i) | Left x = Left (FS x)
strengthen (FS i) | Right x = Right (FS x)
strengthen f = Left f
||| Add some natural number to a Fin, extending the bound accordingly
||| @ n the previous bound
||| @ m the number to increase the Fin by
public export
shift : (m : Nat) -> Fin n -> Fin (m + n)
shift Z f = f
shift {n=n} (S m) f = FS {k = (m + n)} (shift m f)
||| The largest element of some Fin type
public export
last : {n : _} -> Fin (S n)
last {n=Z} = FZ
last {n=S _} = FS last
public export total
FSinjective : {f : Fin n} -> {f' : Fin n} -> (FS f = FS f') -> f = f'
FSinjective Refl = Refl
export
implementation Ord (Fin n) where
compare FZ FZ = EQ
compare FZ (FS _) = LT
compare (FS _) FZ = GT
compare (FS x) (FS y) = compare x y
-- Construct a Fin from an integer literal which must fit in the given Fin
public export
natToFin : Nat -> (n : Nat) -> Maybe (Fin n)
natToFin Z (S j) = Just FZ
natToFin (S k) (S j)
= case natToFin k j of
Just k' => Just (FS k')
Nothing => Nothing
natToFin _ _ = Nothing
||| Convert an `Integer` to a `Fin`, provided the integer is within bounds.
||| @n The upper bound of the Fin
public export
integerToFin : Integer -> (n : Nat) -> Maybe (Fin n)
integerToFin x Z = Nothing -- make sure 'n' is concrete, to save reduction!
integerToFin x n = if x >= 0 then natToFin (fromInteger x) n else Nothing
||| Allow overloading of Integer literals for Fin.
||| @ x the Integer that the user typed
||| @ prf an automatically-constructed proof that `x` is in bounds
public export
fromInteger : (x : Integer) -> {n : Nat} ->
{auto prf : (IsJust (integerToFin x n))} ->
Fin n
fromInteger {n} x {prf} with (integerToFin x n)
fromInteger {n} x {prf = ItIsJust} | Just y = y
||| Convert an Integer to a Fin in the required bounds/
||| This is essentially a composition of `mod` and `fromInteger`
public export
restrict : (n : Nat) -> Integer -> Fin (S n)
restrict n val = let val' = assert_total (abs (mod val (cast (S n)))) in
-- reasoning about primitives, so we need the
-- 'believe_me'. It's fine because val' must be
-- in the right range
fromInteger {n = S n} val'
{prf = believe_me {a=IsJust (Just val')} ItIsJust}
--------------------------------------------------------------------------------
-- DecEq
--------------------------------------------------------------------------------
export total
FZNotFS : {f : Fin n} -> FZ {k = n} = FS f -> Void
FZNotFS Refl impossible
export
implementation DecEq (Fin n) where
decEq FZ FZ = Yes Refl
decEq FZ (FS f) = No FZNotFS
decEq (FS f) FZ = No $ negEqSym FZNotFS
decEq (FS f) (FS f')
= case decEq f f' of
Yes p => Yes $ cong FS p
No p => No $ \h => p $ FSinjective {f = f} {f' = f'} h

View File

@ -0,0 +1,84 @@
module Data.IOArray
import Data.List
-- Implemented externally
data ArrayData : Type -> Type where
-- 'unsafe' primitive access, backend dependent
-- get and set assume that the bounds have been checked. Behavious is undefined
-- otherwise.
%extern prim__newArray : forall a . Int -> a -> PrimIO (ArrayData a)
%extern prim__arrayGet : forall a . ArrayData a -> Int -> PrimIO a
%extern prim__arraySet : forall a . ArrayData a -> Int -> a -> PrimIO ()
export
record IOArray elem where
constructor MkIOArray
max : Int
content : ArrayData (Maybe elem)
export
newArray : Int -> IO (IOArray elem)
newArray size
= pure (MkIOArray size !(primIO (prim__newArray size Nothing)))
export
writeArray : IOArray elem -> Int -> elem -> IO ()
writeArray arr pos el
= if pos < 0 || pos >= max arr
then pure ()
else primIO (prim__arraySet (content arr) pos (Just el))
export
readArray : IOArray elem -> Int -> IO (Maybe elem)
readArray arr pos
= if pos < 0 || pos >= max arr
then pure Nothing
else primIO (prim__arrayGet (content arr) pos)
-- Make a new array of the given size with the elements copied from the
-- other array
export
newArrayCopy : (newsize : Int) -> IOArray elem -> IO (IOArray elem)
newArrayCopy newsize arr
= do let newsize' = if newsize < max arr then max arr else newsize
arr' <- newArray newsize'
copyFrom (content arr) (content arr') (max arr - 1)
pure arr'
where
copyFrom : ArrayData (Maybe elem) ->
ArrayData (Maybe elem) ->
Int -> IO ()
copyFrom old new pos
= if pos < 0
then pure ()
else do el <- primIO $ prim__arrayGet old pos
primIO $ prim__arraySet new pos el
assert_total (copyFrom old new (pos - 1))
export
toList : IOArray elem -> IO (List (Maybe elem))
toList arr = iter 0 (max arr) []
where
iter : Int -> Int -> List (Maybe elem) -> IO (List (Maybe elem))
iter pos end acc
= if pos >= end
then pure (reverse acc)
else do el <- readArray arr pos
assert_total (iter (pos + 1) end (el :: acc))
export
fromList : List (Maybe elem) -> IO (IOArray elem)
fromList ns
= do arr <- newArray (cast (length ns))
addToArray 0 ns arr
pure arr
where
addToArray : Int -> List (Maybe elem) -> IOArray elem -> IO ()
addToArray loc [] arr = pure ()
addToArray loc (Nothing :: ns) arr
= assert_total (addToArray (loc + 1) ns arr)
addToArray loc (Just el :: ns) arr
= do primIO $ prim__arraySet (content arr) loc (Just el)
assert_total (addToArray (loc + 1) ns arr)

34
libs/base/Data/IORef.idr Normal file
View File

@ -0,0 +1,34 @@
module Data.IORef
-- Implemented externally
-- e.g., in Scheme, passed around as a box
data Mut : Type -> Type where [external]
%extern prim__newIORef : forall a . a -> (1 x : %World) -> IORes (Mut a)
%extern prim__readIORef : forall a . Mut a -> (1 x : %World) -> IORes a
%extern prim__writeIORef : forall a . Mut a -> (1 val : a) -> (1 x : %World) -> IORes ()
export
data IORef : Type -> Type where
MkRef : Mut a -> IORef a
export
newIORef : a -> IO (IORef a)
newIORef val
= do m <- primIO (prim__newIORef val)
pure (MkRef m)
export
readIORef : IORef a -> IO a
readIORef (MkRef m) = primIO (prim__readIORef m)
export
writeIORef : IORef a -> (1 val : a) -> IO ()
writeIORef (MkRef m) val = primIO (prim__writeIORef m val)
export
modifyIORef : IORef a -> (a -> a) -> IO ()
modifyIORef ref f
= do val <- readIORef ref
writeIORef ref (f val)

585
libs/base/Data/List.idr Normal file
View File

@ -0,0 +1,585 @@
module Data.List
import Decidable.Equality
public export
isNil : List a -> Bool
isNil [] = True
isNil (x::xs) = False
public export
isCons : List a -> Bool
isCons [] = False
isCons (x::xs) = True
public export
length : List a -> Nat
length [] = Z
length (x::xs) = S (length xs)
public export
take : Nat -> List a -> List a
take Z xs = []
take (S k) [] = []
take (S k) (x :: xs) = x :: take k xs
public export
drop : (n : Nat) -> (xs : List a) -> List a
drop Z xs = xs
drop (S n) [] = []
drop (S n) (x::xs) = drop n xs
public export
takeWhile : (p : a -> Bool) -> List a -> List a
takeWhile p [] = []
takeWhile p (x::xs) = if p x then x :: takeWhile p xs else []
public export
dropWhile : (p : a -> Bool) -> List a -> List a
dropWhile p [] = []
dropWhile p (x::xs) = if p x then dropWhile p xs else x::xs
public export
filter : (p : a -> Bool) -> List a -> List a
filter p [] = []
filter p (x :: xs)
= if p x
then x :: filter p xs
else filter p xs
||| Find associated information in a list using a custom comparison.
public export
lookupBy : (a -> a -> Bool) -> a -> List (a, b) -> Maybe b
lookupBy p e [] = Nothing
lookupBy p e (x::xs) =
let (l, r) = x in
if p e l then
Just r
else
lookupBy p e xs
||| Find associated information in a list using Boolean equality.
public export
lookup : Eq a => a -> List (a, b) -> Maybe b
lookup = lookupBy (==)
||| Check if something is a member of a list using a custom comparison.
public export
elemBy : (a -> a -> Bool) -> a -> List a -> Bool
elemBy p e [] = False
elemBy p e (x::xs) =
if p e x then
True
else
elemBy p e xs
public export
nubBy : (a -> a -> Bool) -> List a -> List a
nubBy = nubBy' []
where
nubBy' : List a -> (a -> a -> Bool) -> List a -> List a
nubBy' acc p [] = []
nubBy' acc p (x::xs) =
if elemBy p x acc then
nubBy' acc p xs
else
x :: nubBy' (x::acc) p xs
||| O(n^2). The nub function removes duplicate elements from a list. In
||| particular, it keeps only the first occurrence of each element. It is a
||| special case of nubBy, which allows the programmer to supply their own
||| equality test.
|||
||| ```idris example
||| nub (the (List _) [1,2,1,3])
||| ```
public export
nub : Eq a => List a -> List a
nub = nubBy (==)
||| The deleteBy function behaves like delete, but takes a user-supplied equality predicate.
public export
deleteBy : (a -> a -> Bool) -> a -> List a -> List a
deleteBy _ _ [] = []
deleteBy eq x (y::ys) = if x `eq` y then ys else y :: deleteBy eq x ys
||| `delete x` removes the first occurrence of `x` from its list argument. For
||| example,
|||
|||````idris example
|||delete 'a' ['b', 'a', 'n', 'a', 'n', 'a']
|||````
|||
||| It is a special case of deleteBy, which allows the programmer to supply
||| their own equality test.
public export
delete : Eq a => a -> List a -> List a
delete = deleteBy (==)
||| The unionBy function returns the union of two lists by user-supplied equality predicate.
public export
unionBy : (a -> a -> Bool) -> List a -> List a -> List a
unionBy eq xs ys = xs ++ foldl (flip (deleteBy eq)) (nubBy eq ys) xs
||| Compute the union of two lists according to their `Eq` implementation.
|||
||| ```idris example
||| union ['d', 'o', 'g'] ['c', 'o', 'w']
||| ```
|||
public export
union : Eq a => List a -> List a -> List a
union = unionBy (==)
public export
span : (a -> Bool) -> List a -> (List a, List a)
span p [] = ([], [])
span p (x::xs) =
if p x then
let (ys, zs) = span p xs in
(x::ys, zs)
else
([], x::xs)
public export
break : (a -> Bool) -> List a -> (List a, List a)
break p xs = span (not . p) xs
public export
split : (a -> Bool) -> List a -> List (List a)
split p xs =
case break p xs of
(chunk, []) => [chunk]
(chunk, (c :: rest)) => chunk :: split p rest
public export
splitAt : (n : Nat) -> (xs : List a) -> (List a, List a)
splitAt Z xs = ([], xs)
splitAt (S k) [] = ([], [])
splitAt (S k) (x :: xs)
= let (tk, dr) = splitAt k xs in
(x :: tk, dr)
public export
partition : (a -> Bool) -> List a -> (List a, List a)
partition p [] = ([], [])
partition p (x::xs) =
let (lefts, rights) = partition p xs in
if p x then
(x::lefts, rights)
else
(lefts, x::rights)
||| The inits function returns all initial segments of the argument, shortest
||| first. For example,
|||
||| ```idris example
||| inits [1,2,3]
||| ```
public export
inits : List a -> List (List a)
inits xs = [] :: case xs of
[] => []
x :: xs' => map (x ::) (inits xs')
||| The tails function returns all final segments of the argument, longest
||| first. For example,
|||
||| ```idris example
||| tails [1,2,3] == [[1,2,3], [2,3], [3], []]
|||```
public export
tails : List a -> List (List a)
tails xs = xs :: case xs of
[] => []
_ :: xs' => tails xs'
||| Split on the given element.
|||
||| ```idris example
||| splitOn 0 [1,0,2,0,0,3]
||| ```
|||
public export
splitOn : Eq a => a -> List a -> List (List a)
splitOn a = split (== a)
||| Replaces all occurences of the first argument with the second argument in a list.
|||
||| ```idris example
||| replaceOn '-' ',' ['1', '-', '2', '-', '3']
||| ```
|||
public export
replaceOn : Eq a => a -> a -> List a -> List a
replaceOn a b l = map (\c => if c == a then b else c) l
public export
reverseOnto : List a -> List a -> List a
reverseOnto acc [] = acc
reverseOnto acc (x::xs) = reverseOnto (x::acc) xs
public export
reverse : List a -> List a
reverse = reverseOnto []
||| Construct a list with `n` copies of `x`.
||| @ n how many copies
||| @ x the element to replicate
public export
replicate : (n : Nat) -> (x : a) -> List a
replicate Z _ = []
replicate (S n) x = x :: replicate n x
||| Compute the intersect of two lists by user-supplied equality predicate.
export
intersectBy : (a -> a -> Bool) -> List a -> List a -> List a
intersectBy eq xs ys = [x | x <- xs, any (eq x) ys]
||| Compute the intersect of two lists according to the `Eq` implementation for the elements.
export
intersect : Eq a => List a -> List a -> List a
intersect = intersectBy (==)
||| Combine two lists elementwise using some function.
|||
||| If the lists are different lengths, the result is truncated to the
||| length of the shortest list.
export
zipWith : (a -> b -> c) -> List a -> List b -> List c
zipWith _ [] _ = []
zipWith _ _ [] = []
zipWith f (x::xs) (y::ys) = f x y :: zipWith f xs ys
||| Combine two lists elementwise into pairs.
|||
||| If the lists are different lengths, the result is truncated to the
||| length of the shortest list.
export
zip : List a -> List b -> List (a, b)
zip = zipWith \x, y => (x, y)
export
zipWith3 : (a -> b -> c -> d) -> List a -> List b -> List c -> List d
zipWith3 _ [] _ _ = []
zipWith3 _ _ [] _ = []
zipWith3 _ _ _ [] = []
zipWith3 f (x::xs) (y::ys) (z::zs) = f x y z :: zipWith3 f xs ys zs
||| Combine three lists elementwise into tuples.
|||
||| If the lists are different lengths, the result is truncated to the
||| length of the shortest list.
export
zip3 : List a -> List b -> List c -> List (a, b, c)
zip3 = zipWith3 \x, y, z => (x, y, z)
public export
data NonEmpty : (xs : List a) -> Type where
IsNonEmpty : NonEmpty (x :: xs)
export
Uninhabited (NonEmpty []) where
uninhabited IsNonEmpty impossible
||| Get the head of a non-empty list.
||| @ ok proof the list is non-empty
public export
head : (l : List a) -> {auto ok : NonEmpty l} -> a
head [] impossible
head (x :: xs) = x
||| Get the tail of a non-empty list.
||| @ ok proof the list is non-empty
public export
tail : (l : List a) -> {auto ok : NonEmpty l} -> List a
tail [] impossible
tail (x :: xs) = xs
||| Attempt to get the head of a list. If the list is empty, return `Nothing`.
export
head' : List a -> Maybe a
head' [] = Nothing
head' (x::xs) = Just x
||| Attempt to get the tail of a list. If the list is empty, return `Nothing`.
export
tail' : List a -> Maybe (List a)
tail' [] = Nothing
tail' (x::xs) = Just xs
||| Convert any Foldable structure to a list.
export
toList : Foldable t => t a -> List a
toList = foldr (::) []
||| Prefix every element in the list with the given element
|||
||| ```idris example
||| with List (mergeReplicate '>' ['a', 'b', 'c', 'd', 'e'])
||| ```
|||
export
mergeReplicate : a -> List a -> List a
mergeReplicate sep [] = []
mergeReplicate sep (y::ys) = sep :: y :: mergeReplicate sep ys
||| Insert some separator between the elements of a list.
|||
||| ````idris example
||| with List (intersperse ',' ['a', 'b', 'c', 'd', 'e'])
||| ````
|||
export
intersperse : a -> List a -> List a
intersperse sep [] = []
intersperse sep (x::xs) = x :: mergeReplicate sep xs
||| Given a separator list and some more lists, produce a new list by
||| placing the separator between each of the lists.
|||
||| @ sep the separator
||| @ xss the lists between which the separator will be placed
|||
||| ```idris example
||| intercalate [0, 0, 0] [ [1, 2, 3], [4, 5, 6], [7, 8, 9] ]
||| ```
export
intercalate : (sep : List a) -> (xss : List (List a)) -> List a
intercalate sep xss = concat $ intersperse sep xss
||| Apply a partial function to the elements of a list, keeping the ones at which
||| it is defined.
export
mapMaybe : (a -> Maybe b) -> List a -> List b
mapMaybe f [] = []
mapMaybe f (x::xs) =
case f x of
Nothing => mapMaybe f xs
Just j => j :: mapMaybe f xs
--------------------------------------------------------------------------------
-- Sorting
--------------------------------------------------------------------------------
||| Check whether a list is sorted with respect to the default ordering for the type of its elements.
export
sorted : Ord a => List a -> Bool
sorted [] = True
sorted (x::xs) =
case xs of
Nil => True
(y::ys) => x <= y && sorted (y::ys)
||| Merge two sorted lists using an arbitrary comparison
||| predicate. Note that the lists must have been sorted using this
||| predicate already.
export
mergeBy : (a -> a -> Ordering) -> List a -> List a -> List a
mergeBy order [] right = right
mergeBy order left [] = left
mergeBy order (x::xs) (y::ys) =
-- The code below will emit `y` before `x` whenever `x == y`.
-- If you change this, `sortBy` will stop being stable, unless you fix `sortBy`, too.
case order x y of
LT => x :: mergeBy order xs (y::ys)
_ => y :: mergeBy order (x::xs) ys
||| Merge two sorted lists using the default ordering for the type of their elements.
export
merge : Ord a => List a -> List a -> List a
merge left right = mergeBy compare left right
||| Sort a list using some arbitrary comparison predicate.
|||
||| @ cmp how to compare elements
||| @ xs the list to sort
export
sortBy : (cmp : a -> a -> Ordering) -> (xs : List a) -> List a
sortBy cmp [] = []
sortBy cmp [x] = [x]
sortBy cmp xs = let (x, y) = split xs in
mergeBy cmp
(sortBy cmp (assert_smaller xs x))
(sortBy cmp (assert_smaller xs y)) -- not structurally smaller, hence assert
where
splitRec : List b -> List a -> (List a -> List a) -> (List a, List a)
splitRec (_::_::xs) (y::ys) zs = splitRec xs ys (zs . ((::) y))
splitRec _ ys zs = (ys, zs [])
-- In the above base-case clause, we put `ys` on the LHS to get a stable sort.
-- This is because `mergeBy` prefers taking elements from its RHS operand
-- if both heads are equal, and all elements in `zs []` precede all elements of `ys`
-- in the original list.
split : List a -> (List a, List a)
split xs = splitRec xs xs id
||| Sort a list using the default ordering for the type of its elements.
export
sort : Ord a => List a -> List a
sort = sortBy compare
export
isPrefixOfBy : (eq : a -> a -> Bool) -> (left, right : List a) -> Bool
isPrefixOfBy p [] right = True
isPrefixOfBy p left [] = False
isPrefixOfBy p (x::xs) (y::ys) =
if p x y then
isPrefixOfBy p xs ys
else
False
||| The isPrefixOf function takes two lists and returns True iff the first list is a prefix of the second.
export
isPrefixOf : Eq a => List a -> List a -> Bool
isPrefixOf = isPrefixOfBy (==)
export
isSuffixOfBy : (a -> a -> Bool) -> List a -> List a -> Bool
isSuffixOfBy p left right = isPrefixOfBy p (reverse left) (reverse right)
||| The isSuffixOf function takes two lists and returns True iff the first list is a suffix of the second.
export
isSuffixOf : Eq a => List a -> List a -> Bool
isSuffixOf = isSuffixOfBy (==)
||| The isInfixOf function takes two lists and returns True iff the first list
||| is contained, wholly and intact, anywhere within the second.
|||
||| ```idris example
||| isInfixOf ['b','c'] ['a', 'b', 'c', 'd']
||| ```
||| ```idris example
||| isInfixOf ['b','d'] ['a', 'b', 'c', 'd']
||| ```
|||
export
isInfixOf : Eq a => List a -> List a -> Bool
isInfixOf n h = any (isPrefixOf n) (tails h)
||| Transposes rows and columns of a list of lists.
|||
||| ```idris example
||| with List transpose [[1, 2], [3, 4]]
||| ```
|||
||| This also works for non square scenarios, thus
||| involution does not always hold:
|||
||| transpose [[], [1, 2]] = [[1], [2]]
||| transpose (transpose [[], [1, 2]]) = [[1, 2]]
export
transpose : List (List a) -> List (List a)
transpose [] = []
transpose (heads :: tails) = spreadHeads heads (transpose tails) where
spreadHeads : List a -> List (List a) -> List (List a)
spreadHeads [] tails = tails
spreadHeads (head :: heads) [] = [head] :: spreadHeads heads []
spreadHeads (head :: heads) (tail :: tails) = (head :: tail) :: spreadHeads heads tails
--------------------------------------------------------------------------------
-- Properties
--------------------------------------------------------------------------------
export
Uninhabited ([] = Prelude.(::) x xs) where
uninhabited Refl impossible
export
Uninhabited (Prelude.(::) x xs = []) where
uninhabited Refl impossible
--
-- ||| (::) is injective
-- consInjective : {x : a} -> {xs : List a} -> {y : b} -> {ys : List b} ->
-- (x :: xs) = (y :: ys) -> (x = y, xs = ys)
-- consInjective Refl = (Refl, Refl)
--
-- ||| Two lists are equal, if their heads are equal and their tails are equal.
-- consCong2 : {x : a} -> {xs : List a} -> {y : b} -> {ys : List b} ->
-- x = y -> xs = ys -> x :: xs = y :: ys
-- consCong2 Refl Refl = Refl
--
-- ||| Appending pairwise equal lists gives equal lists
-- appendCong2 : {x1 : List a} -> {x2 : List a} ->
-- {y1 : List b} -> {y2 : List b} ->
-- x1 = y1 -> x2 = y2 -> x1 ++ x2 = y1 ++ y2
-- appendCong2 {x1=[]} {y1=(_ :: _)} Refl _ impossible
-- appendCong2 {x1=(_ :: _)} {y1=[]} Refl _ impossible
-- appendCong2 {x1=[]} {y1=[]} _ eq2 = eq2
-- appendCong2 {x1=(_ :: _)} {y1=(_ :: _)} eq1 eq2 =
-- consCong2
-- (fst $ consInjective eq1)
-- (appendCong2 (snd $ consInjective eq1) eq2)
--
-- ||| List.map is distributive over appending.
-- mapAppendDistributive : (f : a -> b) -> (x : List a) -> (y : List a) ->
-- map f (x ++ y) = map f x ++ map f y
-- mapAppendDistributive _ [] _ = Refl
-- mapAppendDistributive f (_ :: xs) y = cong $ mapAppendDistributive f xs y
--
||| The empty list is a right identity for append.
export
appendNilRightNeutral : (l : List a) ->
l ++ [] = l
appendNilRightNeutral [] = Refl
appendNilRightNeutral (x::xs) =
let inductiveHypothesis = appendNilRightNeutral xs in
rewrite inductiveHypothesis in Refl
||| Appending lists is associative.
export
appendAssociative : (l : List a) -> (c : List a) -> (r : List a) ->
l ++ (c ++ r) = (l ++ c) ++ r
appendAssociative [] c r = Refl
appendAssociative (x::xs) c r =
let inductiveHypothesis = appendAssociative xs c r in
rewrite inductiveHypothesis in Refl
revOnto : (xs, vs : _) -> reverseOnto xs vs = reverse vs ++ xs
revOnto xs [] = Refl
revOnto xs (v :: vs)
= rewrite revOnto (v :: xs) vs in
rewrite appendAssociative (reverse vs) [v] xs in
rewrite revOnto [v] vs in Refl
export
revAppend : (vs, ns : List a) -> reverse ns ++ reverse vs = reverse (vs ++ ns)
revAppend [] ns = rewrite appendNilRightNeutral (reverse ns) in Refl
revAppend (v :: vs) ns
= rewrite revOnto [v] vs in
rewrite revOnto [v] (vs ++ ns) in
rewrite sym (revAppend vs ns) in
rewrite appendAssociative (reverse ns) (reverse vs) [v] in
Refl
public export
lemma_val_not_nil : {x : t} -> {xs : List t} -> ((x :: xs) = Prelude.Nil {a = t} -> Void)
lemma_val_not_nil Refl impossible
public export
lemma_x_eq_xs_neq : {x : t} -> {xs : List t} -> {y : t} -> {ys : List t} -> (x = y) -> (xs = ys -> Void) -> ((x :: xs) = (y :: ys) -> Void)
lemma_x_eq_xs_neq Refl p Refl = p Refl
public export
lemma_x_neq_xs_eq : {x : t} -> {xs : List t} -> {y : t} -> {ys : List t} -> (x = y -> Void) -> (xs = ys) -> ((x :: xs) = (y :: ys) -> Void)
lemma_x_neq_xs_eq p Refl Refl = p Refl
public export
lemma_x_neq_xs_neq : {x : t} -> {xs : List t} -> {y : t} -> {ys : List t} -> (x = y -> Void) -> (xs = ys -> Void) -> ((x :: xs) = (y :: ys) -> Void)
lemma_x_neq_xs_neq p p' Refl = p Refl
public export
implementation DecEq a => DecEq (List a) where
decEq [] [] = Yes Refl
decEq (x :: xs) [] = No lemma_val_not_nil
decEq [] (x :: xs) = No (negEqSym lemma_val_not_nil)
decEq (x :: xs) (y :: ys) with (decEq x y)
decEq (x :: xs) (x :: ys) | Yes Refl with (decEq xs ys)
decEq (x :: xs) (x :: xs) | (Yes Refl) | (Yes Refl) = Yes Refl
decEq (x :: xs) (x :: ys) | (Yes Refl) | (No p) = No (\eq => lemma_x_eq_xs_neq Refl p eq)
decEq (x :: xs) (y :: ys) | No p with (decEq xs ys)
decEq (x :: xs) (y :: xs) | (No p) | (Yes Refl) = No (\eq => lemma_x_neq_xs_eq p Refl eq)
decEq (x :: xs) (y :: ys) | (No p) | (No p') = No (\eq => lemma_x_neq_xs_neq p p' eq)

View File

@ -0,0 +1,76 @@
module Data.List.Elem
import Decidable.Equality
--------------------------------------------------------------------------------
-- List membership proof
--------------------------------------------------------------------------------
||| A proof that some element is found in a list.
public export
data Elem : a -> List a -> Type where
||| A proof that the element is at the head of the list
Here : Elem x (x :: xs)
||| A proof that the element is in the tail of the list
There : Elem x xs -> Elem x (y :: xs)
export
Uninhabited (Here = There e) where
uninhabited Refl impossible
export
Uninhabited (There e = Here) where
uninhabited Refl impossible
export
thereInjective : {0 e1, e2 : Elem x xs} -> There e1 = There e2 -> e1 = e2
thereInjective Refl = Refl
export
DecEq (Elem x xs) where
decEq Here Here = Yes Refl
decEq Here (There later) = No absurd
decEq (There later) Here = No absurd
decEq (There this) (There that) with (decEq this that)
decEq (There this) (There this) | Yes Refl = Yes Refl
decEq (There this) (There that) | No contra = No (contra . thereInjective)
export
Uninhabited (Elem {a} x []) where
uninhabited Here impossible
uninhabited (There p) impossible
||| An item not in the head and not in the tail is not in the list at all.
export
neitherHereNorThere : Not (x = y) -> Not (Elem x xs) -> Not (Elem x (y :: xs))
neitherHereNorThere xny xnxs Here = xny Refl
neitherHereNorThere xny xnxs (There xxs) = xnxs xxs
||| Check whether the given element is a member of the given list.
export
isElem : DecEq a => (x : a) -> (xs : List a) -> Dec (Elem x xs)
isElem x [] = No absurd
isElem x (y :: xs) with (decEq x y)
isElem x (x :: xs) | Yes Refl = Yes Here
isElem x (y :: xs) | No xny with (isElem x xs)
isElem x (y :: xs) | No xny | Yes xxs = Yes (There xxs)
isElem x (y :: xs) | No xny | No xnxs = No (neitherHereNorThere xny xnxs)
||| Remove the element at the given position.
public export
dropElem : (xs : List a) -> (p : Elem x xs) -> List a
dropElem (x :: ys) Here = ys
dropElem (x :: ys) (There p) = x :: dropElem ys p
||| Erase the indices, returning the numeric position of the element
public export
elemToNat : Elem x xs -> Nat
elemToNat Here = Z
elemToNat (There p) = S (elemToNat p)
||| Find the element with a proof at a given position, if it is valid
public export
indexElem : Nat -> (xs : List a) -> Maybe (x ** Elem x xs)
indexElem _ [] = Nothing
indexElem Z (y :: ys) = Just (y ** Here)
indexElem (S n) (y :: ys) = map (\(x ** p) => (x ** There p)) (indexElem n ys)

View File

@ -0,0 +1,95 @@
module Data.List.Views
import Control.WellFounded
import Data.List
import Data.Nat.Views
lengthSuc : (xs : List a) -> (y : a) -> (ys : List a) ->
length (xs ++ (y :: ys)) = S (length (xs ++ ys))
lengthSuc [] _ _ = Refl
lengthSuc (x :: xs) y ys = cong S (lengthSuc xs y ys)
lengthLT : (xs : List a) -> (ys : List a) ->
LTE (length xs) (length (ys ++ xs))
lengthLT xs [] = lteRefl
lengthLT xs (x :: ys) = lteSuccRight (lengthLT _ _)
smallerLeft : (ys : List a) -> (y : a) -> (zs : List a) ->
LTE (S (S (length ys))) (S (length (ys ++ (y :: zs))))
smallerLeft [] y zs = LTESucc (LTESucc LTEZero)
smallerLeft (z :: ys) y zs = LTESucc (smallerLeft ys _ _)
smallerRight : (ys : List a) -> (zs : List a) ->
LTE (S (S (length zs))) (S (length (ys ++ (y :: zs))))
smallerRight {y} ys zs = rewrite lengthSuc ys y zs in
(LTESucc (LTESucc (lengthLT _ _)))
||| View for splitting a list in half, non-recursively
public export
data Split : List a -> Type where
SplitNil : Split []
SplitOne : (x : a) -> Split [x]
SplitPair : (x : a) -> (xs : List a) ->
(y : a) -> (ys : List a) ->
Split (x :: xs ++ y :: ys)
splitHelp : (head : a) ->
(xs : List a) ->
(counter : List a) -> Split (head :: xs)
splitHelp head [] counter = SplitOne _
splitHelp head (x :: xs) [] = SplitPair head [] x xs
splitHelp head (x :: xs) [y] = SplitPair head [] x xs
splitHelp head (x :: xs) (_ :: _ :: ys)
= case splitHelp head xs ys of
SplitOne x => SplitPair x [] _ []
SplitPair x' xs y' ys => SplitPair x' (x :: xs) y' ys
||| Covering function for the `Split` view
||| Constructs the view in linear time
export
split : (xs : List a) -> Split xs
split [] = SplitNil
split (x :: xs) = splitHelp x xs xs
public export
data SplitRec : List a -> Type where
SplitRecNil : SplitRec []
SplitRecOne : (x : a) -> SplitRec [x]
SplitRecPair : (lefts, rights : List a) -> -- Explicit, don't erase
(lrec : Lazy (SplitRec lefts)) ->
(rrec : Lazy (SplitRec rights)) -> SplitRec (lefts ++ rights)
||| Covering function for the `SplitRec` view
||| Constructs the view in O(n lg n)
public export total
splitRec : (xs : List a) -> SplitRec xs
splitRec xs with (sizeAccessible xs)
splitRec xs | acc with (split xs)
splitRec [] | acc | SplitNil = SplitRecNil
splitRec [x] | acc | SplitOne x = SplitRecOne x
splitRec (y :: ys ++ z :: zs) | Access acc | SplitPair y ys z zs
= SplitRecPair _ _
(splitRec (y :: ys) | acc _ (smallerLeft ys z zs))
(splitRec (z :: zs) | acc _ (smallerRight ys zs))
||| View for traversing a list backwards
public export
data SnocList : List a -> Type where
Empty : SnocList []
Snoc : (x : a) -> (xs : List a) ->
(rec : SnocList xs) -> SnocList (xs ++ [x])
snocListHelp : {input : _} ->
SnocList input -> (rest : List a) -> SnocList (input ++ rest)
snocListHelp snoc [] = rewrite appendNilRightNeutral input in snoc
snocListHelp snoc (x :: xs)
= rewrite appendAssociative input [x] xs in
snocListHelp (Snoc x input snoc) xs
||| Covering function for the `SnocList` view
||| Constructs the view in linear time
export
snocList : (xs : List a) -> SnocList xs
snocList xs = snocListHelp Empty xs

57
libs/base/Data/Maybe.idr Normal file
View File

@ -0,0 +1,57 @@
module Data.Maybe
public export
isNothing : Maybe a -> Bool
isNothing Nothing = True
isNothing (Just j) = False
public export
isJust : Maybe a -> Bool
isJust Nothing = False
isJust (Just j) = True
||| Proof that some `Maybe` is actually `Just`
public export
data IsJust : Maybe a -> Type where
ItIsJust : IsJust (Just x)
export
Uninhabited (IsJust Nothing) where
uninhabited ItIsJust impossible
||| Decide whether a 'Maybe' is 'Just'
public export
isItJust : (v : Maybe a) -> Dec (IsJust v)
isItJust (Just x) = Yes ItIsJust
isItJust Nothing = No absurd
||| Convert a `Maybe a` value to an `a` value by providing a default `a` value
||| in the case that the `Maybe` value is `Nothing`.
public export
fromMaybe : (Lazy a) -> Maybe a -> a
fromMaybe def Nothing = def
fromMaybe def (Just j) = j
||| Returns `Just` the given value if the conditional is `True`
||| and `Nothing` if the conditional is `False`.
public export
toMaybe : Bool -> Lazy a -> Maybe a
toMaybe True j = Just j
toMaybe False j = Nothing
export
justInjective : {x : t} -> {y : t} -> (Just x = Just y) -> x = y
justInjective Refl = Refl
||| Convert a `Maybe a` value to an `a` value, using `neutral` in the case
||| that the `Maybe` value is `Nothing`.
public export
lowerMaybe : Monoid a => Maybe a -> a
lowerMaybe Nothing = neutral
lowerMaybe (Just x) = x
||| Returns `Nothing` when applied to `neutral`, and `Just` the value otherwise.
export
raiseToMaybe : (Monoid a, Eq a) => a -> Maybe a
raiseToMaybe x = if x == neutral then Nothing else Just x

View File

@ -0,0 +1,91 @@
module Data.Morphisms
public export
record Morphism a b where
constructor Mor
applyMor : a -> b
infixr 1 ~>
export
(~>) : Type -> Type -> Type
(~>) = Morphism
public export
record Endomorphism a where
constructor Endo
applyEndo : a -> a
public export
record Kleislimorphism (f : Type -> Type) a b where
constructor Kleisli
applyKleisli : a -> f b
export
Functor (Morphism r) where
map f (Mor a) = Mor $ f . a
export
Applicative (Morphism r) where
pure a = Mor $ const a
(Mor f) <*> (Mor a) = Mor $ \r => f r $ a r
export
Monad (Morphism r) where
(Mor h) >>= f = Mor $ \r => applyMor (f $ h r) r
export
Semigroup a => Semigroup (Morphism r a) where
f <+> g = Mor $ \r => (applyMor f) r <+> (applyMor g) r
export
Monoid a => Monoid (Morphism r a) where
neutral = Mor \r => neutral
export
Semigroup (Endomorphism a) where
(Endo f) <+> (Endo g) = Endo $ g . f
export
Monoid (Endomorphism a) where
neutral = Endo id
export
Functor f => Functor (Kleislimorphism f a) where
map f (Kleisli g) = Kleisli (map f . g)
export
Applicative f => Applicative (Kleislimorphism f a) where
pure a = Kleisli $ const $ pure a
(Kleisli f) <*> (Kleisli a) = Kleisli $ \r => f r <*> a r
export
Monad f => Monad (Kleislimorphism f a) where
(Kleisli f) >>= g = Kleisli $ \r => do
k1 <- f r
applyKleisli (g k1) r
-- Applicative is a bit too strong, but there is no suitable superclass
export
(Semigroup a, Applicative f) => Semigroup (Kleislimorphism f r a) where
f <+> g = Kleisli \r => (<+>) <$> (applyKleisli f) r <*> (applyKleisli g) r
export
(Monoid a, Applicative f) => Monoid (Kleislimorphism f r a) where
neutral = Kleisli \r => pure neutral
export
Cast (Endomorphism a) (Morphism a a) where
cast (Endo f) = Mor f
export
Cast (Morphism a a) (Endomorphism a) where
cast (Mor f) = Endo f
export
Cast (Morphism a (f b)) (Kleislimorphism f a b) where
cast (Mor f) = Kleisli f
export
Cast (Kleislimorphism f a b) (Morphism a (f b)) where
cast (Kleisli f) = Mor f

417
libs/base/Data/Nat.idr Normal file
View File

@ -0,0 +1,417 @@
module Data.Nat
export
Uninhabited (Z = S n) where
uninhabited Refl impossible
export
Uninhabited (S n = Z) where
uninhabited Refl impossible
public export
isZero : Nat -> Bool
isZero Z = True
isZero (S n) = False
public export
isSucc : Nat -> Bool
isSucc Z = False
isSucc (S n) = True
public export
data IsSucc : (n : Nat) -> Type where
ItIsSucc : IsSucc (S n)
export
Uninhabited (IsSucc Z) where
uninhabited ItIsSucc impossible
public export
isItSucc : (n : Nat) -> Dec (IsSucc n)
isItSucc Z = No absurd
isItSucc (S n) = Yes ItIsSucc
public export
power : Nat -> Nat -> Nat
power base Z = S Z
power base (S exp) = base * (power base exp)
public export
hyper : Nat -> Nat -> Nat -> Nat
hyper Z a b = S b
hyper (S Z) a Z = a
hyper (S(S Z)) a Z = Z
hyper n a Z = S Z
hyper (S pn) a (S pb) = hyper pn a (hyper (S pn) a pb)
public export
pred : Nat -> Nat
pred Z = Z
pred (S n) = n
-- Comparisons
public export
data NotBothZero : (n, m : Nat) -> Type where
LeftIsNotZero : NotBothZero (S n) m
RightIsNotZero : NotBothZero n (S m)
public export
data LTE : (n, m : Nat) -> Type where
LTEZero : LTE Z right
LTESucc : LTE left right -> LTE (S left) (S right)
export
Uninhabited (LTE (S n) Z) where
uninhabited LTEZero impossible
public export
GTE : Nat -> Nat -> Type
GTE left right = LTE right left
public export
LT : Nat -> Nat -> Type
LT left right = LTE (S left) right
public export
GT : Nat -> Nat -> Type
GT left right = LT right left
export
succNotLTEzero : Not (LTE (S m) Z)
succNotLTEzero LTEZero impossible
export
fromLteSucc : LTE (S m) (S n) -> LTE m n
fromLteSucc (LTESucc x) = x
export
isLTE : (m, n : Nat) -> Dec (LTE m n)
isLTE Z n = Yes LTEZero
isLTE (S k) Z = No succNotLTEzero
isLTE (S k) (S j)
= case isLTE k j of
No contra => No (contra . fromLteSucc)
Yes prf => Yes (LTESucc prf)
export
lteRefl : {n : Nat} -> LTE n n
lteRefl {n = Z} = LTEZero
lteRefl {n = S k} = LTESucc lteRefl
export
lteSuccRight : LTE n m -> LTE n (S m)
lteSuccRight LTEZero = LTEZero
lteSuccRight (LTESucc x) = LTESucc (lteSuccRight x)
export
lteSuccLeft : LTE (S n) m -> LTE n m
lteSuccLeft (LTESucc x) = lteSuccRight x
export
lteTransitive : LTE n m -> LTE m p -> LTE n p
lteTransitive LTEZero y = LTEZero
lteTransitive (LTESucc x) (LTESucc y) = LTESucc (lteTransitive x y)
export
lteAddRight : (n : Nat) -> LTE n (n + m)
lteAddRight Z = LTEZero
lteAddRight (S k) {m} = LTESucc (lteAddRight {m} k)
export
notLTImpliesGTE : {a, b : _} -> Not (LT a b) -> GTE a b
notLTImpliesGTE {b = Z} _ = LTEZero
notLTImpliesGTE {a = Z} {b = S k} notLt = absurd (notLt (LTESucc LTEZero))
notLTImpliesGTE {a = S k} {b = S j} notLt = LTESucc (notLTImpliesGTE (notLt . LTESucc))
public export
lte : Nat -> Nat -> Bool
lte Z right = True
lte left Z = False
lte (S left) (S right) = lte left right
public export
gte : Nat -> Nat -> Bool
gte left right = lte right left
public export
lt : Nat -> Nat -> Bool
lt left right = lte (S left) right
public export
gt : Nat -> Nat -> Bool
gt left right = lt right left
public export
minimum : Nat -> Nat -> Nat
minimum Z m = Z
minimum (S n) Z = Z
minimum (S n) (S m) = S (minimum n m)
public export
maximum : Nat -> Nat -> Nat
maximum Z m = m
maximum (S n) Z = S n
maximum (S n) (S m) = S (maximum n m)
-- Proofs on S
export
eqSucc : (left : Nat) -> (right : Nat) -> (p : left = right) ->
S left = S right
eqSucc left _ Refl = Refl
export
succInjective : (left : Nat) -> (right : Nat) -> (p : S left = S right) ->
left = right
succInjective left _ Refl = Refl
export
SIsNotZ : (S x = Z) -> Void
SIsNotZ Refl impossible
export
modNatNZ : Nat -> (y: Nat) -> Not (y = Z) -> Nat
modNatNZ left Z p = void (p Refl)
modNatNZ left (S right) _ = mod' left left right
where
mod' : Nat -> Nat -> Nat -> Nat
mod' Z centre right = centre
mod' (S left) centre right =
if lte centre right then
centre
else
mod' left (minus centre (S right)) right
export
modNat : Nat -> Nat -> Nat
modNat left (S right) = modNatNZ left (S right) SIsNotZ
export
divNatNZ : Nat -> (y: Nat) -> Not (y = Z) -> Nat
divNatNZ left Z p = void (p Refl)
divNatNZ left (S right) _ = div' left left right
where
div' : Nat -> Nat -> Nat -> Nat
div' Z centre right = Z
div' (S left) centre right =
if lte centre right then
Z
else
S (div' left (minus centre (S right)) right)
export
divNat : Nat -> Nat -> Nat
divNat left (S right) = divNatNZ left (S right) SIsNotZ
export
divCeilNZ : Nat -> (y: Nat) -> Not (y = Z) -> Nat
divCeilNZ x y p = case (modNatNZ x y p) of
Z => divNatNZ x y p
S _ => S (divNatNZ x y p)
export
divCeil : Nat -> Nat -> Nat
divCeil x (S y) = divCeilNZ x (S y) SIsNotZ
public export
Integral Nat where
div = divNat
mod = modNat
export
gcd : (a: Nat) -> (b: Nat) -> {auto ok: NotBothZero a b} -> Nat
gcd a Z = a
gcd Z b = b
gcd a (S b) = gcd (S b) (modNatNZ a (S b) SIsNotZ)
export
lcm : Nat -> Nat -> Nat
lcm _ Z = Z
lcm Z _ = Z
lcm a (S b) = divNat (a * (S b)) (gcd a (S b))
--------------------------------------------------------------------------------
-- An informative comparison view
--------------------------------------------------------------------------------
public export
data CmpNat : Nat -> Nat -> Type where
CmpLT : (y : _) -> CmpNat x (x + S y)
CmpEQ : CmpNat x x
CmpGT : (x : _) -> CmpNat (y + S x) y
export
total cmp : (x, y : Nat) -> CmpNat x y
cmp Z Z = CmpEQ
cmp Z (S k) = CmpLT _
cmp (S k) Z = CmpGT _
cmp (S x) (S y) with (cmp x y)
cmp (S x) (S (x + (S k))) | CmpLT k = CmpLT k
cmp (S x) (S x) | CmpEQ = CmpEQ
cmp (S (y + (S k))) (S y) | CmpGT k = CmpGT k
-- Proofs on
export
plusZeroLeftNeutral : (right : Nat) -> 0 + right = right
plusZeroLeftNeutral right = Refl
export
plusZeroRightNeutral : (left : Nat) -> left + 0 = left
plusZeroRightNeutral Z = Refl
plusZeroRightNeutral (S n) =
let inductiveHypothesis = plusZeroRightNeutral n in
rewrite inductiveHypothesis in Refl
export
plusSuccRightSucc : (left : Nat) -> (right : Nat) ->
S (left + right) = left + (S right)
plusSuccRightSucc Z right = Refl
plusSuccRightSucc (S left) right =
let inductiveHypothesis = plusSuccRightSucc left right in
rewrite inductiveHypothesis in Refl
export
plusCommutative : (left : Nat) -> (right : Nat) ->
left + right = right + left
plusCommutative Z right = rewrite plusZeroRightNeutral right in Refl
plusCommutative (S left) right =
let inductiveHypothesis = plusCommutative left right in
rewrite inductiveHypothesis in
rewrite plusSuccRightSucc right left in Refl
export
plusAssociative : (left : Nat) -> (centre : Nat) -> (right : Nat) ->
left + (centre + right) = (left + centre) + right
plusAssociative Z centre right = Refl
plusAssociative (S left) centre right =
let inductiveHypothesis = plusAssociative left centre right in
rewrite inductiveHypothesis in Refl
export
plusConstantRight : (left : Nat) -> (right : Nat) -> (c : Nat) ->
(p : left = right) -> left + c = right + c
plusConstantRight left _ c Refl = Refl
export
plusConstantLeft : (left : Nat) -> (right : Nat) -> (c : Nat) ->
(p : left = right) -> c + left = c + right
plusConstantLeft left _ c Refl = Refl
export
plusOneSucc : (right : Nat) -> 1 + right = S right
plusOneSucc n = Refl
export
plusLeftCancel : (left : Nat) -> (right : Nat) -> (right' : Nat) ->
(p : left + right = left + right') -> right = right'
plusLeftCancel Z right right' p = p
plusLeftCancel (S left) right right' p =
let inductiveHypothesis = plusLeftCancel left right right' in
inductiveHypothesis (succInjective _ _ p)
export
plusRightCancel : (left : Nat) -> (left' : Nat) -> (right : Nat) ->
(p : left + right = left' + right) -> left = left'
plusRightCancel left left' Z p = rewrite sym (plusZeroRightNeutral left) in
rewrite sym (plusZeroRightNeutral left') in
p
plusRightCancel left left' (S right) p =
plusRightCancel left left' right
(succInjective _ _ (rewrite plusSuccRightSucc left right in
rewrite plusSuccRightSucc left' right in p))
export
plusLeftLeftRightZero : (left : Nat) -> (right : Nat) ->
(p : left + right = left) -> right = Z
plusLeftLeftRightZero Z right p = p
plusLeftLeftRightZero (S left) right p =
plusLeftLeftRightZero left right (succInjective _ _ p)
-- Proofs on *
export
multZeroLeftZero : (right : Nat) -> Z * right = Z
multZeroLeftZero right = Refl
export
multZeroRightZero : (left : Nat) -> left * Z = Z
multZeroRightZero Z = Refl
multZeroRightZero (S left) = multZeroRightZero left
export
multRightSuccPlus : (left : Nat) -> (right : Nat) ->
left * (S right) = left + (left * right)
multRightSuccPlus Z right = Refl
multRightSuccPlus (S left) right =
let inductiveHypothesis = multRightSuccPlus left right in
rewrite inductiveHypothesis in
rewrite plusAssociative left right (left * right) in
rewrite plusAssociative right left (left * right) in
rewrite plusCommutative right left in
Refl
export
multLeftSuccPlus : (left : Nat) -> (right : Nat) ->
(S left) * right = right + (left * right)
multLeftSuccPlus left right = Refl
export
multCommutative : (left : Nat) -> (right : Nat) ->
left * right = right * left
multCommutative Z right = rewrite multZeroRightZero right in Refl
multCommutative (S left) right =
let inductiveHypothesis = multCommutative left right in
rewrite inductiveHypothesis in
rewrite multRightSuccPlus right left in
Refl
export
multDistributesOverPlusRight : (left : Nat) -> (centre : Nat) -> (right : Nat) ->
left * (centre + right) = (left * centre) + (left * right)
multDistributesOverPlusRight Z centre right = Refl
multDistributesOverPlusRight (S left) centre right =
let inductiveHypothesis = multDistributesOverPlusRight left centre right in
rewrite inductiveHypothesis in
rewrite plusAssociative (centre + (left * centre)) right (left * right) in
rewrite sym (plusAssociative centre (left * centre) right) in
rewrite plusCommutative (left * centre) right in
rewrite plusAssociative centre right (left * centre) in
rewrite plusAssociative (centre + right) (left * centre) (left * right) in
Refl
export
multDistributesOverPlusLeft : (left : Nat) -> (centre : Nat) -> (right : Nat) ->
(left + centre) * right = (left * right) + (centre * right)
multDistributesOverPlusLeft Z centre right = Refl
multDistributesOverPlusLeft (S left) centre right =
let inductiveHypothesis = multDistributesOverPlusLeft left centre right in
rewrite inductiveHypothesis in
rewrite plusAssociative right (left * right) (centre * right) in
Refl
export
multAssociative : (left : Nat) -> (centre : Nat) -> (right : Nat) ->
left * (centre * right) = (left * centre) * right
multAssociative Z centre right = Refl
multAssociative (S left) centre right =
let inductiveHypothesis = multAssociative left centre right in
rewrite inductiveHypothesis in
rewrite multDistributesOverPlusLeft centre (left * centre) right in
Refl
export
multOneLeftNeutral : (right : Nat) -> 1 * right = right
multOneLeftNeutral Z = Refl
multOneLeftNeutral (S right) =
let inductiveHypothesis = multOneLeftNeutral right in
rewrite inductiveHypothesis in
Refl
export
multOneRightNeutral : (left : Nat) -> left * 1 = left
multOneRightNeutral Z = Refl
multOneRightNeutral (S left) =
let inductiveHypothesis = multOneRightNeutral left in
rewrite inductiveHypothesis in
Refl

View File

@ -0,0 +1,36 @@
module Data.Nat.Views
import Control.WellFounded
||| View for dividing a Nat in half
public export
data Half : Nat -> Type where
HalfOdd : (n : Nat) -> Half (S (n + n))
HalfEven : (n : Nat) -> Half (n + n)
||| View for dividing a Nat in half, recursively
public export
data HalfRec : Nat -> Type where
HalfRecZ : HalfRec Z
HalfRecEven : (n : Nat) -> (rec : Lazy (HalfRec n)) -> HalfRec (n + n)
HalfRecOdd : (n : Nat) -> (rec : Lazy (HalfRec n)) -> HalfRec (S (n + n))
||| Covering function for the `Half` view
public export
half : (n : Nat) -> Half n
half Z = HalfEven Z
half (S k) with (half k)
half (S (S (n + n))) | HalfOdd n = rewrite plusSuccRightSucc (S n) n in
HalfEven (S n)
half (S (n + n)) | HalfEven n = HalfOdd n
public export total
halfRec : (n : Nat) -> HalfRec n
halfRec n with (sizeAccessible n)
halfRec Z | acc = HalfRecZ
halfRec (S n) | acc with (half n)
halfRec (S (S (k + k))) | Access acc | HalfOdd k
= rewrite plusSuccRightSucc (S k) k
in HalfRecEven _ (halfRec (S k) | acc (S k) (LTESucc (LTESucc (lteAddRight _))))
halfRec (S (k + k)) | Access acc | HalfEven k
= HalfRecOdd _ (halfRec k | acc k (LTESucc (lteAddRight _)))

View File

@ -0,0 +1,82 @@
module Data.Primitives.Views
-- We need all the believe_mes and asserts throughout this file because we're
-- working with primitive here! We also have separate implementations per
-- primitive, rather than using interfaces, because we're only going to trust
-- the primitive implementations.
namespace IntegerV
||| View for expressing a number as a multiplication + a remainder
public export
data Divides : Integer -> (d : Integer) -> Type where
DivByZero : Divides x 0
DivBy : (div, rem : _) ->
(prf : rem >= 0 && rem < d = True) ->
Divides ((d * div) + rem) d
||| Covering function for the `Divides` view
public export
divides : (val : Integer) -> (d : Integer) -> Divides val d
divides val 0 = DivByZero
divides val d
= assert_total $
let dividend = if d < 0 then -(val `div` abs d)
else val `div` d
remainder = abs (val - dividend * d) in
believe_me (DivBy {d} dividend remainder
(believe_me (Refl {x = True})))
||| View for recursion over Integers
public export
data IntegerRec : Integer -> Type where
IntegerZ : IntegerRec 0
IntegerSucc : IntegerRec (n - 1) -> IntegerRec n
IntegerPred : IntegerRec ((-n) + 1) -> IntegerRec (-n)
||| Covering function for `IntegerRec`
public export
integerRec : (x : Integer) -> IntegerRec x
integerRec 0 = IntegerZ
integerRec x = if x > 0 then IntegerSucc (assert_total (integerRec (x - 1)))
else believe_me (IntegerPred {n = -x}
(assert_total (believe_me (integerRec (x + 1)))))
namespace IntV
||| View for expressing a number as a multiplication + a remainder
public export
data Divides : Int -> (d : Int) -> Type where
DivByZero : IntV.Divides x 0
DivBy : (div, rem : _) ->
(prf : rem >= 0 && rem < d = True) ->
IntV.Divides ((d * div) + rem) d
-- I have assumed, but not actually verified, that this will still
-- give a right result (i.e. still adding up) when the Ints overflow.
-- TODO: Someone please check this and fix if necessary...
||| Covering function for the `Divides` view
public export
divides : (val : Int) -> (d : Int) -> Divides val d
divides val 0 = DivByZero
divides val d
= assert_total $
let dividend = if d < 0 then -(val `div` abs d)
else val `div` d
remainder = abs (val - dividend * d) in
believe_me (DivBy {d} dividend remainder
(believe_me (Refl {x = True})))
||| View for recursion over Ints
public export
data IntRec : Int -> Type where
IntZ : IntRec 0
IntSucc : IntRec (n - 1) -> IntRec n
IntPred : IntRec ((-n) + 1) -> IntRec (-n)
||| Covering function for `IntRec`
public export
intRec : (x : Int) -> IntRec x
intRec 0 = IntZ
intRec x = if x > 0 then IntSucc (assert_total (intRec (x - 1)))
else believe_me (IntPred {n = -x}
(assert_total (believe_me (intRec (x + 1)))))

25
libs/base/Data/So.idr Normal file
View File

@ -0,0 +1,25 @@
module Data.So
||| Ensure that some run-time Boolean test has been performed.
|||
||| This lifts a Boolean predicate to the type level. See the function `choose`
||| if you need to perform a Boolean test and convince the type checker of this
||| fact.
|||
||| If you find yourself using `So` for something other than primitive types,
||| it may be appropriate to define a type of evidence for the property that you
||| care about instead.
public export
data So : Bool -> Type where
Oh : So True
export
implementation Uninhabited (So False) where
uninhabited Oh impossible
||| Perform a case analysis on a Boolean, providing clients with a `So` proof
export
choose : (b : Bool) -> Either (So b) (So (not b))
choose True = Left Oh
choose False = Right Oh

116
libs/base/Data/Stream.idr Normal file
View File

@ -0,0 +1,116 @@
module Data.Stream
import Data.List
||| The first element of an infinite stream
public export
head : Stream a -> a
head (x::xs) = x
||| Drop the first n elements from the stream
||| @ n how many elements to drop
public export
drop : (n : Nat) -> Stream a -> Stream a
drop Z xs = xs
drop (S k) (x::xs) = drop k xs
||| An infinite stream of repetitions of the same thing
public export
repeat : a -> Stream a
repeat x = x :: repeat x
||| Generate an infinite stream by repeatedly applying a function
||| @ f the function to iterate
||| @ x the initial value that will be the head of the stream
public export
iterate : (f : a -> a) -> (x : a) -> Stream a
iterate f x = x :: iterate f (f x)
||| Get the nth element of a stream
public export
index : Nat -> Stream a -> a
index Z (x::xs) = x
index (S k) (x::xs) = index k xs
||| Combine two streams element-wise using a function.
|||
||| @ f the function to combine elements with
||| @ xs the first stream of elements
||| @ ys the second stream of elements
export
zipWith : (f : a -> b -> c) -> (xs : Stream a) -> (ys : Stream b) -> Stream c
zipWith f (x::xs) (y::ys) = f x y :: zipWith f xs ys
||| Combine three streams by applying a function element-wise along them
export
zipWith3 : (a -> b -> c -> d) -> Stream a -> Stream b -> Stream c -> Stream d
zipWith3 f (x::xs) (y::ys) (z::zs) = f x y z :: zipWith3 f xs ys zs
||| Create a stream of pairs from two streams
export
zip : Stream a -> Stream b -> Stream (a, b)
zip = zipWith (\x,y => (x,y))
||| Combine three streams into a stream of tuples elementwise
export
zip3 : Stream a -> Stream b -> Stream c -> Stream (a, b, c)
zip3 = zipWith3 (\x,y,z => (x,y,z))
||| Create a pair of streams from a stream of pairs
export
unzip : Stream (a, b) -> (Stream a, Stream b)
unzip xs = (map fst xs, map snd xs)
||| Split a stream of three-element tuples into three streams
export
unzip3 : Stream (a, b, c) -> (Stream a, Stream b, Stream c)
unzip3 xs = (map (\(x,_,_) => x) xs, map (\(_,x,_) => x) xs, map (\(_,_,x) => x) xs)
||| Return the diagonal elements of a stream of streams
export
diag : Stream (Stream a) -> Stream a
diag ((x::xs)::xss) = x :: diag (map tail xss)
||| Produce a Stream of left folds of prefixes of the given Stream
||| @ f the combining function
||| @ acc the initial value
||| @ xs the Stream to process
export
scanl : (f : a -> b -> a) -> (acc : a) -> (xs : Stream b) -> Stream a
scanl f acc (x :: xs) = acc :: scanl f (f acc x) xs
||| Produce a Stream repeating a sequence
||| @ xs the sequence to repeat
||| @ ok proof that the list is non-empty
export
cycle : (xs : List a) -> {auto ok : NonEmpty xs} -> Stream a
cycle {a} (x :: xs) {ok = IsNonEmpty} = x :: cycle' xs
where cycle' : List a -> Stream a
cycle' [] = x :: cycle' xs
cycle' (y :: ys) = y :: cycle' ys
public export
partial
takeUntil : (n -> Bool) -> Stream n -> List n
takeUntil p (x :: xs)
= if p x
then [x]
else x :: takeUntil p xs
public export
partial
takeBefore : (n -> Bool) -> Stream n -> List n
takeBefore p (x :: xs)
= if p x
then []
else x :: takeBefore p xs
export
Applicative Stream where
pure = repeat
(<*>) = zipWith apply
export
Monad Stream where
s >>= f = diag (map f s)

319
libs/base/Data/Strings.idr Normal file
View File

@ -0,0 +1,319 @@
module Data.Strings
import Data.List
export
singleton : Char -> String
singleton c = strCons c ""
partial
foldr1 : (a -> a -> a) -> List a -> a
foldr1 _ [x] = x
foldr1 f (x::xs) = f x (foldr1 f xs)
partial
foldl1 : (a -> a -> a) -> List a -> a
foldl1 f (x::xs) = foldl f x xs
-- This works quickly because when string-append builds the result, it allocates
-- enough room in advance so there's only one allocation, rather than lots!
export
fastAppend : List String -> String
fastAppend xs = unsafePerformIO (schemeCall String "string-append" (toFArgs xs))
where
toFArgs : List String -> FArgList
toFArgs [] = []
toFArgs (x :: xs) = x :: toFArgs xs
||| Splits a character list into a list of whitespace separated character lists.
|||
||| ```idris example
||| words' (unpack " A B C D E ")
||| ```
words' : List Char -> List (List Char)
words' s = case dropWhile isSpace s of
[] => []
s' => let (w, s'') = break isSpace s'
in w :: words' s''
||| Splits a string into a list of whitespace separated strings.
|||
||| ```idris example
||| words " A B C D E "
||| ```
export
words : String -> List String
words s = map pack (words' (unpack s))
||| Joins the character lists by spaces into a single character list.
|||
||| ```idris example
||| unwords' [['A'], ['B', 'C'], ['D'], ['E']]
||| ```
unwords' : List (List Char) -> List Char
unwords' [] = []
unwords' ws = assert_total (foldr1 addSpace ws) where
addSpace : List Char -> List Char -> List Char
addSpace w s = w ++ (' ' :: s)
||| Joins the strings by spaces into a single string.
|||
||| ```idris example
||| unwords ["A", "BC", "D", "E"]
||| ```
export
unwords : List String -> String
unwords = pack . unwords' . map unpack
||| Splits a character list into a list of newline separated character lists.
|||
||| ```idris example
||| lines' (unpack "\rA BC\nD\r\nE\n")
||| ```
lines' : List Char -> List (List Char)
lines' s = case dropWhile isNL s of
[] => []
s' => let (w, s'') = break isNL s'
in w :: lines' s''
||| Splits a string into a list of newline separated strings.
|||
||| ```idris example
||| lines "\rA BC\nD\r\nE\n"
||| ```
export
lines : String -> List String
lines s = map pack (lines' (unpack s))
||| Joins the character lists by newlines into a single character list.
|||
||| ```idris example
||| unlines' [['l','i','n','e'], ['l','i','n','e','2'], ['l','n','3'], ['D']]
||| ```
unlines' : List (List Char) -> List Char
unlines' [] = []
unlines' (l::ls) = l ++ '\n' :: unlines' ls
||| Joins the strings by newlines into a single string.
|||
||| ```idris example
||| unlines ["line", "line2", "ln3", "D"]
||| ```
export
unlines : List String -> String
unlines = pack . unlines' . map unpack
export
ltrim : String -> String
ltrim xs = pack (ltrimChars (unpack xs))
where
ltrimChars : List Char -> List Char
ltrimChars [] = []
ltrimChars (x :: xs) = if isSpace x then ltrimChars xs else (x :: xs)
export
trim : String -> String
trim = ltrim . reverse . ltrim . reverse
||| Splits the string into a part before the predicate
||| returns False and the rest of the string.
|||
||| ```idris example
||| span (/= 'C') "ABCD"
||| ```
||| ```idris example
||| span (/= 'C') "EFGH"
||| ```
export
span : (Char -> Bool) -> String -> (String, String)
span p xs
= case span p (unpack xs) of
(x, y) => (pack x, pack y)
||| Splits the string into a part before the predicate
||| returns True and the rest of the string.
|||
||| ```idris example
||| break (== 'C') "ABCD"
||| ```
||| ```idris example
||| break (== 'C') "EFGH"
||| ```
public export
break : (Char -> Bool) -> String -> (String, String)
break p = span (not . p)
||| Splits the string into parts with the predicate
||| indicating separator characters.
|||
||| ```idris example
||| split (== '.') ".AB.C..D"
||| ```
public export
split : (Char -> Bool) -> String -> List String
split p xs = map pack (split p (unpack xs))
export
stringToNatOrZ : String -> Nat
stringToNatOrZ = fromInteger . prim__cast_StringInteger
export
toUpper : String -> String
toUpper str = pack (map toUpper (unpack str))
export
toLower : String -> String
toLower str = pack (map toLower (unpack str))
export
strIndex : String -> Int -> Char
strIndex = prim__strIndex
export
strTail : String -> String
strTail = prim__strTail
export
isPrefixOf : String -> String -> Bool
isPrefixOf a b = isPrefixOf (unpack a) (unpack b)
export
isSuffixOf : String -> String -> Bool
isSuffixOf a b = isSuffixOf (unpack a) (unpack b)
export
isInfixOf : String -> String -> Bool
isInfixOf a b = isInfixOf (unpack a) (unpack b)
public export
data StrM : String -> Type where
StrNil : StrM ""
StrCons : (x : Char) -> (xs : String) -> StrM (strCons x xs)
public export -- primitives, so assert_total and believe_me needed
strM : (x : String) -> StrM x
strM "" = StrNil
strM x
= assert_total $ believe_me $
StrCons (prim__strHead x) (prim__strTail x)
parseNumWithoutSign : List Char -> Integer -> Maybe Integer
parseNumWithoutSign [] acc = Just acc
parseNumWithoutSign (c :: cs) acc =
if (c >= '0' && c <= '9')
then parseNumWithoutSign cs ((acc * 10) + (cast ((ord c) - (ord '0'))))
else Nothing
||| Convert a positive number string to a Num.
|||
||| ```idris example
||| parsePositive "123"
||| ```
||| ```idris example
||| parsePositive {a=Int} " +123"
||| ```
public export
parsePositive : Num a => String -> Maybe a
parsePositive s = parsePosTrimmed (trim s)
where
parsePosTrimmed : String -> Maybe a
parsePosTrimmed s with (strM s)
parsePosTrimmed "" | StrNil = Nothing
parsePosTrimmed (strCons '+' xs) | (StrCons '+' xs) =
map fromInteger (parseNumWithoutSign (unpack xs) 0)
parsePosTrimmed (strCons x xs) | (StrCons x xs) =
if (x >= '0' && x <= '9')
then map fromInteger (parseNumWithoutSign (unpack xs) (cast (ord x - ord '0')))
else Nothing
||| Convert a number string to a Num.
|||
||| ```idris example
||| parseInteger " 123"
||| ```
||| ```idris example
||| parseInteger {a=Int} " -123"
||| ```
public export
parseInteger : (Num a, Neg a) => String -> Maybe a
parseInteger s = parseIntTrimmed (trim s)
where
parseIntTrimmed : String -> Maybe a
parseIntTrimmed s with (strM s)
parseIntTrimmed "" | StrNil = Nothing
parseIntTrimmed (strCons x xs) | (StrCons x xs) =
if (x == '-')
then map (\y => negate (fromInteger y)) (parseNumWithoutSign (unpack xs) 0)
else if (x == '+')
then map fromInteger (parseNumWithoutSign (unpack xs) (cast {from=Int} 0))
else if (x >= '0' && x <= '9')
then map fromInteger (parseNumWithoutSign (unpack xs) (cast (ord x - ord '0')))
else Nothing
||| Convert a number string to a Double.
|||
||| ```idris example
||| parseDouble "+123.123e-2"
||| ```
||| ```idris example
||| parseDouble {a=Int} " -123.123E+2"
||| ```
||| ```idris example
||| parseDouble {a=Int} " +123.123"
||| ```
export -- it's a bit too slow at compile time
parseDouble : String -> Maybe Double
parseDouble = mkDouble . wfe . trim
where
intPow : Integer -> Integer -> Double
intPow base exp = assert_total $ if exp > 0 then (num base exp) else 1 / (num base exp)
where
num : Integer -> Integer -> Double
num base 0 = 1
num base e = if e < 0
then cast base * num base (e + 1)
else cast base * num base (e - 1)
natpow : Double -> Nat -> Double
natpow x Z = 1
natpow x (S n) = x * (natpow x n)
mkDouble : Maybe (Double, Double, Integer) -> Maybe Double
mkDouble (Just (w, f, e)) = let ex = intPow 10 e in
Just $ (w * ex + f * ex)
mkDouble Nothing = Nothing
wfe : String -> Maybe (Double, Double, Integer)
wfe cs = case split (== '.') cs of
(wholeAndExp :: []) =>
case split (\c => c == 'e' || c == 'E') wholeAndExp of
(whole::exp::[]) =>
do
w <- cast {from=Integer} <$> parseInteger whole
e <- parseInteger exp
pure (w, 0, e)
(whole::[]) =>
do
w <- cast {from=Integer} <$> parseInteger whole
pure (w, 0, 0)
_ => Nothing
(whole::fracAndExp::[]) =>
case split (\c => c == 'e' || c == 'E') fracAndExp of
(""::exp::[]) => Nothing
(frac::exp::[]) =>
do
w <- cast {from=Integer} <$> parseInteger whole
f <- (/ (natpow 10 (length frac))) <$>
(cast <$> parseNumWithoutSign (unpack frac) 0)
e <- parseInteger exp
pure (w, if w < 0 then (-f) else f, e)
(frac::[]) =>
do
w <- cast {from=Integer} <$> parseInteger whole
f <- (/ (natpow 10 (length frac))) <$>
(cast <$> parseNumWithoutSign (unpack frac) 0)
pure (w, if w < 0 then (-f) else f, 0)
_ => Nothing
_ => Nothing

900
libs/base/Data/Vect.idr Normal file
View File

@ -0,0 +1,900 @@
module Data.Vect
import Data.List
import Data.Nat
import public Data.Fin
import Decidable.Equality
public export
data Vect : (len : Nat) -> (elem : Type) -> Type where
||| Empty vector
Nil : Vect Z elem
||| A non-empty vector of length `S len`, consisting of a head element and
||| the rest of the list, of length `len`.
(::) : (x : elem) -> (xs : Vect len elem) -> Vect (S len) elem
-- Hints for interactive editing
%name Vect xs,ys,zs,ws
public export
length : (xs : Vect len elem) -> Nat
length [] = 0
length (x::xs) = 1 + length xs
||| Show that the length function on vectors in fact calculates the length
private
lengthCorrect : (len : Nat) -> (xs : Vect len elem) -> length xs = len
lengthCorrect Z [] = Refl
lengthCorrect (S n) (x :: xs) = rewrite lengthCorrect n xs in Refl
--------------------------------------------------------------------------------
-- Indexing into vectors
--------------------------------------------------------------------------------
||| All but the first element of the vector
|||
||| ```idris example
||| tail [1,2,3,4]
||| ```
public export
tail : Vect (S len) elem -> Vect len elem
tail (x::xs) = xs
||| Only the first element of the vector
|||
||| ```idris example
||| head [1,2,3,4]
||| ```
public export
head : Vect (S len) elem -> elem
head (x::xs) = x
||| The last element of the vector
|||
||| ```idris example
||| last [1,2,3,4]
||| ```
public export
last : Vect (S len) elem -> elem
last (x::[]) = x
last (x::y::ys) = last $ y::ys
||| All but the last element of the vector
|||
||| ```idris example
||| init [1,2,3,4]
||| ```
public export
init : Vect (S len) elem -> Vect len elem
init (x::[]) = []
init (x::y::ys) = x :: init (y::ys)
||| Extract a particular element from a vector
|||
||| ```idris example
||| index 1 [1,2,3,4]
||| ```
public export
index : Fin len -> Vect len elem -> elem
index FZ (x::xs) = x
index (FS k) (x::xs) = index k xs
||| Insert an element at a particular index
|||
||| ```idris example
||| insertAt 1 8 [1,2,3,4]
||| ```
public export
insertAt : Fin (S len) -> elem -> Vect len elem -> Vect (S len) elem
insertAt FZ y xs = y :: xs
insertAt (FS k) y (x::xs) = x :: insertAt k y xs
insertAt (FS k) y [] = absurd k
||| Construct a new vector consisting of all but the indicated element
|||
||| ```idris example
||| deleteAt 1 [1,2,3,4]
||| ```
public export
deleteAt : {len : _} -> Fin (S len) -> Vect (S len) elem -> Vect len elem
deleteAt FZ (x::xs) = xs
deleteAt {len = S m} (FS k) (x::xs) = x :: deleteAt k xs
deleteAt {len = Z} (FS k) (x::xs) impossible
deleteAt _ [] impossible
||| Replace an element at a particlar index with another
|||
||| ```idris example
||| replaceAt 1 8 [1,2,3,4]
||| ```
public export
replaceAt : Fin len -> elem -> Vect len elem -> Vect len elem
replaceAt FZ y (x::xs) = y :: xs
replaceAt (FS k) y (x::xs) = x :: replaceAt k y xs
||| Replace the element at a particular index with the result of applying a function to it
||| @ i the index to replace at
||| @ f the update function
||| @ xs the vector to replace in
|||
||| ```idris example
||| updateAt 1 (+10) [1,2,3,4]
||| ```
public export
updateAt : (i : Fin len) -> (f : elem -> elem) -> (xs : Vect len elem) -> Vect len elem
updateAt FZ f (x::xs) = f x :: xs
updateAt (FS k) f (x::xs) = x :: updateAt k f xs
||| Append two vectors
|||
||| ```idris example
||| [1,2,3,4] ++ [5,6]
||| ```
public export
(++) : (xs : Vect m elem) -> (ys : Vect n elem) -> Vect (m + n) elem
(++) [] ys = ys
(++) (x::xs) ys = x :: xs ++ ys
||| Repeate some value some number of times.
|||
||| @ len the number of times to repeat it
||| @ x the value to repeat
|||
||| ```idris example
||| replicate 4 1
||| ```
public export
replicate : (len : Nat) -> (x : elem) -> Vect len elem
replicate Z x = []
replicate (S k) x = x :: replicate k x
||| Merge two ordered vectors
|||
||| ```idris example
||| mergeBy compare (fromList [1,3,5]) (fromList [2,3,4,5,6])
||| ```
export
mergeBy : (elem -> elem -> Ordering) -> (xs : Vect n elem) -> (ys : Vect m elem) -> Vect (n + m) elem
mergeBy order [] [] = []
mergeBy order [] (x :: xs) = x :: xs
mergeBy {n = S k} order (x :: xs) [] = rewrite plusZeroRightNeutral (S k) in
x :: xs
mergeBy {n = S k} {m = S k'} order (x :: xs) (y :: ys)
= case order x y of
LT => x :: mergeBy order xs (y :: ys)
_ => rewrite sym (plusSuccRightSucc k k') in
y :: mergeBy order (x :: xs) ys
export
merge : Ord elem => Vect n elem -> Vect m elem -> Vect (n + m) elem
merge = mergeBy compare
--------------------------------------------------------------------------------
-- Transformations
--------------------------------------------------------------------------------
||| Reverse the order of the elements of a vector
|||
||| ```idris example
||| reverse [1,2,3,4]
||| ```
public export
reverse : Vect len elem -> Vect len elem
reverse xs = go [] xs
where go : Vect n elem -> Vect m elem -> Vect (n+m) elem
go {n} acc [] = rewrite plusZeroRightNeutral n in acc
go {n} {m=S m} acc (x :: xs) = rewrite sym $ plusSuccRightSucc n m
in go (x::acc) xs
||| Alternate an element between the other elements of a vector
||| @ sep the element to intersperse
||| @ xs the vector to separate with `sep`
|||
||| ```idris example
||| intersperse 0 [1,2,3,4]
||| ```
export
intersperse : (sep : elem) -> (xs : Vect len elem) -> Vect (len + pred len) elem
intersperse sep [] = []
intersperse sep (x::xs) = x :: intersperse' sep xs
where
intersperse' : elem -> Vect n elem -> Vect (n + n) elem
intersperse' sep [] = []
intersperse' {n=S n} sep (x::xs) = rewrite sym $ plusSuccRightSucc n n
in sep :: x :: intersperse' sep xs
--------------------------------------------------------------------------------
-- Conversion from list (toList is provided by Foldable)
--------------------------------------------------------------------------------
public export
fromList' : Vect len elem -> (l : List elem) -> Vect (length l + len) elem
fromList' ys [] = ys
fromList' {len} ys (x::xs) =
rewrite (plusSuccRightSucc (length xs) len) in
fromList' (x::ys) xs
||| Convert a list to a vector.
|||
||| The length of the list should be statically known.
|||
||| ```idris example
||| fromList [1,2,3,4]
||| ```
public export
fromList : (l : List elem) -> Vect (length l) elem
fromList l =
rewrite (sym $ plusZeroRightNeutral (length l)) in
reverse $ fromList' [] l
--------------------------------------------------------------------------------
-- Zips and unzips
--------------------------------------------------------------------------------
||| Combine two equal-length vectors pairwise with some function.
|||
||| @ f the function to combine elements with
||| @ xs the first vector of elements
||| @ ys the second vector of elements
|||
||| ```idris example
||| zipWith (+) (fromList [1,2,3,4]) (fromList [5,6,7,8])
||| ```
public export
zipWith : (f : a -> b -> c) -> (xs : Vect n a) -> (ys : Vect n b) -> Vect n c
zipWith f [] [] = []
zipWith f (x::xs) (y::ys) = f x y :: zipWith f xs ys
||| Combine three equal-length vectors into a vector with some function
|||
||| ```idris example
||| zipWith3 (\x,y,z => x+y+z) (fromList [1,2,3,4]) (fromList [5,6,7,8]) (fromList [1,1,1,1])
||| ```
public export
zipWith3 : (a -> b -> c -> d) -> (xs : Vect n a) -> (ys : Vect n b) -> (zs : Vect n c) -> Vect n d
zipWith3 f [] [] [] = []
zipWith3 f (x::xs) (y::ys) (z::zs) = f x y z :: zipWith3 f xs ys zs
||| Combine two equal-length vectors pairwise
|||
||| ```idris example
||| zip (fromList [1,2,3,4]) (fromList [1,2,3,4])
||| ```
public export
zip : (xs : Vect n a) -> (ys : Vect n b) -> Vect n (a, b)
zip = zipWith (\x,y => (x,y))
||| Combine three equal-length vectors elementwise into a vector of tuples
|||
||| ```idris example
||| zip3 (fromList [1,2,3,4]) (fromList [1,2,3,4]) (fromList [1,2,3,4])
||| ```
public export
zip3 : (xs : Vect n a) -> (ys : Vect n b) -> (zs : Vect n c) -> Vect n (a, b, c)
zip3 = zipWith3 (\x,y,z => (x,y,z))
||| Convert a vector of pairs to a pair of vectors
|||
||| ```idris example
||| unzip (fromList [(1,2), (1,2)])
||| ```
public export
unzip : (xs : Vect n (a, b)) -> (Vect n a, Vect n b)
unzip [] = ([], [])
unzip ((l, r)::xs) with (unzip xs)
unzip ((l, r)::xs) | (lefts, rights) = (l::lefts, r::rights)
||| Convert a vector of three-tuples to a triplet of vectors
|||
||| ```idris example
||| unzip3 (fromList [(1,2,3), (1,2,3)])
||| ```
public export
unzip3 : (xs : Vect n (a, b, c)) -> (Vect n a, Vect n b, Vect n c)
unzip3 [] = ([], [], [])
unzip3 ((l,c,r)::xs) with (unzip3 xs)
unzip3 ((l,c,r)::xs) | (lefts, centers, rights)
= (l::lefts, c::centers, r::rights)
--------------------------------------------------------------------------------
-- Equality
--------------------------------------------------------------------------------
public export
implementation (Eq elem) => Eq (Vect len elem) where
(==) [] [] = True
(==) (x::xs) (y::ys) = x == y && xs == ys
--------------------------------------------------------------------------------
-- Order
--------------------------------------------------------------------------------
public export
implementation Ord elem => Ord (Vect len elem) where
compare [] [] = EQ
compare (x::xs) (y::ys)
= case compare x y of
EQ => compare xs ys
x => x
--------------------------------------------------------------------------------
-- Maps
--------------------------------------------------------------------------------
public export
implementation Functor (Vect n) where
map f [] = []
map f (x::xs) = f x :: map f xs
||| Map a partial function across a vector, returning those elements for which
||| the function had a value.
|||
||| The first projection of the resulting pair (ie the length) will always be
||| at most the length of the input vector. This is not, however, guaranteed
||| by the type.
|||
||| @ f the partial function (expressed by returning `Maybe`)
||| @ xs the vector to check for results
|||
||| ```idris example
||| mapMaybe ((find (=='a')) . unpack) (fromList ["abc","ade","bgh","xyz"])
||| ```
export
mapMaybe : (f : a -> Maybe b) -> (xs : Vect len a) -> (m : Nat ** Vect m b)
mapMaybe f [] = (_ ** [])
mapMaybe f (x::xs) =
let (len ** ys) = mapMaybe f xs
in case f x of
Just y => (S len ** y :: ys)
Nothing => ( len ** ys)
--------------------------------------------------------------------------------
-- Folds
--------------------------------------------------------------------------------
public export
foldrImpl : (t -> acc -> acc) -> acc -> (acc -> acc) -> Vect n t -> acc
foldrImpl f e go [] = go e
foldrImpl f e go (x::xs) = foldrImpl f e (go . (f x)) xs
public export
implementation Foldable (Vect n) where
foldr f e xs = foldrImpl f e id xs
--------------------------------------------------------------------------------
-- Special folds
--------------------------------------------------------------------------------
||| Flatten a vector of equal-length vectors
|||
||| ```idris example
||| concat [[1,2,3], [4,5,6]]
||| ```
public export
concat : (xss : Vect m (Vect n elem)) -> Vect (m * n) elem
concat [] = []
concat (v::vs) = v ++ Vect.concat vs
||| Foldr without seeding the accumulator
|||
||| ```idris example
||| foldr1 (-) (fromList [1,2,3])
||| ```
public export
foldr1 : (t -> t -> t) -> Vect (S n) t -> t
foldr1 f [x] = x
foldr1 f (x::y::xs) = f x (foldr1 f (y::xs))
||| Foldl without seeding the accumulator
|||
||| ```idris example
||| foldl1 (-) (fromList [1,2,3])
||| ```
public export
foldl1 : (t -> t -> t) -> Vect (S n) t -> t
foldl1 f (x::xs) = foldl f x xs
--------------------------------------------------------------------------------
-- Scans
--------------------------------------------------------------------------------
||| The scanl function is similar to foldl, but returns all the intermediate
||| accumulator states in the form of a vector.
|||
||| ```idris example
||| scanl (-) 0 (fromList [1,2,3])
||| ```
public export
scanl : (res -> elem -> res) -> res -> Vect len elem -> Vect (S len) res
scanl f q [] = [q]
scanl f q (x::xs) = q :: scanl f (f q x) xs
||| The scanl1 function is a variant of scanl that doesn't require an explicit
||| starting value.
||| It assumes the first element of the vector to be the starting value and then
||| starts the fold with the element following it.
|||
||| ```idris example
||| scanl1 (-) (fromList [1,2,3])
||| ```
public export
scanl1 : (elem -> elem -> elem) -> Vect len elem -> Vect len elem
scanl1 f [] = []
scanl1 f (x::xs) = scanl f x xs
--------------------------------------------------------------------------------
-- Membership tests
--------------------------------------------------------------------------------
||| Search for an item using a user-provided test
||| @ p the equality test
||| @ e the item to search for
||| @ xs the vector to search in
|||
||| ```idris example
||| elemBy (==) 2 [1,2,3,4]
||| ```
public export
elemBy : (p : elem -> elem -> Bool) -> (e : elem) -> (xs : Vect len elem) -> Bool
elemBy p e [] = False
elemBy p e (x::xs) = p e x || elemBy p e xs
||| Use the default Boolean equality on elements to search for an item
||| @ x what to search for
||| @ xs where to search
|||
||| ```idris example
||| elem 3 [1,2,3,4]
||| ```
public export
elem : Eq elem => (x : elem) -> (xs : Vect len elem) -> Bool
elem = elemBy (==)
||| Find the association of some key with a user-provided comparison
||| @ p the comparison operator for keys (True if they match)
||| @ e the key to look for
|||
||| ```idris example
||| lookupBy (==) 2 [(1, 'a'), (2, 'b'), (3, 'c')]
||| ```
public export
lookupBy : (p : key -> key -> Bool) -> (e : key) -> (xs : Vect n (key, val)) -> Maybe val
lookupBy p e [] = Nothing
lookupBy p e ((l, r)::xs) = if p e l then Just r else lookupBy p e xs
||| Find the assocation of some key using the default Boolean equality test
|||
||| ```idris example
||| lookup 3 [(1, 'a'), (2, 'b'), (3, 'c')]
||| ```
public export
lookup : Eq key => key -> Vect n (key, val) -> Maybe val
lookup = lookupBy (==)
||| Check if any element of xs is found in elems by a user-provided comparison
||| @ p the comparison operator
||| @ elems the vector to search
||| @ xs what to search for
|||
||| ```idris example
||| hasAnyBy (==) [2,5] [1,2,3,4]
||| ```
public export
hasAnyBy : (p : elem -> elem -> Bool) -> (elems : Vect m elem) -> (xs : Vect len elem) -> Bool
hasAnyBy p elems [] = False
hasAnyBy p elems (x::xs) = elemBy p x elems || hasAnyBy p elems xs
||| Check if any element of xs is found in elems using the default Boolean equality test
|||
||| ```idris example
||| hasAny [2,5] [1,2,3,4]
||| ```
public export
hasAny : Eq elem => Vect m elem -> Vect len elem -> Bool
hasAny = hasAnyBy (==)
--------------------------------------------------------------------------------
-- Searching with a predicate
--------------------------------------------------------------------------------
||| Find the first element of the vector that satisfies some test
||| @ p the test to satisfy
|||
||| ```idris example
||| find (== 3) [1,2,3,4]
||| ```
public export
find : (p : elem -> Bool) -> (xs : Vect len elem) -> Maybe elem
find p [] = Nothing
find p (x::xs) = if p x then Just x else find p xs
||| Find the index of the first element of the vector that satisfies some test
|||
||| ```idris example
||| findIndex (== 3) [1,2,3,4]
||| ```
public export
findIndex : (elem -> Bool) -> Vect len elem -> Maybe (Fin len)
findIndex p [] = Nothing
findIndex p (x :: xs) = if p x then Just FZ else map FS (findIndex p xs)
||| Find the indices of all elements that satisfy some test
|||
||| ```idris example
||| findIndices (< 3) [1,2,3,4]
||| ```
public export
findIndices : (elem -> Bool) -> Vect m elem -> List (Fin m)
findIndices p [] = []
findIndices p (x :: xs)
= let is = map FS $ findIndices p xs in
if p x then FZ :: is else is
||| Find the index of the first element of the vector that satisfies some test
|||
||| ```idris example
||| elemIndexBy (==) 3 [1,2,3,4]
||| ```
public export
elemIndexBy : (elem -> elem -> Bool) -> elem -> Vect m elem -> Maybe (Fin m)
elemIndexBy p e = findIndex $ p e
||| Find the index of the first element of the vector equal to the given one.
|||
||| ```idris example
||| elemIndex 3 [1,2,3,4]
||| ```
public export
elemIndex : Eq elem => elem -> Vect m elem -> Maybe (Fin m)
elemIndex = elemIndexBy (==)
||| Find the indices of all elements that satisfy some test
|||
||| ```idris example
||| elemIndicesBy (<=) 3 [1,2,3,4]
||| ```
public export
elemIndicesBy : (elem -> elem -> Bool) -> elem -> Vect m elem -> List (Fin m)
elemIndicesBy p e = findIndices $ p e
||| Find the indices of all elements uquals to the given one
|||
||| ```idris example
||| elemIndices 3 [1,2,3,4,3]
||| ```
public export
elemIndices : Eq elem => elem -> Vect m elem -> List (Fin m)
elemIndices = elemIndicesBy (==)
--------------------------------------------------------------------------------
-- Filters
--------------------------------------------------------------------------------
||| Find all elements of a vector that satisfy some test
|||
||| ```idris example
||| filter (< 3) (fromList [1,2,3,4])
||| ```
public export
filter : (elem -> Bool) -> Vect len elem -> (p ** Vect p elem)
filter p [] = ( _ ** [] )
filter p (x::xs) =
let (_ ** tail) = filter p xs
in if p x then
(_ ** x::tail)
else
(_ ** tail)
||| Make the elements of some vector unique by some test
|||
||| ```idris example
||| nubBy (==) (fromList [1,2,2,3,4,4])
||| ```
public export
nubBy : (elem -> elem -> Bool) -> Vect len elem -> (p ** Vect p elem)
nubBy = nubBy' []
where
nubBy' : forall len . Vect m elem -> (elem -> elem -> Bool) -> Vect len elem -> (p ** Vect p elem)
nubBy' acc p [] = (_ ** [])
nubBy' acc p (x::xs) with (elemBy p x acc)
nubBy' acc p (x :: xs) | True = nubBy' acc p xs
nubBy' acc p (x :: xs) | False with (nubBy' (x::acc) p xs)
nubBy' acc p (x :: xs) | False | (_ ** tail) = (_ ** x::tail)
||| Make the elements of some vector unique by the default Boolean equality
|||
||| ```idris example
||| nub (fromList [1,2,2,3,4,4])
||| ```
public export
nub : Eq elem => Vect len elem -> (p ** Vect p elem)
nub = nubBy (==)
||| Delete first element from list according to some test
|||
||| ```idris example
||| deleteBy (<) 3 (fromList [1,2,2,3,4,4])
||| ```
public export
deleteBy : {len : _} -> -- needed for the dependent pair
(elem -> elem -> Bool) -> elem -> Vect len elem -> (p ** Vect p elem)
deleteBy _ _ [] = (_ ** [])
deleteBy eq x (y::ys) =
let (len ** zs) = deleteBy eq x ys
in if x `eq` y then (_ ** ys) else (S len ** y ::zs)
||| Delete first element from list equal to the given one
|||
||| ```idris example
||| delete 2 (fromList [1,2,2,3,4,4])
||| ```
public export
delete : {len : _} ->
Eq elem => elem -> Vect len elem -> (p ** Vect p elem)
delete = deleteBy (==)
--------------------------------------------------------------------------------
-- Splitting and breaking lists
--------------------------------------------------------------------------------
||| A tuple where the first element is a `Vect` of the `n` first elements and
||| the second element is a `Vect` of the remaining elements of the original.
||| It is equivalent to `(take n xs, drop n xs)` (`splitAtTakeDrop`),
||| but is more efficient.
|||
||| @ n the index to split at
||| @ xs the `Vect` to split in two
|||
||| ```idris example
||| splitAt 2 (fromList [1,2,3,4])
||| ```
public export
splitAt : (n : Nat) -> (xs : Vect (n + m) elem) -> (Vect n elem, Vect m elem)
splitAt Z xs = ([], xs)
splitAt (S k) (x :: xs) with (splitAt k {m} xs)
splitAt (S k) (x :: xs) | (tk, dr) = (x :: tk, dr)
||| A tuple where the first element is a `Vect` of the `n` elements passing given test
||| and the second element is a `Vect` of the remaining elements of the original.
|||
||| ```idris example
||| partition (== 2) (fromList [1,2,3,2,4])
||| ```
public export
partition : (elem -> Bool) -> Vect len elem -> ((p ** Vect p elem), (q ** Vect q elem))
partition p [] = ((_ ** []), (_ ** []))
partition p (x::xs) =
let ((leftLen ** lefts), (rightLen ** rights)) = partition p xs in
if p x then
((S leftLen ** x::lefts), (rightLen ** rights))
else
((leftLen ** lefts), (S rightLen ** x::rights))
--------------------------------------------------------------------------------
-- Predicates
--------------------------------------------------------------------------------
||| Verify vector prefix
|||
||| ```idris example
||| isPrefixOfBy (==) (fromList [1,2]) (fromList [1,2,3,4])
||| ```
public export
isPrefixOfBy : (elem -> elem -> Bool) -> Vect m elem -> Vect len elem -> Bool
isPrefixOfBy p [] right = True
isPrefixOfBy p left [] = False
isPrefixOfBy p (x::xs) (y::ys) = p x y && isPrefixOfBy p xs ys
||| Verify vector prefix
|||
||| ```idris example
||| isPrefixOf (fromList [1,2]) (fromList [1,2,3,4])
||| ```
public export
isPrefixOf : Eq elem => Vect m elem -> Vect len elem -> Bool
isPrefixOf = isPrefixOfBy (==)
||| Verify vector suffix
|||
||| ```idris example
||| isSuffixOfBy (==) (fromList [3,4]) (fromList [1,2,3,4])
||| ```
public export
isSuffixOfBy : (elem -> elem -> Bool) -> Vect m elem -> Vect len elem -> Bool
isSuffixOfBy p left right = isPrefixOfBy p (reverse left) (reverse right)
||| Verify vector suffix
|||
||| ```idris example
||| isSuffixOf (fromList [3,4]) (fromList [1,2,3,4])
||| ```
public export
isSuffixOf : Eq elem => Vect m elem -> Vect len elem -> Bool
isSuffixOf = isSuffixOfBy (==)
--------------------------------------------------------------------------------
-- Conversions
--------------------------------------------------------------------------------
||| Convert Maybe type into Vect
|||
||| ```idris example
||| maybeToVect (Just 2)
||| ```
public export
maybeToVect : Maybe elem -> (p ** Vect p elem)
maybeToVect Nothing = (_ ** [])
maybeToVect (Just j) = (_ ** [j])
||| Convert first element of Vect (if exists) into Maybe.
|||
||| ```idris example
||| vectToMaybe [2]
||| ```
public export
vectToMaybe : Vect len elem -> Maybe elem
vectToMaybe [] = Nothing
vectToMaybe (x::xs) = Just x
--------------------------------------------------------------------------------
-- Misc
--------------------------------------------------------------------------------
||| Filter out Nothings from Vect
|||
||| ```idris example
||| catMaybes [Just 1, Just 2, Nothing, Nothing, Just 5]
||| ```
public export
catMaybes : Vect n (Maybe elem) -> (p ** Vect p elem)
catMaybes [] = (_ ** [])
catMaybes (Nothing::xs) = catMaybes xs
catMaybes ((Just j)::xs) =
let (_ ** tail) = catMaybes xs
in (_ ** j::tail)
||| Get diagonal elements
|||
||| ```idris example
||| diag [[1,2,3], [4,5,6], [7,8,9]]
||| ```
public export
diag : Vect len (Vect len elem) -> Vect len elem
diag [] = []
diag ((x::xs)::xss) = x :: diag (map tail xss)
public export
range : {len : Nat} -> Vect len (Fin len)
range {len=Z} = []
range {len=S _} = FZ :: map FS range
||| Transpose a `Vect` of `Vect`s, turning rows into columns and vice versa.
|||
||| This is like zipping all the inner `Vect`s together and is equivalent to `traverse id` (`transposeTraverse`).
|||
||| As the types ensure rectangularity, this is an involution, unlike `Prelude.List.transpose`.
|||
||| ```idris example
||| transpose [[1,2], [3,4], [5,6], [7,8]]
||| ```
public export
transpose : {n : _} -> Vect m (Vect n elem) -> Vect n (Vect m elem)
transpose [] = replicate _ [] -- = [| [] |]
transpose (x :: xs) = zipWith (::) x (transpose xs) -- = [| x :: xs |]
--------------------------------------------------------------------------------
-- Applicative/Monad/Traversable
--------------------------------------------------------------------------------
-- These only work if the length is known at run time!
implementation {k : Nat} -> Applicative (Vect k) where
pure = replicate _
fs <*> vs = zipWith apply fs vs
-- ||| This monad is different from the List monad, (>>=)
-- ||| uses the diagonal.
implementation {k : Nat} -> Monad (Vect k) where
m >>= f = diag (map f m)
public export
implementation Traversable (Vect k) where
traverse f [] = pure []
traverse f (x :: xs) = [| f x :: traverse f xs |]
--------------------------------------------------------------------------------
-- Elem
--------------------------------------------------------------------------------
||| A proof that some element is found in a vector
public export
data Elem : a -> Vect k a -> Type where
Here : Elem x (x::xs)
There : (later : Elem x xs) -> Elem x (y::xs)
||| Nothing can be in an empty Vect
export
noEmptyElem : forall x . Elem x [] -> Void
noEmptyElem Here impossible
export
Uninhabited (Elem x []) where
uninhabited = noEmptyElem
||| An item not in the head and not in the tail is not in the Vect at all
export
neitherHereNorThere : {x, y : a} -> {xs : Vect n a} -> Not (x = y) -> Not (Elem x xs) -> Not (Elem x (y :: xs))
neitherHereNorThere xneqy xninxs Here = xneqy Refl
neitherHereNorThere xneqy xninxs (There xinxs) = xninxs xinxs
||| A decision procedure for Elem
public export
isElem : DecEq a => (x : a) -> (xs : Vect n a) -> Dec (Elem x xs)
isElem x [] = No noEmptyElem
isElem x (y :: xs) with (decEq x y)
isElem x (x :: xs) | (Yes Refl) = Yes Here
isElem x (y :: xs) | (No xneqy) with (isElem x xs)
isElem x (y :: xs) | (No xneqy) | (Yes xinxs) = Yes (There xinxs)
isElem x (y :: xs) | (No xneqy) | (No xninxs) = No (neitherHereNorThere xneqy xninxs)
public export
replaceElem : (xs : Vect k t) -> Elem x xs -> (y : t) -> (ys : Vect k t ** Elem y ys)
replaceElem (x::xs) Here y = (y :: xs ** Here)
replaceElem (x::xs) (There xinxs) y with (replaceElem xs xinxs y)
replaceElem (x::xs) (There xinxs) y | (ys ** yinys) = (x :: ys ** There yinys)
public export
replaceByElem : (xs : Vect k t) -> Elem x xs -> t -> Vect k t
replaceByElem (x::xs) Here y = y :: xs
replaceByElem (x::xs) (There xinxs) y = x :: replaceByElem xs xinxs y
public export
mapElem : {0 xs : Vect k t} -> {0 f : t -> u} ->
Elem x xs -> Elem (f x) (map f xs)
mapElem Here = Here
mapElem (There e) = There (mapElem e)
||| Remove the element at the given position.
|||
||| @xs The vector to be removed from
||| @p A proof that the element to be removed is in the vector
public export
dropElem : {k : _} -> (xs : Vect (S k) t) -> (p : Elem x xs) -> Vect k t
dropElem (x :: ys) Here = ys
dropElem {k = (S k)} (x :: ys) (There later) = x :: dropElem ys later
--------------------------------------------------------------------------------
-- Show
--------------------------------------------------------------------------------
export
implementation Show elem => Show (Vect len elem) where
show = show . toList
-- Some convenience functions for testing lengths
-- Needs to be Maybe rather than Dec, because if 'n' is unequal to m, we
-- only know we don't know how to make a Vect n a, not that one can't exist.
export
exactLength : {m : Nat} -> -- expected at run-time
(len : Nat) -> (xs : Vect m a) -> Maybe (Vect len a)
exactLength {m} len xs with (decEq m len)
exactLength {m = m} m xs | (Yes Refl) = Just xs
exactLength {m = m} len xs | (No contra) = Nothing
||| If the given Vect is at least the required length, return a Vect with
||| at least that length in its type, otherwise return Nothing
||| @len the required length
||| @xs the vector with the desired length
export
overLength : {m : Nat} -> -- expected at run-time
(len : Nat) -> (xs : Vect m a) -> Maybe (p ** Vect (plus p len) a)
overLength {m} n xs with (cmp m n)
overLength {m = m} (plus m (S y)) xs | (CmpLT y) = Nothing
overLength {m = m} m xs | CmpEQ
= Just (0 ** xs)
overLength {m = plus n (S x)} n xs | (CmpGT x)
= Just (S x ** rewrite plusCommutative (S x) n in xs)

View File

@ -0,0 +1,8 @@
module Debug.Trace
import Prelude
import PrimIO
export
trace : (msg : String) -> (result : a) -> a
trace x val = unsafePerformIO (do putStrLn x; pure val)

View File

@ -0,0 +1,199 @@
module Decidable.Equality
import Data.Maybe
import Data.Nat
--------------------------------------------------------------------------------
-- Decidable equality
--------------------------------------------------------------------------------
||| Decision procedures for propositional equality
public export
interface DecEq t where
||| Decide whether two elements of `t` are propositionally equal
total decEq : (x1 : t) -> (x2 : t) -> Dec (x1 = x2)
--------------------------------------------------------------------------------
-- Utility lemmas
--------------------------------------------------------------------------------
||| The negation of equality is symmetric (follows from symmetry of equality)
export total
negEqSym : forall a, b . (a = b -> Void) -> (b = a -> Void)
negEqSym p h = p (sym h)
||| Everything is decidably equal to itself
export total
decEqSelfIsYes : DecEq a => {x : a} -> decEq x x = Yes Refl
decEqSelfIsYes {x} with (decEq x x)
decEqSelfIsYes {x} | Yes Refl = Refl
decEqSelfIsYes {x} | No contra = absurd $ contra Refl
--------------------------------------------------------------------------------
--- Unit
--------------------------------------------------------------------------------
export
implementation DecEq () where
decEq () () = Yes Refl
--------------------------------------------------------------------------------
-- Booleans
--------------------------------------------------------------------------------
total trueNotFalse : True = False -> Void
trueNotFalse Refl impossible
export
implementation DecEq Bool where
decEq True True = Yes Refl
decEq False False = Yes Refl
decEq True False = No trueNotFalse
decEq False True = No (negEqSym trueNotFalse)
--------------------------------------------------------------------------------
-- Nat
--------------------------------------------------------------------------------
total ZnotS : Z = S n -> Void
ZnotS Refl impossible
export
implementation DecEq Nat where
decEq Z Z = Yes Refl
decEq Z (S _) = No ZnotS
decEq (S _) Z = No (negEqSym ZnotS)
decEq (S n) (S m) with (decEq n m)
decEq (S n) (S m) | Yes p = Yes $ cong S p
decEq (S n) (S m) | No p = No $ \h : (S n = S m) => p $ succInjective n m h
--------------------------------------------------------------------------------
-- Maybe
--------------------------------------------------------------------------------
total nothingNotJust : {x : t} -> (Nothing {ty = t} = Just x) -> Void
nothingNotJust Refl impossible
export
implementation (DecEq t) => DecEq (Maybe t) where
decEq Nothing Nothing = Yes Refl
decEq (Just x') (Just y') with (decEq x' y')
decEq (Just x') (Just y') | Yes p = Yes $ cong Just p
decEq (Just x') (Just y') | No p
= No $ \h : Just x' = Just y' => p $ justInjective h
decEq Nothing (Just _) = No nothingNotJust
decEq (Just _) Nothing = No (negEqSym nothingNotJust)
-- TODO: Other prelude data types
-- For the primitives, we have to cheat because we don't have access to their
-- internal implementations. We use believe_me for the inequality proofs
-- because we don't them to reduce (and they should never be needed anyway...)
-- A postulate would be better, but erasure analysis may think they're needed
-- for computation in a higher order setting.
--------------------------------------------------------------------------------
-- Tuple
--------------------------------------------------------------------------------
lemma_both_neq : (x = x' -> Void) -> (y = y' -> Void) -> ((x, y) = (x', y') -> Void)
lemma_both_neq p_x_not_x' p_y_not_y' Refl = p_x_not_x' Refl
lemma_snd_neq : (x = x) -> (y = y' -> Void) -> ((x, y) = (x, y') -> Void)
lemma_snd_neq Refl p Refl = p Refl
lemma_fst_neq_snd_eq : (x = x' -> Void) ->
(y = y') ->
((x, y) = (x', y) -> Void)
lemma_fst_neq_snd_eq p_x_not_x' Refl Refl = p_x_not_x' Refl
export
implementation (DecEq a, DecEq b) => DecEq (a, b) where
decEq (a, b) (a', b') with (decEq a a')
decEq (a, b) (a, b') | (Yes Refl) with (decEq b b')
decEq (a, b) (a, b) | (Yes Refl) | (Yes Refl) = Yes Refl
decEq (a, b) (a, b') | (Yes Refl) | (No p) = No (\eq => lemma_snd_neq Refl p eq)
decEq (a, b) (a', b') | (No p) with (decEq b b')
decEq (a, b) (a', b) | (No p) | (Yes Refl) = No (\eq => lemma_fst_neq_snd_eq p Refl eq)
decEq (a, b) (a', b') | (No p) | (No p') = No (\eq => lemma_both_neq p p' eq)
--------------------------------------------------------------------------------
-- List
--------------------------------------------------------------------------------
lemma_val_not_nil : (the (List _) (x :: xs) = Prelude.Nil {a = t} -> Void)
lemma_val_not_nil Refl impossible
lemma_x_eq_xs_neq : (x = y) -> (xs = ys -> Void) -> (the (List _) (x :: xs) = (y :: ys) -> Void)
lemma_x_eq_xs_neq Refl p Refl = p Refl
lemma_x_neq_xs_eq : (x = y -> Void) -> (xs = ys) -> (the (List _) (x :: xs) = (y :: ys) -> Void)
lemma_x_neq_xs_eq p Refl Refl = p Refl
lemma_x_neq_xs_neq : (x = y -> Void) -> (xs = ys -> Void) -> (the (List _) (x :: xs) = (y :: ys) -> Void)
lemma_x_neq_xs_neq p p' Refl = p Refl
implementation DecEq a => DecEq (List a) where
decEq [] [] = Yes Refl
decEq (x :: xs) [] = No lemma_val_not_nil
decEq [] (x :: xs) = No (negEqSym lemma_val_not_nil)
decEq (x :: xs) (y :: ys) with (decEq x y)
decEq (x :: xs) (x :: ys) | Yes Refl with (decEq xs ys)
decEq (x :: xs) (x :: xs) | (Yes Refl) | (Yes Refl) = Yes Refl
decEq (x :: xs) (x :: ys) | (Yes Refl) | (No p) = No (\eq => lemma_x_eq_xs_neq Refl p eq)
decEq (x :: xs) (y :: ys) | No p with (decEq xs ys)
decEq (x :: xs) (y :: xs) | (No p) | (Yes Refl) = No (\eq => lemma_x_neq_xs_eq p Refl eq)
decEq (x :: xs) (y :: ys) | (No p) | (No p') = No (\eq => lemma_x_neq_xs_neq p p' eq)
--------------------------------------------------------------------------------
-- Int
--------------------------------------------------------------------------------
export
implementation DecEq Int where
decEq x y = case x == y of -- Blocks if x or y not concrete
True => Yes primitiveEq
False => No primitiveNotEq
where primitiveEq : forall x, y . x = y
primitiveEq = believe_me (Refl {x})
primitiveNotEq : forall x, y . x = y -> Void
primitiveNotEq prf = believe_me {b = Void} ()
--------------------------------------------------------------------------------
-- Char
--------------------------------------------------------------------------------
export
implementation DecEq Char where
decEq x y = case x == y of -- Blocks if x or y not concrete
True => Yes primitiveEq
False => No primitiveNotEq
where primitiveEq : forall x, y . x = y
primitiveEq = believe_me (Refl {x})
primitiveNotEq : forall x, y . x = y -> Void
primitiveNotEq prf = believe_me {b = Void} ()
--------------------------------------------------------------------------------
-- Integer
--------------------------------------------------------------------------------
export
implementation DecEq Integer where
decEq x y = case x == y of -- Blocks if x or y not concrete
True => Yes primitiveEq
False => No primitiveNotEq
where primitiveEq : forall x, y . x = y
primitiveEq = believe_me (Refl {x})
primitiveNotEq : forall x, y . x = y -> Void
primitiveNotEq prf = believe_me {b = Void} ()
--------------------------------------------------------------------------------
-- String
--------------------------------------------------------------------------------
export
implementation DecEq String where
decEq x y = case x == y of -- Blocks if x or y not concrete
True => Yes primitiveEq
False => No primitiveNotEq
where primitiveEq : forall x, y . x = y
primitiveEq = believe_me (Refl {x})
primitiveNotEq : forall x, y . x = y -> Void
primitiveNotEq prf = believe_me {b = Void} ()

8
libs/base/Makefile Normal file
View File

@ -0,0 +1,8 @@
all:
${IDRIS2} --build base.ipkg
install:
${IDRIS2} --install base.ipkg
clean:
rm -rf build

129
libs/base/System.idr Normal file
View File

@ -0,0 +1,129 @@
module System
import Data.So
import Data.List
import Data.Strings
support : String -> String
support fn = "C:" ++ fn ++ ", libidris2_support"
libc : String -> String
libc fn = "C:" ++ fn ++ ", libc 6"
%foreign support "idris2_sleep"
prim_sleep : Int -> PrimIO ()
%foreign support "idris2_usleep"
prim_usleep : Int -> PrimIO ()
export
sleep : Int -> IO ()
sleep sec = primIO (prim_sleep sec)
export
usleep : (x : Int) -> So (x >= 0) => IO ()
usleep sec = primIO (prim_usleep sec)
-- This one is going to vary for different back ends. Probably needs a
-- better convention. Will revisit...
%foreign "scheme:blodwen-args"
prim__getArgs : PrimIO (List String)
export
getArgs : IO (List String)
getArgs = primIO prim__getArgs
%foreign libc "getenv"
prim_getEnv : String -> PrimIO (Ptr String)
%foreign support "idris2_getEnvPair"
prim_getEnvPair : Int -> PrimIO (Ptr String)
%foreign libc "setenv"
prim_setEnv : String -> String -> Int -> PrimIO Int
%foreign libc "setenv"
prim_unsetEnv : String -> PrimIO Int
export
getEnv : String -> IO (Maybe String)
getEnv var
= do env <- primIO $ prim_getEnv var
if prim__nullPtr env /= 0
then pure Nothing
else pure (Just (prim__getString env))
export
getEnvironment : IO (List (String, String))
getEnvironment = getAllPairs 0 []
where
splitEq : String -> (String, String)
splitEq str
= let (k, v) = break (== '=') str
(_, v') = break (/= '=') v in
(k, v')
getAllPairs : Int -> List String -> IO (List (String, String))
getAllPairs n acc = do
envPair <- primIO $ prim_getEnvPair n
if prim__nullPtr envPair /= 0
then pure $ reverse $ map splitEq acc
else getAllPairs (n + 1) (prim__getString envPair :: acc)
export
setEnv : String -> String -> Bool -> IO Bool
setEnv var val overwrite
= do ok <- primIO $ prim_setEnv var val (if overwrite then 1 else 0)
if ok == 0
then pure True
else pure False
export
unsetEnv : String -> IO Bool
unsetEnv var
= do ok <- primIO $ prim_unsetEnv var
if ok == 0
then pure True
else pure False
%foreign libc "system"
"scheme:blodwen-system"
prim_system : String -> PrimIO Int
export
system : String -> IO Int
system cmd = primIO (prim_system cmd)
%foreign support "idris2_time"
"scheme:blodwen-time"
prim_time : PrimIO Int
export
time : IO Integer
time = pure $ cast !(primIO prim_time)
%foreign libc "exit"
prim_exit : Int -> PrimIO ()
||| Programs can either terminate successfully, or end in a caught
||| failure.
public export
data ExitCode : Type where
||| Terminate successfully.
ExitSuccess : ExitCode
||| Program terminated for some prescribed reason.
|||
||| @errNo A non-zero numerical value indicating failure.
||| @prf Proof that the int value is non-zero.
ExitFailure : (errNo : Int) -> (So (not $ errNo == 0)) => ExitCode
export
exitWith : ExitCode -> IO a
exitWith ExitSuccess = believe_me $ primIO $ prim_exit 0
exitWith (ExitFailure code) = believe_me $ primIO $ prim_exit code
||| Exit the program indicating failure.
export
exitFailure : IO a
exitFailure = exitWith (ExitFailure 1)
||| Exit the program after a successful run.
export
exitSuccess : IO a
exitSuccess = exitWith ExitSuccess

150
libs/base/System/Clock.idr Normal file
View File

@ -0,0 +1,150 @@
module System.Clock
import PrimIO
||| The various types of system clock available.
public export
data ClockType
= UTC -- The time elapsed since the "epoch:" 00:00:00 UTC, January 1, 1970
| Monotonic -- The time elapsed since some arbitrary point in the past
| Duration -- Representing a time duration.
| Process -- The amount of CPU time used by the current process.
| Thread -- The amount of CPU time used by the current thread.
| GCCPU -- The current process's CPU time consumed by the garbage collector.
| GCReal -- The current process's real time consumed by the garbage collector.
export
Show ClockType where
show UTC = "UTC"
show Monotonic = "monotonic"
show Duration = "duration"
show Process = "process"
show Thread = "thread"
show GCCPU = "gcCPU"
show GCReal = "gcReal"
public export
data Clock : (type : ClockType) -> Type where
MkClock
: {type : ClockType}
-> (seconds : Integer)
-> (nanoseconds : Integer)
-> Clock type
public export
Eq (Clock type) where
(MkClock seconds1 nanoseconds1) == (MkClock seconds2 nanoseconds2) =
seconds1 == seconds2 && nanoseconds1 == nanoseconds2
public export
Ord (Clock type) where
compare (MkClock seconds1 nanoseconds1) (MkClock seconds2 nanoseconds2) =
case compare seconds1 seconds2 of
LT => LT
GT => GT
EQ => compare nanoseconds1 nanoseconds2
public export
Show (Clock type) where
show (MkClock {type} seconds nanoseconds) =
show type ++ ": " ++ show seconds ++ "s " ++ show nanoseconds ++ "ns"
||| A helper to deconstruct a Clock.
public export
seconds : Clock type -> Integer
seconds (MkClock s _) = s
||| A helper to deconstruct a Clock.
public export
nanoseconds : Clock type -> Integer
nanoseconds (MkClock _ ns) = ns
||| Make a duration value.
public export
makeDuration : Integer -> Integer -> Clock Duration
makeDuration = MkClock
||| Opaque clock value manipulated by the back-end.
data OSClock : Type where [external]
||| Note: Back-ends are required to implement UTC, monotonic time, CPU time
||| in current process/thread, and time duration; the rest are optional.
data ClockTypeMandatory
= Mandatory
| Optional
public export
isClockMandatory : ClockType -> ClockTypeMandatory
isClockMandatory GCCPU = Optional
isClockMandatory GCReal = Optional
isClockMandatory _ = Mandatory
prim_clockTimeMonotonic : IO OSClock
prim_clockTimeMonotonic = schemeCall OSClock "blodwen-clock-time-monotonic" []
fetchOSClock : ClockType -> IO OSClock
fetchOSClock UTC = schemeCall OSClock "blodwen-clock-time-utc" []
fetchOSClock Monotonic = prim_clockTimeMonotonic
fetchOSClock Process = schemeCall OSClock "blodwen-clock-time-process" []
fetchOSClock Thread = schemeCall OSClock "blodwen-clock-time-thread" []
fetchOSClock GCCPU = schemeCall OSClock "blodwen-clock-time-gccpu" []
fetchOSClock GCReal = schemeCall OSClock "blodwen-clock-time-gcreal" []
fetchOSClock Duration = prim_clockTimeMonotonic
||| A test to determine the status of optional clocks.
osClockValid : OSClock -> IO Int
osClockValid clk = schemeCall Int "blodwen-is-time?" [clk]
fromOSClock : {type : ClockType} -> OSClock -> IO (Clock type)
fromOSClock clk =
pure $
MkClock
{type}
!(schemeCall Integer "blodwen-clock-second" [clk])
!(schemeCall Integer "blodwen-clock-nanosecond" [clk])
public export
clockTimeReturnType : (typ : ClockType) -> Type
clockTimeReturnType typ with (isClockMandatory typ)
clockTimeReturnType typ | Optional = Maybe (Clock typ)
clockTimeReturnType typ | Mandatory = Clock typ
||| Fetch the system clock of a given kind. If the clock is mandatory,
||| we return a (Clock type) else (Maybe (Clock type)).
public export
clockTime : (typ : ClockType) -> IO (clockTimeReturnType typ)
clockTime clockType with (isClockMandatory clockType)
clockTime clockType | Mandatory = fetchOSClock clockType >>= fromOSClock
clockTime clockType | Optional = do
clk <- fetchOSClock clockType
valid <- map (== 1) $ osClockValid clk
if valid
then map Just $ fromOSClock clk
else pure Nothing
toNano : Clock type -> Integer
toNano (MkClock seconds nanoseconds) =
let scale = 1000000000
in scale * seconds + nanoseconds
fromNano : {type : ClockType} -> Integer -> Clock type
fromNano n =
let scale = 1000000000
seconds = n `div` scale
nanoseconds = n `mod` scale
in MkClock seconds nanoseconds
||| Compute difference between two clocks of the same type.
public export
timeDifference : Clock type -> Clock type -> Clock Duration
timeDifference clock duration = fromNano $ toNano clock - toNano duration
||| Add a duration to a clock value.
public export
addDuration : {type : ClockType} -> Clock type -> Clock Duration -> Clock type
addDuration clock duration = fromNano $ toNano clock + toNano duration
||| Subtract a duration from a clock value.
public export
subtractDuration : {type : ClockType} -> Clock type -> Clock Duration -> Clock type
subtractDuration clock duration = fromNano $ toNano clock - toNano duration

View File

@ -0,0 +1,83 @@
module System.Concurrency.Raw
-- At the moment this is pretty fundamentally tied to the Scheme RTS
-- Given that different back ends will have entirely different threading
-- models, it might be unavoidable, but we might want to think about possible
-- primitives that back ends should support.
%foreign "scheme:blodwen-thisthread"
prim__threadID : PrimIO ThreadID
%foreign "scheme:blodwen-set-thread-data"
prim__setThreadData : {a : Type} -> a -> PrimIO ()
%foreign "scheme:blodwen-get-thread-data"
prim__getThreadData : (a : Type) -> PrimIO a
-- Mutexes and condition variables.
export
threadID : IO ThreadID
threadID = primIO prim__threadID
export
setThreadData : {a : Type} -> a -> IO ()
setThreadData val = primIO (prim__setThreadData val)
export
getThreadData : (a : Type) -> IO a
getThreadData a = primIO (prim__getThreadData a)
export
data Mutex : Type where [external]
%foreign "scheme:blodwen-mutex"
prim__makeMutex : PrimIO Mutex
%foreign "scheme:blodwen-lock"
prim__mutexAcquire : Mutex -> PrimIO ()
%foreign "scheme:blodwen-unlock"
prim__mutexRelease : Mutex -> PrimIO ()
export
makeMutex : IO Mutex
makeMutex = primIO prim__makeMutex
export
mutexAcquire : Mutex -> IO ()
mutexAcquire m = primIO (prim__mutexAcquire m)
export
mutexRelease : Mutex -> IO ()
mutexRelease m = primIO (prim__mutexRelease m)
export
data Condition : Type where [external]
%foreign "scheme:blodwen-condition"
prim__makeCondition : PrimIO Condition
%foreign "scheme:blodwen-condition-wait"
prim__conditionWait : Condition -> Mutex -> PrimIO ()
%foreign "scheme:blodwen-condition-wait-timoue"
prim__conditionWaitTimeout : Condition -> Mutex -> Int -> PrimIO ()
%foreign "scheme:blodwen-condition-signal"
prim__conditionSignal : Condition -> PrimIO ()
%foreign "scheme:blodwen-condition-broadcast"
prim__conditionBroadcast : Condition -> PrimIO ()
export
makeCondition : IO Condition
makeCondition = primIO prim__makeCondition
export
conditionWait : Condition -> Mutex -> IO ()
conditionWait c m = primIO (prim__conditionWait c m)
export
conditionWaitTimeout : Condition -> Mutex -> Int -> IO ()
conditionWaitTimeout c m t = primIO (prim__conditionWaitTimeout c m t)
export
conditionSignal : Condition -> IO ()
conditionSignal c = primIO (prim__conditionSignal c)
export
conditionBroadcast : Condition -> IO ()
conditionBroadcast c = primIO (prim__conditionBroadcast c)

View File

@ -0,0 +1,91 @@
module System.Directory
import public System.File
public export
DirPtr : Type
DirPtr = AnyPtr
support : String -> String
support fn = "C:" ++ fn ++ ", libidris2_support"
%foreign support "idris2_fileErrno"
prim_fileErrno : PrimIO Int
returnError : IO (Either FileError a)
returnError
= do err <- primIO prim_fileErrno
case err of
0 => pure $ Left FileReadError
1 => pure $ Left FileWriteError
2 => pure $ Left FileNotFound
3 => pure $ Left PermissionDenied
4 => pure $ Left FileExists
_ => pure $ Left (GenericFileError (err-5))
ok : a -> IO (Either FileError a)
ok x = pure (Right x)
%foreign support "idris2_currentDirectory"
prim_currentDir : PrimIO (Ptr String)
%foreign support "idris2_changeDir"
prim_changeDir : String -> PrimIO Int
%foreign support "idris2_createDir"
prim_createDir : String -> PrimIO Int
%foreign support "idris2_dirOpen"
prim_openDir : String -> PrimIO DirPtr
%foreign support "idris2_dirClose"
prim_closeDir : DirPtr -> PrimIO ()
%foreign support "idris2_nextDirEntry"
prim_dirEntry : DirPtr -> PrimIO (Ptr String)
export
data Directory : Type where
MkDir : DirPtr -> Directory
export
createDir : String -> IO (Either FileError ())
createDir dir
= do res <- primIO (prim_createDir dir)
if res == 0
then ok ()
else returnError
export
changeDir : String -> IO Bool
changeDir dir
= do ok <- primIO (prim_changeDir dir)
pure (ok == 0)
export
currentDir : IO (Maybe String)
currentDir
= do res <- primIO prim_currentDir
if prim__nullPtr res /= 0
then pure Nothing
else pure (Just (prim__getString res))
export
dirOpen : String -> IO (Either FileError Directory)
dirOpen d
= do res <- primIO (prim_openDir d)
if prim__nullAnyPtr res /= 0
then returnError
else ok (MkDir res)
export
dirClose : Directory -> IO ()
dirClose (MkDir d) = primIO (prim_closeDir d)
export
dirEntry : Directory -> IO (Either FileError String)
dirEntry (MkDir d)
= do res <- primIO (prim_dirEntry d)
if prim__nullPtr res /= 0
then returnError
else ok (prim__getString res)

35
libs/base/System/FFI.idr Normal file
View File

@ -0,0 +1,35 @@
-- Additional FFI help for more interesting C types
-- Some assumptions are made about the existence of this module in
-- Compiler.CompileExpr
module System.FFI
export
data Struct : String -> -- C struct name
List (String, Type) -> -- field names and types
Type where
public export
data FieldType : String -> Type -> List (String, Type) -> Type where
First : FieldType n t ((n, t) :: ts)
Later : FieldType n t ts -> FieldType n t (f :: ts)
%extern
prim__getField : {s : _} -> forall fs, ty .
Struct s fs -> (n : String) ->
FieldType n ty fs -> ty
%extern
prim__setField : {s : _} -> forall fs, ty .
Struct s fs -> (n : String) ->
FieldType n ty fs -> ty -> PrimIO ()
public export %inline
getField : {s : _} -> Struct s fs -> (n : String) ->
{auto fieldok : FieldType n ty fs} -> ty
getField s n = prim__getField s n fieldok
public export %inline
setField : {s : _} -> Struct s fs -> (n : String) ->
{auto fieldok : FieldType n ty fs} -> ty -> IO ()
setField s n val = primIO (prim__setField s n fieldok val)

285
libs/base/System/File.idr Normal file
View File

@ -0,0 +1,285 @@
module System.File
import Data.List
import Data.Strings
public export
data Mode = Read | WriteTruncate | Append | ReadWrite | ReadWriteTruncate | ReadAppend
public export
FilePtr : Type
FilePtr = AnyPtr
support : String -> String
support fn = "C:" ++ fn ++ ", libidris2_support"
libc : String -> String
libc fn = "C:" ++ fn ++ ", libc 6"
%foreign support "idris2_openFile"
prim__open : String -> String -> Int -> PrimIO FilePtr
%foreign support "idris2_closeFile"
prim__close : FilePtr -> PrimIO ()
%foreign support "idris2_fileError"
prim_error : FilePtr -> PrimIO Int
%foreign support "idris2_fileErrno"
prim_fileErrno : PrimIO Int
%foreign support "idris2_readLine"
prim__readLine : FilePtr -> PrimIO (Ptr String)
%foreign support "idris2_readChars"
prim__readChars : Int -> FilePtr -> PrimIO (Ptr String)
%foreign support "idris2_writeLine"
prim__writeLine : FilePtr -> String -> PrimIO Int
%foreign support "idris2_eof"
prim__eof : FilePtr -> PrimIO Int
%foreign "C:fflush,libc 6"
prim__flush : FilePtr -> PrimIO Int
%foreign support "idris2_fileRemove"
prim__fileRemove : String -> PrimIO Int
%foreign support "idris2_fileSize"
prim__fileSize : FilePtr -> PrimIO Int
%foreign support "idris2_fileSize"
prim__fPoll : FilePtr -> PrimIO Int
%foreign support "idris2_fileAccessTime"
prim__fileAccessTime : FilePtr -> PrimIO Int
%foreign support "idris2_fileModifiedTime"
prim__fileModifiedTime : FilePtr -> PrimIO Int
%foreign support "idris2_fileStatusTime"
prim__fileStatusTime : FilePtr -> PrimIO Int
%foreign support "idris2_stdin"
prim__stdin : FilePtr
%foreign support "idris2_stdout"
prim__stdout : FilePtr
%foreign support "idris2_stderr"
prim__stderr : FilePtr
%foreign libc "chmod"
prim__chmod : String -> Int -> PrimIO Int
modeStr : Mode -> String
modeStr Read = "r"
modeStr WriteTruncate = "w"
modeStr Append = "a"
modeStr ReadWrite = "r+"
modeStr ReadWriteTruncate = "w+"
modeStr ReadAppend = "a+"
public export
data FileError = GenericFileError Int -- errno
| FileReadError
| FileWriteError
| FileNotFound
| PermissionDenied
| FileExists
returnError : IO (Either FileError a)
returnError
= do err <- primIO prim_fileErrno
case err of
0 => pure $ Left FileReadError
1 => pure $ Left FileWriteError
2 => pure $ Left FileNotFound
3 => pure $ Left PermissionDenied
4 => pure $ Left FileExists
_ => pure $ Left (GenericFileError (err-5))
export
Show FileError where
show (GenericFileError errno) = "File error: " ++ show errno
show FileReadError = "File Read Error"
show FileWriteError = "File Write Error"
show FileNotFound = "File Not Found"
show PermissionDenied = "Permission Denied"
show FileExists = "File Exists"
ok : a -> IO (Either FileError a)
ok x = pure (Right x)
public export
data File : Type where
FHandle : FilePtr -> File
export
stdin : File
stdin = FHandle prim__stdin
export
stdout : File
stdout = FHandle prim__stdout
export
stderr : File
stderr = FHandle prim__stderr
export
openFile : String -> Mode -> IO (Either FileError File)
openFile f m
= do res <- primIO (prim__open f (modeStr m) 0)
if prim__nullAnyPtr res /= 0
then returnError
else ok (FHandle res)
export
closeFile : File -> IO ()
closeFile (FHandle f) = primIO (prim__close f)
export
fGetLine : (h : File) -> IO (Either FileError String)
fGetLine (FHandle f)
= do res <- primIO (prim__readLine f)
if prim__nullPtr res /= 0
then returnError
else ok (prim__getString res)
export
fGetChars : (h : File) -> Int -> IO (Either FileError String)
fGetChars (FHandle f) max
= do res <- primIO (prim__readChars max f)
if prim__nullPtr res /= 0
then returnError
else ok (prim__getString res)
export
fPutStr : (h : File) -> String -> IO (Either FileError ())
fPutStr (FHandle f) str
= do res <- primIO (prim__writeLine f str)
if res == 0
then returnError
else ok ()
export
fPutStrLn : (h : File) -> String -> IO (Either FileError ())
fPutStrLn f str = fPutStr f (str ++ "\n")
export
fEOF : (h : File) -> IO Bool
fEOF (FHandle f)
= do res <- primIO (prim__eof f)
pure (res /= 0)
export
fflush : (h : File) -> IO ()
fflush (FHandle f)
= do primIO (prim__flush f)
pure ()
export
fileAccessTime : (h : File) -> IO (Either FileError Int)
fileAccessTime (FHandle f)
= do res <- primIO (prim__fileAccessTime f)
if res > 0
then ok res
else returnError
export
fileModifiedTime : (h : File) -> IO (Either FileError Int)
fileModifiedTime (FHandle f)
= do res <- primIO (prim__fileModifiedTime f)
if res > 0
then ok res
else returnError
export
fileStatusTime : (h : File) -> IO (Either FileError Int)
fileStatusTime (FHandle f)
= do res <- primIO (prim__fileStatusTime f)
if res > 0
then ok res
else returnError
export
fileRemove : String -> IO (Either FileError ())
fileRemove fname
= do res <- primIO (prim__fileRemove fname)
if res == 0
then ok ()
else returnError
export
fileSize : (h : File) -> IO (Either FileError Int)
fileSize (FHandle f)
= do res <- primIO (prim__fileSize f)
if res >= 0
then ok res
else returnError
export
fPoll : File -> IO Bool
fPoll (FHandle f)
= do p <- primIO (prim__fPoll f)
pure (p > 0)
export
readFile : String -> IO (Either FileError String)
readFile file
= do Right h <- openFile file Read
| Left err => returnError
Right content <- read [] h
| Left err => do closeFile h
returnError
closeFile h
pure (Right (fastAppend content))
where
read : List String -> File -> IO (Either FileError (List String))
read acc h
= do eof <- fEOF h
if eof
then pure (Right (reverse acc))
else
do Right str <- fGetLine h
| Left err => returnError
read (str :: acc) h
||| Write a string to a file
export
writeFile : (filepath : String) -> (contents : String) ->
IO (Either FileError ())
writeFile fn contents = do
Right h <- openFile fn WriteTruncate
| Left err => pure (Left err)
Right () <- fPutStr h contents
| Left err => do closeFile h
pure (Left err)
closeFile h
pure (Right ())
namespace FileMode
public export
data FileMode = Read | Write | Execute
public export
record Permissions where
constructor MkPermissions
user : List FileMode
group : List FileMode
others : List FileMode
mkMode : Permissions -> Int
mkMode p
= getMs (user p) * 64 + getMs (group p) * 8 + getMs (others p)
where
getM : FileMode -> Int
getM Read = 4
getM Write = 2
getM Execute = 1
getMs : List FileMode -> Int
getMs = sum . map getM
export
chmodRaw : String -> Int -> IO (Either FileError ())
chmodRaw fname p
= do ok <- primIO $ prim__chmod fname p
if ok == 0
then pure (Right ())
else returnError
export
chmod : String -> Permissions -> IO (Either FileError ())
chmod fname p = chmodRaw fname (mkMode p)

12
libs/base/System/Info.idr Normal file
View File

@ -0,0 +1,12 @@
module System.Info
%extern prim__os : String
%extern prim__codegen : String
export
os : String
os = prim__os
export
codegen : String
codegen = prim__codegen

36
libs/base/System/REPL.idr Normal file
View File

@ -0,0 +1,36 @@
module System.REPL
import System.File
||| A basic read-eval-print loop, maintaining a state
||| @ state the input state
||| @ prompt the prompt to show
||| @ onInput the function to run on reading input, returning a String to
||| output and a new state. Returns Nothing if the repl should exit
export
replWith : (state : a) -> (prompt : String) ->
(onInput : a -> String -> Maybe (String, a)) -> IO ()
replWith acc prompt fn
= do eof <- fEOF stdin
if eof
then pure ()
else do putStr prompt
fflush stdout
x <- getLine
case fn acc x of
Just (out, acc') =>
do putStr out
replWith acc' prompt fn
Nothing => pure ()
||| A basic read-eval-print loop
||| @ prompt the prompt to show
||| @ onInput the function to run on reading input, returning a String to
||| output
export
repl : (prompt : String) ->
(onInput : String -> String) -> IO ()
repl prompt fn
= replWith () prompt (\x, s => Just (fn s, ()))

44
libs/base/base.ipkg Normal file
View File

@ -0,0 +1,44 @@
package base
modules = Control.App,
Control.App.Console,
Control.App.FileIO,
Control.Monad.Identity,
Control.Monad.State,
Control.Monad.Trans,
Control.WellFounded,
Data.Buffer,
Data.Either,
Data.Fin,
Data.IOArray,
Data.IORef,
Data.List,
Data.List.Elem,
Data.List.Views,
Data.Maybe,
Data.Morphisms,
Data.Nat,
Data.Nat.Views,
Data.Primitives.Views,
Data.So,
Data.Stream,
Data.Strings,
Data.Vect,
Debug.Trace,
Decidable.Equality,
Language.Reflection,
Language.Reflection.TT,
Language.Reflection.TTImp,
System,
System.Concurrency.Raw,
System.Directory,
System.File,
System.FFI,
System.Info,
System.REPL

View File

@ -0,0 +1,14 @@
||| Utilities functions for contitionally delaying values.
module Control.Delayed
||| Type-level function for a conditionally infinite type.
public export
inf : Bool -> Type -> Type
inf False t = t
inf True t = Inf t
||| Type-level function for a conditionally lazy type.
public export
lazy : Bool -> Type -> Type
lazy False t = t
lazy True t = Lazy t

View File

@ -0,0 +1,36 @@
module Data.Bool.Extra
public export
andSameNeutral : (x : Bool) -> x && x = x
andSameNeutral False = Refl
andSameNeutral True = Refl
public export
andFalseFalse : (x : Bool) -> x && False = False
andFalseFalse False = Refl
andFalseFalse True = Refl
public export
andTrueNeutral : (x : Bool) -> x && True = x
andTrueNeutral False = Refl
andTrueNeutral True = Refl
public export
orSameNeutral : (x : Bool) -> x || x = x
orSameNeutral False = Refl
orSameNeutral True = Refl
public export
orFalseNeutral : (x : Bool) -> x || False = x
orFalseNeutral False = Refl
orFalseNeutral True = Refl
public export
orTrueTrue : (x : Bool) -> x || True = True
orTrueTrue False = Refl
orTrueTrue True = Refl
public export
orSameAndRightNeutral : (x, right : Bool) -> x || (x && right) = x
orSameAndRightNeutral False _ = Refl
orSameAndRightNeutral True _ = Refl

View File

@ -0,0 +1,394 @@
||| Contains:
|||
||| 1. Tail recursive versions of the list processing functions from
||| Data.List.
|||
||| 2. Extensional equality proofs that these variants are
||| (extensionally) equivalent to their non-tail-recursive
||| counterparts.
|||
||| Note:
|||
||| 1. Written in one sitting, feel free to refactor
|||
||| 2. The proofs below also work on non-publicly exported
||| definitions. This could be due to a bug, and will need to be
||| moved elsewhere if it's fixed.
module Data.List.TailRec
import Syntax.PreorderReasoning
import Syntax.WithProof
import Data.List
import Data.Vect
import Data.Nat
length_aux : List a -> Nat -> Nat
length_aux [] len = len
length_aux (_::xs) len = length_aux xs (S len)
export
length : List a -> Nat
length xs = length_aux xs Z
export
length_ext : (xs : List a) -> Data.List.length xs = Data.List.TailRec.length xs
length_ext xs = Calculate [
Data.List.length xs
==| sym (plusZeroRightNeutral (Data.List.length xs)),
Data.List.length xs + 0
==| lemma 0 xs,
Data.List.TailRec.length_aux xs 0
==| Refl,
Data.List.TailRec.length xs
==| QED
]
where
lemma : (n : Nat) -> (xs : List a) ->
Data.List.length xs + n = length_aux xs n
lemma n [] = Refl
lemma n (_ :: xs) =
let length_xs : Nat
length_xs = Data.List.length xs in
Calculate [
1 + (length_xs + n)
-- Hopefully we could Frex these two steps one day
==| plusAssociative 1 length_xs n ,
(1 + length_xs) + n
==| cong (+n) (plusCommutative 1 length_xs) ,
(length_xs + 1) + n
==| sym (plusAssociative length_xs 1 n),
(length_xs) + (1 + n)
==| lemma (1 + n) xs,
length_aux xs (1 + n)
==| QED
]
take_aux : Nat -> List a -> List a -> List a
take_aux Z xs acc = reverseOnto [] acc
take_aux (S n) [] acc = reverseOnto [] acc
take_aux (S n) (x :: xs) acc = take_aux n xs (x :: acc)
export
take : Nat -> List a -> List a
take n xs = take_aux n xs []
export
take_ext :
(n : Nat) -> (xs : List a) ->
Data.List.take n xs = Data.List.TailRec.take n xs
take_ext n xs = Calculate [
Data.List.take n xs
==| Refl ,
reverse [] ++ (Data.List.take n xs)
==| sym (revOnto (Data.List.take n xs) []) ,
reverseOnto (Data.List.take n xs) []
==| sym (lemma n xs []) ,
take_aux n xs []
==| QED
]
where
lemma : (n : Nat) -> (xs, acc : List a) ->
take_aux n xs acc = reverseOnto (Data.List.take n xs) acc
lemma Z xs acc = Refl
lemma (S n) [] acc = Refl
lemma (S n) (x::xs) acc = lemma n xs (x :: acc)
span_aux : (a -> Bool) -> List a -> List a -> (List a, List a)
span_aux p [] acc = (reverseOnto [] acc, [])
span_aux p (x::xs) acc =
if p x then
Data.List.TailRec.span_aux p xs (x :: acc)
else
(reverseOnto [] acc, x::xs)
export
span : (a -> Bool) -> List a -> (List a, List a)
span p xs = span_aux p xs []
coe : (f : (i : a) -> Type) -> i = i' -> f i -> f i'
coe f Refl x = x
span_aux_ext : (p : a -> Bool) -> (xs, acc : List a) ->
(reverseOnto (fst $ Data.List.span p xs) acc, snd $ Data.List.span p xs)
=
span_aux p xs acc
span_aux_ext p [] acc = Refl
-- This is disgusting. Please teach me a better way.
span_aux_ext p (x::xs) acc with (@@(p x), @@(Data.List.span p xs))
span_aux_ext p (x::xs) acc | ((True ** px_tru), ((pre, rest)**dl_pf)) =
rewrite px_tru in
rewrite dl_pf in
let u = span_aux_ext p xs (x::acc) in
coe (\u => (reverseOnto (x :: fst u) acc, snd u) =
Data.List.TailRec.span_aux p xs (x :: acc)) dl_pf u
span_aux_ext p (x::xs) acc | ((False**px_fls), ((pre,rest)**dl_pf)) =
rewrite px_fls in
Refl
export
span_ext : (p : a -> Bool) -> (xs : List a) ->
Data.List.span p xs = Data.List.TailRec.span p xs
span_ext p xs with (@@(Data.List.span p xs))
span_ext p xs | ((pre, rest) ** pf) =
rewrite pf in
let u = span_aux_ext p xs [] in
coe (\u => (fst u, snd u) = span_aux p xs []) pf u
export
break : (a -> Bool) -> List a -> (List a, List a)
break p xs = Data.List.TailRec.span (not . p) xs
export
break_ext : (p : a -> Bool) -> (xs : List a) ->
Data.List.break p xs = Data.List.TailRec.break p xs
break_ext p xs = span_ext (not . p) xs
splitOnto : List (List a) -> (a -> Bool) -> List a -> List (List a)
splitOnto acc p xs =
case Data.List.break p xs of
(chunk, [] ) => reverseOnto [chunk] acc
(chunk, (c::rest)) => splitOnto (chunk::acc) p rest
export
split : (a -> Bool) -> List a -> List (List a)
split p xs = splitOnto [] p xs
splitOnto_ext : (acc : List (List a)) -> (p : a -> Bool) -> (xs : List a) ->
reverseOnto (Data.List.split p xs) acc
= Data.List.TailRec.splitOnto acc p xs
splitOnto_ext acc p xs with (@@(Data.List.break p xs))
splitOnto_ext acc p xs | ((chunk, [] )**pf) =
rewrite pf in
Refl
splitOnto_ext acc p xs | ((chunk, c::rest)**pf) =
rewrite pf in
rewrite splitOnto_ext (chunk::acc) p rest in
Refl
export
split_ext : (p : a -> Bool) -> (xs : List a) ->
Data.List.split p xs = Data.List.TailRec.split p xs
split_ext p xs = splitOnto_ext [] p xs
splitAtOnto : List a -> (n : Nat) -> (xs : List a) -> (List a, List a)
splitAtOnto acc Z xs = (reverseOnto [] acc, xs)
splitAtOnto acc (S n) [] = (reverseOnto [] acc, [])
splitAtOnto acc (S n) (x::xs) = splitAtOnto (x::acc) n xs
export
splitAt : (n : Nat) -> (xs : List a) -> (List a, List a)
splitAt n xs = splitAtOnto [] n xs
splitAtOnto_ext : (acc : List a) -> (n : Nat) -> (xs : List a) ->
(reverseOnto (fst $ Data.List.splitAt n xs) acc, snd $ Data.List.splitAt n xs)
= splitAtOnto acc n xs
splitAtOnto_ext acc Z xs = Refl
splitAtOnto_ext acc (S n) [] = Refl
splitAtOnto_ext acc (S n) (x::xs) with (@@(Data.List.splitAt n xs))
splitAtOnto_ext acc (S n) (x::xs) | ((tk, dr)**pf) =
rewrite pf in
let u = splitAtOnto_ext (x::acc) n xs in
coe (\u => (reverseOnto (x :: fst u) acc, snd u) =
splitAtOnto (x :: acc) n xs) pf u
export
splitAt_ext : (n : Nat) -> (xs : List a) ->
Data.List.splitAt n xs =
Data.List.TailRec.splitAt n xs
splitAt_ext n xs with (@@(Data.List.splitAt n xs))
splitAt_ext n xs | ((tk, dr)**pf) =
rewrite pf in
rewrite sym $ splitAtOnto_ext [] n xs in
rewrite pf in
Refl
partitionOnto : List a -> List a -> (a -> Bool) -> List a -> (List a, List a)
partitionOnto lfts rgts p [] = (reverseOnto [] lfts, reverseOnto [] rgts)
partitionOnto lfts rgts p (x::xs) =
if p x then
partitionOnto (x :: lfts) rgts p xs
else
partitionOnto lfts (x::rgts) p xs
export
partition : (a -> Bool) -> List a -> (List a, List a)
partition p xs = partitionOnto [] [] p xs
partitionOnto_ext : (lfts, rgts : List a) -> (p : a -> Bool) -> (xs : List a) ->
(reverseOnto (fst $ Data.List.partition p xs) lfts
,reverseOnto (snd $ Data.List.partition p xs) rgts)
= Data.List.TailRec.partitionOnto lfts rgts p xs
partitionOnto_ext lfts rgts p [] = Refl
partitionOnto_ext lfts rgts p (x::xs) with (@@(p x), @@(Data.List.partition p xs))
partitionOnto_ext lfts rgts p (x::xs) | ((True **px_tru), ((dl_l, dl_r)**dl_pf))
= rewrite px_tru in
rewrite dl_pf in
rewrite px_tru in
let u = partitionOnto_ext (x :: lfts) rgts p xs in
coe (\u => (reverseOnto (x :: fst u) lfts
,reverseOnto ( snd u) rgts)
= partitionOnto (x :: lfts) rgts p xs) dl_pf u
partitionOnto_ext lfts rgts p (x::xs) | ((False**px_fls), ((dl_l, dl_r)**dl_pf))
= rewrite px_fls in
rewrite dl_pf in
rewrite px_fls in
let u = partitionOnto_ext lfts (x :: rgts) p xs in
coe (\u => (reverseOnto ( fst u) lfts
,reverseOnto (x :: snd u) rgts)
= partitionOnto lfts (x::rgts) p xs) dl_pf u
mergeReplicate_aux : a -> List a -> List a -> List a
mergeReplicate_aux sep [] acc = reverseOnto [] acc
mergeReplicate_aux sep (x::xs) acc = mergeReplicate_aux sep xs (x :: sep :: acc)
mergeReplicate_ext : (sep : a) -> (xs : List a) -> (acc : List a) ->
mergeReplicate_aux sep xs acc =
reverseOnto (mergeReplicate sep xs) acc
mergeReplicate_ext sep [] acc = Refl
mergeReplicate_ext sep (x::xs) acc = Calculate [
mergeReplicate_aux sep xs (x :: sep :: acc)
==| mergeReplicate_ext sep xs (x :: sep :: acc) ,
reverseOnto (sep :: x :: mergeReplicate sep xs) acc
==| QED
]
export
intersperse : a -> List a -> List a
intersperse sep [] = []
intersperse sep (y::ys) = y :: mergeReplicate_aux sep ys []
export
intersperse_ext : (sep : a) -> (xs : List a) ->
Data.List.intersperse sep xs =
Data.List.TailRec.intersperse sep xs
intersperse_ext sep [] = Refl
intersperse_ext sep (y::ys) = cong (y::) (sym $ mergeReplicate_ext sep ys [])
mapMaybeOnto : List b -> (a -> Maybe b) -> List a -> List b
mapMaybeOnto acc f [] = reverseOnto [] acc
mapMaybeOnto acc f (x::xs) =
case f x of
Nothing => mapMaybeOnto acc f xs
Just y => mapMaybeOnto (y :: acc) f xs
export
mapMaybe : (a -> Maybe b) -> List a -> List b
mapMaybe f xs = mapMaybeOnto [] f xs
mapMaybeOnto_ext : (acc : List b) -> (f : a -> Maybe b) -> (xs : List a) ->
reverseOnto (Data.List.mapMaybe f xs) acc
=
mapMaybeOnto acc f xs
mapMaybeOnto_ext acc f [] = Refl
mapMaybeOnto_ext acc f (x::xs) with (f x)
mapMaybeOnto_ext acc f (x::xs) | Nothing = mapMaybeOnto_ext acc f xs
mapMaybeOnto_ext acc f (x::xs) | Just y = mapMaybeOnto_ext (y :: acc) f xs
export
mapMaybe_ext : (f : a -> Maybe b) -> (xs : List a) ->
Data.List.mapMaybe f xs = Data.List.TailRec.mapMaybe f xs
mapMaybe_ext f xs = mapMaybeOnto_ext [] f xs
export
sorted : Ord a => List a -> Bool
sorted [ ] = True
sorted [x] = True
sorted (x :: xs@(y :: ys)) = case (x <= y) of
False => False
True => Data.List.TailRec.sorted xs
export
sorted_ext : Ord a => (xs : List a) ->
Data.List.sorted xs = Data.List.TailRec.sorted xs
sorted_ext [] = Refl
sorted_ext [x] = Refl
sorted_ext (x :: y :: ys) with (x <= y)
sorted_ext (x :: y :: ys) | False = Refl
sorted_ext (x :: y :: ys) | True = sorted_ext (y::ys)
mergeByOnto : List a -> (a -> a -> Ordering) -> List a -> List a -> List a
mergeByOnto acc order [] right = reverseOnto right acc
mergeByOnto acc order left [] = reverseOnto left acc
mergeByOnto acc order left@(x::xs) right@(y::ys) =
case order x y of
LT => mergeByOnto (x :: acc) order xs right
_ => mergeByOnto (y :: acc) order left ys
mergeByOnto_ext : (acc : List a)
-> (order : a -> a -> Ordering)
-> (left, right : List a)
-> reverseOnto (mergeBy order left right) acc
= mergeByOnto acc order left right
mergeByOnto_ext acc order [] right = Refl
mergeByOnto_ext acc order left [] = ?bug139
mergeByOnto_ext acc order left@(x::xs) right@(y::ys) with (order x y)
mergeByOnto_ext acc order left@(x::xs) right@(y::ys) | LT =
mergeByOnto_ext (x :: acc) order xs (y::ys)
-- We need to exhaust the two other cases explicitly to convince the
-- typecheker. See #140
mergeByOnto_ext acc order left@(x::xs) right@(y::ys) | EQ =
mergeByOnto_ext (y :: acc) order (x::xs) ys
mergeByOnto_ext acc order left@(x::xs) right@(y::ys) | GT =
mergeByOnto_ext (y :: acc) order (x::xs) ys
export
mergeBy : (a -> a -> Ordering) -> List a -> List a -> List a
mergeBy order left right = mergeByOnto [] order left right
export
mergeBy_ext : (order : a -> a -> Ordering) -> (left, right : List a) ->
Data.List.mergeBy order left right =
Data.List.TailRec.mergeBy order left right
mergeBy_ext order left right = mergeByOnto_ext [] order left right
export
merge : Ord a => List a -> List a -> List a
merge = Data.List.TailRec.mergeBy compare
export
merge_ext : Ord a => (left, right : List a) ->
Data.List.merge left right = Data.List.TailRec.merge left right
merge_ext left right = mergeBy_ext compare left right
-- Not quite finished due to a bug.
sortBy_splitRec : List a -> List a -> (List a -> List a) -> (List a, List a)
sortBy_splitRec (_::_::xs) (y::ys) zs = sortBy_splitRec xs ys (zs . ((::) y))
sortBy_splitRec _ ys zs = (zs [], ys)
sortBy_split : List a -> (List a, List a)
sortBy_split xs = sortBy_splitRec xs xs id
export
sortBy : (cmp : a -> a -> Ordering) -> (xs : List a) -> List a
sortBy cmp [] = []
sortBy cmp [x] = [x]
sortBy cmp xs = let (x, y) = sortBy_split xs in
Data.List.TailRec.mergeBy cmp
(Data.List.TailRec.sortBy cmp (assert_smaller xs x))
(Data.List.TailRec.sortBy cmp (assert_smaller xs y))
{-- -- This code seems to make Idris2 loop
export
sortBy_ext : (cmp : a -> a -> Ordering) -> (xs : List a) ->
Data.List.sortBy cmp xs = Data.List.TailRec.sortBy cmp xs
sortBy_ext cmp [] = Refl
sortBy_ext cmp [x] = Refl
sortBy_ext cmp xs@(y::ys) = let (x, y) = split xs in
Calculate [
Data.List.mergeBy cmp
(Data.List.sortBy cmp (assert_smaller xs x))
(Data.List.sortBy cmp (assert_smaller xs y))
==| ?help1 ,
Data.List.TailRec.mergeBy cmp
(Data.List.TailRec.sortBy cmp (assert_smaller xs x))
(Data.List.TailRec.sortBy cmp (assert_smaller xs y))
==| QED
]
--}

View File

@ -0,0 +1,332 @@
module Data.SortedMap
-- TODO: write split
private
data Tree : Nat -> (k : Type) -> Type -> Ord k -> Type where
Leaf : k -> v -> Tree Z k v o
Branch2 : Tree n k v o -> k -> Tree n k v o -> Tree (S n) k v o
Branch3 : Tree n k v o -> k -> Tree n k v o -> k -> Tree n k v o -> Tree (S n) k v o
branch4 :
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o ->
Tree (S (S n)) k v o
branch4 a b c d e f g =
Branch2 (Branch2 a b c) d (Branch2 e f g)
branch5 :
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o ->
Tree (S (S n)) k v o
branch5 a b c d e f g h i =
Branch2 (Branch2 a b c) d (Branch3 e f g h i)
branch6 :
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o ->
Tree (S (S n)) k v o
branch6 a b c d e f g h i j k =
Branch3 (Branch2 a b c) d (Branch2 e f g) h (Branch2 i j k)
branch7 :
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o -> k ->
Tree n k v o ->
Tree (S (S n)) k v o
branch7 a b c d e f g h i j k l m =
Branch3 (Branch3 a b c d e) f (Branch2 g h i) j (Branch2 k l m)
merge1 : Tree n k v o -> k -> Tree (S n) k v o -> k -> Tree (S n) k v o -> Tree (S (S n)) k v o
merge1 a b (Branch2 c d e) f (Branch2 g h i) = branch5 a b c d e f g h i
merge1 a b (Branch2 c d e) f (Branch3 g h i j k) = branch6 a b c d e f g h i j k
merge1 a b (Branch3 c d e f g) h (Branch2 i j k) = branch6 a b c d e f g h i j k
merge1 a b (Branch3 c d e f g) h (Branch3 i j k l m) = branch7 a b c d e f g h i j k l m
merge2 : Tree (S n) k v o -> k -> Tree n k v o -> k -> Tree (S n) k v o -> Tree (S (S n)) k v o
merge2 (Branch2 a b c) d e f (Branch2 g h i) = branch5 a b c d e f g h i
merge2 (Branch2 a b c) d e f (Branch3 g h i j k) = branch6 a b c d e f g h i j k
merge2 (Branch3 a b c d e) f g h (Branch2 i j k) = branch6 a b c d e f g h i j k
merge2 (Branch3 a b c d e) f g h (Branch3 i j k l m) = branch7 a b c d e f g h i j k l m
merge3 : Tree (S n) k v o -> k -> Tree (S n) k v o -> k -> Tree n k v o -> Tree (S (S n)) k v o
merge3 (Branch2 a b c) d (Branch2 e f g) h i = branch5 a b c d e f g h i
merge3 (Branch2 a b c) d (Branch3 e f g h i) j k = branch6 a b c d e f g h i j k
merge3 (Branch3 a b c d e) f (Branch2 g h i) j k = branch6 a b c d e f g h i j k
merge3 (Branch3 a b c d e) f (Branch3 g h i j k) l m = branch7 a b c d e f g h i j k l m
treeLookup : Ord k => k -> Tree n k v o -> Maybe v
treeLookup k (Leaf k' v) =
if k == k' then
Just v
else
Nothing
treeLookup k (Branch2 t1 k' t2) =
if k <= k' then
treeLookup k t1
else
treeLookup k t2
treeLookup k (Branch3 t1 k1 t2 k2 t3) =
if k <= k1 then
treeLookup k t1
else if k <= k2 then
treeLookup k t2
else
treeLookup k t3
treeInsert' : Ord k => k -> v -> Tree n k v o -> Either (Tree n k v o) (Tree n k v o, k, Tree n k v o)
treeInsert' k v (Leaf k' v') =
case compare k k' of
LT => Right (Leaf k v, k, Leaf k' v')
EQ => Left (Leaf k v)
GT => Right (Leaf k' v', k', Leaf k v)
treeInsert' k v (Branch2 t1 k' t2) =
if k <= k' then
case treeInsert' k v t1 of
Left t1' => Left (Branch2 t1' k' t2)
Right (a, b, c) => Left (Branch3 a b c k' t2)
else
case treeInsert' k v t2 of
Left t2' => Left (Branch2 t1 k' t2')
Right (a, b, c) => Left (Branch3 t1 k' a b c)
treeInsert' k v (Branch3 t1 k1 t2 k2 t3) =
if k <= k1 then
case treeInsert' k v t1 of
Left t1' => Left (Branch3 t1' k1 t2 k2 t3)
Right (a, b, c) => Right (Branch2 a b c, k1, Branch2 t2 k2 t3)
else
if k <= k2 then
case treeInsert' k v t2 of
Left t2' => Left (Branch3 t1 k1 t2' k2 t3)
Right (a, b, c) => Right (Branch2 t1 k1 a, b, Branch2 c k2 t3)
else
case treeInsert' k v t3 of
Left t3' => Left (Branch3 t1 k1 t2 k2 t3')
Right (a, b, c) => Right (Branch2 t1 k1 t2, k2, Branch2 a b c)
treeInsert : Ord k => k -> v -> Tree n k v o -> Either (Tree n k v o) (Tree (S n) k v o)
treeInsert k v t =
case treeInsert' k v t of
Left t' => Left t'
Right (a, b, c) => Right (Branch2 a b c)
delType : Nat -> (k : Type) -> Type -> Ord k -> Type
delType Z k v o = ()
delType (S n) k v o = Tree n k v o
treeDelete : Ord k => (n : Nat) -> k -> Tree n k v o -> Either (Tree n k v o) (delType n k v o)
treeDelete _ k (Leaf k' v) =
if k == k' then
Right ()
else
Left (Leaf k' v)
treeDelete (S Z) k (Branch2 t1 k' t2) =
if k <= k' then
case treeDelete Z k t1 of
Left t1' => Left (Branch2 t1' k' t2)
Right () => Right t2
else
case treeDelete Z k t2 of
Left t2' => Left (Branch2 t1 k' t2')
Right () => Right t1
treeDelete (S Z) k (Branch3 t1 k1 t2 k2 t3) =
if k <= k1 then
case treeDelete Z k t1 of
Left t1' => Left (Branch3 t1' k1 t2 k2 t3)
Right () => Left (Branch2 t2 k2 t3)
else if k <= k2 then
case treeDelete Z k t2 of
Left t2' => Left (Branch3 t1 k1 t2' k2 t3)
Right () => Left (Branch2 t1 k1 t3)
else
case treeDelete Z k t3 of
Left t3' => Left (Branch3 t1 k1 t2 k2 t3')
Right () => Left (Branch2 t1 k1 t2)
treeDelete (S (S n)) k (Branch2 t1 k' t2) =
if k <= k' then
case treeDelete (S n) k t1 of
Left t1' => Left (Branch2 t1' k' t2)
Right t1' =>
case t2 of
Branch2 a b c => Right (Branch3 t1' k' a b c)
Branch3 a b c d e => Left (branch4 t1' k' a b c d e)
else
case treeDelete (S n) k t2 of
Left t2' => Left (Branch2 t1 k' t2')
Right t2' =>
case t1 of
Branch2 a b c => Right (Branch3 a b c k' t2')
Branch3 a b c d e => Left (branch4 a b c d e k' t2')
treeDelete (S (S n)) k (Branch3 t1 k1 t2 k2 t3) =
if k <= k1 then
case treeDelete (S n) k t1 of
Left t1' => Left (Branch3 t1' k1 t2 k2 t3)
Right t1' => Left (merge1 t1' k1 t2 k2 t3)
else if k <= k2 then
case treeDelete (S n) k t2 of
Left t2' => Left (Branch3 t1 k1 t2' k2 t3)
Right t2' => Left (merge2 t1 k1 t2' k2 t3)
else
case treeDelete (S n) k t3 of
Left t3' => Left (Branch3 t1 k1 t2 k2 t3')
Right t3' => Left (merge3 t1 k1 t2 k2 t3')
treeToList : Tree n k v o -> List (k, v)
treeToList = treeToList' (:: [])
where
-- explicit quantification to avoid conflation with {n} from treeToList
treeToList' : {0 n : Nat} -> ((k, v) -> List (k, v)) -> Tree n k v o -> List (k, v)
treeToList' cont (Leaf k v) = cont (k, v)
treeToList' cont (Branch2 t1 _ t2) = treeToList' (:: treeToList' cont t2) t1
treeToList' cont (Branch3 t1 _ t2 _ t3) = treeToList' (:: treeToList' (:: treeToList' cont t3) t2) t1
export
data SortedMap : Type -> Type -> Type where
Empty : Ord k => SortedMap k v
M : (o : Ord k) => (n:Nat) -> Tree n k v o -> SortedMap k v
export
empty : Ord k => SortedMap k v
empty = Empty
export
lookup : k -> SortedMap k v -> Maybe v
lookup _ Empty = Nothing
lookup k (M _ t) = treeLookup k t
export
insert : k -> v -> SortedMap k v -> SortedMap k v
insert k v Empty = M Z (Leaf k v)
insert k v (M _ t) =
case treeInsert k v t of
Left t' => (M _ t')
Right t' => (M _ t')
export
insertFrom : Foldable f => f (k, v) -> SortedMap k v -> SortedMap k v
insertFrom = flip $ foldl $ flip $ uncurry insert
export
delete : k -> SortedMap k v -> SortedMap k v
delete _ Empty = Empty
delete k (M Z t) =
case treeDelete Z k t of
Left t' => (M _ t')
Right () => Empty
delete k (M (S n) t) =
case treeDelete (S n) k t of
Left t' => (M _ t')
Right t' => (M _ t')
export
fromList : Ord k => List (k, v) -> SortedMap k v
fromList l = foldl (flip (uncurry insert)) empty l
export
toList : SortedMap k v -> List (k, v)
toList Empty = []
toList (M _ t) = treeToList t
||| Gets the keys of the map.
export
keys : SortedMap k v -> List k
keys = map fst . toList
||| Gets the values of the map. Could contain duplicates.
export
values : SortedMap k v -> List v
values = map snd . toList
export
null : SortedMap k v -> Bool
null Empty = True
null (M _ _) = False
treeMap : (a -> b) -> Tree n k a o -> Tree n k b o
treeMap f (Leaf k v) = Leaf k (f v)
treeMap f (Branch2 t1 k t2) = Branch2 (treeMap f t1) k (treeMap f t2)
treeMap f (Branch3 t1 k1 t2 k2 t3)
= Branch3 (treeMap f t1) k1 (treeMap f t2) k2 (treeMap f t3)
treeTraverse : Applicative f => (a -> f b) -> Tree n k a o -> f (Tree n k b o)
treeTraverse f (Leaf k v) = Leaf k <$> f v
treeTraverse f (Branch2 t1 k t2) =
Branch2
<$> treeTraverse f t1
<*> pure k
<*> treeTraverse f t2
treeTraverse f (Branch3 t1 k1 t2 k2 t3) =
Branch3
<$> treeTraverse f t1
<*> pure k1
<*> treeTraverse f t2
<*> pure k2
<*> treeTraverse f t3
export
implementation Functor (SortedMap k) where
map _ Empty = Empty
map f (M n t) = M _ (treeMap f t)
export
implementation Foldable (SortedMap k) where
foldr f z = foldr f z . values
export
implementation Traversable (SortedMap k) where
traverse _ Empty = pure Empty
traverse f (M n t) = M _ <$> treeTraverse f t
||| Merge two maps. When encountering duplicate keys, using a function to combine the values.
||| Uses the ordering of the first map given.
export
mergeWith : (v -> v -> v) -> SortedMap k v -> SortedMap k v -> SortedMap k v
mergeWith f x y = insertFrom inserted x where
inserted : List (k, v)
inserted = do
(k, v) <- toList y
let v' = (maybe id f $ lookup k x) v
pure (k, v')
||| Merge two maps using the Semigroup (and by extension, Monoid) operation.
||| Uses mergeWith internally, so the ordering of the left map is kept.
export
merge : Semigroup v => SortedMap k v -> SortedMap k v -> SortedMap k v
merge = mergeWith (<+>)
||| Left-biased merge, also keeps the ordering specified by the left map.
export
mergeLeft : SortedMap k v -> SortedMap k v -> SortedMap k v
mergeLeft = mergeWith const
export
(Show k, Show v) => Show (SortedMap k v) where
show m = "fromList " ++ (show $ toList m)
-- TODO: is this the right variant of merge to use for this? I think it is, but
-- I could also see the advantages of using `mergeLeft`. The current approach is
-- strictly more powerful I believe, because `mergeLeft` can be emulated with
-- the `First` monoid. However, this does require more code to do the same
-- thing.
export
Semigroup v => Semigroup (SortedMap k v) where
(<+>) = merge
||| For `neutral <+> y`, y is rebuilt in `Ord k`, so this is not a "strict" Monoid.
||| However, semantically, it should be equal.
export
(Ord k, Semigroup v) => Monoid (SortedMap k v) where
neutral = empty

View File

@ -0,0 +1,67 @@
module Data.SortedSet
import Data.Maybe
import Data.SortedMap
export
data SortedSet k = SetWrapper (Data.SortedMap.SortedMap k ())
export
empty : Ord k => SortedSet k
empty = SetWrapper Data.SortedMap.empty
export
insert : k -> SortedSet k -> SortedSet k
insert k (SetWrapper m) = SetWrapper (Data.SortedMap.insert k () m)
export
delete : k -> SortedSet k -> SortedSet k
delete k (SetWrapper m) = SetWrapper (Data.SortedMap.delete k m)
export
contains : k -> SortedSet k -> Bool
contains k (SetWrapper m) = isJust (Data.SortedMap.lookup k m)
export
fromList : Ord k => List k -> SortedSet k
fromList l = SetWrapper (Data.SortedMap.fromList (map (\i => (i, ())) l))
export
toList : SortedSet k -> List k
toList (SetWrapper m) = map (\(i, _) => i) (Data.SortedMap.toList m)
export
Foldable SortedSet where
foldr f e xs = foldr f e (Data.SortedSet.toList xs)
||| Set union. Inserts all elements of x into y
export
union : (x, y : SortedSet k) -> SortedSet k
union x y = foldr insert x y
||| Set difference. Delete all elments in y from x
export
difference : (x, y : SortedSet k) -> SortedSet k
difference x y = foldr delete x y
||| Set symmetric difference. Uses the union of the differences.
export
symDifference : (x, y : SortedSet k) -> SortedSet k
symDifference x y = union (difference x y) (difference y x)
||| Set intersection. Implemented as the difference of the union and the symetric difference.
export
intersection : (x, y : SortedSet k) -> SortedSet k
intersection x y = difference x (difference x y)
export
Ord k => Semigroup (SortedSet k) where
(<+>) = union
export
Ord k => Monoid (SortedSet k) where
neutral = empty
export
keySet : SortedMap k v -> SortedSet k
keySet = SetWrapper . map (const ())

View File

@ -0,0 +1,102 @@
module Data.String.Extra
import Data.Strings
%default total
infixl 5 +>
infixr 5 <+
||| Adds a character to the end of the specified string.
|||
||| ```idris example
||| strSnoc "AB" 'C'
||| ```
||| ```idris example
||| strSnoc "" 'A'
||| ```
public export
strSnoc : String -> Char -> String
strSnoc s c = s ++ (singleton c)
||| Alias of `strSnoc`
|||
||| ```idris example
||| "AB" +> 'C'
||| ```
public export
(+>) : String -> Char -> String
(+>) = strSnoc
||| Alias of `strCons`
|||
||| ```idris example
||| 'A' <+ "AB"
||| ```
public export
(<+) : Char -> String -> String
(<+) = strCons
||| Take the first `n` characters from a string. Returns the whole string
||| if it's too short.
public export
take : (n : Nat) -> (input : String) -> String
take n str = substr Z n str
||| Take the last `n` characters from a string. Returns the whole string
||| if it's too short.
public export
takeLast : (n : Nat) -> (input : String) -> String
takeLast n str with (length str)
takeLast n str | len with (isLTE n len)
takeLast n str | len | Yes prf = substr (len `minus` n) len str
takeLast n str | len | No contra = str
||| Remove the first `n` characters from a string. Returns the empty string if
||| the input string is too short.
public export
drop : (n : Nat) -> (input : String) -> String
drop n str = substr n (length str) str
||| Remove the last `n` characters from a string. Returns the empty string if
||| the input string is too short.
public export
dropLast : (n : Nat) -> (input : String) -> String
dropLast n str = reverse (drop n (reverse str))
||| Remove the first and last `n` characters from a string. Returns the empty
||| string if the input string is too short.
public export
shrink : (n : Nat) -> (input : String) -> String
shrink n str = dropLast n (drop n str)
||| Concatenate the strings from a `Foldable` containing strings, separated by
||| the given string.
public export
join : (sep : String) -> Foldable t => (xs : t String) -> String
join sep xs = drop (length sep)
(foldl (\acc, x => acc ++ sep ++ x) "" xs)
||| Get a character from a string if the string is long enough.
public export
index : (n : Nat) -> (input : String) -> Maybe Char
index n str with (unpack str)
index n str | [] = Nothing
index Z str | (x :: xs) = Just x
index (S n) str | (x :: xs) = index n str | xs
||| Produce a string by repeating the character `c` `n` times.
public export
replicate : (n : Nat) -> (c : Char) -> String
replicate n c = pack $ replicate n c
||| Indent a given string by `n` spaces.
public export
indent : (n : Nat) -> String -> String
indent n x = replicate n ' ' ++ x
||| Indent each line of a given string by `n` spaces.
public export
indentLines : (n : Nat) -> String -> String
indentLines n str = unlines $ map (indent n) $ lines str

View File

@ -0,0 +1,14 @@
||| The JSON language, as described at https://json.org/
module Language.JSON
import Language.JSON.Lexer
import Language.JSON.Parser
import public Language.JSON.Data
%default total
||| Parse a JSON string.
export
parse : String -> Maybe JSON
parse x = parseJSON !(lexJSON x)

View File

@ -0,0 +1,126 @@
module Language.JSON.Data
import Data.String.Extra
import Data.Strings
import Data.List
%default total
public export
data JSON
= JNull
| JBoolean Bool
| JNumber Double
| JString String
| JArray (List JSON)
| JObject (List (String, JSON))
%name JSON json
private
intToHexString : Int -> String
intToHexString n =
case n of
0 => "0"
1 => "1"
2 => "2"
3 => "3"
4 => "4"
5 => "5"
6 => "6"
7 => "7"
8 => "8"
9 => "9"
10 => "A"
11 => "B"
12 => "C"
13 => "D"
14 => "E"
15 => "F"
other => intToHexString (shiftR n 4) ++ intToHexString (mod n 16)
private
showChar : Char -> String
showChar c
= case c of
'\b' => "\\b"
'\f' => "\\f"
'\n' => "\\n"
'\r' => "\\r"
'\t' => "\\t"
'\\' => "\\\\"
'"' => "\\\""
c => if isControl c || c >= '\127'
-- then "\\u" ++ b16ToHexString (fromInteger (cast (ord c)))
then "\\u" ++ intToHexString (ord c)-- quick hack until b16ToHexString is available in idris2
else singleton c
private
showString : String -> String
showString x = "\"" ++ concatMap showChar (unpack x) ++ "\""
||| Convert a JSON value into its string representation.
||| No whitespace is added.
private
stringify : JSON -> String
stringify JNull = "null"
stringify (JBoolean x) = if x then "true" else "false"
stringify (JNumber x) = show x
stringify (JString x) = showString x
stringify (JArray xs) = "[" ++ stringifyValues xs ++ "]"
where
stringifyValues : List JSON -> String
stringifyValues [] = ""
stringifyValues (x :: xs) = stringify x
++ if isNil xs
then ""
else "," ++ stringifyValues xs
stringify (JObject xs) = "{" ++ stringifyProps xs ++ "}"
where
stringifyProp : (String, JSON) -> String
stringifyProp (key, value) = showString key ++ ":" ++ stringify value
stringifyProps : List (String, JSON) -> String
stringifyProps [] = ""
stringifyProps (x :: xs) = stringifyProp x
++ if isNil xs
then ""
else "," ++ stringifyProps xs
export
Show JSON where
show = stringify
||| Format a JSON value, indenting by `n` spaces per nesting level.
|||
||| @curr The current indentation amount, measured in spaces.
||| @n The amount of spaces to indent per nesting level.
export
format : {default 0 curr : Nat} -> (n : Nat) -> JSON -> String
format {curr} n json = indent curr $ formatValue curr n json
where
formatValue : (curr, n : Nat) -> JSON -> String
formatValue _ _ (JArray []) = "[]"
formatValue curr n (JArray xs@(_ :: _)) = "[\n" ++ formatValues xs
++ indent curr "]"
where
formatValues : (xs : List JSON) -> {auto ok : NonEmpty xs} -> String
formatValues (x :: xs) = format {curr=(curr + n)} n x
++ case xs of
_ :: _ => ",\n" ++ formatValues xs
[] => "\n"
formatValue _ _ (JObject []) = "{}"
formatValue curr n (JObject xs@(_ :: _)) = "{\n" ++ formatProps xs
++ indent curr "}"
where
formatProp : (String, JSON) -> String
formatProp (key, value) = indent (curr + n) (showString key ++ ": ")
++ formatValue (curr + n) n value
formatProps : (xs : List (String, JSON)) -> {auto ok : NonEmpty xs} -> String
formatProps (x :: xs) = formatProp x
++ case xs of
_ :: _ => ",\n" ++ formatProps xs
[] => "\n"
formatValue _ _ x = stringify x

View File

@ -0,0 +1,41 @@
module Language.JSON.Lexer
import Language.JSON.String
import Text.Lexer
import Text.Token
import public Language.JSON.Tokens
%default total
private
numberLit : Lexer
numberLit
= let sign = is '-'
whole = is '0' <|> range '1' '9' <+> many digit
frac = is '.' <+> digits
exp = like 'e' <+> opt (oneOf "+-") <+> digits in
opt sign <+> whole <+> opt frac <+> opt exp
private
jsonTokenMap : TokenMap JSONToken
jsonTokenMap = toTokenMap $
[ (spaces, JTIgnore)
, (is ',', JTPunct Comma)
, (is ':', JTPunct Colon)
, (is '[', JTPunct $ Square Open)
, (is ']', JTPunct $ Square Close)
, (is '{', JTPunct $ Curly Open)
, (is '}', JTPunct $ Curly Close)
, (exact "null", JTNull)
, (exact strTrue <|> exact strFalse, JTBoolean)
, (numberLit, JTNumber)
, (permissiveStringLit, JTString)
]
export
lexJSON : String -> Maybe (List JSONToken)
lexJSON str
= case lex jsonTokenMap str of
(tokens, _, _, "") => Just $ map TokenData.tok tokens
_ => Nothing

View File

@ -0,0 +1,80 @@
module Language.JSON.Parser
import Language.JSON.Data
import Text.Parser
import Text.Token
import Data.List
import public Language.JSON.Tokens
%default total
private
punct : Punctuation -> Grammar JSONToken True ()
punct p = match $ JTPunct p
private
rawString : Grammar JSONToken True String
rawString = do mstr <- match JTString
the (Grammar _ False _) $
case mstr of
Just str => pure str
Nothing => fail "invalid string"
mutual
private
json : Grammar JSONToken True JSON
json = object
<|> array
<|> string
<|> boolean
<|> number
<|> null
private
object : Grammar JSONToken True JSON
object = do punct $ Curly Open
commit
props <- properties
punct $ Curly Close
pure $ JObject props
where
properties : Grammar JSONToken False (List (String, JSON))
properties = sepBy (punct Comma) $
do key <- rawString
punct Colon
value <- json
pure (key, value)
private
array : Grammar JSONToken True JSON
array = do punct (Square Open)
commit
vals <- values
punct (Square Close)
pure (JArray vals)
where
values : Grammar JSONToken False (List JSON)
values = sepBy (punct Comma) json
private
string : Grammar JSONToken True JSON
string = map JString rawString
private
boolean : Grammar JSONToken True JSON
boolean = map JBoolean $ match JTBoolean
private
number : Grammar JSONToken True JSON
number = map JNumber $ match JTNumber
private
null : Grammar JSONToken True JSON
null = map (const JNull) $ match JTNull
export
parseJSON : List JSONToken -> Maybe JSON
parseJSON toks = case parse json $ filter (not . ignored) toks of
Right (j, []) => Just j
_ => Nothing

View File

@ -0,0 +1,17 @@
module Language.JSON.String
import Language.JSON.String.Lexer
import Language.JSON.String.Parser
import Language.JSON.String.Tokens
import Text.Lexer
%default total
export
permissiveStringLit : Lexer
permissiveStringLit
= quo <+> manyUntil quo (esc any <|> any) <+> opt quo
export
stringValue : String -> Maybe String
stringValue x = parseString !(lexString x)

View File

@ -0,0 +1,41 @@
module Language.JSON.String.Lexer
import Language.JSON.String.Tokens
import Text.Lexer
%default total
export
quo : Lexer
quo = is '"'
export
esc : Lexer -> Lexer
esc = escape '\\'
private
unicodeEscape : Lexer
unicodeEscape = esc $ is 'u' <+> count (exactly 4) hexDigit
private
simpleEscape : Lexer
simpleEscape = esc $ oneOf "\"\\/bfnrt"
private
legalChar : Lexer
legalChar = non (quo <|> is '\\' <|> control)
private
jsonStringTokenMap : TokenMap JSONStringToken
jsonStringTokenMap = toTokenMap $
[ (quo, JSTQuote)
, (unicodeEscape, JSTUnicodeEscape)
, (simpleEscape, JSTSimpleEscape)
, (legalChar, JSTChar)
]
export
lexString : String -> Maybe (List JSONStringToken)
lexString x = case lex jsonStringTokenMap x of
(toks, _, _, "") => Just $ map TokenData.tok toks
_ => Nothing

View File

@ -0,0 +1,26 @@
module Language.JSON.String.Parser
import Language.JSON.String.Tokens
import Text.Lexer
import Text.Parser
%default total
private
stringChar : Grammar JSONStringToken True Char
stringChar = match JSTChar
<|> match JSTSimpleEscape
<|> match JSTUnicodeEscape
private
quotedString : Grammar JSONStringToken True String
quotedString = let q = match JSTQuote in
do chars <- between q q (many stringChar)
eof
pure $ pack chars
export
parseString : List JSONStringToken -> Maybe String
parseString toks = case parse quotedString toks of
Right (str, []) => Just str
_ => Nothing

View File

@ -0,0 +1,63 @@
module Language.JSON.String.Tokens
import Data.String.Extra
import Text.Token
%default total
public export
data JSONStringTokenKind
= JSTQuote
| JSTChar
| JSTSimpleEscape
| JSTUnicodeEscape
public export
JSONStringToken : Type
JSONStringToken = Token JSONStringTokenKind
public export
Eq JSONStringTokenKind where
(==) JSTQuote JSTQuote = True
(==) JSTChar JSTChar = True
(==) JSTSimpleEscape JSTSimpleEscape = True
(==) JSTUnicodeEscape JSTUnicodeEscape = True
(==) _ _ = False
private
charValue : String -> Char
charValue x = case index 0 x of
Nothing => '\NUL'
Just c => c
private
simpleEscapeValue : String -> Char
simpleEscapeValue x
= case index 1 x of
Nothing => '\NUL'
Just c => case c of
'"' => '"'
'\\' => '\\'
'/' => '/'
'b' => '\b'
'f' => '\f'
'n' => '\n'
'r' => '\r'
't' => '\t'
_ => '\NUL'
private
unicodeEscapeValue : String -> Char
unicodeEscapeValue x = chr $ cast ("0x" ++ drop 2 x)
public export
TokenKind JSONStringTokenKind where
TokType JSTQuote = ()
TokType JSTChar = Char
TokType JSTSimpleEscape = Char
TokType JSTUnicodeEscape = Char
tokValue JSTQuote = const ()
tokValue JSTChar = charValue
tokValue JSTSimpleEscape = simpleEscapeValue
tokValue JSTUnicodeEscape = unicodeEscapeValue

View File

@ -0,0 +1,81 @@
module Language.JSON.Tokens
import Language.JSON.String
import Text.Token
%default total
public export
strTrue : String
strTrue = "true"
public export
strFalse : String
strFalse = "false"
public export
data Bracket = Open | Close
public export
Eq Bracket where
(==) Open Open = True
(==) Close Close = True
(==) _ _ = False
public export
data Punctuation
= Comma
| Colon
| Square Bracket
| Curly Bracket
public export
Eq Punctuation where
(==) Comma Comma = True
(==) Colon Colon = True
(==) (Square b1) (Square b2) = b1 == b2
(==) (Curly b1) (Curly b2) = b1 == b2
(==) _ _ = False
public export
data JSONTokenKind
= JTBoolean
| JTNumber
| JTString
| JTNull
| JTPunct Punctuation
| JTIgnore
public export
JSONToken : Type
JSONToken = Token JSONTokenKind
public export
Eq JSONTokenKind where
(==) JTBoolean JTBoolean = True
(==) JTNumber JTNumber = True
(==) JTString JTString = True
(==) JTNull JTNull = True
(==) (JTPunct p1) (JTPunct p2) = p1 == p2
(==) _ _ = False
public export
TokenKind JSONTokenKind where
TokType JTBoolean = Bool
TokType JTNumber = Double
TokType JTString = Maybe String
TokType JTNull = ()
TokType (JTPunct _) = ()
TokType JTIgnore = ()
tokValue JTBoolean x = x == strTrue
tokValue JTNumber x = cast x
tokValue JTString x = stringValue x
tokValue JTNull _ = ()
tokValue (JTPunct _) _ = ()
tokValue JTIgnore _ = ()
export
ignored : JSONToken -> Bool
ignored (Tok JTIgnore _) = True
ignored _ = False

8
libs/contrib/Makefile Normal file
View File

@ -0,0 +1,8 @@
all:
${IDRIS2} --build contrib.ipkg
install:
${IDRIS2} --install contrib.ipkg
clean:
rm -rf build

View File

@ -0,0 +1,42 @@
||| Until Idris2 starts supporting the 'syntax' keyword, here's a
||| poor-man's equational reasoning
module Syntax.PreorderReasoning
||| Deep embedding of equation derivation sequences.
||| Using the Nil, (::) constructors lets us use list syntax.
public export
data Derivation : (x : a) -> (y : b) -> Type where
Nil : Derivation x x
(::) : (x = y) -> Derivation y z -> Derivation x z
infix 1 ==|
||| Explicate the term under consideration and the justification for
||| the next step.
export
(==|) : (x : a) -> (x = y) -> x = y
(==|) x pf = pf
||| Finishg the derivation.
||| A bit klunky, but this /is/ a poor-man's version.
export
QED : {x : a} -> x = x
QED = Refl
export
Calculate : Derivation x y -> x = y
Calculate [] = Refl
Calculate (Refl :: der) = Calculate der
{--
||| Requires Data.Nata
example : (x : Nat) -> (x + 1) + 0 = 1 + x
example x = Calculate [
(x + 1) + 0
==| plusZeroRightNeutral (x + 1) ,
x + 1
==| plusCommutative x 1 ,
1 + x
==| QED
]
--}

View File

@ -0,0 +1,8 @@
module Syntax.WithProof
||| Until Idris2 supports the 'with (...) proof p' construct, here's a
||| poor-man's replacement.
prefix 10 @@
public export
(@@) : (t : a ) -> (u : a ** t = u)
(@@) x = ( x ** Refl)

View File

@ -0,0 +1,150 @@
module System.Clock
import PrimIO
||| The various types of system clock available.
public export
data ClockType
= UTC -- The time elapsed since the "epoch:" 00:00:00 UTC, January 1, 1970
| Monotonic -- The time elapsed since some arbitrary point in the past
| Duration -- Representing a time duration.
| Process -- The amount of CPU time used by the current process.
| Thread -- The amount of CPU time used by the current thread.
| GCCPU -- The current process's CPU time consumed by the garbage collector.
| GCReal -- The current process's real time consumed by the garbage collector.
export
Show ClockType where
show UTC = "UTC"
show Monotonic = "monotonic"
show Duration = "duration"
show Process = "process"
show Thread = "thread"
show GCCPU = "gcCPU"
show GCReal = "gcReal"
public export
data Clock : (type : ClockType) -> Type where
MkClock
: {type : ClockType}
-> (seconds : Integer)
-> (nanoseconds : Integer)
-> Clock type
public export
Eq (Clock type) where
(MkClock seconds1 nanoseconds1) == (MkClock seconds2 nanoseconds2) =
seconds1 == seconds2 && nanoseconds1 == nanoseconds2
public export
Ord (Clock type) where
compare (MkClock seconds1 nanoseconds1) (MkClock seconds2 nanoseconds2) =
case compare seconds1 seconds2 of
LT => LT
GT => GT
EQ => compare nanoseconds1 nanoseconds2
public export
Show (Clock type) where
show (MkClock {type} seconds nanoseconds) =
show type ++ ": " ++ show seconds ++ "s " ++ show nanoseconds ++ "ns"
||| A helper to deconstruct a Clock.
public export
seconds : Clock type -> Integer
seconds (MkClock s _) = s
||| A helper to deconstruct a Clock.
public export
nanoseconds : Clock type -> Integer
nanoseconds (MkClock _ ns) = ns
||| Make a duration value.
public export
makeDuration : Integer -> Integer -> Clock Duration
makeDuration = MkClock
||| Opaque clock value manipulated by the back-end.
data OSClock : Type where [external]
||| Note: Back-ends are required to implement UTC, monotonic time, CPU time
||| in current process/thread, and time duration; the rest are optional.
data ClockTypeMandatory
= Mandatory
| Optional
public export
isClockMandatory : ClockType -> ClockTypeMandatory
isClockMandatory GCCPU = Optional
isClockMandatory GCReal = Optional
isClockMandatory _ = Mandatory
prim_clockTimeMonotonic : IO OSClock
prim_clockTimeMonotonic = schemeCall OSClock "blodwen-clock-time-monotonic" []
fetchOSClock : ClockType -> IO OSClock
fetchOSClock UTC = schemeCall OSClock "blodwen-clock-time-utc" []
fetchOSClock Monotonic = prim_clockTimeMonotonic
fetchOSClock Process = schemeCall OSClock "blodwen-clock-time-process" []
fetchOSClock Thread = schemeCall OSClock "blodwen-clock-time-thread" []
fetchOSClock GCCPU = schemeCall OSClock "blodwen-clock-time-gccpu" []
fetchOSClock GCReal = schemeCall OSClock "blodwen-clock-time-gcreal" []
fetchOSClock Duration = prim_clockTimeMonotonic
||| A test to determine the status of optional clocks.
osClockValid : OSClock -> IO Int
osClockValid clk = schemeCall Int "blodwen-is-time?" [clk]
fromOSClock : {type : ClockType} -> OSClock -> IO (Clock type)
fromOSClock clk =
pure $
MkClock
{type}
!(schemeCall Integer "blodwen-clock-second" [clk])
!(schemeCall Integer "blodwen-clock-nanosecond" [clk])
public export
clockTimeReturnType : (typ : ClockType) -> Type
clockTimeReturnType typ with (isClockMandatory typ)
clockTimeReturnType typ | Optional = Maybe (Clock typ)
clockTimeReturnType typ | Mandatory = Clock typ
||| Fetch the system clock of a given kind. If the clock is mandatory,
||| we return a (Clock type) else (Maybe (Clock type)).
public export
clockTime : (typ : ClockType) -> IO (clockTimeReturnType typ)
clockTime clockType with (isClockMandatory clockType)
clockTime clockType | Mandatory = fetchOSClock clockType >>= fromOSClock
clockTime clockType | Optional = do
clk <- fetchOSClock clockType
valid <- map (== 1) $ osClockValid clk
if valid
then map Just $ fromOSClock clk
else pure Nothing
toNano : Clock type -> Integer
toNano (MkClock seconds nanoseconds) =
let scale = 1000000000
in scale * seconds + nanoseconds
fromNano : {type : ClockType} -> Integer -> Clock type
fromNano n =
let scale = 1000000000
seconds = n `div` scale
nanoseconds = n `mod` scale
in MkClock seconds nanoseconds
||| Compute difference between two clocks of the same type.
public export
timeDifference : Clock type -> Clock type -> Clock Duration
timeDifference clock duration = fromNano $ toNano clock - toNano duration
||| Add a duration to a clock value.
public export
addDuration : {type : ClockType} -> Clock type -> Clock Duration -> Clock type
addDuration clock duration = fromNano $ toNano clock + toNano duration
||| Subtract a duration from a clock value.
public export
subtractDuration : {type : ClockType} -> Clock type -> Clock Duration -> Clock type
subtractDuration clock duration = fromNano $ toNano clock - toNano duration

View File

@ -0,0 +1,68 @@
module System.Random
import Data.Fin
import Data.Vect
import Data.List
public export
interface Random a where
randomIO : IO a
-- Takes a range (lo, hi), and returns a random value uniformly
-- distributed in the closed interval [lo, hi]. It is unspecified what
-- happens if lo > hi.
randomRIO : (a, a) -> IO a
prim_randomInt : Int -> IO Int
prim_randomInt upperBound = schemeCall Int "blodwen-random" [upperBound]
public export
Random Int where
-- Generate a random value within [-2^31, 2^31-1].
randomIO =
let maxInt = shiftL 1 31 - 1
minInt = negate $ shiftL 1 31
range = maxInt - minInt
in map (+ minInt) $ prim_randomInt range
-- Generate a random value within [lo, hi].
randomRIO (lo, hi) =
let range = hi - lo + 1
in map (+ lo) $ prim_randomInt range
prim_randomDouble : IO Double
prim_randomDouble = schemeCall Double "blodwen-random" []
public export
Random Double where
-- Generate a random value within [0, 1].
randomIO = prim_randomDouble
-- Generate a random value within [lo, hi].
randomRIO (lo, hi) = map ((+ lo) . (* (hi - lo))) prim_randomDouble
||| Sets the random seed
export
srand : Integer -> IO ()
srand n = schemeCall () "blodwen-random-seed" [n]
||| Generate a random number in Fin (S `k`)
|||
||| Note that rndFin k takes values 0, 1, ..., k.
public export
rndFin : (n : Nat) -> IO (Fin (S n))
rndFin 0 = pure FZ
rndFin (S k) = do
let intBound = the Int (cast (S k))
randomInt <- randomRIO (0, intBound)
pure $ restrict (S k) (cast randomInt)
||| Select a random element from a vector
public export
rndSelect' : {k : Nat} -> Vect (S k) a -> IO a
rndSelect' {k} xs = pure $ Vect.index !(rndFin k) xs
||| Select a random element from a non-empty list
public export
rndSelect : (elems : List a) -> {auto prf : NonEmpty elems} -> IO a
rndSelect (x :: xs) {prf = IsNonEmpty} = rndSelect' $ fromList (x :: xs)

375
libs/contrib/Text/Lexer.idr Normal file
View File

@ -0,0 +1,375 @@
module Text.Lexer
import Data.Bool.Extra
import Data.List
import Data.Nat
import public Text.Lexer.Core
import public Text.Quantity
import public Text.Token
export
toTokenMap : List (Lexer, k) -> TokenMap (Token k)
toTokenMap = map $ \(l, kind) => (l, Tok kind)
||| Recognise any character.
||| /./
export
any : Lexer
any = pred (const True)
||| Recognise a lexer or recognise no input. This is not guaranteed
||| to consume input.
||| /`l`?/
export
opt : (l : Lexer) -> Recognise False
opt l = l <|> empty
||| Recognise any character if the sub-lexer `l` fails.
||| /(?!`l`)./
export
non : (l : Lexer) -> Lexer
non l = reject l <+> any
||| Produce recognisers by applying a function to elements of a container, and
||| recognise the first match. Consumes input if the function produces consuming
||| recognisers. Fails if the container is empty.
export
choiceMap : {c : Bool} ->
Foldable t => (a -> Recognise c) -> t a -> Recognise c
choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in
f x <|> acc)
fail xs
||| Recognise the first matching recogniser in a container. Consumes input if
||| recognisers in the list consume. Fails if the container is empty.
export
choice : {c : _} ->
Foldable t => t (Recognise c) -> Recognise c
choice = Lexer.choiceMap id
||| Sequence a list of recognisers. Guaranteed to consume input if the list is
||| non-empty and the recognisers consume.
export
concat : {c : _} ->
(xs : List (Recognise c)) -> Recognise (c && isCons xs)
concat = Lexer.Core.concatMap id
||| Recognise a specific character.
||| /[`x`]/
export
is : (x : Char) -> Lexer
is x = pred (==x)
||| Recognise anything but the given character.
||| /[\^`x`]/
export
isNot : (x : Char) -> Lexer
isNot x = pred (/=x)
||| Recognise a specific character (case-insensitive).
||| /[`x`]/i
export
like : (x : Char) -> Lexer
like x = pred (\y => toUpper x == toUpper y)
||| Recognise anything but the given character (case-insensitive).
||| /[\^`x`]/i
export
notLike : (x : Char) -> Lexer
notLike x = pred (\y => toUpper x /= toUpper y)
||| Recognise a specific string.
||| Fails if the string is empty.
||| /`str`/
export
exact : (str : String) -> Lexer
exact str = case unpack str of
[] => fail
(x :: xs) => concatMap is (x :: xs)
||| Recognise a specific string (case-insensitive).
||| Fails if the string is empty.
||| /`str`/i
export
approx : (str : String) -> Lexer
approx str = case unpack str of
[] => fail
(x :: xs) => concatMap like (x :: xs)
||| Recognise any of the characters in the given string.
||| /[`chars`]/
export
oneOf : (chars : String) -> Lexer
oneOf chars = pred (\x => x `elem` unpack chars)
||| Recognise a character range. Also works in reverse!
||| /[`start`-`end`]/
export
range : (start : Char) -> (end : Char) -> Lexer
range start end = pred (\x => (x >= min start end)
&& (x <= max start end))
mutual
||| Recognise a sequence of at least one sub-lexers
||| /`l`+/
export
some : Lexer -> Lexer
some l = l <+> many l
||| Recognise a sequence of at zero or more sub-lexers. This is not
||| guaranteed to consume input
||| /`l`\*/
export
many : Lexer -> Recognise False
many l = opt (some l)
||| Repeat the sub-lexer `l` zero or more times until the lexer
||| `stopBefore` is encountered. `stopBefore` will not be consumed.
||| Not guaranteed to consume input.
||| /((?!`stopBefore`)`l`)\*/
export
manyUntil : (stopBefore : Recognise c) -> (l : Lexer) -> Recognise False
manyUntil stopBefore l = many (reject stopBefore <+> l)
||| Repeat the sub-lexer `l` zero or more times until the lexer
||| `stopAfter` is encountered, and consume it. Guaranteed to
||| consume if `stopAfter` consumes.
||| /`l`\*?`stopAfter`/
export
manyThen : (stopAfter : Recognise c) -> (l : Lexer) -> Recognise c
manyThen stopAfter l = manyUntil stopAfter l <+> stopAfter
||| Recognise many instances of `l` until an instance of `end` is
||| encountered.
|||
||| Useful for defining comments.
export
manyTill : (l : Lexer) -> (end : Lexer) -> Recognise False
manyTill l end = end <|> opt (l <+> manyTill l end)
||| Recognise a sub-lexer repeated as specified by `q`. Fails if `q` has
||| `min` and `max` in the wrong order. Consumes input unless `min q` is zero.
||| /`l`{`q`}/
export
count : (q : Quantity) -> (l : Lexer) -> Recognise (isSucc (min q))
count (Qty Z Nothing) l = many l
count (Qty Z (Just Z)) _ = empty
count (Qty Z (Just (S max))) l = opt $ l <+> count (atMost max) l
count (Qty (S min) Nothing) l = l <+> count (atLeast min) l
count (Qty (S min) (Just Z)) _ = fail
count (Qty (S min) (Just (S max))) l = l <+> count (between min max) l
||| Recognise a single digit 0-9
||| /[0-9]/
export
digit : Lexer
digit = pred isDigit
||| Recognise one or more digits
||| /[0-9]+/
export
digits : Lexer
digits = some digit
||| Recognise a single hexidecimal digit
||| /[0-9A-Fa-f]/
export
hexDigit : Lexer
hexDigit = pred isHexDigit
||| Recognise one or more hexidecimal digits
||| /[0-9A-Fa-f]+/
export
hexDigits : Lexer
hexDigits = some hexDigit
||| Recognise a single octal digit
||| /[0-8]/
export
octDigit : Lexer
octDigit = pred isHexDigit
||| Recognise one or more octal digits
||| /[0-8]+/
export
octDigits : Lexer
octDigits = some hexDigit
||| Recognise a single alpha character
||| /[A-Za-z]/
export
alpha : Lexer
alpha = pred isAlpha
||| Recognise one or more alpha characters
||| /[A-Za-z]+/
export
alphas : Lexer
alphas = some alpha
||| Recognise a lowercase alpha character
||| /[a-z]/
export
lower : Lexer
lower = pred isLower
||| Recognise one or more lowercase alpha characters
||| /[a-z]+/
export
lowers : Lexer
lowers = some lower
||| Recognise an uppercase alpha character
||| /[A-Z]/
export
upper : Lexer
upper = pred isUpper
||| Recognise one or more uppercase alpha characters
||| /[A-Z]+/
export
uppers : Lexer
uppers = some upper
||| Recognise an alphanumeric character
||| /[A-Za-z0-9]/
export
alphaNum : Lexer
alphaNum = pred isAlphaNum
||| Recognise one or more alphanumeric characters
||| /[A-Za-z0-9]+/
export
alphaNums : Lexer
alphaNums = some alphaNum
||| Recognise a single whitespace character
||| /\\s/
export
space : Lexer
space = pred isSpace
||| Recognise one or more whitespace characters
||| /\\s+/
export
spaces : Lexer
spaces = some space
||| Recognise a single newline sequence. Understands CRLF, CR, and LF
||| /\\r\\n|[\\r\\n]/
export
newline : Lexer
newline = let crlf = "\r\n" in
exact crlf <|> oneOf crlf
||| Recognise one or more newline sequences. Understands CRLF, CR, and LF
||| /(\\r\\n|[\\r\\n])+)/
export
newlines : Lexer
newlines = some newline
||| Recognise a single non-whitespace, non-alphanumeric character
||| /[\^\\sA-Za-z0-9]/
export
symbol : Lexer
symbol = pred (\x => not (isSpace x || isAlphaNum x))
||| Recognise one or more non-whitespace, non-alphanumeric characters
||| /[\^\\sA-Za-z0-9]+/
export
symbols : Lexer
symbols = some symbol
||| Recognise a single control character
||| /[\\x00-\\x1f\\x7f-\\x9f]/
export
control : Lexer
control = pred isControl
||| Recognise one or more control characters
||| /[\\x00-\\x1f\\x7f-\\x9f]+/
export
controls : Lexer
controls = some control
||| Recognise zero or more occurrences of a sub-lexer between
||| delimiting lexers
||| /`start`(`l`)\*?`end`/
export
surround : (start : Lexer) -> (end : Lexer) -> (l : Lexer) -> Lexer
surround start end l = start <+> manyThen end l
||| Recognise zero or more occurrences of a sub-lexer surrounded
||| by the same quote lexer on both sides (useful for strings)
||| /`q`(`l`)\*?`q`/
export
quote : (q : Lexer) -> (l : Lexer) -> Lexer
quote q l = surround q q l
||| Recognise an escape character (often '\\') followed by a sub-lexer
||| /[`esc`]`l`/
export
escape : (esc : Char) -> Lexer -> Lexer
escape esc l = is esc <+> l
||| Recognise a string literal, including escaped characters.
||| (Note: doesn't yet handle escape sequences such as \123)
||| /"(\\\\.|.)\*?"/
export
stringLit : Lexer
stringLit = quote (is '"') (escape '\\' any <|> any)
||| Recognise a character literal, including escaped characters.
||| (Note: doesn't yet handle escape sequences such as \123)
||| /'(\\\\.|[\^'])'/
export
charLit : Lexer
charLit = let q = '\'' in
is q <+> (escape '\\' (control <|> any) <|> isNot q) <+> is q
where
lexStr : List String -> Lexer
lexStr [] = fail
lexStr (t :: ts) = exact t <|> lexStr ts
control : Lexer
control = lexStr ["NUL", "SOH", "STX", "ETX", "EOT", "ENQ", "ACK", "BEL",
"BS", "HT", "LF", "VT", "FF", "CR", "SO", "SI",
"DLE", "DC1", "DC2", "DC3", "DC4", "NAK", "SYN", "ETB",
"CAN", "EM", "SUB", "ESC", "FS", "GS", "RS", "US",
"SP", "DEL"]
<|> (is 'x' <+> hexDigits)
<|> (is 'o' <+> octDigits)
<|> digits
||| Recognise an integer literal (possibly with a '-' prefix)
||| /-?[0-9]+/
export
intLit : Lexer
intLit = opt (is '-') <+> digits
||| Recognise a hexidecimal literal, prefixed by "0x" or "0X"
||| /0[Xx][0-9A-Fa-f]+/
export
hexLit : Lexer
hexLit = approx "0x" <+> hexDigits
||| Recognise `start`, then recognise all input until a newline is encountered,
||| and consume the newline. Will succeed if end-of-input is encountered before
||| a newline.
||| /`start`[\^\\r\\n]+(\\r\\n|[\\r\\n])?/
export
lineComment : (start : Lexer) -> Lexer
lineComment start = start <+> manyUntil newline any <+> opt newline
||| Recognise all input between `start` and `end` lexers.
||| Supports balanced nesting.
|||
||| For block comments that don't support nesting (such as C-style comments),
||| use `surround`
export
blockComment : (start : Lexer) -> (end : Lexer) -> Lexer
blockComment start end = start <+> middle <+> end
where
middle : Recognise False
middle = manyUntil end (blockComment start end <|> any)

View File

@ -0,0 +1,212 @@
module Text.Lexer.Core
import public Control.Delayed
import Data.Bool.Extra
import Data.List
import Data.Nat
import Data.Strings
||| A language of token recognisers.
||| @ consumes If `True`, this recogniser is guaranteed to consume at
||| least one character of input when it succeeds.
public export
data Recognise : (consumes : Bool) -> Type where
Empty : Recognise False
Fail : Recognise c
Lookahead : (positive : Bool) -> Recognise c -> Recognise False
Pred : (Char -> Bool) -> Recognise True
SeqEat : Recognise True -> Inf (Recognise e) -> Recognise True
SeqEmpty : Recognise e1 -> Recognise e2 -> Recognise (e1 || e2)
Alt : Recognise e1 -> Recognise e2 -> Recognise (e1 && e2)
||| A token recogniser. Guaranteed to consume at least one character.
public export
Lexer : Type
Lexer = Recognise True
-- %allow_overloads (<+>)
||| Sequence two recognisers. If either consumes a character, the sequence
||| is guaranteed to consume a character.
export %inline
(<+>) : {c1 : Bool} ->
Recognise c1 -> inf c1 (Recognise c2) -> Recognise (c1 || c2)
(<+>) {c1 = False} = SeqEmpty
(<+>) {c1 = True} = SeqEat
%allow_overloads (<|>)
||| Alternative recognisers. If both consume, the combination is guaranteed
||| to consumer a character.
%inline
export
(<|>) : Recognise c1 -> Recognise c2 -> Recognise (c1 && c2)
(<|>) = Alt
||| A recogniser that always fails.
export
fail : Recognise c
fail = Fail
||| Recognise no input (doesn't consume any input)
%inline
export
empty : Recognise False
empty = Empty
||| Recognise a character that matches a predicate
%inline
export
pred : (Char -> Bool) -> Lexer
pred = Pred
||| Positive lookahead. Never consumes input.
export
expect : Recognise c -> Recognise False
expect = Lookahead True
||| Negative lookahead. Never consumes input.
export
reject : Recognise c -> Recognise False
reject = Lookahead False
%allow_overloads concatMap
||| Sequence the recognisers resulting from applying a function to each element
||| of a list. The resulting recogniser will consume input if the produced
||| recognisers consume and the list is non-empty.
export
concatMap : (a -> Recognise c) -> (xs : List a) ->
Recognise (c && (isCons xs))
concatMap {c} _ [] = rewrite andFalseFalse c in Empty
concatMap {c} f (x :: xs)
= rewrite andTrueNeutral c in
rewrite sym (orSameAndRightNeutral c (isCons xs)) in
SeqEmpty (f x) (Core.concatMap f xs)
data StrLen : Type where
MkStrLen : String -> Nat -> StrLen
Show StrLen where
show (MkStrLen str n) = str ++ "(" ++ show n ++ ")"
getString : StrLen -> String
getString (MkStrLen str n) = str
strIndex : StrLen -> Nat -> Maybe Char
strIndex (MkStrLen str len) i
= if cast {to = Integer} i >= cast len then Nothing
else Just (assert_total (prim__strIndex str (cast i)))
mkStr : String -> StrLen
mkStr str = MkStrLen str (length str)
export
strTail : Nat -> StrLen -> StrLen
strTail start (MkStrLen str len)
= MkStrLen (substr start len str) (minus len start)
isJust : Maybe a -> Bool
isJust Nothing = False
isJust (Just x) = True
export
unpack' : String -> List Char
unpack' str
= case strUncons str of
Nothing => []
Just (x, xs) => x :: unpack' xs
-- If the string is recognised, returns the index at which the token
-- ends
scan : Recognise c -> List Char -> List Char -> Maybe (List Char, List Char)
scan Empty tok str = pure (tok, str)
scan Fail tok str = Nothing
scan (Lookahead positive r) tok str
= if isJust (scan r tok str) == positive
then pure (tok, str)
else Nothing
scan (Pred f) tok [] = Nothing
scan (Pred f) tok (c :: str)
= if f c
then Just (c :: tok, str)
else Nothing
scan (SeqEat r1 r2) tok str
= do (tok', rest) <- scan r1 tok str
-- TODO: Can we prove totality instead by showing idx has increased?
assert_total (scan r2 tok' rest)
scan (SeqEmpty r1 r2) tok str
= do (tok', rest) <- scan r1 tok str
scan r2 tok' rest
scan (Alt r1 r2) tok str
= maybe (scan r2 tok str) Just (scan r1 tok str)
||| A mapping from lexers to the tokens they produce.
||| This is a list of pairs `(Lexer, String -> tokenType)`
||| For each Lexer in the list, if a substring in the input matches, run
||| the associated function to produce a token of type `tokenType`
public export
TokenMap : (tokenType : Type) -> Type
TokenMap tokenType = List (Lexer, String -> tokenType)
||| A token, and the line and column where it was in the input
public export
record TokenData a where
constructor MkToken
line : Int
col : Int
tok : a
export
Show a => Show (TokenData a) where
show t = show (line t) ++ ":" ++ show (col t) ++ ":" ++ show (tok t)
tokenise : (TokenData a -> Bool) ->
(line : Int) -> (col : Int) ->
List (TokenData a) -> TokenMap a ->
List Char -> (List (TokenData a), (Int, Int, List Char))
tokenise pred line col acc tmap str
= case getFirstToken tmap str of
Just (tok, line', col', rest) =>
-- assert total because getFirstToken must consume something
if pred tok
then (reverse acc, (line, col, []))
else assert_total (tokenise pred line' col' (tok :: acc) tmap rest)
Nothing => (reverse acc, (line, col, str))
where
countNLs : List Char -> Nat
countNLs str = List.length (filter (== '\n') str)
getCols : List Char -> Int -> Int
getCols x c
= case span (/= '\n') (reverse x) of
(incol, []) => c + cast (length incol)
(incol, _) => cast (length incol)
getFirstToken : TokenMap a -> List Char ->
Maybe (TokenData a, Int, Int, List Char)
getFirstToken [] str = Nothing
getFirstToken ((lex, fn) :: ts) str
= case scan lex [] str of
Just (tok, rest) => Just (MkToken line col (fn (pack (reverse tok))),
line + cast (countNLs tok),
getCols tok col, rest)
Nothing => getFirstToken ts str
||| Given a mapping from lexers to token generating functions (the
||| TokenMap a) and an input string, return a list of recognised tokens,
||| and the line, column, and remainder of the input at the first point in the
||| string where there are no recognised tokens.
export
lex : TokenMap a -> String -> (List (TokenData a), (Int, Int, String))
lex tmap str
= let (ts, (l, c, str')) = tokenise (const False) 0 0 [] tmap (unpack str) in
(ts, (l, c, pack str'))
export
lexTo : (TokenData a -> Bool) ->
TokenMap a -> String -> (List (TokenData a), (Int, Int, String))
lexTo pred tmap str
= let (ts, (l, c, str')) = tokenise pred 0 0 [] tmap (unpack str) in
(ts, (l, c, pack str'))

View File

@ -0,0 +1,260 @@
module Text.Parser
import Data.Bool.Extra
import Data.List
import Data.Nat
import public Text.Parser.Core
import public Text.Quantity
import public Text.Token
||| Parse a terminal based on a kind of token.
export
match : (Eq k, TokenKind k) =>
(kind : k) ->
Grammar (Token k) True (TokType kind)
match kind = terminal "Unrecognised input" $
\(Tok kind' text) => if kind' == kind
then Just $ tokValue kind text
else Nothing
||| Optionally parse a thing, with a default value if the grammar doesn't
||| match. May match the empty input.
export
option : {c : Bool} ->
(def : a) -> (p : Grammar tok c a) ->
Grammar tok False a
option {c = False} def p = p <|> pure def
option {c = True} def p = p <|> pure def
||| Optionally parse a thing.
||| To provide a default value, use `option` instead.
export
optional : {c : _} ->
(p : Grammar tok c a) ->
Grammar tok False (Maybe a)
optional p = option Nothing (map Just p)
||| Try to parse one thing or the other, producing a value that indicates
||| which option succeeded. If both would succeed, the left option
||| takes priority.
export
choose : {c1, c2 : _} ->
(l : Grammar tok c1 a) ->
(r : Grammar tok c2 b) ->
Grammar tok (c1 && c2) (Either a b)
choose l r = map Left l <|> map Right r
||| Produce a grammar by applying a function to each element of a container,
||| then try each resulting grammar until the first one succeeds. Fails if the
||| container is empty.
export
choiceMap : {c : Bool} ->
(a -> Grammar tok c b) ->
Foldable t => t a ->
Grammar tok c b
choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in
f x <|> acc)
(fail "No more options") xs
%hide Prelude.(>>=)
||| Try each grammar in a container until the first one succeeds.
||| Fails if the container is empty.
export
choice : {c : _} ->
Foldable t => t (Grammar tok c a) ->
Grammar tok c a
choice = Parser.choiceMap id
mutual
||| Parse one or more things
export
some : Grammar tok True a ->
Grammar tok True (List a)
some p = pure (!p :: !(many p))
||| Parse zero or more things (may match the empty input)
export
many : Grammar tok True a ->
Grammar tok False (List a)
many p = option [] (some p)
||| Parse one or more instances of `p`, returning the parsed items and proof
||| that the resulting list is non-empty.
export
some' : (p : Grammar tok True a) ->
Grammar tok True (xs : List a ** NonEmpty xs)
some' p = pure (!p :: !(many p) ** IsNonEmpty)
mutual
private
count1 : (q : Quantity) ->
(p : Grammar tok True a) ->
Grammar tok True (List a)
count1 q p = do x <- p
seq (count q p)
(\xs => pure (x :: xs))
||| Parse `p`, repeated as specified by `q`, returning the list of values.
export
count : (q : Quantity) ->
(p : Grammar tok True a) ->
Grammar tok (isSucc (min q)) (List a)
count (Qty Z Nothing) p = many p
count (Qty Z (Just Z)) _ = pure []
count (Qty Z (Just (S max))) p = option [] $ count1 (atMost max) p
count (Qty (S min) Nothing) p = count1 (atLeast min) p
count (Qty (S min) (Just Z)) _ = fail "Quantity out of order"
count (Qty (S min) (Just (S max))) p = count1 (between (S min) max) p
mutual
||| Parse one or more instances of `p` until `end` succeeds, returning the
||| list of values from `p`. Guaranteed to consume input.
export
someTill : {c : Bool} ->
(end : Grammar tok c e) ->
(p : Grammar tok True a) ->
Grammar tok True (List a)
someTill {c} end p = do x <- p
seq (manyTill end p)
(\xs => pure (x :: xs))
||| Parse zero or more instances of `p` until `end` succeeds, returning the
||| list of values from `p`. Guaranteed to consume input if `end` consumes.
export
manyTill : {c : Bool} ->
(end : Grammar tok c e) ->
(p : Grammar tok True a) ->
Grammar tok c (List a)
manyTill {c} end p = rewrite sym (andTrueNeutral c) in
map (const []) end <|> someTill end p
||| Parse one or more instances of `p` until `end` succeeds, returning the
||| list of values from `p`, along with a proof that the resulting list is
||| non-empty.
export
someTill' : {c : Bool} ->
(end : Grammar tok c e) ->
(p : Grammar tok True a) ->
Grammar tok True (xs : List a ** NonEmpty xs)
someTill' end p
= do x <- p
seq (manyTill end p)
(\xs => pure (x :: xs ** IsNonEmpty))
mutual
||| Parse one or more instance of `skip` until `p` is encountered,
||| returning its value.
export
afterSome : {c : _} ->
(skip : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok True a
afterSome skip p = do skip
afterMany skip p
||| Parse zero or more instance of `skip` until `p` is encountered,
||| returning its value.
export
afterMany : {c : Bool} ->
(skip : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok c a
afterMany {c} skip p = rewrite sym (andTrueNeutral c) in
p <|> afterSome skip p
||| Parse one or more things, each separated by another thing.
export
sepBy1 : {c : Bool} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok c (List a)
sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
[| p :: many (sep *> p) |]
||| Parse zero or more things, each separated by another thing. May
||| match the empty input.
export
sepBy : {c : _} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok False (List a)
sepBy sep p = option [] $ sepBy1 sep p
||| Parse one or more instances of `p` separated by `sep`, returning the
||| parsed items and proof that the resulting list is non-empty.
export
sepBy1' : {c : Bool} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok c (xs : List a ** NonEmpty xs)
sepBy1' {c} sep p
= rewrite sym (orFalseNeutral c) in
seq p (\x => do xs <- many (sep *> p)
pure (x :: xs ** IsNonEmpty))
||| Parse one or more instances of `p` separated by and optionally terminated by
||| `sep`.
export
sepEndBy1 : {c : Bool} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok c (List a)
sepEndBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
sepBy1 sep p <* optional sep
||| Parse zero or more instances of `p`, separated by and optionally terminated
||| by `sep`. Will not match a separator by itself.
export
sepEndBy : {c : _} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok False (List a)
sepEndBy sep p = option [] $ sepEndBy1 sep p
||| Parse zero or more instances of `p`, separated by and optionally terminated
||| by `sep`, returning the parsed items and a proof that the resulting list
||| is non-empty.
export
sepEndBy1' : {c : Bool} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok c (xs : List a ** NonEmpty xs)
sepEndBy1' {c} sep p = rewrite sym (orFalseNeutral c) in
sepBy1' sep p <* optional sep
||| Parse one or more instances of `p`, separated and terminated by `sep`.
export
endBy1 : {c : Bool} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok True (List a)
endBy1 {c} sep p = some $ rewrite sym (orTrueTrue c) in
p <* sep
export
endBy : {c : _} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok False (List a)
endBy sep p = option [] $ endBy1 sep p
||| Parse zero or more instances of `p`, separated and terminated by `sep`,
||| returning the parsed items and a proof that the resulting list is non-empty.
export
endBy1' : {c : Bool} ->
(sep : Grammar tok True s) ->
(p : Grammar tok c a) ->
Grammar tok True (xs : List a ** NonEmpty xs)
endBy1' {c} sep p = some' $ rewrite sym (orTrueTrue c) in
p <* sep
||| Parse an instance of `p` that is between `left` and `right`.
export
between : {c : _} ->
(left : Grammar tok True l) ->
(right : Grammar tok True r) ->
(p : Grammar tok c a) ->
Grammar tok True a
between left right contents = left *> contents <* right

View File

@ -0,0 +1,280 @@
module Text.Parser.Core
import public Control.Delayed
import Data.List
||| Description of a language's grammar. The `tok` parameter is the type
||| of tokens, and the `consumes` flag is True if the language is guaranteed
||| to be non-empty - that is, successfully parsing the language is guaranteed
||| to consume some input.
public export
data Grammar : (tok : Type) -> (consumes : Bool) -> Type -> Type where
Empty : (val : ty) -> Grammar tok False ty
Terminal : String -> (tok -> Maybe a) -> Grammar tok True a
NextIs : String -> (tok -> Bool) -> Grammar tok False tok
EOF : Grammar tok False ()
Fail : Bool -> String -> Grammar tok c ty
Commit : Grammar tok False ()
MustWork : Grammar tok c a -> Grammar tok c a
SeqEat : {c2 : _} -> Grammar tok True a -> Inf (a -> Grammar tok c2 b) ->
Grammar tok True b
SeqEmpty : {c1, c2 : _} -> Grammar tok c1 a -> (a -> Grammar tok c2 b) ->
Grammar tok (c1 || c2) b
Alt : {c1, c2 : _} -> Grammar tok c1 ty -> Grammar tok c2 ty ->
Grammar tok (c1 && c2) ty
||| Sequence two grammars. If either consumes some input, the sequence is
||| guaranteed to consume some input. If the first one consumes input, the
||| second is allowed to be recursive (because it means some input has been
||| consumed and therefore the input is smaller)
public export %inline
(>>=) : {c1, c2 : Bool} ->
Grammar tok c1 a ->
inf c1 (a -> Grammar tok c2 b) ->
Grammar tok (c1 || c2) b
(>>=) {c1 = False} = SeqEmpty
(>>=) {c1 = True} = SeqEat
||| Sequence two grammars. If either consumes some input, the sequence is
||| guaranteed to consume input. This is an explicitly non-infinite version
||| of `>>=`.
export
seq : {c1, c2 : _} -> Grammar tok c1 a ->
(a -> Grammar tok c2 b) ->
Grammar tok (c1 || c2) b
seq = SeqEmpty
||| Sequence a grammar followed by the grammar it returns.
export
join : {c1, c2 : Bool} ->
Grammar tok c1 (Grammar tok c2 a) ->
Grammar tok (c1 || c2) a
join {c1 = False} p = SeqEmpty p id
join {c1 = True} p = SeqEat p id
||| Give two alternative grammars. If both consume, the combination is
||| guaranteed to consume.
export
(<|>) : {c1, c2 : _} ->
Grammar tok c1 ty ->
Grammar tok c2 ty ->
Grammar tok (c1 && c2) ty
(<|>) = Alt
||| Allows the result of a grammar to be mapped to a different value.
export
{c : _} -> Functor (Grammar tok c) where
map f (Empty val) = Empty (f val)
map f (Fail fatal msg) = Fail fatal msg
map f (MustWork g) = MustWork (map f g)
map f (Terminal msg g) = Terminal msg (\t => map f (g t))
map f (Alt x y) = Alt (map f x) (map f y)
map f (SeqEat act next)
= SeqEat act (\val => map f (next val))
map f (SeqEmpty act next)
= SeqEmpty act (\val => map f (next val))
-- The remaining constructors (NextIs, EOF, Commit) have a fixed type,
-- so a sequence must be used.
map {c = False} f p = SeqEmpty p (Empty . f)
||| Sequence a grammar with value type `a -> b` and a grammar
||| with value type `a`. If both succeed, apply the function
||| from the first grammar to the value from the second grammar.
||| Guaranteed to consume if either grammar consumes.
export
(<*>) : {c1, c2 : _} ->
Grammar tok c1 (a -> b) ->
Inf (Grammar tok c2 a) ->
Grammar tok (c1 || c2) b
(<*>) {c1 = False} x y = SeqEmpty x (\f => map f y)
(<*>) {c1 = True } x y = SeqEmpty x (\f => map f (Force y))
||| Sequence two grammars. If both succeed, use the value of the first one.
||| Guaranteed to consume if either grammar consumes.
export
(<*) : {c1, c2 : _} ->
Grammar tok c1 a ->
Inf (Grammar tok c2 b) ->
Grammar tok (c1 || c2) a
(<*) x y = map const x <*> y
||| Sequence two grammars. If both succeed, use the value of the second one.
||| Guaranteed to consume if either grammar consumes.
export
(*>) : {c1, c2 : _} ->
Grammar tok c1 a ->
Inf (Grammar tok c2 b) ->
Grammar tok (c1 || c2) b
(*>) x y = map (const id) x <*> y
||| Produce a grammar that can parse a different type of token by providing a
||| function converting the new token type into the original one.
export
mapToken : (a -> b) -> Grammar b c ty -> Grammar a c ty
mapToken f (Empty val) = Empty val
mapToken f (Terminal msg g) = Terminal msg (g . f)
mapToken f (NextIs msg g) = SeqEmpty (NextIs msg (g . f)) (Empty . f)
mapToken f EOF = EOF
mapToken f (Fail fatal msg) = Fail fatal msg
mapToken f (MustWork g) = MustWork (mapToken f g)
mapToken f Commit = Commit
mapToken f (SeqEat act next) = SeqEat (mapToken f act) (\x => mapToken f (next x))
mapToken f (SeqEmpty act next) = SeqEmpty (mapToken f act) (\x => mapToken f (next x))
mapToken f (Alt x y) = Alt (mapToken f x) (mapToken f y)
||| Always succeed with the given value.
export
pure : (val : ty) -> Grammar tok False ty
pure = Empty
||| Check whether the next token satisfies a predicate
export
nextIs : String -> (tok -> Bool) -> Grammar tok False tok
nextIs = NextIs
||| Look at the next token in the input
export
peek : Grammar tok False tok
peek = nextIs "Unrecognised token" (const True)
||| Succeeds if running the predicate on the next token returns Just x,
||| returning x. Otherwise fails.
export
terminal : String -> (tok -> Maybe a) -> Grammar tok True a
terminal = Terminal
||| Always fail with a message
export
fail : String -> Grammar tok c ty
fail = Fail False
export
fatalError : String -> Grammar tok c ty
fatalError = Fail True
||| Succeed if the input is empty
export
eof : Grammar tok False ()
eof = EOF
||| Commit to an alternative; if the current branch of an alternative
||| fails to parse, no more branches will be tried
export
commit : Grammar tok False ()
commit = Commit
||| If the parser fails, treat it as a fatal error
export
mustWork : Grammar tok c ty -> Grammar tok c ty
mustWork = MustWork
data ParseResult : List tok -> (consumes : Bool) -> Type -> Type where
Failure : {xs : List tok} ->
(committed : Bool) -> (fatal : Bool) ->
(err : String) -> (rest : List tok) -> ParseResult xs c ty
EmptyRes : (committed : Bool) ->
(val : ty) -> (more : List tok) -> ParseResult more False ty
NonEmptyRes : {xs : List tok} ->
(committed : Bool) ->
(val : ty) -> (more : List tok) ->
ParseResult (x :: xs ++ more) c ty
-- Take the result of an alternative branch, reset the commit flag to
-- the commit flag from the outer alternative, and weaken the 'consumes'
-- flag to take both alternatives into account
weakenRes : {whatever, c : Bool} -> {xs : List tok} ->
(com' : Bool) ->
ParseResult xs c ty -> ParseResult xs (whatever && c) ty
weakenRes com' (Failure com fatal msg ts) = Failure com' fatal msg ts
weakenRes {whatever=True} com' (EmptyRes com val xs) = EmptyRes com' val xs
weakenRes {whatever=False} com' (EmptyRes com val xs) = EmptyRes com' val xs
weakenRes com' (NonEmptyRes {xs} com val more) = NonEmptyRes {xs} com' val more
doParse : (commit : Bool) ->
(act : Grammar tok c ty) ->
(xs : List tok) ->
ParseResult xs c ty
-- doParse com xs act with (sizeAccessible xs)
doParse com (Empty val) xs = EmptyRes com val xs
doParse com (Fail fatal str) [] = Failure com fatal str []
doParse com (Fail fatal str) (x :: xs) = Failure com fatal str (x :: xs)
doParse com Commit xs = EmptyRes True () xs
doParse com (MustWork g) xs =
let p' = doParse com g xs in
case p' of
Failure com' _ msg ts => Failure com' True msg ts
res => res
doParse com (Terminal err f) [] = Failure com False "End of input" []
doParse com (Terminal err f) (x :: xs)
= maybe
(Failure com False err (x :: xs))
(\a => NonEmptyRes com {xs=[]} a xs)
(f x)
doParse com EOF [] = EmptyRes com () []
doParse com EOF (x :: xs)
= Failure com False "Expected end of input" (x :: xs)
doParse com (NextIs err f) [] = Failure com False "End of input" []
doParse com (NextIs err f) (x :: xs)
= if f x
then EmptyRes com x (x :: xs)
else Failure com False err (x :: xs)
doParse com (Alt {c1} {c2} x y) xs
= let p' = doParse False x xs in
case p' of
Failure com' fatal msg ts
=> if com' || fatal
-- If the alternative had committed, don't try the
-- other branch (and reset commit flag)
then Failure com fatal msg ts
else weakenRes {whatever = c1} com (doParse False y xs)
-- Successfully parsed the first option, so use the outer commit flag
EmptyRes _ val xs => EmptyRes com val xs
NonEmptyRes {xs=xs'} _ val more => NonEmptyRes {xs=xs'} com val more
doParse com (SeqEmpty {c1} {c2} act next) xs
= let p' = assert_total (doParse {c = c1} com act xs) in
case p' of
Failure com fatal msg ts => Failure com fatal msg ts
EmptyRes com val xs =>
case assert_total (doParse com (next val) xs) of
Failure com' fatal msg ts => Failure com' fatal msg ts
EmptyRes com' val xs => EmptyRes com' val xs
NonEmptyRes {xs=xs'} com' val more =>
NonEmptyRes {xs=xs'} com' val more
NonEmptyRes {x} {xs=ys} com val more =>
case (assert_total (doParse com (next val) more)) of
Failure com' fatal msg ts => Failure com' fatal msg ts
EmptyRes com' val _ => NonEmptyRes {xs=ys} com' val more
NonEmptyRes {x=x1} {xs=xs1} com' val more' =>
rewrite appendAssociative (x :: ys) (x1 :: xs1) more' in
NonEmptyRes {xs = ys ++ (x1 :: xs1)} com' val more'
doParse com (SeqEat act next) xs with (doParse com act xs)
doParse com (SeqEat act next) xs | Failure com' fatal msg ts
= Failure com' fatal msg ts
doParse com (SeqEat act next) (x :: (ys ++ more)) | (NonEmptyRes {xs=ys} com' val more)
= let p' = assert_total (doParse com' (next val) more) in
case p' of
Failure com' fatal msg ts => Failure com' fatal msg ts
EmptyRes com' val _ => NonEmptyRes {xs=ys} com' val more
NonEmptyRes {x=x1} {xs=xs1} com' val more' =>
rewrite appendAssociative (x :: ys) (x1 :: xs1) more' in
NonEmptyRes {xs = ys ++ (x1 :: xs1)} com' val more'
-- This next line is not strictly necessary, but it stops the coverage
-- checker taking a really long time and eating lots of memory...
-- doParse _ _ _ = Failure True True "Help the coverage checker!" []
public export
data ParseError tok = Error String (List tok)
||| Parse a list of tokens according to the given grammar. If successful,
||| returns a pair of the parse result and the unparsed tokens (the remaining
||| input).
export
parse : {c : _} -> (act : Grammar tok c ty) -> (xs : List tok) ->
Either (ParseError tok) (ty, List tok)
parse act xs
= case doParse False act xs of
Failure _ _ msg ts => Left (Error msg ts)
EmptyRes _ val rest => pure (val, rest)
NonEmptyRes _ val rest => pure (val, rest)

View File

@ -0,0 +1,42 @@
module Text.Quantity
public export
record Quantity where
constructor Qty
min : Nat
max : Maybe Nat
public export
Show Quantity where
show (Qty Z Nothing) = "*"
show (Qty Z (Just (S Z))) = "?"
show (Qty (S Z) Nothing) = "+"
show (Qty min max) = "{" ++ show min ++ showMax ++ "}"
where
showMax : String
showMax = case max of
Nothing => ","
Just max' => if min == max'
then ""
else "," ++ show max'
public export
between : Nat -> Nat -> Quantity
between min max = Qty min (Just max)
public export
atLeast : Nat -> Quantity
atLeast min = Qty min Nothing
public export
atMost : Nat -> Quantity
atMost max = Qty 0 (Just max)
public export
exactly : Nat -> Quantity
exactly n = Qty n (Just n)
public export
inOrder : Quantity -> Bool
inOrder (Qty min Nothing) = True
inOrder (Qty min (Just max)) = min <= max

View File

@ -0,0 +1,38 @@
module Text.Token
||| For a type `kind`, specify a way of converting the recognised
||| string into a value.
|||
||| ```idris example
||| data SimpleKind = SKString | SKInt | SKComma
|||
||| TokenKind SimpleKind where
||| TokType SKString = String
||| TokType SKInt = Int
||| TokType SKComma = ()
|||
||| tokValue SKString x = x
||| tokValue SKInt x = cast x
||| tokValue SKComma x = ()
||| ```
public export
interface TokenKind (k : Type) where
||| The type that a token of this kind converts to.
TokType : k -> Type
||| Convert a recognised string into a value. The type is determined
||| by the kind of token.
tokValue : (kind : k) -> String -> TokType kind
||| A token of a particular kind and the text that was recognised.
public export
record Token k where
constructor Tok
kind : k
text : String
||| Get the value of a `Token k`. The resulting type depends upon
||| the kind of token.
public export
value : TokenKind k => (t : Token k) -> TokType (kind t)
value (Tok k x) = tokValue k x

31
libs/contrib/contrib.ipkg Normal file
View File

@ -0,0 +1,31 @@
package contrib
modules = Control.Delayed,
Data.List.TailRec,
Data.Bool.Extra,
Data.SortedMap,
Data.SortedSet,
Data.String.Extra,
Language.JSON,
Language.JSON.Data,
Language.JSON.Lexer,
Language.JSON.Parser,
Language.JSON.String,
Language.JSON.String.Lexer,
Language.JSON.String.Parser,
Language.JSON.String.Tokens,
Language.JSON.Tokens,
Text.Token,
Text.Quantity,
Text.Parser,
Text.Lexer,
Text.Parser.Core,
Text.Lexer.Core,
System.Clock,
System.Random,
Syntax.WithProof,
Syntax.PreorderReasoning

8
libs/network/.gitignore vendored Normal file
View File

@ -0,0 +1,8 @@
/network-tests
*.d
*.o
*.obj
*.so
*.dylib
*.dll

58
libs/network/Echo.idr Normal file
View File

@ -0,0 +1,58 @@
module Main
import System
import Network.Socket
import Network.Socket.Data
import Network.Socket.Raw
%cg chez libidris_net
runServer : IO (Either String (Port, ThreadID))
runServer = do
Right sock <- socket AF_INET Stream 0
| Left fail => pure (Left $ "Failed to open socket: " ++ show fail)
res <- bind sock (Just (Hostname "localhost")) 0
if res /= 0
then pure (Left $ "Failed to bind socket with error: " ++ show res)
else do
port <- getSockPort sock
res <- listen sock
if res /= 0
then pure (Left $ "Failed to listen on socket with error: " ++ show res)
else do
forked <- fork (serve port sock)
pure $ Right (port, forked)
where
serve : Port -> Socket -> IO ()
serve port sock = do
Right (s, _) <- accept sock
| Left err => putStrLn ("Failed to accept on socket with error: " ++ show err)
Right (str, _) <- recv s 1024
| Left err => putStrLn ("Failed to accept on socket with error: " ++ show err)
putStrLn ("Received: " ++ str)
Right n <- send s ("echo: " ++ str)
| Left err => putStrLn ("Server failed to send data with error: " ++ show err)
pure ()
runClient : Port -> IO ()
runClient serverPort = do
Right sock <- socket AF_INET Stream 0
| Left fail => putStrLn ("Failed to open socket: " ++ show fail)
res <- connect sock (Hostname "localhost") serverPort
if res /= 0
then putStrLn ("Failed to connect client to port " ++ show serverPort ++ ": " ++ show res)
else do
Right n <- send sock ("hello world!")
| Left err => putStrLn ("Client failed to send data with error: " ++ show err)
Right (str, _) <- recv sock 1024
| Left err => putStrLn ("Client failed to receive on socket with error: " ++ show err)
-- assuming that stdout buffers get flushed in between system calls, this is "guaranteed"
-- to be printed after the server prints its own message
putStrLn ("Received: " ++ str)
main : IO ()
main = do
Right (serverPort, tid) <- runServer
| Left err => putStrLn $ "[server] " ++ err
runClient serverPort

54
libs/network/Makefile Normal file
View File

@ -0,0 +1,54 @@
include $(IDRIS2_CURDIR)/config.mk
INSTALLDIR = `${IDRIS2} --libdir`/network/lib
IDRIS_SRCS = Network/Socket.idr Network/Socket/Data.idr Network/Socket/Raw.idr
TARGET = idris_net
SRCS = $(wildcard *.c)
OBJS = $(SRCS:.c=.o)
DEPS = $(OBJS:.o=.d)
all: $(TARGET)$(SHLIB_SUFFIX)
${IDRIS2} --build network.ipkg
build: $(TARGET)$(SHLIB_SUFFIX) $(IDRIS_SRCS)
@if [ -z "${IDRIS2}" ]; then echo 'variable $$IDRIS2 is not set, aborting'; exit 1; fi
${IDRIS2} --build network.ipkg
$(TARGET)$(SHLIB_SUFFIX): $(OBJS)
$(CC) -shared $(LDFLAGS) -o $@ $^
-include $(DEPS)
%.d: %.c
@$(CPP) $(CFLAGS) $< -MM -MT $(@:.d=.o) >$@
.PHONY: clean
clean :
rm -f $(OBJS) $(TARGET)$(SHLIB_SUFFIX)
rm -rf build
cleandep: clean
rm -f $(DEPS)
.PHONY: install
install:
@if [ -z "${IDRIS2}" ]; then echo 'variable $$IDRIS2 is not set, aborting'; exit 1; fi
${IDRIS2} --install network.ipkg
@if ! [ -d $(INSTALLDIR) ]; then mkdir -p $(INSTALLDIR); fi
install $(TARGET)$(SHLIB_SUFFIX) $(wildcard *.h) $(INSTALLDIR)
test: build test.c
$(CC) -o network-tests -L. -I. test.c $(TARGET)$(SHLIB_SUFFIX)
LD_LIBRARY_PATH=. ./network-tests
@rm -f ./network-tests test.o

View File

@ -0,0 +1,233 @@
||| Low-Level C Sockets bindings for Idris. Used by higher-level, cleverer things.
|||
||| Original (C) SimonJF, MIT Licensed, 2014
||| Modified (C) The Idris Community, 2015, 2016, 2019
module Network.Socket
import public Network.Socket.Data
import Network.Socket.Raw
import Data.List
-- ----------------------------------------------------- [ Network Socket API. ]
||| Creates a UNIX socket with the given family, socket type and protocol
||| number. Returns either a socket or an error.
export
socket : (fam : SocketFamily)
-> (ty : SocketType)
-> (pnum : ProtocolNumber)
-> IO (Either SocketError Socket)
socket sf st pn = do
socket_res <- cCall Int "idrnet_socket" [toCode sf, toCode st, pn]
if socket_res == -1
then map Left getErrno
else pure $ Right (MkSocket socket_res sf st pn)
||| Close a socket
export
close : Socket -> IO ()
close sock = cCall () "close" [descriptor sock]
||| Binds a socket to the given socket address and port.
||| Returns 0 on success, an error code otherwise.
export
bind : (sock : Socket)
-> (addr : Maybe SocketAddress)
-> (port : Port)
-> IO Int
bind sock addr port = do
bind_res <- cCall Int "idrnet_bind"
[ descriptor sock, toCode $ family sock
, toCode $ socketType sock, saString addr, port
]
if bind_res == (-1)
then getErrno
else pure 0
where
saString : Maybe SocketAddress -> String
saString (Just sa) = show sa
saString Nothing = ""
||| Connects to a given address and port.
||| Returns 0 on success, and an error number on error.
export
connect : (sock : Socket)
-> (addr : SocketAddress)
-> (port : Port)
-> IO ResultCode
connect sock addr port = do
conn_res <- cCall Int "idrnet_connect"
[ descriptor sock, toCode $ family sock, toCode $ socketType sock, show addr, port]
if conn_res == (-1)
then getErrno
else pure 0
||| Listens on a bound socket.
|||
||| @sock The socket to listen on.
export
listen : (sock : Socket) -> IO Int
listen sock = do
listen_res <- cCall Int "listen" [ descriptor sock, BACKLOG ]
if listen_res == (-1)
then getErrno
else pure 0
||| Accept a connection on the provided socket.
|||
||| Returns on failure a `SocketError`
||| Returns on success a pairing of:
||| + `Socket` :: The socket representing the connection.
||| + `SocketAddress` :: The
|||
||| @sock The socket used to establish connection.
export
accept : (sock : Socket)
-> IO (Either SocketError (Socket, SocketAddress))
accept sock = do
-- We need a pointer to a sockaddr structure. This is then passed into
-- idrnet_accept and populated. We can then query it for the SocketAddr and free it.
sockaddr_ptr <- cCall AnyPtr "idrnet_create_sockaddr" []
accept_res <- cCall Int "idrnet_accept" [ descriptor sock, sockaddr_ptr ]
if accept_res == (-1)
then map Left getErrno
else do
let (MkSocket _ fam ty p_num) = sock
sockaddr <- getSockAddr (SAPtr sockaddr_ptr)
sockaddr_free (SAPtr sockaddr_ptr)
pure $ Right ((MkSocket accept_res fam ty p_num), sockaddr)
||| Send data on the specified socket.
|||
||| Returns on failure a `SocketError`.
||| Returns on success the `ResultCode`.
|||
||| @sock The socket on which to send the message.
||| @msg The data to send.
export
send : (sock : Socket)
-> (msg : String)
-> IO (Either SocketError ResultCode)
send sock dat = do
send_res <- cCall Int "idrnet_send" [ descriptor sock, dat ]
if send_res == (-1)
then map Left getErrno
else pure $ Right send_res
||| Receive data on the specified socket.
|||
||| Returns on failure a `SocketError`
||| Returns on success a pairing of:
||| + `String` :: The payload.
||| + `ResultCode` :: The result of the underlying function.
|||
||| @sock The socket on which to receive the message.
||| @len How much of the data to receive.
export
recv : (sock : Socket)
-> (len : ByteLength)
-> IO (Either SocketError (String, ResultCode))
recv sock len = do
-- Firstly make the request, get some kind of recv structure which
-- contains the result of the recv and possibly the retrieved payload
recv_struct_ptr <- cCall AnyPtr "idrnet_recv" [ descriptor sock, len]
recv_res <- cCall Int "idrnet_get_recv_res" [ recv_struct_ptr ]
if recv_res == (-1)
then do
errno <- getErrno
freeRecvStruct (RSPtr recv_struct_ptr)
pure $ Left errno
else
if recv_res == 0
then do
freeRecvStruct (RSPtr recv_struct_ptr)
pure $ Left 0
else do
payload <- cCall String "idrnet_get_recv_payload" [ recv_struct_ptr ]
freeRecvStruct (RSPtr recv_struct_ptr)
pure $ Right (payload, recv_res)
||| Receive all the remaining data on the specified socket.
|||
||| Returns on failure a `SocketError`
||| Returns on success the payload `String`
|||
||| @sock The socket on which to receive the message.
export
recvAll : (sock : Socket) -> IO (Either SocketError String)
recvAll sock = recvRec sock [] 64
where
partial
recvRec : Socket -> List String -> ByteLength -> IO (Either SocketError String)
recvRec sock acc n = do res <- recv sock n
case res of
Left 0 => pure (Right $ concat $ reverse acc)
Left c => pure (Left c)
Right (str, _) => let n' = min (n * 2) 65536 in
recvRec sock (str :: acc) n'
||| Send a message.
|||
||| Returns on failure a `SocketError`
||| Returns on success the `ResultCode`
|||
||| @sock The socket on which to send the message.
||| @addr Address of the recipient.
||| @port The port on which to send the message.
||| @msg The message to send.
export
sendTo : (sock : Socket)
-> (addr : SocketAddress)
-> (port : Port)
-> (msg : String)
-> IO (Either SocketError ByteLength)
sendTo sock addr p dat = do
sendto_res <- cCall Int "idrnet_sendto"
[ descriptor sock, dat, show addr, p ,toCode $ family sock ]
if sendto_res == (-1)
then map Left getErrno
else pure $ Right sendto_res
||| Receive a message.
|||
||| Returns on failure a `SocketError`.
||| Returns on success a triple of
||| + `UDPAddrInfo` :: The address of the sender.
||| + `String` :: The payload.
||| + `Int` :: Result value from underlying function.
|||
||| @sock The channel on which to receive.
||| @len Size of the expected message.
|||
export
recvFrom : (sock : Socket)
-> (len : ByteLength)
-> IO (Either SocketError (UDPAddrInfo, String, ResultCode))
recvFrom sock bl = do
recv_ptr <- cCall AnyPtr "idrnet_recvfrom"
[ descriptor sock, bl ]
let recv_ptr' = RFPtr recv_ptr
isNull <- (nullPtr recv_ptr)
if isNull
then map Left getErrno
else do
result <- cCall Int "idrnet_get_recvfrom_res" [ recv_ptr ]
if result == -1
then do
freeRecvfromStruct recv_ptr'
map Left getErrno
else do
payload <- foreignGetRecvfromPayload recv_ptr'
port <- foreignGetRecvfromPort recv_ptr'
addr <- foreignGetRecvfromAddr recv_ptr'
freeRecvfromStruct recv_ptr'
pure $ Right (MkUDPAddrInfo addr port, payload, result)

View File

@ -0,0 +1,213 @@
||| Low-Level C Sockets bindings for Idris. Used by higher-level, cleverer things.
||| Types used by Network.Socket.Raw and Network.Socket.
|||
||| Original (C) SimonJF, MIT Licensed, 2014
||| Modified (C) The Idris Community, 2015, 2016, 2019
module Network.Socket.Data
import Data.List
import Data.Strings
-- ------------------------------------------------------------ [ Type Aliases ]
%cg chez "libidris_net"
public export
ByteLength : Type
ByteLength = Int
public export
ResultCode : Type
ResultCode = Int
||| Protocol Number.
|||
||| Generally good enough to just set it to 0.
public export
ProtocolNumber : Type
ProtocolNumber = Int
||| SocketError: Error thrown by a socket operation
public export
SocketError : Type
SocketError = Int
||| SocketDescriptor: Native C Socket Descriptor
public export
SocketDescriptor : Type
SocketDescriptor = Int
public export
Port : Type
Port = Int
-- --------------------------------------------------------------- [ Constants ]
||| Backlog used within listen() call -- number of incoming calls
export
BACKLOG : Int
BACKLOG = 20
export
EAGAIN : Int
EAGAIN =
-- I'm sorry
-- maybe
unsafePerformIO $ cCall Int "idrnet_geteagain" []
-- ---------------------------------------------------------------- [ Error Code ]
export
getErrno : IO SocketError
getErrno = cCall Int "idrnet_errno" []
export
nullPtr : AnyPtr -> IO Bool
nullPtr p = cCall Bool "isNull" [p]
-- -------------------------------------------------------------- [ Interfaces ]
public export
interface ToCode a where
toCode : a -> Int
-- --------------------------------------------------------- [ Socket Families ]
||| Socket Families
|||
||| The ones that people might actually use. We're not going to need US
||| Government proprietary ones.
public export
data SocketFamily : Type where
||| Unspecified
AF_UNSPEC : SocketFamily
||| Unix type sockets
AF_UNIX : SocketFamily
||| IP / UDP etc. IPv4
AF_INET : SocketFamily
||| IP / UDP etc. IPv6
AF_INET6 : SocketFamily
Show SocketFamily where
show AF_UNSPEC = "AF_UNSPEC"
show AF_UNIX = "AF_UNIX"
show AF_INET = "AF_INET"
show AF_INET6 = "AF_INET6"
export
ToCode SocketFamily where
-- Don't know how to read a constant value from C code in idris2...
-- gotta to hardcode those for now
toCode AF_UNSPEC = 0 -- unsafePerformIO (cMacro "#AF_UNSPEC" Int)
toCode AF_UNIX = 1
toCode AF_INET = 2
toCode AF_INET6 = 10
export
getSocketFamily : Int -> Maybe SocketFamily
getSocketFamily i =
lookup i [ (toCode AF_UNSPEC, AF_UNSPEC)
, (toCode AF_UNIX, AF_UNIX)
, (toCode AF_INET, AF_INET)
, (toCode AF_INET6, AF_INET6)
]
-- ------------------------------------------------------------ [ Socket Types ]
||| Socket Types.
public export
data SocketType : Type where
||| Not a socket, used in certain operations
NotASocket : SocketType
||| TCP
Stream : SocketType
||| UDP
Datagram : SocketType
||| Raw sockets
RawSocket : SocketType
export
Show SocketType where
show NotASocket = "Not a socket"
show Stream = "Stream"
show Datagram = "Datagram"
show RawSocket = "Raw"
export
ToCode SocketType where
toCode NotASocket = 0
toCode Stream = 1
toCode Datagram = 2
toCode RawSocket = 3
-- --------------------------------------------------------------- [ Addresses ]
||| Network Addresses
public export
data SocketAddress : Type where
IPv4Addr : Int -> Int -> Int -> Int -> SocketAddress
||| Not implemented (yet)
IPv6Addr : SocketAddress
Hostname : String -> SocketAddress
||| Used when there's a parse error
InvalidAddress : SocketAddress
export
Show SocketAddress where
show (IPv4Addr i1 i2 i3 i4) = concat $ intersperse "." (map show [i1, i2, i3, i4])
show IPv6Addr = "NOT IMPLEMENTED YET"
show (Hostname host) = host
show InvalidAddress = "Invalid"
||| Parses a textual representation of an IPv4 address into a SocketAddress
export
parseIPv4 : String -> SocketAddress
parseIPv4 str =
case splitted of
(i1 :: i2 :: i3 :: i4 :: _) => IPv4Addr i1 i2 i3 i4
otherwise => InvalidAddress
where
toInt' : String -> Integer
toInt' = cast
toInt : String -> Int
toInt s = fromInteger $ toInt' s
splitted : List Int
splitted = map toInt (split (\c => c == '.') str)
-- --------------------------------------------------------- [ UDP Information ]
-- TODO: Expand to non-string payloads
public export
record UDPRecvData where
constructor MkUDPRecvData
remote_addr : SocketAddress
remote_port : Port
recv_data : String
data_len : Int
public export
record UDPAddrInfo where
constructor MkUDPAddrInfo
remote_addr : SocketAddress
remote_port : Port
-- ----------------------------------------------------------------- [ Sockets ]
||| The metadata about a socket
public export
record Socket where
constructor MkSocket
descriptor : SocketDescriptor
family : SocketFamily
socketType : SocketType
protocolNumber : ProtocolNumber

View File

@ -0,0 +1,202 @@
||| Low-Level C Sockets bindings for Idris. Used by higher-level, cleverer things.
||| Type-unsafe parts. Use Network.Socket for a safe variant.
|||
||| Original (C) SimonJF, MIT Licensed, 2014
||| Modified (C) The Idris Community, 2015, 2016
module Network.Socket.Raw
import public Network.Socket.Data
-- ---------------------------------------------------------------- [ Pointers ]
public export
data RecvStructPtr = RSPtr AnyPtr
public export
data RecvfromStructPtr = RFPtr AnyPtr
public export
data BufPtr = BPtr AnyPtr
public export
data SockaddrPtr = SAPtr AnyPtr
-- ---------------------------------------------------------- [ Socket Utilies ]
||| Put a value in a buffer
export
sock_poke : BufPtr -> Int -> Int -> IO ()
sock_poke (BPtr ptr) offset val = cCall () "idrnet_poke" [ptr, offset, val]
||| Take a value from a buffer
export
sock_peek : BufPtr -> Int -> IO Int
sock_peek (BPtr ptr) offset = cCall Int "idrnet_peek" [ptr, offset]
||| Frees a given pointer
export
sock_free : BufPtr -> IO ()
sock_free (BPtr ptr) = cCall () "idrnet_free" [ptr]
export
sockaddr_free : SockaddrPtr -> IO ()
sockaddr_free (SAPtr ptr) = cCall () "idrnet_free" [ptr]
||| Allocates an amount of memory given by the ByteLength parameter.
|||
||| Used to allocate a mutable pointer to be given to the Recv functions.
export
sock_alloc : ByteLength -> IO BufPtr
sock_alloc bl = map BPtr $ cCall AnyPtr "idrnet_malloc" [bl]
||| Retrieves the port the given socket is bound to
export
getSockPort : Socket -> IO Port
getSockPort sock = cCall Int "idrnet_sockaddr_port" [descriptor sock]
||| Retrieves a socket address from a sockaddr pointer
export
getSockAddr : SockaddrPtr -> IO SocketAddress
getSockAddr (SAPtr ptr) = do
addr_family_int <- cCall Int "idrnet_sockaddr_family" [ptr]
-- ASSUMPTION: Foreign call returns a valid int
assert_total (case getSocketFamily addr_family_int of
Just AF_INET => do
ipv4_addr <- cCall String "idrnet_sockaddr_ipv4" [ptr]
pure $ parseIPv4 ipv4_addr
Just AF_INET6 => pure IPv6Addr
Just AF_UNSPEC => pure InvalidAddress)
export
freeRecvStruct : RecvStructPtr -> IO ()
freeRecvStruct (RSPtr p) = cCall () "idrnet_free_recv_struct" [p]
||| Utility to extract data.
export
freeRecvfromStruct : RecvfromStructPtr -> IO ()
freeRecvfromStruct (RFPtr p) = cCall () "idrnet_free_recvfrom_struct" [p]
||| Sends the data in a given memory location
|||
||| Returns on failure a `SocketError`
||| Returns on success the `ResultCode`
|||
||| @sock The socket on which to send the message.
||| @ptr The location containing the data to send.
||| @len How much of the data to send.
export
sendBuf : (sock : Socket)
-> (ptr : BufPtr)
-> (len : ByteLength)
-> IO (Either SocketError ResultCode)
sendBuf sock (BPtr ptr) len = do
send_res <- cCall Int "idrnet_send_buf" [ descriptor sock, ptr, len]
if send_res == (-1)
then map Left getErrno
else pure $ Right send_res
||| Receive data from a given memory location.
|||
||| Returns on failure a `SocketError`
||| Returns on success the `ResultCode`
|||
||| @sock The socket on which to receive the message.
||| @ptr The location containing the data to receive.
||| @len How much of the data to receive.
export
recvBuf : (sock : Socket)
-> (ptr : BufPtr)
-> (len : ByteLength)
-> IO (Either SocketError ResultCode)
recvBuf sock (BPtr ptr) len = do
recv_res <- cCall Int "idrnet_recv_buf" [ descriptor sock, ptr, len ]
if (recv_res == (-1))
then map Left getErrno
else pure $ Right recv_res
||| Send a message stored in some buffer.
|||
||| Returns on failure a `SocketError`
||| Returns on success the `ResultCode`
|||
||| @sock The socket on which to send the message.
||| @addr Address of the recipient.
||| @port The port on which to send the message.
||| @ptr A Pointer to the buffer containing the message.
||| @len The size of the message.
export
sendToBuf : (sock : Socket)
-> (addr : SocketAddress)
-> (port : Port)
-> (ptr : BufPtr)
-> (len : ByteLength)
-> IO (Either SocketError ResultCode)
sendToBuf sock addr p (BPtr dat) len = do
sendto_res <- cCall Int "idrnet_sendto_buf"
[ descriptor sock, dat, len, show addr, p, toCode $ family sock ]
if sendto_res == (-1)
then map Left getErrno
else pure $ Right sendto_res
||| Utility function to get the payload of the sent message as a `String`.
export
foreignGetRecvfromPayload : RecvfromStructPtr -> IO String
foreignGetRecvfromPayload (RFPtr p) = cCall String "idrnet_get_recvfrom_payload" [ p ]
||| Utility function to return senders socket address.
export
foreignGetRecvfromAddr : RecvfromStructPtr -> IO SocketAddress
foreignGetRecvfromAddr (RFPtr p) = do
sockaddr_ptr <- map SAPtr $ cCall AnyPtr "idrnet_get_recvfrom_sockaddr" [p]
getSockAddr sockaddr_ptr
||| Utility function to return sender's IPV4 port.
export
foreignGetRecvfromPort : RecvfromStructPtr -> IO Port
foreignGetRecvfromPort (RFPtr p) = do
sockaddr_ptr <- cCall AnyPtr "idrnet_get_recvfrom_sockaddr" [p]
port <- cCall Int "idrnet_sockaddr_ipv4_port" [sockaddr_ptr]
pure port
||| Receive a message placed on a 'known' buffer.
|||
||| Returns on failure a `SocketError`.
||| Returns on success a pair of
||| + `UDPAddrInfo` :: The address of the sender.
||| + `Int` :: Result value from underlying function.
|||
||| @sock The channel on which to receive.
||| @ptr Pointer to the buffer to place the message.
||| @len Size of the expected message.
|||
export
recvFromBuf : (sock : Socket)
-> (ptr : BufPtr)
-> (len : ByteLength)
-> IO (Either SocketError (UDPAddrInfo, ResultCode))
recvFromBuf sock (BPtr ptr) bl = do
recv_ptr <- cCall AnyPtr "idrnet_recvfrom_buf" [ descriptor sock, ptr, bl]
let recv_ptr' = RFPtr recv_ptr
isnull <- nullPtr recv_ptr
if isnull
then map Left getErrno
else do
result <- cCall Int "idrnet_get_recvfrom_res" [recv_ptr]
if result == -1
then do
freeRecvfromStruct recv_ptr'
map Left getErrno
else do
port <- foreignGetRecvfromPort recv_ptr'
addr <- foreignGetRecvfromAddr recv_ptr'
freeRecvfromStruct recv_ptr'
pure $ Right (MkUDPAddrInfo addr port, result + 1)

2
libs/network/expected Normal file
View File

@ -0,0 +1,2 @@
Received: hello world!
Received: echo: hello world!

393
libs/network/idris_net.c Normal file
View File

@ -0,0 +1,393 @@
// C-Side of the Idris network library
// (C) Simon Fowler, 2014
// MIT Licensed. Have fun!
#include "idris_net.h"
#include <errno.h>
#include <stdbool.h>
#include <stdlib.h>
#include <stdio.h>
#include <string.h>
#ifndef _WIN32
#include <netinet/in.h>
#include <arpa/inet.h>
#else
static int socket_inited = 0;
static WSADATA wsa_data;
static void clean_sockets(void) {
WSACleanup();
}
static int check_init(void) {
if (!socket_inited) {
int result = WSAStartup(MAKEWORD(2, 2), &wsa_data);
if (result == NO_ERROR) {
socket_inited = 1;
atexit(clean_sockets);
}
}
return socket_inited;
}
#endif
void buf_htonl(void* buf, int len) {
int* buf_i = (int*) buf;
int i;
for (i = 0; i < (len / sizeof(int)) + 1; i++) {
buf_i[i] = htonl(buf_i[i]);
}
}
void buf_ntohl(void* buf, int len) {
int* buf_i = (int*) buf;
int i;
for (i = 0; i < (len / sizeof(int)) + 1; i++) {
buf_i[i] = ntohl(buf_i[i]);
}
}
void* idrnet_malloc(int size) {
return malloc(size);
}
void idrnet_free(void* ptr) {
free(ptr);
}
unsigned int idrnet_peek(void *ptr, unsigned int offset) {
unsigned char *buf_c = (unsigned char*) ptr;
return (unsigned int) buf_c[offset];
}
void idrnet_poke(void *ptr, unsigned int offset, char val) {
char *buf_c = (char*)ptr;
buf_c[offset] = val;
}
int idrnet_socket(int domain, int type, int protocol) {
#ifdef _WIN32
if (!check_init()) {
return -1;
}
#endif
return socket(domain, type, protocol);
}
// We call this from quite a few functions. Given a textual host and an int port,
// populates a struct addrinfo.
int idrnet_getaddrinfo(struct addrinfo** address_res, char* host, int port,
int family, int socket_type) {
struct addrinfo hints;
// Convert port into string
char str_port[8];
sprintf(str_port, "%d", port);
// Set up hints structure
memset(&hints, 0, sizeof(hints)); // zero out hints
hints.ai_family = family;
hints.ai_socktype = socket_type;
// If the length of the hostname is 0 (i.e, it was set to Nothing in Idris)
// then we want to instruct the C library to fill in the IP automatically
if (strlen(host) == 0) {
hints.ai_flags = AI_PASSIVE; // fill in IP automatically
return getaddrinfo(NULL, str_port, &hints, address_res);
}
return getaddrinfo(host, str_port, &hints, address_res);
}
int idrnet_bind(int sockfd, int family, int socket_type, char* host, int port) {
struct addrinfo *address_res;
int addr_res = idrnet_getaddrinfo(&address_res, host, port, family, socket_type);
if (addr_res != 0) {
//printf("Lib err: bind getaddrinfo\n");
return -1;
}
int bind_res = bind(sockfd, address_res->ai_addr, address_res->ai_addrlen);
if (bind_res == -1) {
//freeaddrinfo(address_res);
//printf("Lib err: bind\n");
return -1;
}
return 0;
}
// to retrieve information about a socket bound to port 0
int idrnet_getsockname(int sockfd, void *address, void *len) {
int res = getsockname(sockfd, address, len);
if(res != 0) {
return -1;
}
return 0;
}
int idrnet_sockaddr_port(int sockfd) {
struct sockaddr address;
socklen_t addrlen = sizeof(struct sockaddr);
int res = getsockname(sockfd, &address, &addrlen);
if(res < 0) {
return -1;
}
switch(address.sa_family) {
case AF_INET:
return ntohs(((struct sockaddr_in*)&address)->sin_port);
case AF_INET6:
return ntohs(((struct sockaddr_in6*)&address)->sin6_port);
default:
return -1;
}
}
int idrnet_connect(int sockfd, int family, int socket_type, char* host, int port) {
struct addrinfo* remote_host;
int addr_res = idrnet_getaddrinfo(&remote_host, host, port, family, socket_type);
if (addr_res != 0) {
return -1;
}
int connect_res = connect(sockfd, remote_host->ai_addr, remote_host->ai_addrlen);
if (connect_res == -1) {
freeaddrinfo(remote_host);
return -1;
}
freeaddrinfo(remote_host);
return 0;
}
int idrnet_sockaddr_family(void* sockaddr) {
struct sockaddr* addr = (struct sockaddr*) sockaddr;
return (int) addr->sa_family;
}
char* idrnet_sockaddr_ipv4(void* sockaddr) {
struct sockaddr_in* addr = (struct sockaddr_in*) sockaddr;
char* ip_addr = (char*) malloc(sizeof(char) * INET_ADDRSTRLEN);
inet_ntop(AF_INET, &(addr->sin_addr), ip_addr, INET_ADDRSTRLEN);
//printf("Lib: ip_addr: %s\n", ip_addr);
return ip_addr;
}
int idrnet_sockaddr_ipv4_port(void* sockaddr) {
struct sockaddr_in* addr = (struct sockaddr_in*) sockaddr;
return ((int) ntohs(addr->sin_port));
}
void* idrnet_create_sockaddr() {
return malloc(sizeof(struct sockaddr_storage));
}
int idrnet_accept(int sockfd, void* sockaddr) {
struct sockaddr* addr = (struct sockaddr*) sockaddr;
socklen_t addr_size = sizeof(struct sockaddr_storage);
return accept(sockfd, addr, &addr_size);
}
int idrnet_send(int sockfd, char* data) {
int len = strlen(data); // For now.
return send(sockfd, (void*) data, len, 0);
}
int idrnet_send_buf(int sockfd, void* data, int len) {
void* buf_cpy = malloc(len);
memset(buf_cpy, 0, len);
memcpy(buf_cpy, data, len);
buf_htonl(buf_cpy, len);
int res = send(sockfd, buf_cpy, len, 0);
free(buf_cpy);
return res;
}
void* idrnet_recv(int sockfd, int len) {
idrnet_recv_result* res_struct =
(idrnet_recv_result*) malloc(sizeof(idrnet_recv_result));
char* buf = malloc(len + 1);
memset(buf, 0, len + 1);
int recv_res = recv(sockfd, buf, len, 0);
res_struct->result = recv_res;
if (recv_res > 0) { // Data was received
buf[recv_res + 1] = 0x00; // Null-term, so Idris can interpret it
}
res_struct->payload = buf;
return (void*) res_struct;
}
int idrnet_recv_buf(int sockfd, void* buf, int len) {
int recv_res = recv(sockfd, buf, len, 0);
if (recv_res != -1) {
buf_ntohl(buf, len);
}
return recv_res;
}
int idrnet_get_recv_res(void* res_struct) {
return (((idrnet_recv_result*) res_struct)->result);
}
char* idrnet_get_recv_payload(void* res_struct) {
return (((idrnet_recv_result*) res_struct)->payload);
}
void idrnet_free_recv_struct(void* res_struct) {
idrnet_recv_result* i_res_struct =
(idrnet_recv_result*) res_struct;
if (i_res_struct->payload != NULL) {
free(i_res_struct->payload);
}
free(res_struct);
}
int idrnet_errno() {
return errno;
}
int idrnet_sendto(int sockfd, char* data, char* host, int port, int family) {
struct addrinfo* remote_host;
int addr_res = idrnet_getaddrinfo(&remote_host, host, port, family, SOCK_DGRAM);
if (addr_res != 0) {
return -1;
}
int send_res = sendto(sockfd, data, strlen(data), 0,
remote_host->ai_addr, remote_host->ai_addrlen);
freeaddrinfo(remote_host);
return send_res;
}
int idrnet_sendto_buf(int sockfd, void* buf, int buf_len, char* host, int port, int family) {
struct addrinfo* remote_host;
int addr_res = idrnet_getaddrinfo(&remote_host, host, port, family, SOCK_DGRAM);
if (addr_res != 0) {
//printf("lib err: sendto getaddrinfo \n");
return -1;
}
buf_htonl(buf, buf_len);
int send_res = sendto(sockfd, buf, buf_len, 0,
remote_host->ai_addr, remote_host->ai_addrlen);
if (send_res == -1) {
perror("lib err: sendto \n");
}
//freeaddrinfo(remote_host);
return send_res;
}
void* idrnet_recvfrom(int sockfd, int len) {
/*
* int recvfrom(int sockfd, void *buf, int len, unsigned int flags,
struct sockaddr *from, int *fromlen);
*/
// Allocate the required structures, and nuke them
struct sockaddr_storage* remote_addr =
(struct sockaddr_storage*) malloc(sizeof(struct sockaddr_storage));
char* buf = (char*) malloc(len + 1);
idrnet_recvfrom_result* ret =
(idrnet_recvfrom_result*) malloc(sizeof(idrnet_recvfrom_result));
memset(remote_addr, 0, sizeof(struct sockaddr_storage));
memset(buf, 0, len + 1);
memset(ret, 0, sizeof(idrnet_recvfrom_result));
socklen_t fromlen = sizeof(struct sockaddr_storage);
int recv_res = recvfrom(sockfd, buf, len, 0, (struct sockaddr*) remote_addr, &fromlen);
ret->result = recv_res;
// Check for failure...
if (recv_res == -1) {
free(buf);
free(remote_addr);
} else {
// If data was received, process and populate
ret->result = recv_res;
ret->remote_addr = remote_addr;
// Ensure the last byte is NULL, since in this mode we're sending strings
buf[len] = 0x00;
ret->payload = (void*) buf;
}
return ret;
}
void* idrnet_recvfrom_buf(int sockfd, void* buf, int len) {
// Allocate the required structures, and nuke them
struct sockaddr_storage* remote_addr =
(struct sockaddr_storage*) malloc(sizeof(struct sockaddr_storage));
idrnet_recvfrom_result* ret =
(idrnet_recvfrom_result*) malloc(sizeof(idrnet_recvfrom_result));
memset(remote_addr, 0, sizeof(struct sockaddr_storage));
memset(ret, 0, sizeof(idrnet_recvfrom_result));
socklen_t fromlen = 0;
int recv_res = recvfrom(sockfd, buf, len, 0, (struct sockaddr*) remote_addr, &fromlen);
// Check for failure... But don't free the buffer! Not our job.
ret->result = recv_res;
if (recv_res == -1) {
free(remote_addr);
}
// Payload will be NULL -- since it's been put into the user-specified buffer. We
// still need the return struct to get our hands on the remote address, though.
if (recv_res > 0) {
buf_ntohl(buf, len);
ret->payload = NULL;
ret->remote_addr = remote_addr;
}
return ret;
}
int idrnet_get_recvfrom_res(void* res_struct) {
return (((idrnet_recvfrom_result*) res_struct)->result);
}
char* idrnet_get_recvfrom_payload(void* res_struct) {
return (((idrnet_recvfrom_result*) res_struct)->payload);
}
void* idrnet_get_recvfrom_sockaddr(void* res_struct) {
idrnet_recvfrom_result* recv_struct = (idrnet_recvfrom_result*) res_struct;
return recv_struct->remote_addr;
}
int idrnet_get_recvfrom_port(void* res_struct) {
idrnet_recvfrom_result* recv_struct = (idrnet_recvfrom_result*) res_struct;
if (recv_struct->remote_addr != NULL) {
struct sockaddr_in* remote_addr_in =
(struct sockaddr_in*) recv_struct->remote_addr;
return ((int) ntohs(remote_addr_in->sin_port));
}
return -1;
}
void idrnet_free_recvfrom_struct(void* res_struct) {
idrnet_recvfrom_result* recv_struct = (idrnet_recvfrom_result*) res_struct;
if (recv_struct != NULL) {
if (recv_struct->payload != NULL) {
free(recv_struct->payload);
}
if (recv_struct->remote_addr != NULL) {
free(recv_struct->remote_addr);
}
}
}
int idrnet_geteagain() {
return EAGAIN;
}
int isNull(void* ptr) {
return ptr==NULL;
}

100
libs/network/idris_net.h Normal file
View File

@ -0,0 +1,100 @@
#ifndef IDRISNET_H
#define IDRISNET_H
// Includes used by the idris-file.
#ifdef _WIN32
#include <winsock2.h>
#include <Ws2tcpip.h>
#else
#include <netdb.h>
#include <unistd.h>
#include <sys/types.h>
#include <sys/socket.h>
#endif
struct sockaddr_storage;
struct addrinfo;
typedef struct idrnet_recv_result {
int result;
void* payload;
} idrnet_recv_result;
// Same type of thing as idrnet_recv_result, but for UDP, so stores some
// address information
typedef struct idrnet_recvfrom_result {
int result;
void* payload;
struct sockaddr_storage* remote_addr;
} idrnet_recvfrom_result;
// Memory management functions
void* idrnet_malloc(int size);
void idrnet_free(void* ptr);
unsigned int idrnet_peek(void *ptr, unsigned int offset);
void idrnet_poke(void *ptr, unsigned int offset, char val);
// Gets value of errno
int idrnet_errno();
int idrnet_socket(int domain, int type, int protocol);
// Bind
int idrnet_bind(int sockfd, int family, int socket_type, char* host, int port);
// Retrieve information about socket
int idrnet_getsockname(int sockfd, void *address, void *len);
// Connect
int idrnet_connect(int sockfd, int family, int socket_type, char* host, int port);
// Accessor functions for struct_sockaddr
int idrnet_sockaddr_family(void* sockaddr);
char* idrnet_sockaddr_ipv4(void* sockaddr);
int idrnet_sockaddr_ipv4_port(void* sockaddr);
void* idrnet_create_sockaddr();
int idrnet_sockaddr_port(int sockfd);
// Accept
int idrnet_accept(int sockfd, void* sockaddr);
// Send
int idrnet_send(int sockfd, char* data);
int idrnet_send_buf(int sockfd, void* data, int len);
// Receive
// Creates a result structure containing result and payload
void* idrnet_recv(int sockfd, int len);
// Receives directly into a buffer
int idrnet_recv_buf(int sockfd, void* buf, int len);
// UDP Send
int idrnet_sendto(int sockfd, char* data, char* host, int port, int family);
int idrnet_sendto_buf(int sockfd, void* buf, int buf_len, char* host, int port, int family);
// UDP Receive
void* idrnet_recvfrom(int sockfd, int len);
void* idrnet_recvfrom_buf(int sockfd, void* buf, int len);
// Receive structure accessors
int idrnet_get_recv_res(void* res_struct);
char* idrnet_get_recv_payload(void* res_struct);
void idrnet_free_recv_struct(void* res_struct);
// Recvfrom structure accessors
int idrnet_get_recvfrom_res(void* res_struct);
char* idrnet_get_recvfrom_payload(void* res_struct);
void* idrnet_get_recvfrom_sockaddr(void* res_struct);
void idrnet_free_recvfrom_struct(void* res_struct);
int idrnet_getaddrinfo(struct addrinfo** address_res, char* host,
int port, int family, int socket_type);
int idrnet_geteagain();
int isNull(void* ptr);
#endif

View File

@ -0,0 +1,5 @@
package network
modules = Network.Socket,
Network.Socket.Data,
Network.Socket.Raw

56
libs/network/test.c Normal file
View File

@ -0,0 +1,56 @@
#include <assert.h>
#include <stdio.h>
#include <netinet/in.h>
#include <arpa/inet.h>
#include "idris_net.h"
void test_sockaddr_port_returns_random_port_when_bind_port_is_0() {
int sock = idrnet_socket(AF_INET, 1, 0);
assert(sock > 0);
int res = idrnet_bind(sock, AF_INET, 1, "127.0.0.1", 0);
assert(res == 0);
res = idrnet_sockaddr_port(sock);
assert(res > 0);
close(sock);
}
void test_sockaddr_port_returns_explicitly_assigned_port() {
int sock = idrnet_socket(AF_INET, 1, 0);
assert(sock > 0);
int res = idrnet_bind(sock, AF_INET, 1, "127.0.0.1", 34567);
assert(res == 0);
res = idrnet_sockaddr_port(sock);
assert(res == 34567);
close(sock);
}
void test_peek_and_poke_buffer() {
void *buf = idrnet_malloc(100);
assert(buf > 0);
for (int i = 0; i < 100; i++) {
idrnet_poke(buf,i,7*i);
}
for(int i = 0; i < 100; i++) {
assert (idrnet_peek(buf,i) == (0xff & 7*i));
}
idrnet_free(buf);
}
int main(void) {
test_sockaddr_port_returns_explicitly_assigned_port();
test_sockaddr_port_returns_random_port_when_bind_port_is_0();
test_peek_and_poke_buffer();
printf("network library tests SUCCESS\n");
return 0;
}

146
libs/prelude/Builtin.idr Normal file
View File

@ -0,0 +1,146 @@
module Builtin
-- The most primitive data types; things which are used by desugaring
-- Totality assertions
||| Assert to the totality checker that the given expression will always
||| terminate.
%inline
public export
assert_total : {0 a : _} -> a -> a
assert_total x = x
||| Assert to the totality checker that y is always structurally smaller than x
||| (which is typically a pattern argument, and *must* be in normal form for
||| this to work).
||| @ x the larger value (typically a pattern argument)
||| @ y the smaller value (typically an argument to a recursive call)
%inline
public export
assert_smaller : {0 a, b : _} -> (x : a) -> (y : b) -> b
assert_smaller x y = y
-- Unit type and pairs
||| The canonical single-element type, also known as the trivially true
||| proposition.
public export
data Unit =
||| The trivial constructor for `()`.
MkUnit
||| The non-dependent pair type, also known as conjunction.
public export
data Pair : Type -> Type -> Type where
||| A pair of elements.
||| @ a the left element of the pair
||| @ b the right element of the pair
MkPair : {0 a, b : Type} -> (1 x : a) -> (1 y : b) -> Pair a b
||| Return the first element of a pair.
public export
fst : {0 a, b : Type} -> (a, b) -> a
fst (x, y) = x
||| Return the second element of a pair.
public export
snd : {0 a, b : Type} -> (a, b) -> b
snd (x, y) = y
-- This directive tells auto implicit search what to use to look inside pairs
%pair Pair fst snd
||| Dependent pairs aid in the construction of dependent types by providing
||| evidence that some value resides in the type.
|||
||| Formally, speaking, dependent pairs represent existential quantification -
||| they consist of a witness for the existential claim and a proof that the
||| property holds for it.
|||
||| @ a the value to place in the type.
||| @ p the dependent type that requires the value.
namespace DPair
public export
record DPair a (p : a -> Type) where
constructor MkDPair
fst : a
snd : p fst
-- The empty type
||| The empty type, also known as the trivially false proposition.
|||
||| Use `void` or `absurd` to prove anything if you have a variable of type
||| `Void` in scope.
public export
data Void : Type where
-- Equality
public export
data Equal : forall a, b . a -> b -> Type where
Refl : {0 x : a} -> Equal x x
%name Equal prf
infix 9 ===, ~=~
-- An equality type for when you want to assert that each side of the
-- equality has the same type, but there's not other evidence available
-- to help with unification
public export
(===) : (x : a) -> (y : a) -> Type
(===) = Equal
||| Explicit heterogeneous ("John Major") equality. Use this when Idris
||| incorrectly chooses homogeneous equality for `(=)`.
||| @ a the type of the left side
||| @ b the type of the right side
||| @ x the left side
||| @ y the right side
public export
(~=~) : (x : a) -> (y : b) -> Type
(~=~) = Equal
||| Perform substitution in a term according to some equality.
|||
||| Like `replace`, but with an explicit predicate, and applying the rewrite in
||| the other direction, which puts it in a form usable by the `rewrite` tactic
||| and term.
%inline
public export
rewrite__impl : {0 x, y : a} -> (0 p : _) ->
(0 rule : x = y) -> (1 val : p y) -> p x
rewrite__impl p Refl prf = prf
%rewrite Equal rewrite__impl
||| Perform substitution in a term according to some equality.
%inline
public export
replace : forall x, y, p . (0 rule : x = y) -> p x -> p y
replace Refl prf = prf
||| Symmetry of propositional equality.
%inline
public export
sym : (0 rule : x = y) -> y = x
sym Refl = Refl
||| Transitivity of propositional equality.
%inline
public export
trans : forall a, b, c . (0 l : a = b) -> (0 r : b = c) -> a = c
trans Refl Refl = Refl
||| Subvert the type checker. This function is abstract, so it will not reduce
||| in the type checker. Use it with care - it can result in segfaults or
||| worse!
public export
believe_me : a -> b
believe_me = prim__believe_me _ _
export partial
idris_crash : String -> a
idris_crash = prim__crash _

8
libs/prelude/Makefile Normal file
View File

@ -0,0 +1,8 @@
all:
${IDRIS2} --build prelude.ipkg
install:
${IDRIS2} --install prelude.ipkg
clean:
rm -rf build

1698
libs/prelude/Prelude.idr Normal file

File diff suppressed because it is too large Load Diff

160
libs/prelude/PrimIO.idr Normal file
View File

@ -0,0 +1,160 @@
module PrimIO
import Builtin
public export
data IORes : Type -> Type where
MkIORes : (result : a) -> (1 x : %World) -> IORes a
||| Idris's primitive IO, for building abstractions on top of.
public export
PrimIO : Type -> Type
PrimIO a = (1 x : %World) -> IORes a
export
data IO : Type -> Type where
MkIO : (1 fn : PrimIO a) -> IO a
export
prim_io_pure : a -> PrimIO a
prim_io_pure x = \w => MkIORes x w
%inline
export
io_pure : a -> IO a
io_pure x = MkIO (\w => MkIORes x w)
export
prim_io_bind : (1 act : PrimIO a) -> (1 k : a -> PrimIO b) -> PrimIO b
prim_io_bind fn k w
= let MkIORes x' w' = fn w in k x' w'
-- There's a special case for inlining this is Compiler.Inline, because
-- the inliner is cautious about case blocks at the moment. Once it's less
-- cautious, add an explicit %inline directive and take out the special case.
-- See also dealing with the newtype optimisation via %World in
-- Compiler.CompileExpr
export
io_bind : (1 act : IO a) -> (1 k : a -> IO b) -> IO b
io_bind (MkIO fn) k
= MkIO (\w => let MkIORes x' w' = fn w
MkIO res = k x' in
res w')
%foreign "C:putchar,libc 6"
prim__putChar : Char -> (1 x : %World) -> IORes ()
%foreign "C:getchar,libc 6"
%extern prim__getChar : (1 x : %World) -> IORes Char
-- A pointer representing a given parameter type
-- The parameter is a phantom type, to help differentiate between
-- different pointer types
public export
data Ptr : Type -> Type where [external]
-- A pointer to any type (representing a void* in foreign calls)
public export
data AnyPtr : Type where [external]
public export
data ThreadID : Type where [external]
public export
data FArgList : Type where
Nil : FArgList
(::) : {a : Type} -> (1 arg : a) -> (1 args : FArgList) -> FArgList
%extern prim__cCall : (ret : Type) -> String -> (1 args : FArgList) ->
(1 x : %World) -> IORes ret
%extern prim__schemeCall : (ret : Type) -> String -> (1 args : FArgList) ->
(1 x : %World) -> IORes ret
export %inline
primIO : (1 fn : (1 x : %World) -> IORes a) -> IO a
primIO op = MkIO op
export %inline
toPrim : (1 act : IO a) -> PrimIO a
toPrim (MkIO fn) = fn
export %inline
schemeCall : (ret : Type) -> String -> (1 args : FArgList) -> IO ret
schemeCall ret fn args = primIO (prim__schemeCall ret fn args)
export %inline
cCall : (ret : Type) -> String -> FArgList -> IO ret
cCall ret fn args = primIO (prim__cCall ret fn args)
%foreign "C:idris2_isNull, libidris2_support"
export
prim__nullAnyPtr : AnyPtr -> Int
prim__forgetPtr : Ptr t -> AnyPtr
prim__forgetPtr = believe_me
export %inline
prim__nullPtr : Ptr t -> Int -- can't pass 'type' to a foreign function
prim__nullPtr p = prim__nullAnyPtr (prim__forgetPtr p)
%foreign "C:idris2_getString, libidris2_support"
export
prim__getString : Ptr String -> String
%foreign "C:idris2_getStr,libidris2_support"
prim__getStr : PrimIO String
%foreign "C:idris2_putStr,libidris2_support"
prim__putStr : String -> PrimIO ()
||| Output a string to stdout without a trailing newline.
export
putStr : String -> IO ()
putStr str = primIO (prim__putStr str)
||| Output a string to stdout with a trailing newline.
export
putStrLn : String -> IO ()
putStrLn str = putStr (prim__strAppend str "\n")
||| Read one line of input from stdin, without the trailing newline.
export
getLine : IO String
getLine = primIO prim__getStr
||| Write a single character to stdout.
export
putChar : Char -> IO ()
putChar c = primIO (prim__putChar c)
||| Write a single character to stdout, with a trailing newline.
export
putCharLn : Char -> IO ()
putCharLn c = putStrLn (prim__cast_CharString c)
||| Read a single character from stdin.
export
getChar : IO Char
getChar = primIO prim__getChar
export
fork : (1 prog : IO ()) -> IO ThreadID
fork (MkIO act) = schemeCall ThreadID "blodwen-thread" [act]
export
prim_fork : (1 prog : PrimIO ()) -> PrimIO ThreadID
prim_fork act w = prim__schemeCall ThreadID "blodwen-thread" [act] w
%foreign "C:idris2_readString, libidris2_support"
export
prim__getErrno : Int
unsafeCreateWorld : (1 f : (1 x : %World) -> a) -> a
unsafeCreateWorld f = f %MkWorld
unsafeDestroyWorld : (1 x : %World) -> a -> a
unsafeDestroyWorld %MkWorld x = x
export
unsafePerformIO : IO a -> a
unsafePerformIO (MkIO f)
= unsafeCreateWorld (\w => case f w of
MkIORes res w' => unsafeDestroyWorld w' res)

View File

@ -0,0 +1,8 @@
package prelude
opts = "--no-prelude"
modules = Builtin,
PrimIO,
Prelude

View File

@ -175,7 +175,7 @@ writeTTCFile : (HasNames extra, TTC extra) =>
Ref Bin Binary -> TTCFile extra -> Core ()
writeTTCFile b file_in
= do file <- toFullNames file_in
toBuf b "TTC"
toBuf b "TT2"
toBuf b (version file)
toBuf b (ifaceHash file)
toBuf b (importHashes file)
@ -202,7 +202,7 @@ readTTCFile : TTC extra =>
readTTCFile modns as b
= do hdr <- fromBuf b
chunk <- get Bin
when (hdr /= "TTC") $ corrupt ("TTC header in " ++ show modns ++ " " ++ show hdr)
when (hdr /= "TT2") $ corrupt ("TTC header in " ++ show modns ++ " " ++ show hdr)
ver <- fromBuf b
checkTTCVersion (show modns) ver ttcVersion
ifaceHash <- fromBuf b
@ -460,7 +460,7 @@ getImportHashes : String -> Ref Bin Binary ->
Core (List (List String, Int))
getImportHashes file b
= do hdr <- fromBuf {a = String} b
when (hdr /= "TTC") $ corrupt ("TTC header in " ++ file ++ " " ++ show hdr)
when (hdr /= "TT2") $ corrupt ("TTC header in " ++ file ++ " " ++ show hdr)
ver <- fromBuf {a = Int} b
checkTTCVersion file ver ttcVersion
ifaceHash <- fromBuf {a = Int} b
@ -469,7 +469,7 @@ getImportHashes file b
getHash : String -> Ref Bin Binary -> Core Int
getHash file b
= do hdr <- fromBuf {a = String} b
when (hdr /= "TTC") $ corrupt ("TTC header in " ++ file ++ " " ++ show hdr)
when (hdr /= "TT2") $ corrupt ("TTC header in " ++ file ++ " " ++ show hdr)
ver <- fromBuf {a = Int} b
checkTTCVersion file ver ttcVersion
fromBuf b

View File

@ -33,7 +33,7 @@ import IdrisPaths
%default covering
yprefix : String
yprefix = "/home/edwin/.idris2" -- TODO! unsafePerformIO (foreign FFI_C "getIdris2_prefix" (IO String))
yprefix = "/home/edwin/.idris2sh" -- TODO! unsafePerformIO (foreign FFI_C "getIdris2_prefix" (IO String))
findInput : List CLOpt -> Maybe String
findInput [] = Nothing

View File

@ -161,7 +161,7 @@ TTC Int where
= do chunk <- get Bin
if avail chunk >= 4
then
do coreLift $ setInt (buf chunk) (loc chunk) (val `mod` 2147483648)
do coreLift $ setInt (buf chunk) (loc chunk) val
put Bin (appended 4 chunk)
else do chunk' <- extendBinary 4 chunk
coreLift $ setInt (buf chunk') (loc chunk') val