Merge github.com:idris-lang/Idris2 into tweak-cg

This commit is contained in:
Edwin Brady 2021-05-11 11:37:12 +01:00
commit f21a495711
181 changed files with 5895 additions and 1486 deletions

16
.gitignore vendored
View File

@ -17,10 +17,11 @@ idris2docs_venv
/libs/**/build /libs/**/build
/tests/**/build /tests/**/build
/tests/**/output /tests/**/output*
/tests/**/*.so /tests/**/*.so
/tests/**/*.dylib /tests/**/*.dylib
/tests/**/*.dll /tests/**/*.dll
/tests/failures
/benchmark/**/build /benchmark/**/build
/benchmark/*.csv /benchmark/*.csv
@ -39,9 +40,16 @@ idris2docs_venv
/result /result
# Editor/IDE Related # Editor/IDE Related
*~ # Vim swap file # WARNING:
.\#* # Emacs swap file # do not put comments on the same line as a regex
.vscode/* # VS Code # git seems to ignore the pattern in this case
# Vim swap file
*~
# Emacs swap file
.\#*
# VS Code
.vscode/*
# macOS # macOS
.DS_Store .DS_Store

View File

@ -29,6 +29,26 @@ Compiler changes:
order to not block the Racket runtime when `sleep` is called. order to not block the Racket runtime when `sleep` is called.
* Added `--profile` flag, which generates profile data if supported by a back * Added `--profile` flag, which generates profile data if supported by a back
end. Currently supported by the Chez and Racket back ends. end. Currently supported by the Chez and Racket back ends.
* Javascript codegens now use `Number` to represent up to 32 bit precision
signed and unsigned integers. `Int32` still goes via `BigInt` for
multiplication to avoid precision issues when getting results larger
than `Number.MAX_SAFE_INTEGER`. `Bits32` goes via `BigInt` for
multiplication for the same reason as well as for all bitops, since `Number`
uses signed 32 bit integers for those.
* New code generator: `chez-sep`. This code generator produces many Chez Scheme
files and compiles them separately instead of producing one huge Scheme
program. This significantly reduces the amount of memory needed to build
large programs. Since this backend will skip calling the Chez compiler on
modules that haven't changed, it also leads to shorter compilation times in
large codebases where only some files have changed -- for example when
developing Idris2 code generators. The codegen has a large parallelisation
potential but at the moment, it is significantly slower for a full rebuild of
a large code base (the code generation stage takes about 3x longer).
API changes:
* The API now exposes `Compiler.Separate.getCompilationUnits`, which
can be used for separate code generation by any backend.
Library changes: Library changes:

View File

@ -50,6 +50,7 @@ Rohit Grover
Rui Barreiro Rui Barreiro
Ruslan Feizerahmanov Ruslan Feizerahmanov
Simon Chatterjee Simon Chatterjee
Stefan Höck
tensorknower69 tensorknower69
then0rTh then0rTh
Theo Butler Theo Butler

View File

@ -112,14 +112,23 @@ test: testenv
@echo @echo
@${MAKE} -C tests only=$(only) IDRIS2=${TARGET} IDRIS2_PREFIX=${TEST_PREFIX} @${MAKE} -C tests only=$(only) IDRIS2=${TARGET} IDRIS2_PREFIX=${TEST_PREFIX}
retest: testenv
@echo
@echo "NOTE: \`${MAKE} retest\` does not rebuild Idris or the libraries packaged with it; to do that run \`${MAKE}\`"
@if [ ! -x "${TARGET}" ]; then echo "ERROR: Missing IDRIS2 executable. Cannot run tests!\n"; exit 1; fi
@echo
@${MAKE} -C tests retest only=$(only) IDRIS2=${TARGET} IDRIS2_PREFIX=${TEST_PREFIX}
support: support:
@${MAKE} -C support/c @${MAKE} -C support/c
@${MAKE} -C support/refc @${MAKE} -C support/refc
@${MAKE} -C support/chez
support-clean: support-clean:
@${MAKE} -C support/c clean @${MAKE} -C support/c clean
@${MAKE} -C support/refc clean @${MAKE} -C support/refc clean
@${MAKE} -C support/chez clean
clean-libs: clean-libs:
${MAKE} -C libs/prelude clean ${MAKE} -C libs/prelude clean
@ -151,18 +160,17 @@ endif
install ${TARGETDIR}/${NAME}_app/* ${PREFIX}/bin/${NAME}_app install ${TARGETDIR}/${NAME}_app/* ${PREFIX}/bin/${NAME}_app
install-support: install-support:
mkdir -p ${PREFIX}/${NAME_VERSION}/support/chez
mkdir -p ${PREFIX}/${NAME_VERSION}/support/docs mkdir -p ${PREFIX}/${NAME_VERSION}/support/docs
mkdir -p ${PREFIX}/${NAME_VERSION}/support/racket mkdir -p ${PREFIX}/${NAME_VERSION}/support/racket
mkdir -p ${PREFIX}/${NAME_VERSION}/support/gambit mkdir -p ${PREFIX}/${NAME_VERSION}/support/gambit
mkdir -p ${PREFIX}/${NAME_VERSION}/support/js mkdir -p ${PREFIX}/${NAME_VERSION}/support/js
install support/chez/* ${PREFIX}/${NAME_VERSION}/support/chez
install support/docs/* ${PREFIX}/${NAME_VERSION}/support/docs install support/docs/* ${PREFIX}/${NAME_VERSION}/support/docs
install support/racket/* ${PREFIX}/${NAME_VERSION}/support/racket install support/racket/* ${PREFIX}/${NAME_VERSION}/support/racket
install support/gambit/* ${PREFIX}/${NAME_VERSION}/support/gambit install support/gambit/* ${PREFIX}/${NAME_VERSION}/support/gambit
install support/js/* ${PREFIX}/${NAME_VERSION}/support/js install support/js/* ${PREFIX}/${NAME_VERSION}/support/js
@${MAKE} -C support/c install @${MAKE} -C support/c install
@${MAKE} -C support/refc install @${MAKE} -C support/refc install
@${MAKE} -C support/chez install
install-libs: install-libs:
${MAKE} -C libs/prelude install IDRIS2?=${TARGET} IDRIS2_PATH=${IDRIS2_BOOT_PATH} ${MAKE} -C libs/prelude install IDRIS2?=${TARGET} IDRIS2_PATH=${IDRIS2_BOOT_PATH}

View File

@ -203,7 +203,17 @@ The following keys are available:
``implicit`` ``implicit``
provides a Boolean value that is True if the region is the name of an implicit argument provides a Boolean value that is True if the region is the name of an implicit argument
``decor`` ``decor``
describes the category of a token, which can be ``type``, ``function``, ``data``, ``keyword``, or ``bound``. describes the category of a token, which can be:
``type``: type constructors
``function``: defined functions
``data``: data constructors
``bound``: bound variables, or
``keyword``
``source-loc`` ``source-loc``
states that the region refers to a source code location. Its body is a collection of key-value pairs, with the following possibilities: states that the region refers to a source code location. Its body is a collection of key-value pairs, with the following possibilities:

View File

@ -70,3 +70,32 @@ Here are the specifics for the conversion:
case k of case k of
0 => zexp 0 => zexp
_ => let k' = k - 1 in sexp _ => let k' = k - 1 in sexp
``%builtin NaturalToInteger``
=============================
The ``%builtin NaturalToInteger`` pragma allows O(1) conversion of naturals to ``Integer`` s.
For example
.. code-block:: idris
natToInteger : Nat -> Integer
natToInteger Z = 0
natToInteger (S k) = 1 + natToInteger k
%builtin NaturalToInteger natToInteger
Please note, this only checks the type, not that the type is correct.
This can be used with ``%transform`` to allow many other operations to be O(1) too.
.. code-block:: idris
eqNat : Nat -> Nat -> Bool
eqNat Z Z = True
eqNat (S j) (S k) = eqNat j k
eqNat _ _ = False
%transform "eqNat" eqNat j k = natToInteger j == natToInteger k
For now, any ``NaturalToInteger`` function must have exactly 1 non-erased argument, which must be a natural.

View File

@ -13,6 +13,7 @@ modules =
Compiler.CompileExpr, Compiler.CompileExpr,
Compiler.Inline, Compiler.Inline,
Compiler.LambdaLift, Compiler.LambdaLift,
Compiler.Separate,
Compiler.VMCode, Compiler.VMCode,
Compiler.ES.ES, Compiler.ES.ES,
@ -27,6 +28,7 @@ modules =
Compiler.RefC.RefC, Compiler.RefC.RefC,
Compiler.Scheme.Chez, Compiler.Scheme.Chez,
Compiler.Scheme.ChezSep,
Compiler.Scheme.Racket, Compiler.Scheme.Racket,
Compiler.Scheme.Gambit, Compiler.Scheme.Gambit,
Compiler.Scheme.Common, Compiler.Scheme.Common,
@ -169,7 +171,6 @@ modules =
Parser.Lexer.Package, Parser.Lexer.Package,
Parser.Lexer.Source, Parser.Lexer.Source,
Parser.Rule.Common,
Parser.Rule.Package, Parser.Rule.Package,
Parser.Rule.Source, Parser.Rule.Source,

View File

@ -229,7 +229,7 @@ namespace Equality
||| The actual proof used by cast is irrelevant ||| The actual proof used by cast is irrelevant
export export
congCast : {n, q : Nat} -> {k : Fin m} -> {l : Fin p} -> congCast : {0 n, q : Nat} -> {k : Fin m} -> {l : Fin p} ->
{0 eq1 : m = n} -> {0 eq2 : p = q} -> {0 eq1 : m = n} -> {0 eq2 : p = q} ->
k ~~~ l -> k ~~~ l ->
cast eq1 k ~~~ cast eq2 l cast eq1 k ~~~ cast eq2 l
@ -237,9 +237,8 @@ namespace Equality
||| Last is congruent wrt index equality ||| Last is congruent wrt index equality
export export
congLast : {m, n : Nat} -> (0 _ : m = n) -> last {n=m} ~~~ last {n} congLast : {m : Nat} -> (0 _ : m = n) -> last {n=m} ~~~ last {n}
congLast {m = Z} {n = Z} eq = reflexive congLast Refl = reflexive
congLast {m = S _} {n = S _} eq = FS (congLast (succInjective _ _ eq))
export export
congShift : (m : Nat) -> k ~~~ l -> shift m k ~~~ shift m l congShift : (m : Nat) -> k ~~~ l -> shift m k ~~~ shift m l
@ -263,7 +262,7 @@ namespace Equality
export export
hetPointwiseIsTransport : hetPointwiseIsTransport :
{0 k : Fin m} -> {0 l : Fin n} -> {0 k : Fin m} -> {0 l : Fin n} ->
(eq : m === n) -> k ~~~ l -> k === rewrite eq in l (0 eq : m === n) -> k ~~~ l -> k === rewrite eq in l
hetPointwiseIsTransport Refl = homoPointwiseIsEqual hetPointwiseIsTransport Refl = homoPointwiseIsEqual
export export

View File

@ -459,11 +459,6 @@ toList1' : (l : List a) -> Maybe (List1 a)
toList1' [] = Nothing toList1' [] = Nothing
toList1' (x :: xs) = Just (x ::: xs) toList1' (x :: xs) = Just (x ::: xs)
||| Convert any Foldable structure to a list.
public export
toList : Foldable t => t a -> List a
toList = foldr (::) []
||| Prefix every element in the list with the given element ||| Prefix every element in the list with the given element
||| |||
||| ```idris example ||| ```idris example
@ -631,11 +626,11 @@ transpose (heads :: tails) = spreadHeads heads (transpose tails) where
-------------------------------------------------------------------------------- --------------------------------------------------------------------------------
export export
Uninhabited ([] = Prelude.(::) x xs) where Uninhabited ([] = x :: xs) where
uninhabited Refl impossible uninhabited Refl impossible
export export
Uninhabited (Prelude.(::) x xs = []) where Uninhabited (x :: xs = []) where
uninhabited Refl impossible uninhabited Refl impossible
||| (::) is injective ||| (::) is injective

View File

@ -130,6 +130,7 @@ export
Foldable List1 where Foldable List1 where
foldr c n (x ::: xs) = c x (foldr c n xs) foldr c n (x ::: xs) = c x (foldr c n xs)
null _ = False null _ = False
toList = forget
export export
Traversable List1 where Traversable List1 where

View File

@ -14,8 +14,8 @@ FilePos = (Int, Int)
-- the second 'FilePos' indicates the start of the next term. -- the second 'FilePos' indicates the start of the next term.
public export public export
data FC : Type where data FC : Type where
MkFC : String -> FilePos -> FilePos -> FC MkFC : String -> FilePos -> FilePos -> FC
EmptyFC : FC EmptyFC : FC
public export public export
emptyFC : FC emptyFC : FC

View File

@ -13,23 +13,26 @@ infixr 2 \|/
public export public export
interface Category arr => Arrow (0 arr : Type -> Type -> Type) where interface Category arr => Arrow (0 arr : Type -> Type -> Type) where
||| Converts a function from input to output into a arrow computation.
arrow : (a -> b) -> arr a b arrow : (a -> b) -> arr a b
||| Converts an arrow from `a` to `b` into an arrow on pairs, that applies
||| its argument to the first component and leaves the second component
||| untouched, thus saving its value across a computation.
first : arr a b -> arr (a, c) (b, c) first : arr a b -> arr (a, c) (b, c)
||| Converts an arrow from `a` to `b` into an arrow on pairs, that applies
||| its argument to the second component and leaves the first component
||| untouched, thus saving its value across a computation.
second : arr a b -> arr (c, a) (c, b) second : arr a b -> arr (c, a) (c, b)
second f = arrow {arr = arr} swap >>> first f >>> arrow {arr = arr} swap second f = arrow {arr = arr} swap >>> first f >>> arrow {arr = arr} swap
where where
swap : (x, y) -> (y, x) swap : (x, y) -> (y, x)
swap (a, b) = (b, a) swap (a, b) = (b, a)
||| A combinator which processes both components of a pair.
(***) : arr a b -> arr a' b' -> arr (a, a') (b, b') (***) : arr a b -> arr a' b' -> arr (a, a') (b, b')
f *** g = first f >>> second g f *** g = first f >>> second g
||| A combinator which builds a pair from the results of two arrows.
(&&&) : arr a b -> arr a b' -> arr a (b, b') (&&&) : arr a b -> arr a b' -> arr a (b, b')
f &&& g = arrow dup >>> f *** g f &&& g = arrow dup >>> f *** g
where
dup : x -> (x,x)
dup x = (x,x)
public export public export
implementation Arrow Morphism where implementation Arrow Morphism where
@ -122,3 +125,8 @@ implementation ArrowApply a => Monad (ArrowMonad a) where
public export public export
interface Arrow arr => ArrowLoop (0 arr : Type -> Type -> Type) where interface Arrow arr => ArrowLoop (0 arr : Type -> Type -> Type) where
loop : arr (Pair a c) (Pair b c) -> arr a b loop : arr (Pair a c) (Pair b c) -> arr a b
||| Applying a binary operator to the results of two arrow computations.
public export
liftA2 : Arrow arr => (a -> b -> c) -> arr d a -> arr d b -> arr d c
liftA2 op f g = (f &&& g) >>> arrow (\(a, b) => a `op` b)

View File

@ -213,7 +213,7 @@ finToNatMultHomo {m = S _} (FS x) y = Calc $
~~ finToNat (y + x * y) ...( Refl ) ~~ finToNat (y + x * y) ...( Refl )
~~ finToNat y + finToNat (x * y) ...( finToNatPlusHomo y (x * y) ) ~~ finToNat y + finToNat (x * y) ...( finToNatPlusHomo y (x * y) )
~~ finToNat y + finToNat x * finToNat y ...( cong (finToNat y +) (finToNatMultHomo x y) ) ~~ finToNat y + finToNat x * finToNat y ...( cong (finToNat y +) (finToNatMultHomo x y) )
~~ finToNat (FS x) * finToNat y ...( Refl) ~~ finToNat (FS x) * finToNat y ...( Refl )
-- Relations to `Fin`'s `last` -- Relations to `Fin`'s `last`
@ -292,7 +292,7 @@ plusZeroRightNeutral (FS k) = FS (plusZeroRightNeutral k)
export export
congPlusRight : {m, n, p : Nat} -> {k : Fin (S n)} -> {l : Fin (S p)} -> congPlusRight : {m, n, p : Nat} -> {k : Fin (S n)} -> {l : Fin (S p)} ->
(c : Fin m) -> k ~~~ l -> c + k ~~~ c + l (c : Fin m) -> k ~~~ l -> c + k ~~~ c + l
congPlusRight c FZ congPlusRight c FZ
= transitive (plusZeroRightNeutral c) = transitive (plusZeroRightNeutral c)
(symmetric $ plusZeroRightNeutral c) (symmetric $ plusZeroRightNeutral c)

View File

@ -1,5 +1,7 @@
module Data.SortedMap module Data.SortedMap
%hide Prelude.toList
-- TODO: write split -- TODO: write split
private private

View File

@ -3,6 +3,8 @@ module Data.SortedSet
import Data.Maybe import Data.Maybe
import Data.SortedMap import Data.SortedMap
%hide Prelude.toList
export export
data SortedSet k = SetWrapper (Data.SortedMap.SortedMap k ()) data SortedSet k = SetWrapper (Data.SortedMap.SortedMap k ())

View File

@ -154,3 +154,18 @@ public export
intToBool : Int -> Bool intToBool : Int -> Bool
intToBool 0 = False intToBool 0 = False
intToBool x = True intToBool x = True
--------------
-- LISTS --
--------------
||| Generic lists.
public export
data List a =
||| Empty list
Nil
| ||| A non-empty list, consisting of a head element and the rest of the list.
(::) a (List a)
%name List xs, ys, zs

View File

@ -222,12 +222,17 @@ public export
guard : Alternative f => Bool -> f () guard : Alternative f => Bool -> f ()
guard x = if x then pure () else empty guard x = if x then pure () else empty
||| Conditionally execute an applicative expression. ||| Conditionally execute an applicative expression when the boolean is true.
public export public export
when : Applicative f => Bool -> Lazy (f ()) -> f () when : Applicative f => Bool -> Lazy (f ()) -> f ()
when True f = f when True f = f
when False f = pure () when False f = pure ()
||| Execute an applicative expression unless the boolean is true.
%inline public export
unless : Applicative f => Bool -> Lazy (f ()) -> f ()
unless = when . not
--------------------------- ---------------------------
-- FOLDABLE, TRAVERSABLE -- -- FOLDABLE, TRAVERSABLE --
--------------------------- ---------------------------
@ -266,6 +271,10 @@ interface Foldable t where
foldlM : Monad m => (funcM : acc -> elem -> m acc) -> (init : acc) -> (input : t elem) -> m acc foldlM : Monad m => (funcM : acc -> elem -> m acc) -> (init : acc) -> (input : t elem) -> m acc
foldlM fm a0 = foldl (\ma, b => ma >>= flip fm b) (pure a0) foldlM fm a0 = foldl (\ma, b => ma >>= flip fm b) (pure a0)
||| Produce a list of the elements contained in the parametrised type.
toList : t elem -> List elem
toList = foldr (::) []
||| Maps each element to a value and combine them ||| Maps each element to a value and combine them
public export public export
foldMap : (Foldable t, Monoid m) => (a -> m) -> t a -> m foldMap : (Foldable t, Monoid m) => (a -> m) -> t a -> m

View File

@ -157,6 +157,12 @@ maybe : Lazy b -> Lazy (a -> b) -> Maybe a -> b
maybe n j Nothing = n maybe n j Nothing = n
maybe n j (Just x) = j x maybe n j (Just x) = j x
||| Execute an applicative expression when the Maybe is Just
%inline public export
whenJust : Applicative f => Maybe a -> (a -> f ()) -> f ()
whenJust (Just a) k = k a
whenJust Nothing k = pure ()
public export public export
Eq a => Eq (Maybe a) where Eq a => Eq (Maybe a) where
Nothing == Nothing = True Nothing == Nothing = True
@ -331,17 +337,6 @@ Traversable (Either e) where
-- LISTS -- -- LISTS --
----------- -----------
||| Generic lists.
public export
data List a =
||| Empty list
Nil
| ||| A non-empty list, consisting of a head element and the rest of the list.
(::) a (List a)
%name List xs, ys, zs
public export public export
Eq a => Eq (List a) where Eq a => Eq (List a) where
[] == [] = True [] == [] = True
@ -393,6 +388,8 @@ Foldable List where
null [] = True null [] = True
null (_::_) = False null (_::_) = False
toList = id
public export public export
Applicative List where Applicative List where
pure x = [x] pure x = [x]

View File

@ -43,9 +43,11 @@
||| + `idris2` The path of the executable we are testing. ||| + `idris2` The path of the executable we are testing.
||| + `codegen` The backend to use for code generation. ||| + `codegen` The backend to use for code generation.
||| + `onlyNames` The tests to run relative to the generated executable. ||| + `onlyNames` The tests to run relative to the generated executable.
||| + `onlyFile` The file listing the tests to run relative to the generated executable.
||| + `interactive` Whether to offer to update the expected file or not. ||| + `interactive` Whether to offer to update the expected file or not.
||| + `timing` Whether to display time taken for each test. ||| + `timing` Whether to display time taken for each test.
||| + `threads` The maximum numbers to use (default: number of cores). ||| + `threads` The maximum numbers to use (default: number of cores).
||| + `failureFile` The file in which to write the list of failing tests.
||| |||
||| We provide an options parser (`options`) that takes the list of command line ||| We provide an options parser (`options`) that takes the list of command line
||| arguments and constructs this for you. ||| arguments and constructs this for you.
@ -62,6 +64,7 @@
module Test.Golden module Test.Golden
import Data.Either
import Data.Maybe import Data.Maybe
import Data.List import Data.List
import Data.List1 import Data.List1
@ -93,6 +96,8 @@ record Options where
timing : Bool timing : Bool
||| How many threads should we use? ||| How many threads should we use?
threads : Nat threads : Nat
||| Should we write the list of failing cases from a file?
failureFile : Maybe String
export export
initOptions : String -> Options initOptions : String -> Options
@ -103,6 +108,7 @@ initOptions exe
False False
False False
1 1
Nothing
export export
usage : String -> String usage : String -> String
@ -113,37 +119,48 @@ usage exe = unwords
, "[--interactive]" , "[--interactive]"
, "[--cg CODEGEN]" , "[--cg CODEGEN]"
, "[--threads N]" , "[--threads N]"
, "[--failure-file PATH]"
, "[--only-file PATH]"
, "[--only [NAMES]]" , "[--only [NAMES]]"
] ]
||| Process the command line options.
export export
options : List String -> Maybe Options fail : String -> IO a
options args = case args of
(_ :: exe :: rest) => go rest (initOptions exe)
_ => Nothing
where
go : List String -> Options -> Maybe Options
go rest opts = case rest of
[] => pure opts
("--timing" :: xs) => go xs (record { timing = True} opts)
("--interactive" :: xs) => go xs (record { interactive = True } opts)
("--cg" :: cg :: xs) => go xs (record { codegen = Just cg } opts)
("--threads" :: n :: xs) => do let pos : Nat = !(parsePositive n)
go xs (record { threads = pos } opts)
("--only" :: xs) => pure $ record { onlyNames = xs } opts
_ => Nothing
-- [ Core ]
export
fail : String -> IO ()
fail err fail err
= do putStrLn err = do putStrLn err
exitWith (ExitFailure 1) exitWith (ExitFailure 1)
||| Process the command line options.
export
options : List String -> IO (Maybe Options)
options args = case args of
(_ :: exe :: rest) => mkOptions exe rest
_ => pure Nothing
where
go : List String -> Maybe String -> Options -> Maybe (Maybe String, Options)
go rest only opts = case rest of
[] => pure (only, opts)
("--timing" :: xs) => go xs only (record { timing = True} opts)
("--interactive" :: xs) => go xs only (record { interactive = True } opts)
("--cg" :: cg :: xs) => go xs only (record { codegen = Just cg } opts)
("--threads" :: n :: xs) => do let pos : Nat = !(parsePositive n)
go xs only (record { threads = pos } opts)
("--failure-file" :: p :: xs) => go xs only (record { failureFile = Just p } opts)
("--only" :: xs) => pure (only, record { onlyNames = xs } opts)
("--only-file" :: p :: xs) => go xs (Just p) opts
_ => Nothing
mkOptions : String -> List String -> IO (Maybe Options)
mkOptions exe rest
= do let Just (mfp, opts) = go rest Nothing (initOptions exe)
| Nothing => pure Nothing
let Just fp = mfp
| Nothing => pure (Just opts)
Right only <- readFile fp
| Left err => fail (show err)
pure $ Just $ record { onlyNames $= (forget (lines only) ++) } opts
||| Normalise strings between different OS. ||| Normalise strings between different OS.
||| |||
@ -156,13 +173,18 @@ normalize str =
then pack $ filter (\ch => ch /= '/' && ch /= '\\') (unpack str) then pack $ filter (\ch => ch /= '/' && ch /= '\\') (unpack str)
else str else str
||| The result of a test run
||| `Left` corresponds to a failure, and `Right` to a success
Result : Type
Result = Either String String
||| Run the specified Golden test with the supplied options. ||| Run the specified Golden test with the supplied options.
||| |||
||| See the module documentation for more information. ||| See the module documentation for more information.
||| |||
||| @testPath the directory that contains the test. ||| @testPath the directory that contains the test.
export export
runTest : Options -> String -> IO (Future Bool) runTest : Options -> String -> IO (Future Result)
runTest opts testPath = forkIO $ do runTest opts testPath = forkIO $ do
start <- clockTime Thread start <- clockTime Thread
let cg = case codegen opts of let cg = case codegen opts of
@ -174,16 +196,16 @@ runTest opts testPath = forkIO $ do
Right out <- readFile $ testPath ++ "/output" Right out <- readFile $ testPath ++ "/output"
| Left err => do print err | Left err => do print err
pure False pure (Left testPath)
Right exp <- readFile $ testPath ++ "/expected" Right exp <- readFile $ testPath ++ "/expected"
| Left FileNotFound => do | Left FileNotFound => do
if interactive opts if interactive opts
then mayOverwrite Nothing out then mayOverwrite Nothing out
else print FileNotFound else print FileNotFound
pure False pure (Left testPath)
| Left err => do print err | Left err => do print err
pure False pure (Left testPath)
let result = normalize out == normalize exp let result = normalize out == normalize exp
let time = timeDifference end start let time = timeDifference end start
@ -196,7 +218,7 @@ runTest opts testPath = forkIO $ do
then mayOverwrite (Just exp) out then mayOverwrite (Just exp) out
else putStrLn . unlines $ expVsOut exp out else putStrLn . unlines $ expVsOut exp out
pure result pure $ if result then Right testPath else Left testPath
where where
getAnswer : IO Bool getAnswer : IO Bool
@ -298,6 +320,7 @@ findCG
public export public export
record TestPool where record TestPool where
constructor MkTestPool constructor MkTestPool
poolName : String
constraints : List Requirement constraints : List Requirement
testCases : List String testCases : List String
@ -308,14 +331,43 @@ filterTests opts = case onlyNames opts of
[] => id [] => id
xs => filter (\ name => any (`isInfixOf` name) xs) xs => filter (\ name => any (`isInfixOf` name) xs)
||| The summary of a test pool run
public export
record Summary where
constructor MkSummary
success : List String
failure : List String
export
initSummary : Summary
initSummary = MkSummary [] []
export
updateSummary : List Result -> Summary -> Summary
updateSummary res =
let (ls, ws) = partitionEithers res in
{ success $= (ws ++)
, failure $= (ls ++)
}
export
Semigroup Summary where
MkSummary ws1 ls1 <+> MkSummary ws2 ls2
= MkSummary (ws1 ++ ws2) (ls1 ++ ls2)
export
Monoid Summary where
neutral = initSummary
||| A runner for a test pool ||| A runner for a test pool
export export
poolRunner : Options -> TestPool -> IO (List Bool) poolRunner : Options -> TestPool -> IO Summary
poolRunner opts pool poolRunner opts pool
= do -- check that we indeed want to run some of these tests = do -- check that we indeed want to run some of these tests
let tests = filterTests opts (testCases pool) let tests = filterTests opts (testCases pool)
let (_ :: _) = tests let (_ :: _) = tests
| [] => pure [] | [] => pure initSummary
putStrLn banner
-- if so make sure the constraints are satisfied -- if so make sure the constraints are satisfied
cs <- for (constraints pool) $ \ req => do cs <- for (constraints pool) $ \ req => do
mfp <- checkRequirement req mfp <- checkRequirement req
@ -324,17 +376,23 @@ poolRunner opts pool
Just fp => "Found " ++ show req ++ " at " ++ fp Just fp => "Found " ++ show req ++ " at " ++ fp
pure mfp pure mfp
let Just _ = the (Maybe (List String)) (sequence cs) let Just _ = the (Maybe (List String)) (sequence cs)
| Nothing => pure [] | Nothing => pure initSummary
-- if so run them all! -- if so run them all!
loop [] tests loop initSummary tests
where where
loop : List (List Bool) -> List String -> IO (List Bool)
loop acc [] = pure (concat $ reverse acc) banner : String
banner =
let separator = fastPack $ replicate 72 '-' in
fastUnlines [ "", separator, pool.poolName, separator, "" ]
loop : Summary -> List String -> IO Summary
loop acc [] = pure acc
loop acc tests loop acc tests
= do let (now, later) = splitAt opts.threads tests = do let (now, later) = splitAt opts.threads tests
bs <- map await <$> traverse (runTest opts) now bs <- map await <$> traverse (runTest opts) now
loop (bs :: acc) later loop (updateSummary bs acc) later
||| A runner for a whole test suite ||| A runner for a whole test suite
@ -342,19 +400,36 @@ export
runner : List TestPool -> IO () runner : List TestPool -> IO ()
runner tests runner tests
= do args <- getArgs = do args <- getArgs
let (Just opts) = options args Just opts <- options args
| _ => do print args | _ => do print args
putStrLn (usage "runtests") putStrLn (usage "runtests")
-- if no CG has been set, find a sensible default based on what is available -- if no CG has been set, find a sensible default based on what is available
opts <- case codegen opts of opts <- case codegen opts of
Nothing => pure $ record { codegen = !findCG } opts Nothing => pure $ record { codegen = !findCG } opts
Just _ => pure opts Just _ => pure opts
-- run the tests -- run the tests
res <- concat <$> traverse (poolRunner opts) tests res <- concat <$> traverse (poolRunner opts) tests
putStrLn (show (length (filter id res)) ++ "/" ++ show (length res)
++ " tests successful") -- report the result
if (any not res) let nsucc = length res.success
then exitWith (ExitFailure 1) let nfail = length res.failure
else exitWith ExitSuccess let ntotal = nsucc + nfail
putStrLn (show nsucc ++ "/" ++ show ntotal ++ " tests successful")
-- deal with failures
let list = fastUnlines res.failure
when (nfail > 0) $
do putStrLn "Failing tests:"
putStrLn list
-- always overwrite the failure file, if it is given
whenJust opts.failureFile $ \ path =>
do Right _ <- writeFile path list
| Left err => fail (show err)
pure ()
-- exit
if nfail == 0
then exitWith ExitSuccess
else exitWith (ExitFailure 1)
-- [ EOF ] -- [ EOF ]

View File

@ -342,23 +342,20 @@ exists f
| Left err => pure False | Left err => pure False
closeFile ok closeFile ok
pure True pure True
-- Select the most preferred target from an ordered list of choices and
-- Parse a calling convention into a backend/target for the call, and -- parse the calling convention into a backend/target for the call, and
-- a comma separated list of any other location data. -- a comma separated list of any other location data. For example
-- the chez backend would supply ["scheme,chez", "scheme", "C"]. For a function with
-- more than one string, a string with "scheme" would be preferred over one
-- with "C" and "scheme,chez" would be preferred to both.
-- e.g. "scheme:display" - call the scheme function 'display' -- e.g. "scheme:display" - call the scheme function 'display'
-- "C:puts,libc,stdio.h" - call the C function 'puts' which is in -- "C:puts,libc,stdio.h" - call the C function 'puts' which is in
-- the library libc and the header stdio.h -- the library libc and the header stdio.h
-- Returns Nothing if the string is empty (which a backend can interpret -- Returns Nothing if there is no match.
-- however it likes)
export export
parseCC : String -> Maybe (String, List String) parseCC : List String -> List String -> Maybe (String, List String)
parseCC "" = Nothing parseCC [] _ = Nothing
parseCC str parseCC (target::ts) strs = findTarget target strs <|> parseCC ts strs
= case span (/= ':') str of
(target, "") => Just (trim target, [])
(target, opts) => Just (trim target,
map trim (getOpts
(assert_total (strTail opts))))
where where
getOpts : String -> List String getOpts : String -> List String
getOpts "" = [] getOpts "" = []
@ -366,6 +363,17 @@ parseCC str
= case span (/= ',') str of = case span (/= ',') str of
(opt, "") => [opt] (opt, "") => [opt]
(opt, rest) => opt :: getOpts (assert_total (strTail rest)) (opt, rest) => opt :: getOpts (assert_total (strTail rest))
hasTarget : String -> String -> Bool
hasTarget target str = case span (/= ':') str of
(targetSpec, _) => targetSpec == target
findTarget : String -> List String -> Maybe (String, List String)
findTarget target [] = Nothing
findTarget target (s::xs) = if hasTarget target s
then case span (/= ':') s of
(t, "") => Just (trim t, [])
(t, opts) => Just (trim t, map trim (getOpts
(assert_total (strTail opts))))
else findTarget target xs
export export
dylib_suffix : String dylib_suffix : String

View File

@ -198,8 +198,9 @@ builtinMagic : Ref Ctxt Defs => Core (forall vars. CExp vars -> CExp vars)
builtinMagic = do builtinMagic = do
defs <- get Ctxt defs <- get Ctxt
let b = defs.builtinTransforms let b = defs.builtinTransforms
let nats = concatMap builtinMagicNat $ values $ natTyNames b let nats = foldMap builtinMagicNat $ values $ natTyNames b
pure $ magic $ natHack ++ nats let natToInts = map natToIntMagic $ toList $ natToIntegerFns b
pure $ magic $ natHack ++ nats ++ natToInts
where where
builtinMagicNat : NatBuiltin -> List Magic builtinMagicNat : NatBuiltin -> List Magic
builtinMagicNat cons = builtinMagicNat cons =
@ -208,6 +209,10 @@ builtinMagic = do
, MagicCCon cons.succ 1 , MagicCCon cons.succ 1
(\ fc, [k] => CApp fc (CRef fc (UN "prim__add_Integer")) [CPrimVal fc (BI 1), k]) (\ fc, [k] => CApp fc (CRef fc (UN "prim__add_Integer")) [CPrimVal fc (BI 1), k])
] -- TODO: add builtin pragmas for Nat related functions (to/from Integer, add, mult, minus, compare) ] -- TODO: add builtin pragmas for Nat related functions (to/from Integer, add, mult, minus, compare)
natToIntMagic : (Name, NatToInt) -> Magic
natToIntMagic (fn, MkNatToInt arity natIdx) =
MagicCRef fn arity
(\ _, _, args => index natIdx args)
isNatCon : (zeroMap : NameMap ZERO) -> isNatCon : (zeroMap : NameMap ZERO) ->
(succMap : NameMap SUCC) -> (succMap : NameMap SUCC) ->

View File

@ -120,102 +120,172 @@ jsCrashExp message =
n <- addConstToPreamble "crashExp" "x=>{throw new IdrisError(x)}" n <- addConstToPreamble "crashExp" "x=>{throw new IdrisError(x)}"
pure $ n ++ "("++ message ++ ")" pure $ n ++ "("++ message ++ ")"
jsIntegerOfString : {auto c : Ref ESs ESSt} -> String -> Core String toBigInt : String -> String
jsIntegerOfString x = toBigInt e = "BigInt(" ++ e ++ ")"
fromBigInt : String -> String
fromBigInt e = "Number(" ++ e ++ ")"
useBigInt' : Int -> Bool
useBigInt' = (> 32)
useBigInt : IntKind -> Bool
useBigInt (Signed $ P x) = useBigInt' x
useBigInt (Signed Unlimited) = True
useBigInt (Unsigned x) = useBigInt' x
jsBigIntOfString : {auto c : Ref ESs ESSt} -> String -> Core String
jsBigIntOfString x =
do do
n <- addConstToPreamble "integerOfString" "s=>{const idx = s.indexOf('.'); return idx === -1 ? BigInt(s) : BigInt(s.slice(0, idx));}" n <- addConstToPreamble "bigIntOfString" "s=>{const idx = s.indexOf('.'); return idx === -1 ? BigInt(s) : BigInt(s.slice(0, idx));}"
pure $ n ++ "(" ++ x ++ ")" pure $ n ++ "(" ++ x ++ ")"
jsNumberOfString : {auto c : Ref ESs ESSt} -> String -> Core String
jsNumberOfString x = pure $ "parseFloat(" ++ x ++ ")"
jsIntOfString : {auto c : Ref ESs ESSt} -> IntKind -> String -> Core String
jsIntOfString k = if useBigInt k then jsBigIntOfString else jsNumberOfString
nSpaces : Nat -> String nSpaces : Nat -> String
nSpaces n = pack $ List.replicate n ' ' nSpaces n = pack $ List.replicate n ' '
binOp : String -> String -> String -> String binOp : String -> String -> String -> String
binOp o lhs rhs = "(" ++ lhs ++ " " ++ o ++ " " ++ rhs ++ ")" binOp o lhs rhs = "(" ++ lhs ++ " " ++ o ++ " " ++ rhs ++ ")"
toBigInt : String -> String adjInt : Int -> String -> String
toBigInt e = "BigInt(" ++ e ++ ")" adjInt bits = if useBigInt' bits then toBigInt else id
fromBigInt : String -> String toInt : IntKind -> String -> String
fromBigInt e = "Number(" ++ e ++ ")" toInt k = if useBigInt k then toBigInt else id
jsIntegerOfChar : String -> String fromInt : IntKind -> String -> String
jsIntegerOfChar s = toBigInt (s++ ".codePointAt(0)") fromInt k = if useBigInt k then fromBigInt else id
jsIntegerOfDouble : String -> String jsIntOfChar : IntKind -> String -> String
jsIntegerOfDouble s = toBigInt $ "Math.trunc(" ++ s ++ ")" jsIntOfChar k s = toInt k $ s ++ ".codePointAt(0)"
jsIntOfDouble : IntKind -> String -> String
jsIntOfDouble k s = toInt k $ "Math.trunc(" ++ s ++ ")"
jsAnyToString : String -> String jsAnyToString : String -> String
jsAnyToString s = "(''+" ++ s ++ ")" jsAnyToString s = "(''+" ++ s ++ ")"
-- Valid unicode code poing range is [0,1114111], therefore,
-- we calculate the remainder modulo 1114112 (= 17 * 2^16).
jsCharOfInt : {auto c : Ref ESs ESSt} -> IntKind -> String -> Core String
jsCharOfInt k e =
do fn <- addConstToPreamble
("truncToChar")
("x=>(x >= 0 && x <= 55295) || (x >= 57344 && x <= 1114111) ? x : 0")
pure $ "String.fromCodePoint(" ++ fn ++ "(" ++ fromInt k e ++ "))"
makeIntBound : {auto c : Ref ESs ESSt} -> Int -> Core String makeIntBound : {auto c : Ref ESs ESSt} ->
makeIntBound bits = addConstToPreamble ("int_bound_" ++ show bits) ("BigInt(2) ** BigInt("++ show bits ++") ") (isBigInt : Bool) -> Int -> Core String
makeIntBound isBigInt bits =
let f = if isBigInt then toBigInt else id
name = if isBigInt then "bigint_bound_" else "int_bound_"
in addConstToPreamble (name ++ show bits) (f "2" ++ " ** "++ f (show bits))
truncateInt : {auto c : Ref ESs ESSt} -> Int -> String -> Core String truncateIntWithBitMask : {auto c : Ref ESs ESSt} -> Int -> String -> Core String
truncateInt bits e = truncateIntWithBitMask bits e =
let bs = show bits let bs = show bits
in do f = adjInt bits
mn <- addConstToPreamble ("int_mask_neg_" ++ bs) in do ib <- makeIntBound (useBigInt' bits) bits
("BigInt(-1) << BigInt(" ++ bs ++")") mn <- addConstToPreamble ("int_mask_neg_" ++ bs) ("-" ++ ib)
mp <- addConstToPreamble ("int_mask_pos_" ++ bs) mp <- addConstToPreamble ("int_mask_pos_" ++ bs) (ib ++ " - " ++ f "1")
("(BigInt(1) << BigInt(" ++ bs ++")) - BigInt(1)")
pure $ concat {t = List} pure $ concat {t = List}
[ "((", mn, " & ", e, ") == BigInt(0) ? " [ "((", ib, " & ", e, ") == " ++ ib ++ " ? "
, "(", e, " & ", mp, ") : " , "(", e, " | ", mn, ") : "
, "(", e, " | ", mn, ")" , "(", e, " & ", mp, ")"
, ")" , ")"
] ]
-- Valid unicode code poing range is [0,1114111], therefore, -- We can't determine `isBigInt` from the given number of bits, since
-- we calculate the remainder modulo 1114112 (= 17 * 2^16). -- when casting from BigInt to Number we need to truncate the BigInt
truncChar : {auto c : Ref ESs ESSt} -> String -> Core String -- first, otherwise we might lose precision
truncChar e = boundedInt : {auto c : Ref ESs ESSt} ->
do fn <- addConstToPreamble ("truncToChar") ("x=>(x >= 0 && x <= 55295) || (x >= 57344 && x <= 1114111) ? x : 0") (isBigInt : Bool) -> Int -> String -> Core String
pure $ "String.fromCodePoint(" ++ fn ++ "(" ++ fromBigInt e ++ "))" boundedInt isBigInt bits e =
let name = if isBigInt then "truncToBigInt" else "truncToInt"
in do n <- makeIntBound isBigInt bits
fn <- addConstToPreamble
(name ++ show bits)
("x=>(x<(-" ++ n ++ ")||(x>=" ++ n ++ "))?x%" ++ n ++ ":x")
pure $ fn ++ "(" ++ e ++ ")"
boundedInt : {auto c : Ref ESs ESSt} -> Int -> String -> Core String boundedUInt : {auto c : Ref ESs ESSt} ->
boundedInt bits e = (isBigInt : Bool) -> Int -> String -> Core String
do boundedUInt isBigInt bits e =
n <- makeIntBound bits let name = if isBigInt then "truncToUBigInt" else "truncToUInt"
fn <- addConstToPreamble ("truncToInt"++show bits) ("x=>(x<(-" ++ n ++ ")||(x>=" ++ n ++ "))?x%" ++ n ++ ":x") in do n <- makeIntBound isBigInt bits
pure $ fn ++ "(" ++ e ++ ")" fn <- addConstToPreamble
(name ++ show bits)
("x=>{const m = x%" ++ n ++ ";return m>=0?m:m+" ++ n ++ "}")
pure $ fn ++ "(" ++ e ++ ")"
boundedUInt : {auto c : Ref ESs ESSt} -> Int -> String -> Core String boundedIntOp : {auto c : Ref ESs ESSt} ->
boundedUInt bits e = Int -> String -> String -> String -> Core String
do boundedIntOp bits o lhs rhs =
n <- makeIntBound bits boundedInt (useBigInt' bits) bits (binOp o lhs rhs)
fn <- addConstToPreamble ("truncToUInt"++show bits) ("x=>{const m = x%" ++ n ++ ";return m>=0?m:m+" ++ n ++ "}")
pure $ fn ++ "(" ++ e ++ ")"
boundedIntOp : {auto c : Ref ESs ESSt} -> Int -> String -> String -> String -> Core String boundedIntBitOp : {auto c : Ref ESs ESSt} ->
boundedIntOp bits o lhs rhs = boundedInt bits (binOp o lhs rhs) Int -> String -> String -> String -> Core String
boundedIntBitOp bits o lhs rhs = truncateIntWithBitMask bits (binOp o lhs rhs)
boundedIntBitOp : {auto c : Ref ESs ESSt} -> Int -> String -> String -> String -> Core String boundedUIntOp : {auto c : Ref ESs ESSt} ->
boundedIntBitOp bits o lhs rhs = truncateInt bits (binOp o lhs rhs) Int -> String -> String -> String -> Core String
boundedUIntOp bits o lhs rhs =
boundedUIntOp : {auto c : Ref ESs ESSt} -> Int -> String -> String -> String -> Core String boundedUInt (useBigInt' bits) bits (binOp o lhs rhs)
boundedUIntOp bits o lhs rhs = boundedUInt bits (binOp o lhs rhs)
boolOp : String -> String -> String -> String boolOp : String -> String -> String -> String
boolOp o lhs rhs = "(" ++ binOp o lhs rhs ++ " ? BigInt(1) : BigInt(0))" boolOp o lhs rhs = "(" ++ binOp o lhs rhs ++ " ? BigInt(1) : BigInt(0))"
jsConstant : {auto c : Ref ESs ESSt} -> Constant -> Core String jsConstant : {auto c : Ref ESs ESSt} -> Constant -> Core String
jsConstant (I i) = pure $ show i ++ "n" jsConstant (I i) = pure $ show i ++ "n"
jsConstant (I8 i) = pure $ show i ++ "n" jsConstant (I8 i) = pure $ show i
jsConstant (I16 i) = pure $ show i ++ "n" jsConstant (I16 i) = pure $ show i
jsConstant (I32 i) = pure $ show i ++ "n" jsConstant (I32 i) = pure $ show i
jsConstant (I64 i) = pure $ show i ++ "n" jsConstant (I64 i) = pure $ show i ++ "n"
jsConstant (BI i) = pure $ show i ++ "n" jsConstant (BI i) = pure $ show i ++ "n"
jsConstant (Str s) = pure $ jsString s jsConstant (Str s) = pure $ jsString s
jsConstant (Ch c) = pure $ jsString $ Data.Strings.singleton c jsConstant (Ch c) = pure $ jsString $ Data.Strings.singleton c
jsConstant (Db f) = pure $ show f jsConstant (Db f) = pure $ show f
jsConstant WorldVal = addConstToPreamble "idrisworld" "Symbol('idrisworld')"; jsConstant WorldVal = addConstToPreamble "idrisworld" "Symbol('idrisworld')";
jsConstant (B8 i) = pure $ show i ++ "n" jsConstant (B8 i) = pure $ show i
jsConstant (B16 i) = pure $ show i ++ "n" jsConstant (B16 i) = pure $ show i
jsConstant (B32 i) = pure $ show i ++ "n" jsConstant (B32 i) = pure $ show i
jsConstant (B64 i) = pure $ show i ++ "n" jsConstant (B64 i) = pure $ show i ++ "n"
jsConstant ty = throw (InternalError $ "Unsuported constant " ++ show ty) jsConstant ty = throw (InternalError $ "Unsuported constant " ++ show ty)
-- For multiplication of 32bit integers (signed or unsigned),
-- we go via BigInt and back, otherwise the calculation is
-- susceptible to rounding error, since we might exceed `MAX_SAFE_INTEGER`.
mult : {auto c : Ref ESs ESSt}
-> Maybe IntKind
-> (x : String)
-> (y : String)
-> Core String
mult (Just $ Signed $ P 32) x y =
fromBigInt <$> boundedInt True 31 (binOp "*" (toBigInt x) (toBigInt y))
mult (Just $ Unsigned 32) x y =
fromBigInt <$> boundedUInt True 32 (binOp "*" (toBigInt x) (toBigInt y))
mult (Just $ Signed $ P n) x y = boundedIntOp (n-1) "*" x y
mult (Just $ Unsigned n) x y = boundedUIntOp n "*" x y
mult _ x y = pure $ binOp "*" x y
div : {auto c : Ref ESs ESSt}
-> Maybe IntKind
-> (x : String)
-> (y : String)
-> Core String
div (Just k) x y =
if useBigInt k then pure $ binOp "/" x y
else pure $ jsIntOfDouble k (x ++ " / " ++ y)
div Nothing x y = pure $ binOp "/" x y
-- Creates the definition of a binary arithmetic operation. -- Creates the definition of a binary arithmetic operation.
-- Rounding / truncation behavior is determined from the -- Rounding / truncation behavior is determined from the
-- `IntKind`. -- `IntKind`.
@ -231,6 +301,9 @@ arithOp _ op x y = pure $ binOp op x y
-- Same as `arithOp` but for bitwise operations that might -- Same as `arithOp` but for bitwise operations that might
-- go out of the valid range. -- go out of the valid range.
-- Note: Bitwise operations on `Number` work correctly for
-- 32bit *signed* integers. For `Bits32` we therefore go via `BigInt`
-- in order not having to deal with all kinds of negativity nastiness.
bitOp : {auto c : Ref ESs ESSt} bitOp : {auto c : Ref ESs ESSt}
-> Maybe IntKind -> Maybe IntKind
-> (op : String) -> (op : String)
@ -238,43 +311,75 @@ bitOp : {auto c : Ref ESs ESSt}
-> (y : String) -> (y : String)
-> Core String -> Core String
bitOp (Just $ Signed $ P n) op x y = boundedIntBitOp (n-1) op x y bitOp (Just $ Signed $ P n) op x y = boundedIntBitOp (n-1) op x y
bitOp (Just $ Unsigned 32) op x y =
fromBigInt <$> boundedUInt True 32 (binOp op (toBigInt x) (toBigInt y))
bitOp (Just $ Unsigned n) op x y = boundedUIntOp n op x y bitOp (Just $ Unsigned n) op x y = boundedUIntOp n op x y
bitOp _ op x y = pure $ binOp op x y bitOp _ op x y = pure $ binOp op x y
constPrimitives : {auto c : Ref ESs ESSt} -> ConstantPrimitives constPrimitives : {auto c : Ref ESs ESSt} -> ConstantPrimitives
constPrimitives = MkConstantPrimitives { constPrimitives = MkConstantPrimitives {
charToInt = \k => truncInt k . jsIntegerOfChar charToInt = \k => truncInt (useBigInt k) k . jsIntOfChar k
, intToChar = \_ => truncChar , intToChar = \k => jsCharOfInt k
, stringToInt = \k,s => jsIntegerOfString s >>= truncInt k , stringToInt = \k,s => jsIntOfString k s >>= truncInt (useBigInt k) k
, intToString = \_ => pure . jsAnyToString , intToString = \_ => pure . jsAnyToString
, doubleToInt = \k => truncInt k . jsIntegerOfDouble , doubleToInt = \k => truncInt (useBigInt k) k . jsIntOfDouble k
, intToDouble = \_ => pure . fromBigInt , intToDouble = \k => pure . fromInt k
, intToInt = intImpl , intToInt = intImpl
} }
where truncInt : IntKind -> String -> Core String where truncInt : (isBigInt : Bool) -> IntKind -> String -> Core String
truncInt (Signed Unlimited) = pure truncInt b (Signed Unlimited) = pure
truncInt (Signed $ P n) = boundedInt (n-1) truncInt b (Signed $ P n) = boundedInt b (n-1)
truncInt (Unsigned n) = boundedUInt n truncInt b (Unsigned n) = boundedUInt b n
shrink : IntKind -> IntKind -> String -> String
shrink k1 k2 = case (useBigInt k1, useBigInt k2) of
(True, False) => fromBigInt
_ => id
expand : IntKind -> IntKind -> String -> String
expand k1 k2 = case (useBigInt k1, useBigInt k2) of
(False,True) => toBigInt
_ => id
-- when going from BigInt to Number, we must make
-- sure to first truncate the BigInt, otherwise we
-- might get rounding issues
intImpl : IntKind -> IntKind -> String -> Core String intImpl : IntKind -> IntKind -> String -> Core String
intImpl _ (Signed Unlimited) = pure intImpl k1 k2 s =
intImpl (Signed m) k@(Signed n) = if n >= m then pure else truncInt k let expanded = expand k1 k2 s
intImpl (Signed _) k@(Unsigned n) = truncInt k shrunk = shrink k1 k2 <$> truncInt (useBigInt k1) k2 s
intImpl (Unsigned m) k@(Unsigned n) = if n >= m then pure else truncInt k in case (k1,k2) of
(_, Signed Unlimited) => pure $ expanded
(Signed m, Signed n) =>
if n >= m then pure expanded else shrunk
-- Only if the precision of the target is greater (Signed _, Unsigned n) =>
-- than the one of the source, there is no need to cast. case (useBigInt k1, useBigInt k2) of
intImpl (Unsigned m) k@(Signed n) = if n > P m then pure else truncInt k (False,True) => truncInt True k2 (toBigInt s)
_ => shrunk
(Unsigned m, Unsigned n) =>
if n >= m then pure expanded else shrunk
-- Only if the precision of the target is greater
-- than the one of the source, there is no need to cast.
(Unsigned m, Signed n) =>
if n > P m then pure expanded else shrunk
jsOp : {auto c : Ref ESs ESSt} -> PrimFn arity -> Vect arity String -> Core String jsOp : {auto c : Ref ESs ESSt} -> PrimFn arity -> Vect arity String -> Core String
jsOp (Add ty) [x, y] = arithOp (intKind ty) "+" x y jsOp (Add ty) [x, y] = arithOp (intKind ty) "+" x y
jsOp (Sub ty) [x, y] = arithOp (intKind ty) "-" x y jsOp (Sub ty) [x, y] = arithOp (intKind ty) "-" x y
jsOp (Mul ty) [x, y] = arithOp (intKind ty) "*" x y jsOp (Mul ty) [x, y] = mult (intKind ty) x y
jsOp (Div ty) [x, y] = arithOp (intKind ty) "/" x y jsOp (Div ty) [x, y] = div (intKind ty) x y
jsOp (Mod ty) [x, y] = arithOp (intKind ty) "%" x y jsOp (Mod ty) [x, y] = arithOp (intKind ty) "%" x y
jsOp (Neg ty) [x] = pure $ "(-(" ++ x ++ "))" jsOp (Neg ty) [x] = pure $ "(-(" ++ x ++ "))"
jsOp (ShiftL Int32Type) [x, y] = pure $ binOp "<<" x y
jsOp (ShiftL ty) [x, y] = bitOp (intKind ty) "<<" x y jsOp (ShiftL ty) [x, y] = bitOp (intKind ty) "<<" x y
jsOp (ShiftR Int32Type) [x, y] = pure $ binOp ">>" x y
jsOp (ShiftR ty) [x, y] = bitOp (intKind ty) ">>" x y jsOp (ShiftR ty) [x, y] = bitOp (intKind ty) ">>" x y
jsOp (BAnd Bits32Type) [x, y] = pure . fromBigInt $ binOp "&" (toBigInt x) (toBigInt y)
jsOp (BOr Bits32Type) [x, y] = pure . fromBigInt $ binOp "|" (toBigInt x) (toBigInt y)
jsOp (BXOr Bits32Type) [x, y] = pure . fromBigInt $ binOp "^" (toBigInt x) (toBigInt y)
jsOp (BAnd ty) [x, y] = pure $ binOp "&" x y jsOp (BAnd ty) [x, y] = pure $ binOp "&" x y
jsOp (BOr ty) [x, y] = pure $ binOp "|" x y jsOp (BOr ty) [x, y] = pure $ binOp "|" x y
jsOp (BXOr ty) [x, y] = pure $ binOp "^" x y jsOp (BXOr ty) [x, y] = pure $ binOp "^" x y
@ -307,7 +412,7 @@ jsOp DoubleSqrt [x] = pure $ "Math.sqrt(" ++ x ++ ")"
jsOp DoubleFloor [x] = pure $ "Math.floor(" ++ x ++ ")" jsOp DoubleFloor [x] = pure $ "Math.floor(" ++ x ++ ")"
jsOp DoubleCeiling [x] = pure $ "Math.ceil(" ++ x ++ ")" jsOp DoubleCeiling [x] = pure $ "Math.ceil(" ++ x ++ ")"
jsOp (Cast StringType DoubleType) [x] = pure $ "parseFloat(" ++ x ++ ")" jsOp (Cast StringType DoubleType) [x] = jsNumberOfString x
jsOp (Cast ty StringType) [x] = pure $ jsAnyToString x jsOp (Cast ty StringType) [x] = pure $ jsAnyToString x
jsOp (Cast ty ty2) [x] = castInt constPrimitives ty ty2 x jsOp (Cast ty ty2) [x] = castInt constPrimitives ty ty2 x
jsOp BelieveMe [_,_,x] = pure x jsOp BelieveMe [_,_,x] = pure x

View File

@ -436,7 +436,7 @@ mutual
addRefs ds (CLam _ _ sc) = addRefs ds sc addRefs ds (CLam _ _ sc) = addRefs ds sc
addRefs ds (CLet _ _ _ val sc) = addRefs (addRefs ds val) sc addRefs ds (CLet _ _ _ val sc) = addRefs (addRefs ds val) sc
addRefs ds (CApp _ f args) = addRefsArgs (addRefs ds f) args addRefs ds (CApp _ f args) = addRefsArgs (addRefs ds f) args
addRefs ds (CCon _ _ _ _ args) = addRefsArgs ds args addRefs ds (CCon _ n _ _ args) = addRefsArgs (insert n False ds) args
addRefs ds (COp _ _ args) = addRefsArgs ds (toList args) addRefs ds (COp _ _ args) = addRefsArgs ds (toList args)
addRefs ds (CExtPrim _ _ args) = addRefsArgs ds args addRefs ds (CExtPrim _ _ args) = addRefsArgs ds args
addRefs ds (CForce _ _ e) = addRefs ds e addRefs ds (CForce _ _ e) = addRefs ds e
@ -473,7 +473,7 @@ inlineDef : {auto c : Ref Ctxt Defs} ->
inlineDef n inlineDef n
= do defs <- get Ctxt = do defs <- get Ctxt
Just def <- lookupCtxtExact n (gamma defs) | Nothing => pure () Just def <- lookupCtxtExact n (gamma defs) | Nothing => pure ()
let Just cexpr = compexpr def | Nothing => pure () let Just cexpr = compexpr def | Nothing => pure ()
setCompiled n !(inline n cexpr) setCompiled n !(inline n cexpr)
-- Update the names a function refers to at runtime based on the transformation -- Update the names a function refers to at runtime based on the transformation

View File

@ -30,6 +30,7 @@ import System.Info
%default covering %default covering
export
findChez : IO String findChez : IO String
findChez findChez
= do Nothing <- idrisGetEnv "CHEZ" = do Nothing <- idrisGetEnv "CHEZ"
@ -43,6 +44,7 @@ findChez
-- of the library paths managed by Idris -- of the library paths managed by Idris
-- If it can't be found, we'll assume it's a system library and that chez -- If it can't be found, we'll assume it's a system library and that chez
-- will thus be able to find it. -- will thus be able to find it.
export
findLibs : {auto c : Ref Ctxt Defs} -> findLibs : {auto c : Ref Ctxt Defs} ->
List String -> Core (List (String, String)) List String -> Core (List (String, String))
findLibs ds findLibs ds
@ -55,7 +57,7 @@ findLibs ds
then Just (trim (substr 3 (length d) d)) then Just (trim (substr 3 (length d) d))
else Nothing else Nothing
export
escapeString : String -> String escapeString : String -> String
escapeString s = pack $ foldr escape [] $ unpack s escapeString s = pack $ foldr escape [] $ unpack s
where where
@ -97,6 +99,7 @@ showChezString [] = id
showChezString ('"'::cs) = ("\\\"" ++) . showChezString cs showChezString ('"'::cs) = ("\\\"" ++) . showChezString cs
showChezString (c ::cs) = (showChezChar c) . showChezString cs showChezString (c ::cs) = (showChezChar c) . showChezString cs
export
chezString : String -> String chezString : String -> String
chezString cs = strCons '"' (showChezString (unpack cs) "\"") chezString cs = strCons '"' (showChezString (unpack cs) "\"")
@ -129,6 +132,7 @@ mutual
getFArgs (NmCon fc _ _ (Just 1) [ty, val, rest]) = pure $ (ty, val) :: !(getFArgs rest) getFArgs (NmCon fc _ _ (Just 1) [ty, val, rest]) = pure $ (ty, val) :: !(getFArgs rest)
getFArgs arg = throw (GenericMsg (getFC arg) ("Badly formed c call argument list " ++ show arg)) getFArgs arg = throw (GenericMsg (getFC arg) ("Badly formed c call argument list " ++ show arg))
export
chezExtPrim : Int -> ExtPrim -> List NamedCExp -> Core String chezExtPrim : Int -> ExtPrim -> List NamedCExp -> Core String
chezExtPrim i GetField [NmPrimVal _ (Str s), _, _, struct, chezExtPrim i GetField [NmPrimVal _ (Str s), _, _, struct,
NmPrimVal _ (Str fld), _] NmPrimVal _ (Str fld), _]
@ -162,9 +166,11 @@ mutual
= schExtCommon chezExtPrim chezString i prim args = schExtCommon chezExtPrim chezString i prim args
-- Reference label for keeping track of loaded external libraries -- Reference label for keeping track of loaded external libraries
export
data Loaded : Type where data Loaded : Type where
-- Label for noting which struct types are declared -- Label for noting which struct types are declared
export
data Structs : Type where data Structs : Type where
cftySpec : FC -> CFType -> Core String cftySpec : FC -> CFType -> Core String
@ -272,10 +278,9 @@ schemeCall fc sfn argns ret
useCC : {auto c : Ref Ctxt Defs} -> useCC : {auto c : Ref Ctxt Defs} ->
{auto l : Ref Loaded (List String)} -> {auto l : Ref Loaded (List String)} ->
String -> FC -> List String -> List (Name, CFType) -> CFType -> Core (String, String) String -> FC -> List String -> List (Name, CFType) -> CFType -> Core (String, String)
useCC appdir fc [] args ret = throw (NoForeignCC fc) useCC appdir fc ccs args ret
useCC appdir fc (cc :: ccs) args ret = case parseCC ["scheme,chez", "scheme", "C"] ccs of
= case parseCC cc of Nothing => throw (NoForeignCC fc)
Nothing => useCC appdir fc ccs args ret
Just ("scheme,chez", [sfn]) => Just ("scheme,chez", [sfn]) =>
do body <- schemeCall fc sfn (map fst args) ret do body <- schemeCall fc sfn (map fst args) ret
pure ("", body) pure ("", body)
@ -284,7 +289,7 @@ useCC appdir fc (cc :: ccs) args ret
pure ("", body) pure ("", body)
Just ("C", [cfn, clib]) => cCall appdir fc cfn clib args ret Just ("C", [cfn, clib]) => cCall appdir fc cfn clib args ret
Just ("C", [cfn, clib, chdr]) => cCall appdir fc cfn clib args ret Just ("C", [cfn, clib, chdr]) => cCall appdir fc cfn clib args ret
_ => useCC appdir fc ccs args ret _ => throw (NoForeignCC fc)
-- For every foreign arg type, return a name, and whether to pass it to the -- For every foreign arg type, return a name, and whether to pass it to the
-- foreign call (we don't pass '%World') -- foreign call (we don't pass '%World')
@ -329,14 +334,16 @@ schFgnDef appdir fc n (MkNmForeign cs args ret)
body ++ "))\n") body ++ "))\n")
schFgnDef _ _ _ _ = pure ("", "") schFgnDef _ _ _ _ = pure ("", "")
export
getFgnCall : {auto c : Ref Ctxt Defs} -> getFgnCall : {auto c : Ref Ctxt Defs} ->
{auto l : Ref Loaded (List String)} -> {auto l : Ref Loaded (List String)} ->
{auto s : Ref Structs (List String)} -> {auto s : Ref Structs (List String)} ->
String -> (Name, FC, NamedDef) -> Core (String, String) String -> (Name, FC, NamedDef) -> Core (String, String)
getFgnCall appdir (n, fc, d) = schFgnDef appdir fc n d getFgnCall appdir (n, fc, d) = schFgnDef appdir fc n d
startChez : String -> String -> String export
startChez appdir target = unlines startChezPreamble : String
startChezPreamble = unlines
[ "#!/bin/sh" [ "#!/bin/sh"
, "" , ""
, "set -e # exit on any error" , "set -e # exit on any error"
@ -357,7 +364,12 @@ startChez appdir target = unlines
, "fi " , "fi "
, "" , ""
, "DIR=$(dirname \"$($REALPATH \"$0\")\")" , "DIR=$(dirname \"$($REALPATH \"$0\")\")"
, "export LD_LIBRARY_PATH=\"$DIR/" ++ appdir ++ "\":$LD_LIBRARY_PATH" , "" -- so that the preamble ends with a newline
]
startChez : String -> String -> String
startChez appdir target = startChezPreamble ++ unlines
[ "export LD_LIBRARY_PATH=\"$DIR/" ++ appdir ++ "\":$LD_LIBRARY_PATH"
, "\"$DIR/" ++ target ++ "\" \"$@\"" , "\"$DIR/" ++ target ++ "\" \"$@\""
] ]

View File

@ -0,0 +1,310 @@
module Compiler.Scheme.ChezSep
import Compiler.Common
import Compiler.CompileExpr
import Compiler.Inline
import Compiler.Scheme.Common
import Compiler.Scheme.Chez
import Compiler.Separate
import Core.Core
import Core.Hash
import Core.Context
import Core.Context.Log
import Core.Directory
import Core.Name
import Core.Options
import Core.TT
import Libraries.Utils.Hex
import Libraries.Utils.Path
import Data.List
import Data.List1
import Data.Maybe
import Libraries.Data.NameMap
import Data.Strings
import Data.Vect
import Idris.Env
import System
import System.Directory
import System.File
import System.Info
%default covering
schHeader : List String -> List String -> String
schHeader libs compilationUnits = unlines
[ "(import (chezscheme) (support) "
++ unwords ["(" ++ cu ++ ")" | cu <- compilationUnits]
++ ")"
, "(case (machine-type)"
, " [(i3le ti3le a6le ta6le) (load-shared-object \"libc.so.6\")]"
, " [(i3osx ti3osx a6osx ta6osx) (load-shared-object \"libc.dylib\")]"
, " [(i3nt ti3nt a6nt ta6nt) (load-shared-object \"msvcrt.dll\")"
, " (load-shared-object \"ws2_32.dll\")]"
, " [else (load-shared-object \"libc.so\")]"
, unlines [" (load-shared-object \"" ++ escapeString lib ++ "\")" | lib <- libs]
, ")"
]
schFooter : String
schFooter = "(collect 4)\n(blodwen-run-finalisers)\n"
startChez : String -> String -> String -> String
startChez chez appDirSh targetSh = Chez.startChezPreamble ++ unlines
[ "export LD_LIBRARY_PATH=\"$DIR/" ++ appDirSh ++ "\":$LD_LIBRARY_PATH"
, "\"" ++ chez ++ "\" -q "
++ "--libdirs \"$DIR/" ++ appDirSh ++ "\" "
++ "--program \"$DIR/" ++ targetSh ++ "\" "
++ "\"$@\""
]
startChezCmd : String -> String -> String -> String
startChezCmd chez appDirSh targetSh = unlines
[ "@echo off"
, "set APPDIR=%~dp0"
, "set PATH=%APPDIR%\\" ++ appDirSh ++ ";%PATH%"
, "\"" ++ chez ++ "\" -q "
++ "--libdirs \"%APPDIR%/" ++ appDirSh ++ "\" "
++ "--program \"%APPDIR%/" ++ targetSh ++ "\" "
++ "%*"
]
startChezWinSh : String -> String -> String -> String
startChezWinSh chez appDirSh targetSh = unlines
[ "#!/bin/sh"
, ""
, "set -e # exit on any error"
, ""
, "DIR=$(dirname \"$(realpath \"$0\")\")"
, "CHEZ=$(cygpath \"" ++ chez ++"\")"
, "export PATH=\"$DIR/" ++ appDirSh ++ "\":$PATH"
, "\"$CHEZ\" --program \"$DIR/" ++ targetSh ++ "\" \"$@\""
, "\"$CHEZ\" -q "
++ "--libdirs \"$DIR/" ++ appDirSh ++ "\" "
++ "--program \"$DIR/" ++ targetSh ++ "\" "
++ "\"$@\""
]
-- TODO: parallelise this
compileChezLibraries : (chez : String) -> (libDir : String) -> (ssFiles : List String) -> Core ()
compileChezLibraries chez libDir ssFiles = coreLift_ $ system $ unwords
[ "echo"
, unwords
[ "'(parameterize ([optimize-level 3] [compile-file-message #f]) (compile-library " ++ chezString ssFile ++ "))'"
++ " '(delete-file " ++ chezString ssFile ++ ")'"
-- we must delete the SS file to prevent it from interfering with the SO files
-- we keep the .hash file, though, so we still keep track of what to rebuild
| ssFile <- ssFiles
]
, "|", chez, "-q", "--libdirs", libDir
]
compileChezLibrary : (chez : String) -> (libDir : String) -> (ssFile : String) -> Core ()
compileChezLibrary chez libDir ssFile = coreLift_ $ system $ unwords
[ "echo"
, "'(parameterize ([optimize-level 3] [compile-file-message #f]) (compile-library " ++ chezString ssFile ++ "))'"
, "'(delete-file " ++ chezString ssFile ++ ")'"
, "|", chez, "-q", "--libdirs", libDir
]
compileChezProgram : (chez : String) -> (libDir : String) -> (ssFile : String) -> Core ()
compileChezProgram chez libDir ssFile = coreLift_ $ system $ unwords
[ "echo"
, "'(parameterize ([optimize-level 3] [compile-file-message #f]) (compile-program " ++ chezString ssFile ++ "))'"
, "'(delete-file " ++ chezString ssFile ++ ")'"
, "|", chez, "-q", "--libdirs", libDir
]
chezNS : Namespace -> String
chezNS ns = case showNSWithSep "-" ns of
"" => "unqualified"
nss => nss
-- arbitrarily name the compilation unit
-- after the alphabetically first namespace contained within
chezLibraryName : CompilationUnit def -> String
chezLibraryName cu = chezNS (min1 cu.namespaces)
where
-- the stdlib of the previous stable version
-- has some strange version of List1.foldl1
-- so we reimplement it here for compatibility
min1 : List1 Namespace -> Namespace
min1 (ns ::: nss) = foldl min ns nss
-- TODO: use a proper exec function without shell injection
touch : String -> Core ()
touch s = coreLift_ $ system ("touch \"" ++ s ++ "\"")
record ChezLib where
constructor MkChezLib
name : String
isOutdated : Bool -- needs recompiling
||| Compile a TT expression to a bunch of Chez Scheme files
compileToSS : Ref Ctxt Defs -> String -> String -> ClosedTerm -> Core (Bool, List ChezLib)
compileToSS c chez appdir tm = do
-- process native libraries
ds <- getDirectives Chez
libs <- findLibs ds
traverse_ copyLib libs
-- get the material for compilation
cdata <- getCompileData False Cases tm
let ctm = forget (mainExpr cdata)
let ndefs = namedDefs cdata
let cui = getCompilationUnits ndefs
-- copy the support library
support <- readDataFile "chez/support-sep.ss"
let supportHash = show $ hash support
supportChanged <-
coreLift (File.readFile (appdir </> "support.hash")) >>= \case
Left err => pure True
Right fileHash => pure (fileHash /= supportHash)
when supportChanged $ do
Core.writeFile (appdir </> "support.ss") support
Core.writeFile (appdir </> "support.hash") supportHash
-- TODO: add extraRuntime
-- the problem with this is that it's unclear what to put in the (export) clause of the library
-- extraRuntime <- getExtraRuntime ds
-- for each compilation unit, generate code
chezLibs <- for cui.compilationUnits $ \cu => do
let chezLib = chezLibraryName cu
-- check if the hash has changed
-- TODO: also check that the .so file exists
let cuHash = show (hash cu)
hashChanged <-
coreLift (File.readFile (appdir </> chezLib <.> "hash")) >>= \case
Left err => pure True
Right fileHash => pure (fileHash /= cuHash)
-- generate code only when necessary
when hashChanged $ do
defs <- get Ctxt
l <- newRef {t = List String} Loaded ["libc", "libc 6"]
s <- newRef {t = List String} Structs []
-- create imports + exports + header + footer
let imports = unwords
[ "(" ++
maybe
"unqualified"
chezLibraryName
(SortedMap.lookup cuid cui.byId)
++ ")"
| cuid <- SortedSet.toList cu.dependencies
]
let exports = unwords $ concat
-- constructors don't generate Scheme definitions
[ case d of
MkNmCon _ _ _ => []
_ => [schName dn]
| (dn, fc, d) <- cu.definitions
]
let header =
"(library (" ++ chezLib ++ ")\n"
++ " (export " ++ exports ++ ")\n"
++ " (import (chezscheme) (support) " ++ imports ++ ")\n\n"
let footer = ")"
fgndefs <- traverse (Chez.getFgnCall appdir) cu.definitions
compdefs <- traverse (getScheme Chez.chezExtPrim Chez.chezString) cu.definitions
-- write the files
log "compiler.scheme.chez" 3 $ "Generating code for " ++ chezLib
Core.writeFile (appdir </> chezLib <.> "ss") $ fastAppend $
[header]
++ map snd fgndefs -- definitions using foreign libs
++ compdefs
++ map fst fgndefs -- foreign library load statements
++ [footer]
Core.writeFile (appdir </> chezLib <.> "hash") cuHash
pure (MkChezLib chezLib hashChanged)
-- main module
main <- schExp Chez.chezExtPrim Chez.chezString 0 ctm
Core.writeFile (appdir </> "mainprog.ss") $ unlines $
[ schHeader (map snd libs) [lib.name | lib <- chezLibs]
, "(collect-request-handler (lambda () (collect) (blodwen-run-finalisers)))"
, main
, schFooter
]
pure (supportChanged, chezLibs)
makeSh : String -> String -> String -> String -> Core ()
makeSh chez outShRel appDirSh targetSh =
Core.writeFile outShRel (startChez chez appDirSh targetSh)
||| Make Windows start scripts, one for bash environments and one batch file
makeShWindows : String -> String -> String -> String -> Core ()
makeShWindows chez outShRel appDirSh targetSh = do
let cmdFile = outShRel ++ ".cmd"
Core.writeFile cmdFile (startChezCmd chez appDirSh targetSh)
Core.writeFile outShRel (startChezWinSh chez appDirSh targetSh)
||| Chez Scheme implementation of the `compileExpr` interface.
compileExpr : Bool -> Ref Ctxt Defs -> (tmpDir : String) -> (outputDir : String) ->
ClosedTerm -> (outfile : String) -> Core (Maybe String)
compileExpr makeitso c tmpDir outputDir tm outfile = do
-- set up paths
Just cwd <- coreLift currentDir
| Nothing => throw (InternalError "Can't get current directory")
let appDirSh = outfile ++ "_app" -- relative to the launcher shell script
let appDirRel = outputDir </> appDirSh -- relative to CWD
let appDirAbs = cwd </> appDirRel
coreLift_ $ mkdirAll appDirRel
-- generate the code
chez <- coreLift $ findChez
(supportChanged, chezLibs) <- compileToSS c chez appDirRel tm
-- compile the code
logTime "++ Make SO" $ when makeitso $ do
-- compile the support code
when supportChanged $ do
log "compiler.scheme.chez" 3 $ "Compiling support"
compileChezLibrary chez appDirRel (appDirRel </> "support.ss")
-- compile every compilation unit
compileChezLibraries chez appDirRel
[appDirRel </> lib.name <.> "ss" | lib <- chezLibs, lib.isOutdated]
-- touch them in the right order to make the timestamps right
-- even for the libraries that were not recompiled
for_ chezLibs $ \lib => do
log "compiler.scheme.chez" 3 $ "Touching " ++ lib.name
touch (appDirRel </> lib.name <.> "so")
-- compile the main program
compileChezProgram chez appDirRel (appDirRel </> "mainprog.ss")
-- generate the launch script
let outShRel = outputDir </> outfile
let launchTargetSh = appDirSh </> "mainprog" <.> (if makeitso then "so" else "ss")
if isWindows
then makeShWindows chez outShRel appDirSh launchTargetSh
else makeSh chez outShRel appDirSh launchTargetSh
coreLift_ $ chmodRaw outShRel 0o755
pure (Just outShRel)
||| Chez Scheme implementation of the `executeExpr` interface.
||| This implementation simply runs the usual compiler, saving it to a temp file, then interpreting it.
executeExpr : Ref Ctxt Defs -> (tmpDir : String) -> ClosedTerm -> Core ()
executeExpr c tmpDir tm
= do Just sh <- compileExpr False c tmpDir tmpDir tm "_tmpchez"
| Nothing => throw (InternalError "compileExpr returned Nothing")
coreLift_ $ system sh
||| Codegen wrapper for Chez scheme implementation.
export
codegenChezSep : Codegen
codegenChezSep = MkCG (compileExpr True) executeExpr

View File

@ -301,15 +301,14 @@ schemeCall fc sfn argns ret
useCC : {auto c : Ref Ctxt Defs} -> useCC : {auto c : Ref Ctxt Defs} ->
{auto l : Ref Loaded (List String)} -> {auto l : Ref Loaded (List String)} ->
FC -> List String -> List (Name, CFType) -> CFType -> Core (Maybe String, (String, String)) FC -> List String -> List (Name, CFType) -> CFType -> Core (Maybe String, (String, String))
useCC fc [] args ret = throw (NoForeignCC fc) useCC fc ccs args ret
useCC fc (cc :: ccs) args ret = case parseCC ["scheme,gambit", "scheme", "C"] ccs of
= case parseCC cc of Nothing => throw (NoForeignCC fc)
Nothing => useCC fc ccs args ret
Just ("scheme,gambit", [sfn]) => pure (Nothing, (!(schemeCall fc sfn (map fst args) ret), "")) Just ("scheme,gambit", [sfn]) => pure (Nothing, (!(schemeCall fc sfn (map fst args) ret), ""))
Just ("scheme", [sfn]) => pure (Nothing, (!(schemeCall fc sfn (map fst args) ret), "")) Just ("scheme", [sfn]) => pure (Nothing, (!(schemeCall fc sfn (map fst args) ret), ""))
Just ("C", [cfn, clib]) => pure (Just clib, !(cCall fc cfn (fnWrapName cfn) clib args ret)) Just ("C", [cfn, clib]) => pure (Just clib, !(cCall fc cfn (fnWrapName cfn) clib args ret))
Just ("C", [cfn, clib, chdr]) => pure (Just clib, !(cCall fc cfn (fnWrapName cfn) clib args ret)) Just ("C", [cfn, clib, chdr]) => pure (Just clib, !(cCall fc cfn (fnWrapName cfn) clib args ret))
_ => useCC fc ccs args ret _ => throw (NoForeignCC fc)
where where
fnWrapName : String -> String -> String fnWrapName : String -> String -> String
fnWrapName cfn schemeArgName = schemeArgName ++ "-" ++ cfn ++ "-cFunWrap" fnWrapName cfn schemeArgName = schemeArgName ++ "-" ++ cfn ++ "-cFunWrap"

View File

@ -260,10 +260,9 @@ useCC : {auto f : Ref Done (List String) } ->
{auto c : Ref Ctxt Defs} -> {auto c : Ref Ctxt Defs} ->
{auto l : Ref Loaded (List String)} -> {auto l : Ref Loaded (List String)} ->
String -> FC -> List String -> List (Name, CFType) -> CFType -> Core (String, String) String -> FC -> List String -> List (Name, CFType) -> CFType -> Core (String, String)
useCC appdir fc [] args ret = throw (NoForeignCC fc) useCC appdir fc ccs args ret
useCC appdir fc (cc :: ccs) args ret = case parseCC ["scheme,racket", "scheme", "C"] ccs of
= case parseCC cc of Nothing => throw (NoForeignCC fc)
Nothing => useCC appdir fc ccs args ret
Just ("scheme,racket", [sfn]) => Just ("scheme,racket", [sfn]) =>
do body <- schemeCall fc sfn (map fst args) ret do body <- schemeCall fc sfn (map fst args) ret
pure ("", body) pure ("", body)
@ -272,7 +271,7 @@ useCC appdir fc (cc :: ccs) args ret
pure ("", body) pure ("", body)
Just ("C", [cfn, clib]) => cCall appdir fc cfn clib args ret Just ("C", [cfn, clib]) => cCall appdir fc cfn clib args ret
Just ("C", [cfn, clib, chdr]) => cCall appdir fc cfn clib args ret Just ("C", [cfn, clib, chdr]) => cCall appdir fc cfn clib args ret
_ => useCC appdir fc ccs args ret _ => throw (NoForeignCC fc)
-- For every foreign arg type, return a name, and whether to pass it to the -- For every foreign arg type, return a name, and whether to pass it to the
-- foreign call (we don't pass '%World') -- foreign call (we don't pass '%World')

345
src/Compiler/Separate.idr Normal file
View File

@ -0,0 +1,345 @@
module Compiler.Separate
import public Core.FC
import public Core.Name
import public Core.Name.Namespace
import public Core.CompileExpr
import public Compiler.VMCode
import public Libraries.Data.SortedMap
import public Libraries.Data.SortedSet
import public Libraries.Data.StringMap
import Core.Hash
import Core.TT
import Data.List
import Data.List1
import Data.Vect
import Data.Maybe
-- Compilation unit IDs are intended to be opaque,
-- just to be able to express dependencies via keys in a map and such.
export
record CompilationUnitId where
constructor CUID
int : Int
export
Eq CompilationUnitId where
CUID x == CUID y = x == y
export
Ord CompilationUnitId where
compare (CUID x) (CUID y) = compare x y
export
Hashable CompilationUnitId where
hashWithSalt h (CUID int) = hashWithSalt h int
||| A compilation unit is a set of namespaces.
|||
||| The record is parameterised by the type of the definition,
||| which makes it reusable for various IRs provided by getCompileData.
public export
record CompilationUnit def where
constructor MkCompilationUnit
||| Unique identifier of a compilation unit within a CompilationUnitInfo record.
id : CompilationUnitId
||| Namespaces contained within the compilation unit.
namespaces : List1 Namespace
||| Other units that this unit depends on.
dependencies : SortedSet CompilationUnitId
||| The definitions belonging into this compilation unit.
definitions : List (Name, def)
export
Hashable def => Hashable (CompilationUnit def) where
hashWithSalt h cu =
h `hashWithSalt` SortedSet.toList cu.dependencies
`hashWithSalt` cu.definitions
private
getNS : Name -> Namespace
getNS (NS ns _) = ns
getNS _ = emptyNS
||| Group definitions by namespace.
private
splitByNS : List (Name, def) -> List (Namespace, List (Name, def))
splitByNS = SortedMap.toList . foldl addOne SortedMap.empty
where
addOne
: SortedMap Namespace (List (Name, def))
-> (Name, def)
-> SortedMap Namespace (List (Name, def))
addOne nss ndef@(n, _) =
SortedMap.mergeWith
(++)
(SortedMap.singleton (getNS n) [ndef])
nss
-- Mechanically transcribed from
-- https://en.wikipedia.org/wiki/Tarjan%27s_strongly_connected_components_algorithm#The_algorithm_in_pseudocode
namespace Tarjan
private
record TarjanVertex where
constructor TV
index : Int
lowlink : Int
inStack : Bool
private
record TarjanState cuid where
constructor TS
vertices : SortedMap cuid TarjanVertex
stack : List cuid
nextIndex : Int
components : List (List1 cuid)
impossibleHappened : Bool -- we should get at least some indication of broken assumptions
||| Find strongly connected components in the given graph.
|||
||| Input: map from vertex X to all vertices Y such that there is edge X->Y
||| Output: list of strongly connected components, ordered by output degree descending
export
tarjan : Ord cuid => SortedMap cuid (SortedSet cuid) -> List (List1 cuid)
tarjan {cuid} deps = loop initialState (SortedMap.keys deps)
where
initialState : TarjanState cuid
initialState =
TS
SortedMap.empty
[]
0
[]
False
strongConnect : TarjanState cuid -> cuid -> TarjanState cuid
strongConnect ts v =
let ts'' = case SortedMap.lookup v deps of
Nothing => ts' -- no edges
Just edgeSet => loop ts' (SortedSet.toList edgeSet)
in case SortedMap.lookup v ts''.vertices of
Nothing => record { impossibleHappened = True } ts''
Just vtv =>
if vtv.index == vtv.lowlink
then createComponent ts'' v []
else ts''
where
createComponent : TarjanState cuid -> cuid -> List cuid -> TarjanState cuid
createComponent ts v acc =
case ts.stack of
[] => record { impossibleHappened = True } ts
w :: ws =>
let ts' : TarjanState cuid = record {
vertices $= SortedMap.adjust w record{ inStack = False },
stack = ws
} ts
in if w == v
then record { components $= ((v ::: acc) ::) } ts' -- that's it
else createComponent ts' v (w :: acc)
loop : TarjanState cuid -> List cuid -> TarjanState cuid
loop ts [] = ts
loop ts (w :: ws) =
loop (
case SortedMap.lookup w ts.vertices of
Nothing => let ts' = strongConnect ts w in
case SortedMap.lookup w ts'.vertices of
Nothing => record { impossibleHappened = True } ts'
Just wtv => record { vertices $= SortedMap.adjust v record{ lowlink $= min wtv.lowlink } } ts'
Just wtv => case wtv.inStack of
False => ts -- nothing to do
True => record { vertices $= SortedMap.adjust v record{ lowlink $= min wtv.index } } ts
) ws
ts' : TarjanState cuid
ts' = record {
vertices $= SortedMap.insert v (TV ts.nextIndex ts.nextIndex True),
stack $= (v ::),
nextIndex $= (1+)
} ts
loop : TarjanState cuid -> List cuid -> List (List1 cuid)
loop ts [] =
if ts.impossibleHappened
then []
else ts.components
loop ts (v :: vs) =
case SortedMap.lookup v ts.vertices of
Just _ => loop ts vs -- done, skip
Nothing => loop (strongConnect ts v) vs
public export
interface HasNamespaces a where
||| Return the set of namespaces mentioned within
nsRefs : a -> SortedSet Namespace
-- For now, we have instances only for NamedDef and VMDef.
-- For other IR representations, we'll have to add more instances.
-- This is not hard, just a bit of tedious mechanical work.
mutual
export
HasNamespaces NamedCExp where
nsRefs (NmLocal fc n) = SortedSet.empty
nsRefs (NmRef fc n) = SortedSet.singleton $ getNS n
nsRefs (NmLam fc n rhs) = nsRefs rhs
nsRefs (NmLet fc n val rhs) = nsRefs val <+> nsRefs rhs
nsRefs (NmApp fc f args) = nsRefs f <+> concatMap nsRefs args
nsRefs (NmCon fc cn ci tag args) = concatMap nsRefs args
nsRefs (NmForce fc reason rhs) = nsRefs rhs
nsRefs (NmDelay fc reason rhs) = nsRefs rhs
nsRefs (NmErased fc) = SortedSet.empty
nsRefs (NmPrimVal ft x) = SortedSet.empty
nsRefs (NmOp fc op args) = concatMap nsRefs args
nsRefs (NmExtPrim fc n args) = concatMap nsRefs args
nsRefs (NmConCase fc scrut alts mbDflt) =
nsRefs scrut <+> concatMap nsRefs alts <+> concatMap nsRefs mbDflt
nsRefs (NmConstCase fc scrut alts mbDflt) =
nsRefs scrut <+> concatMap nsRefs alts <+> concatMap nsRefs mbDflt
nsRefs (NmCrash fc msg) = SortedSet.empty
export
HasNamespaces NamedConAlt where
nsRefs (MkNConAlt n ci tag args rhs) = nsRefs rhs
export
HasNamespaces NamedConstAlt where
nsRefs (MkNConstAlt c rhs) = nsRefs rhs
export
HasNamespaces NamedDef where
nsRefs (MkNmFun argNs rhs) = nsRefs rhs
nsRefs (MkNmCon tag arity nt) = SortedSet.empty
nsRefs (MkNmForeign ccs fargs rty) = SortedSet.empty
nsRefs (MkNmError rhs) = nsRefs rhs
export
HasNamespaces VMInst where
nsRefs (DECLARE x) = empty
nsRefs START = empty
nsRefs (ASSIGN x y) = empty
nsRefs (MKCON x tag args) = either (const empty) (singleton . getNS) tag
nsRefs (MKCLOSURE x n missing args) = singleton $ getNS n
nsRefs (MKCONSTANT x y) = empty
nsRefs (APPLY x f a) = empty
nsRefs (CALL x tailpos n args) = singleton $ getNS n
nsRefs (OP x y xs) = empty
nsRefs (EXTPRIM x n xs) = singleton $ getNS n
nsRefs (CASE x alts def) =
maybe empty (concatMap nsRefs) def <+>
concatMap ((concatMap nsRefs) . snd) alts <+>
concatMap ((either (const empty) (singleton . getNS)) . fst) alts
nsRefs (CONSTCASE x alts def) =
maybe empty (concatMap nsRefs) def <+>
concatMap ((concatMap nsRefs) . snd) alts
nsRefs (PROJECT x value pos) = empty
nsRefs (NULL x) = empty
nsRefs (ERROR x) = empty
export
HasNamespaces VMDef where
nsRefs (MkVMFun args is) = concatMap nsRefs is
nsRefs (MkVMError is) = concatMap nsRefs is
-- a slight hack for convenient use with CompileData.namedDefs
export
HasNamespaces a => HasNamespaces (FC, a) where
nsRefs (_, x) = nsRefs x
-- another slight hack for convenient use with CompileData.namedDefs
export
Hashable def => Hashable (FC, def) where
-- ignore FC in hash, like everywhere else
hashWithSalt h (fc, x) = hashWithSalt h x
||| Output of the codegen separation algorithm.
||| Should contain everything you need in a separately compiling codegen.
public export
record CompilationUnitInfo def where
constructor MkCompilationUnitInfo
||| Compilation units computed from the given definitions,
||| ordered topologically, starting from units depending on no other unit.
compilationUnits : List (CompilationUnit def)
||| Mapping from ID to CompilationUnit.
byId : SortedMap CompilationUnitId (CompilationUnit def)
||| Maps each namespace to the compilation unit that contains it.
namespaceMap : SortedMap Namespace CompilationUnitId
||| Group the given definitions into compilation units for separate code generation.
export
getCompilationUnits : HasNamespaces def => List (Name, def) -> CompilationUnitInfo def
getCompilationUnits {def} defs =
let
-- Definitions grouped by namespace.
defsByNS : SortedMap Namespace (List (Name, def))
= SortedMap.fromList $ splitByNS defs
-- Mapping from a namespace to all namespaces mentioned within.
-- Represents graph edges pointing in that direction.
nsDeps : SortedMap Namespace (SortedSet Namespace)
= foldl (SortedMap.mergeWith SortedSet.union) SortedMap.empty
[ SortedMap.singleton (getNS n) (SortedSet.delete (getNS n) (nsRefs d))
| (n, d) <- defs
]
-- Strongly connected components of the NS dep graph,
-- ordered by output degree ascending.
--
-- Each SCC will become a compilation unit.
components : List (List1 Namespace)
= List.reverse $ tarjan nsDeps -- tarjan generates reverse toposort
-- Maps a namespace to the compilation unit that contains it.
nsMap : SortedMap Namespace CompilationUnitId
= SortedMap.fromList [(ns, cuid) | (cuid, nss) <- withCUID components, ns <- List1.forget nss]
-- List of all compilation units, ordered by number of dependencies, ascending.
units : List (CompilationUnit def)
= [mkUnit nsDeps nsMap defsByNS cuid nss | (cuid, nss) <- withCUID components]
in MkCompilationUnitInfo
{ compilationUnits = units
, byId = SortedMap.fromList [(unit.id, unit) | unit <- units]
, namespaceMap = nsMap
}
where
withCUID : List a -> List (CompilationUnitId, a)
withCUID xs = [(CUID $ cast i, x) | (i, x) <- zip [0..length xs] xs]
||| Wrap all information in a compilation unit record.
mkUnit :
SortedMap Namespace (SortedSet Namespace)
-> SortedMap Namespace CompilationUnitId
-> SortedMap Namespace (List (Name, def))
-> CompilationUnitId -> List1 Namespace -> CompilationUnit def
mkUnit nsDeps nsMap defsByNS cuid nss =
MkCompilationUnit
{ id = cuid
, namespaces = nss
, dependencies = SortedSet.delete cuid dependencies
, definitions = definitions
}
where
dependencies : SortedSet CompilationUnitId
dependencies = SortedSet.fromList $ do
ns <- List1.forget nss -- NS contained within
depsNS <- SortedSet.toList $ -- NS we depend on
fromMaybe SortedSet.empty $
SortedMap.lookup ns nsDeps
case SortedMap.lookup depsNS nsMap of
Nothing => []
Just depCUID => [depCUID]
definitions : List (Name, def)
definitions = concat [fromMaybe [] $ SortedMap.lookup ns defsByNS | ns <- nss]

View File

@ -12,6 +12,7 @@ import public Core.TT
import Libraries.Utils.Binary import Libraries.Utils.Binary
import Data.Fin
import Libraries.Data.IntMap import Libraries.Data.IntMap
import Data.IOArray import Data.IOArray
import Data.List import Data.List
@ -654,22 +655,15 @@ data Transform : Type where
||| during codegen. ||| during codegen.
public export public export
data BuiltinType : Type where data BuiltinType : Type where
||| A built-in 'Nat'-like type
||| 'NatLike : [index ->] Type'
||| 'SLike : {0 _ : index} -> NatLike [index] -> NatLike [f index]'
||| 'ZLike : {0 _ : index} -> NatLike [index]'
BuiltinNatural : BuiltinType BuiltinNatural : BuiltinType
-- All the following aren't implemented yet
-- but are here to reduce number of TTC version changes
NaturalPlus : BuiltinType
NaturalMult : BuiltinType
NaturalToInteger : BuiltinType NaturalToInteger : BuiltinType
IntegerToNatural : BuiltinType IntegerToNatural : BuiltinType
export export
Show BuiltinType where Show BuiltinType where
show BuiltinNatural = "Natural" show BuiltinNatural = "Natural"
show _ = "Not yet implemented" show NaturalToInteger = "NaturalToInteger"
show IntegerToNatural = "IntegerToNatural"
-- Token types to make it harder to get the constructor names -- Token types to make it harder to get the constructor names
-- the wrong way round. -- the wrong way round.
@ -683,15 +677,24 @@ record NatBuiltin where
zero : Name zero : Name
succ : Name succ : Name
||| Record containing information about a natToInteger function.
public export
record NatToInt where
constructor MkNatToInt
arity : Nat -- total number of arguments
natIdx : Fin arity -- index into arguments of the 'Nat'-like argument
||| Rewrite rules for %builtin pragmas ||| Rewrite rules for %builtin pragmas
||| Seperate to 'Transform' because it must also modify case statements ||| Seperate to 'Transform' because it must also modify case statements
||| behaviour should remain the same after this transform ||| behaviour should remain the same after this transform
public export public export
record BuiltinTransforms where record BuiltinTransforms where
constructor MkBuiltinTransforms constructor MkBuiltinTransforms
-- Nat
natTyNames : NameMap NatBuiltin -- map from Nat-like names to their constructors natTyNames : NameMap NatBuiltin -- map from Nat-like names to their constructors
natZNames : NameMap ZERO -- map from Z-like names to their type constructor natZNames : NameMap ZERO -- set of Z-like names
natSNames : NameMap SUCC -- map from S-like names to their type constructor natSNames : NameMap SUCC -- set of S-like names
natToIntegerFns : NameMap NatToInt -- set of functions to transform to `believe_me`
-- TODO: After next release remove nat from here and use %builtin pragma instead -- TODO: After next release remove nat from here and use %builtin pragma instead
initBuiltinTransforms : BuiltinTransforms initBuiltinTransforms : BuiltinTransforms
@ -703,6 +706,7 @@ initBuiltinTransforms =
{ natTyNames = singleton type (MkNatBuiltin {zero, succ}) { natTyNames = singleton type (MkNatBuiltin {zero, succ})
, natZNames = singleton zero MkZERO , natZNames = singleton zero MkZERO
, natSNames = singleton succ MkSUCC , natSNames = singleton succ MkSUCC
, natToIntegerFns = empty
} }
export export

View File

@ -463,6 +463,10 @@ export %inline
(<$>) : (a -> b) -> Core a -> Core b (<$>) : (a -> b) -> Core a -> Core b
(<$>) f (MkCore a) = MkCore (map (map f) a) (<$>) f (MkCore a) = MkCore (map (map f) a)
export %inline
(<$) : b -> Core a -> Core b
(<$) = (<$>) . const
export %inline export %inline
ignore : Core a -> Core () ignore : Core a -> Core ()
ignore = map (\ _ => ()) ignore = map (\ _ => ())
@ -615,6 +619,10 @@ traverse_ f [] = pure ()
traverse_ f (x :: xs) traverse_ f (x :: xs)
= Core.do ignore (f x) = Core.do ignore (f x)
traverse_ f xs traverse_ f xs
%inline
export
for_ : List a -> (a -> Core ()) -> Core ()
for_ = flip traverse_
%inline %inline
export export
@ -747,3 +755,18 @@ condC : List (Core Bool, Core a) -> Core a -> Core a
condC [] def = def condC [] def = def
condC ((x, y) :: xs) def condC ((x, y) :: xs) def
= if !x then y else condC xs def = if !x then y else condC xs def
export
writeFile : (fname : String) -> (content : String) -> Core ()
writeFile fname content =
coreLift (File.writeFile fname content) >>= \case
Right () => pure ()
Left err => throw $ FileErr fname err
export
readFile : (fname : String) -> Core String
readFile fname =
coreLift (File.readFile fname) >>= \case
Right content => pure content
Left err => throw $ FileErr fname err

View File

@ -25,7 +25,10 @@ FileName = String
||| file or by the compiler. That makes it useful to have the notion of ||| file or by the compiler. That makes it useful to have the notion of
||| `EmptyFC` as part of the type. ||| `EmptyFC` as part of the type.
public export public export
data FC = MkFC FileName FilePos FilePos data FC = MkFC FileName FilePos FilePos
| ||| Virtual FCs are FC attached to desugared/generated code. They can help with marking
||| errors, but we shouldn't attach semantic highlighting metadata to them.
MkVirtualFC FileName FilePos FilePos
| EmptyFC | EmptyFC
||| A version of a file context that cannot be empty ||| A version of a file context that cannot be empty
@ -45,8 +48,22 @@ justFC (fname, start, end) = MkFC fname start end
export export
isNonEmptyFC : FC -> Maybe NonEmptyFC isNonEmptyFC : FC -> Maybe NonEmptyFC
isNonEmptyFC (MkFC fn start end) = Just (fn, start, end) isNonEmptyFC (MkFC fn start end) = Just (fn, start, end)
isNonEmptyFC (MkVirtualFC fn start end) = Just (fn, start, end)
isNonEmptyFC EmptyFC = Nothing isNonEmptyFC EmptyFC = Nothing
||| A view checking whether an arbitrary FC originates from a source location
export
isConcreteFC : FC -> Maybe NonEmptyFC
isConcreteFC (MkFC fn start end) = Just (fn, start, end)
isConcreteFC _ = Nothing
||| Turn an FC into a virtual one
export
virtualiseFC : FC -> FC
virtualiseFC (MkFC fn start end) = MkVirtualFC fn start end
virtualiseFC fc = fc
export export
defaultFC : NonEmptyFC defaultFC : NonEmptyFC
defaultFC = ("", (0, 0), (0, 0)) defaultFC = ("", (0, 0), (0, 0))
@ -140,6 +157,7 @@ mergeFC _ _ = Nothing
export export
Eq FC where Eq FC where
(==) (MkFC n s e) (MkFC n' s' e') = n == n' && s == s' && e == e' (==) (MkFC n s e) (MkFC n' s' e') = n == n' && s == s' && e == e'
(==) (MkVirtualFC n s e) (MkVirtualFC n' s' e') = n == n' && s == s' && e == e'
(==) EmptyFC EmptyFC = True (==) EmptyFC EmptyFC = True
(==) _ _ = False (==) _ _ = False
@ -149,6 +167,12 @@ Show FC where
show (MkFC file startPos endPos) = file ++ ":" ++ show (MkFC file startPos endPos) = file ++ ":" ++
showPos startPos ++ "--" ++ showPos startPos ++ "--" ++
showPos endPos showPos endPos
show (MkVirtualFC file startPos endPos) = file ++ ":" ++
showPos startPos ++ "--" ++
showPos endPos
prettyPos : FilePos -> Doc ann
prettyPos (l, c) = pretty (l + 1) <+> colon <+> pretty (c + 1)
export export
Pretty FC where Pretty FC where
@ -156,6 +180,6 @@ Pretty FC where
pretty (MkFC file startPos endPos) = pretty file <+> colon pretty (MkFC file startPos endPos) = pretty file <+> colon
<+> prettyPos startPos <+> pretty "--" <+> prettyPos startPos <+> pretty "--"
<+> prettyPos endPos <+> prettyPos endPos
where pretty (MkVirtualFC file startPos endPos) = pretty file <+> colon
prettyPos : FilePos -> Doc ann <+> prettyPos startPos <+> pretty "--"
prettyPos (l, c) = pretty (l + 1) <+> colon <+> pretty (c + 1) <+> prettyPos endPos

View File

@ -2,12 +2,14 @@ module Core.Hash
import Core.CaseTree import Core.CaseTree
import Core.TT import Core.TT
import Core.CompileExpr
import Data.List import Data.List
import Data.List1 import Data.List1
import Libraries.Data.List.Lazy import Libraries.Data.List.Lazy
import Data.Strings import Data.Strings
import Libraries.Data.String.Iterator import Libraries.Data.String.Iterator
import Data.Vect
%default covering %default covering
@ -42,6 +44,11 @@ export
Hashable Char where Hashable Char where
hash = cast hash = cast
export
Hashable a => Hashable (Vect n a) where
hashWithSalt h [] = abs h
hashWithSalt h (x :: xs) = hashWithSalt (h * 33 + hash x) xs
export export
Hashable a => Hashable (List a) where Hashable a => Hashable (List a) where
hashWithSalt h [] = abs h hashWithSalt h [] = abs h
@ -56,10 +63,18 @@ Hashable a => Hashable (Maybe a) where
hashWithSalt h Nothing = abs h hashWithSalt h Nothing = abs h
hashWithSalt h (Just x) = hashWithSalt h x hashWithSalt h (Just x) = hashWithSalt h x
export
Hashable a => Hashable b => Hashable (a, b) where
hashWithSalt h (x, y) = h `hashWithSalt` x `hashWithSalt` y
export export
Hashable String where Hashable String where
hashWithSalt h = String.Iterator.foldl hashWithSalt h hashWithSalt h = String.Iterator.foldl hashWithSalt h
export
Hashable Double where
hash = hash . show
export export
Hashable Namespace where Hashable Namespace where
hashWithSalt h ns = hashWithSalt h (unsafeUnfoldNamespace ns) hashWithSalt h ns = hashWithSalt h (unsafeUnfoldNamespace ns)
@ -174,3 +189,256 @@ mutual
= h `hashWithSalt` 3 `hashWithSalt` (show x) `hashWithSalt` y = h `hashWithSalt` 3 `hashWithSalt` (show x) `hashWithSalt` y
hashWithSalt h (DefaultCase x) hashWithSalt h (DefaultCase x)
= h `hashWithSalt` 4 `hashWithSalt` x = h `hashWithSalt` 4 `hashWithSalt` x
export
Hashable CFType where
hashWithSalt h = \case
CFUnit =>
h `hashWithSalt` 0
CFInt =>
h `hashWithSalt` 1
CFUnsigned8 =>
h `hashWithSalt` 2
CFUnsigned16 =>
h `hashWithSalt` 3
CFUnsigned32 =>
h `hashWithSalt` 4
CFUnsigned64 =>
h `hashWithSalt` 5
CFString =>
h `hashWithSalt` 6
CFDouble =>
h `hashWithSalt` 7
CFChar =>
h `hashWithSalt` 8
CFPtr =>
h `hashWithSalt` 9
CFGCPtr =>
h `hashWithSalt` 10
CFBuffer =>
h `hashWithSalt` 11
CFWorld =>
h `hashWithSalt` 12
CFFun a b =>
h `hashWithSalt` 13 `hashWithSalt` a `hashWithSalt` b
CFIORes a =>
h `hashWithSalt` 14 `hashWithSalt` a
CFStruct n fs =>
h `hashWithSalt` 15 `hashWithSalt` n `hashWithSalt` fs
CFUser n xs =>
h `hashWithSalt` 16 `hashWithSalt` n `hashWithSalt` xs
export
Hashable Constant where
hashWithSalt h = \case
I i =>
h `hashWithSalt` 0 `hashWithSalt` i
BI x =>
h `hashWithSalt` 1 `hashWithSalt` x
B8 x =>
h `hashWithSalt` 2 `hashWithSalt` x
B16 x =>
h `hashWithSalt` 3 `hashWithSalt` x
B32 x =>
h `hashWithSalt` 4 `hashWithSalt` x
B64 x =>
h `hashWithSalt` 5 `hashWithSalt` x
Str x =>
h `hashWithSalt` 6 `hashWithSalt` x
Ch x =>
h `hashWithSalt` 7 `hashWithSalt` x
Db x =>
h `hashWithSalt` 8 `hashWithSalt` x
WorldVal =>
h `hashWithSalt` 9
IntType =>
h `hashWithSalt` 10
IntegerType =>
h `hashWithSalt` 11
Bits8Type =>
h `hashWithSalt` 12
Bits16Type =>
h `hashWithSalt` 13
Bits32Type =>
h `hashWithSalt` 14
Bits64Type =>
h `hashWithSalt` 15
StringType =>
h `hashWithSalt` 16
CharType =>
h `hashWithSalt` 17
DoubleType =>
h `hashWithSalt` 18
WorldType =>
h `hashWithSalt` 19
I8 x => h `hashWithSalt` 20 `hashWithSalt` x
I16 x => h `hashWithSalt` 21 `hashWithSalt` x
I32 x => h `hashWithSalt` 22 `hashWithSalt` x
I64 x => h `hashWithSalt` 23 `hashWithSalt` x
Int8Type => h `hashWithSalt` 24
Int16Type => h `hashWithSalt` 25
Int32Type => h `hashWithSalt` 26
Int64Type => h `hashWithSalt` 27
export
Hashable LazyReason where
hashWithSalt h = \case
LInf => h `hashWithSalt` 0
LLazy => h `hashWithSalt` 1
LUnknown => h `hashWithSalt` 2
export
Hashable (PrimFn arity) where
hashWithSalt h = \case
Add ty =>
h `hashWithSalt` 0 `hashWithSalt` ty
Sub ty =>
h `hashWithSalt` 1 `hashWithSalt` ty
Mul ty =>
h `hashWithSalt` 2 `hashWithSalt` ty
Div ty =>
h `hashWithSalt` 3 `hashWithSalt` ty
Mod ty =>
h `hashWithSalt` 4 `hashWithSalt` ty
Neg ty =>
h `hashWithSalt` 5 `hashWithSalt` ty
ShiftL ty =>
h `hashWithSalt` 6 `hashWithSalt` ty
ShiftR ty =>
h `hashWithSalt` 7 `hashWithSalt` ty
BAnd ty =>
h `hashWithSalt` 8 `hashWithSalt` ty
BOr ty =>
h `hashWithSalt` 9 `hashWithSalt` ty
BXOr ty =>
h `hashWithSalt` 10 `hashWithSalt` ty
LT ty =>
h `hashWithSalt` 11 `hashWithSalt` ty
LTE ty =>
h `hashWithSalt` 12 `hashWithSalt` ty
EQ ty =>
h `hashWithSalt` 13 `hashWithSalt` ty
GTE ty =>
h `hashWithSalt` 14 `hashWithSalt` ty
GT ty =>
h `hashWithSalt` 15 `hashWithSalt` ty
StrLength =>
h `hashWithSalt` 16
StrHead =>
h `hashWithSalt` 17
StrTail =>
h `hashWithSalt` 18
StrIndex =>
h `hashWithSalt` 19
StrCons =>
h `hashWithSalt` 20
StrAppend =>
h `hashWithSalt` 21
StrReverse =>
h `hashWithSalt` 22
StrSubstr =>
h `hashWithSalt` 23
DoubleExp =>
h `hashWithSalt` 24
DoubleLog =>
h `hashWithSalt` 25
DoubleSin =>
h `hashWithSalt` 26
DoubleCos =>
h `hashWithSalt` 27
DoubleTan =>
h `hashWithSalt` 28
DoubleASin =>
h `hashWithSalt` 29
DoubleACos =>
h `hashWithSalt` 30
DoubleATan =>
h `hashWithSalt` 31
DoubleSqrt =>
h `hashWithSalt` 32
DoubleFloor =>
h `hashWithSalt` 33
DoubleCeiling =>
h `hashWithSalt` 34
Cast f t =>
h `hashWithSalt` 35 `hashWithSalt` f `hashWithSalt` t
BelieveMe =>
h `hashWithSalt` 36
Crash =>
h `hashWithSalt` 37
export
Hashable ConInfo where
hashWithSalt h = \case
DATACON => h `hashWithSalt` 0
TYCON => h `hashWithSalt` 1
NIL => h `hashWithSalt` 2
CONS => h `hashWithSalt` 3
ENUM => h `hashWithSalt` 4
mutual
export
Hashable NamedCExp where
hashWithSalt h = \case
NmLocal fc n =>
h `hashWithSalt` 0 `hashWithSalt` n
NmRef fc n =>
h `hashWithSalt` 1 `hashWithSalt` n
NmLam fc x rhs =>
h `hashWithSalt` 2 `hashWithSalt` x `hashWithSalt` rhs
NmLet fc x val rhs =>
h `hashWithSalt` 3 `hashWithSalt` x `hashWithSalt` val `hashWithSalt` rhs
NmApp fc f xs =>
h `hashWithSalt` 4 `hashWithSalt` f `hashWithSalt` xs
NmCon fc cn ci t args =>
h `hashWithSalt` 5 `hashWithSalt` cn `hashWithSalt` ci `hashWithSalt` t `hashWithSalt` args
NmOp fc fn args =>
h `hashWithSalt` 6 `hashWithSalt` fn `hashWithSalt` args
NmExtPrim fc p args =>
h `hashWithSalt` 7 `hashWithSalt` p `hashWithSalt` args
NmForce fc r x =>
h `hashWithSalt` 8 `hashWithSalt` r `hashWithSalt` x
NmDelay fc r x =>
h `hashWithSalt` 9 `hashWithSalt` r `hashWithSalt` x
NmConCase fc sc alts df =>
h `hashWithSalt` 10 `hashWithSalt` sc `hashWithSalt` alts `hashWithSalt` df
NmConstCase fc sc alts df =>
h `hashWithSalt` 11 `hashWithSalt` sc `hashWithSalt` alts `hashWithSalt` df
NmPrimVal fc c =>
h `hashWithSalt` 12 `hashWithSalt` c
NmErased fc =>
h `hashWithSalt` 13
NmCrash fc msg =>
h `hashWithSalt` 14 `hashWithSalt` msg
export
Hashable NamedConAlt where
hashWithSalt h (MkNConAlt n ci tag args rhs) =
h `hashWithSalt` n `hashWithSalt` ci `hashWithSalt` tag `hashWithSalt` args `hashWithSalt` rhs
export
Hashable NamedConstAlt where
hashWithSalt h (MkNConstAlt c rhs) =
h `hashWithSalt` c `hashWithSalt` rhs
export
Hashable NamedDef where
hashWithSalt h = \case
MkNmFun args ce =>
h `hashWithSalt` 0 `hashWithSalt` args `hashWithSalt` ce
MkNmCon tag arity nt =>
h `hashWithSalt` 1 `hashWithSalt` tag `hashWithSalt` arity `hashWithSalt` nt
MkNmForeign ccs fargs cft =>
h `hashWithSalt` 2 `hashWithSalt` ccs `hashWithSalt` fargs `hashWithSalt` cft
MkNmError e =>
h `hashWithSalt` 3 `hashWithSalt` e

View File

@ -17,7 +17,60 @@ import Libraries.Utils.Binary
%default covering %default covering
-- Additional data we keep about the context to support interactive editing public export
data Decoration : Type where
Typ : Decoration
Function : Decoration
Data : Decoration
Keyword : Decoration
Bound : Decoration
export
nameTypeDecoration : NameType -> Decoration
nameTypeDecoration Bound = Bound
nameTypeDecoration Func = Function
nameTypeDecoration (DataCon _ _) = Data
nameTypeDecoration (TyCon _ _ ) = Typ
public export
ASemanticDecoration : Type
ASemanticDecoration = (NonEmptyFC, Decoration, Maybe Name)
public export
SemanticDecorations : Type
SemanticDecorations = List ASemanticDecoration
public export
Eq Decoration where
Typ == Typ = True
Function == Function = True
Data == Data = True
Keyword == Keyword = True
Bound == Bound = True
_ == _ = False
public export
Show Decoration where
show Typ = "type"
show Function = "function"
show Data = "data"
show Keyword = "keyword"
show Bound = "bound"
TTC Decoration where
toBuf b Typ = tag 0
toBuf b Function = tag 1
toBuf b Data = tag 2
toBuf b Keyword = tag 3
toBuf b Bound = tag 4
fromBuf b
= case !getTag of
0 => pure Typ
1 => pure Function
2 => pure Data
3 => pure Keyword
4 => pure Bound
_ => corrupt "Decoration"
public export public export
record Metadata where record Metadata where
@ -43,26 +96,41 @@ record Metadata where
currentLHS : Maybe ClosedTerm currentLHS : Maybe ClosedTerm
holeLHS : List (Name, ClosedTerm) holeLHS : List (Name, ClosedTerm)
nameLocMap : PosMap (NonEmptyFC, Name) nameLocMap : PosMap (NonEmptyFC, Name)
sourcefile : String
-- Semantic Highlighting
-- Posmap of known semantic decorations
semanticHighlighting : PosMap ASemanticDecoration
-- Posmap of aliases (in `with` clauses the LHS disapear during
-- elaboration after making sure that they match their parents'
semanticAliases : PosMap (NonEmptyFC, NonEmptyFC)
Show Metadata where Show Metadata where
show (MkMetadata apps names tydecls currentLHS holeLHS nameLocMap) show (MkMetadata apps names tydecls currentLHS holeLHS nameLocMap
fname semanticHighlighting semanticAliases)
= "Metadata:\n" ++ = "Metadata:\n" ++
" lhsApps: " ++ show apps ++ "\n" ++ " lhsApps: " ++ show apps ++ "\n" ++
" names: " ++ show names ++ "\n" ++ " names: " ++ show names ++ "\n" ++
" type declarations: " ++ show tydecls ++ "\n" ++ " type declarations: " ++ show tydecls ++ "\n" ++
" current LHS: " ++ show currentLHS ++ "\n" ++ " current LHS: " ++ show currentLHS ++ "\n" ++
" holes: " ++ show holeLHS ++ "\n" ++ " holes: " ++ show holeLHS ++ "\n" ++
" nameLocMap: " ++ show nameLocMap " nameLocMap: " ++ show nameLocMap ++ "\n" ++
" sourcefile: " ++ fname ++
" semanticHighlighting: " ++ show semanticHighlighting ++
" semanticAliases: " ++ show semanticAliases
export export
initMetadata : Metadata initMetadata : String -> Metadata
initMetadata = MkMetadata initMetadata fname = MkMetadata
{ lhsApps = [] { lhsApps = []
, names = [] , names = []
, tydecls = [] , tydecls = []
, currentLHS = Nothing , currentLHS = Nothing
, holeLHS = [] , holeLHS = []
, nameLocMap = empty , nameLocMap = empty
, sourcefile = fname
, semanticHighlighting = empty
, semanticAliases = empty
} }
-- A label for metadata in the global state -- A label for metadata in the global state
@ -76,6 +144,9 @@ TTC Metadata where
toBuf b (tydecls m) toBuf b (tydecls m)
toBuf b (holeLHS m) toBuf b (holeLHS m)
toBuf b (nameLocMap m) toBuf b (nameLocMap m)
toBuf b (sourcefile m)
toBuf b (semanticHighlighting m)
toBuf b (semanticAliases m)
fromBuf b fromBuf b
= do apps <- fromBuf b = do apps <- fromBuf b
@ -83,7 +154,10 @@ TTC Metadata where
tys <- fromBuf b tys <- fromBuf b
hlhs <- fromBuf b hlhs <- fromBuf b
dlocs <- fromBuf b dlocs <- fromBuf b
pure (MkMetadata apps ns tys Nothing hlhs dlocs) fname <- fromBuf b
semhl <- fromBuf b
semal <- fromBuf b
pure (MkMetadata apps ns tys Nothing hlhs dlocs fname semhl semal)
export export
addLHS : {vars : _} -> addLHS : {vars : _} ->
@ -213,6 +287,31 @@ findHoleLHS hn
= do meta <- get MD = do meta <- get MD
pure (lookupBy (\x, y => dropNS x == dropNS y) hn (holeLHS meta)) pure (lookupBy (\x, y => dropNS x == dropNS y) hn (holeLHS meta))
export
addSemanticAlias : {auto m : Ref MD Metadata} ->
NonEmptyFC -> NonEmptyFC -> Core ()
addSemanticAlias from to
= do meta <- get MD
put MD $ { semanticAliases $= insert (from, to) } meta
export
addSemanticDecorations : {auto m : Ref MD Metadata} ->
{auto c : Ref Ctxt Defs} ->
SemanticDecorations -> Core ()
addSemanticDecorations decors
= do meta <- get MD
let posmap = meta.semanticHighlighting
let (newDecors,droppedDecors) = span
((meta.sourcefile ==)
. (\((fn, _), _) => fn))
decors
unless (isNil droppedDecors)
$ log "ide-mode.highlight" 19 $ "ignored adding decorations to "
++ meta.sourcefile ++ ": " ++ show droppedDecors
put MD $ record {semanticHighlighting
= (fromList newDecors) `union` posmap} meta
-- Normalise all the types of the names, since they might have had holes -- Normalise all the types of the names, since they might have had holes
-- when added and the holes won't necessarily get saved -- when added and the holes won't necessarily get saved
normaliseTypes : {auto m : Ref MD Metadata} -> normaliseTypes : {auto m : Ref MD Metadata} ->
@ -249,13 +348,14 @@ TTC TTMFile where
pure (MkTTMFile ver md) pure (MkTTMFile ver md)
HasNames Metadata where HasNames Metadata where
full gam (MkMetadata lhs ns tys clhs hlhs dlocs) full gam md
= pure $ MkMetadata !(traverse fullLHS lhs) = pure $ record { lhsApps = !(traverse fullLHS $ md.lhsApps)
!(traverse fullTy ns) , names = !(traverse fullTy $ md.names)
!(traverse fullTy tys) , tydecls = !(traverse fullTy $ md.tydecls)
Nothing , currentLHS = Nothing
!(traverse fullHLHS hlhs) , holeLHS = !(traverse fullHLHS $ md.holeLHS)
(fromList !(traverse fullDecls (toList dlocs))) , nameLocMap = fromList !(traverse fullDecls (toList $ md.nameLocMap))
} md
where where
fullLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm)) fullLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm))
fullLHS (fc, (i, tm)) = pure (fc, (i, !(full gam tm))) fullLHS (fc, (i, tm)) = pure (fc, (i, !(full gam tm)))
@ -269,13 +369,16 @@ HasNames Metadata where
fullDecls : (NonEmptyFC, Name) -> Core (NonEmptyFC, Name) fullDecls : (NonEmptyFC, Name) -> Core (NonEmptyFC, Name)
fullDecls (fc, n) = pure (fc, !(full gam n)) fullDecls (fc, n) = pure (fc, !(full gam n))
resolved gam (MkMetadata lhs ns tys clhs hlhs dlocs) resolved gam (MkMetadata lhs ns tys clhs hlhs dlocs fname semhl semal)
= pure $ MkMetadata !(traverse resolvedLHS lhs) = pure $ MkMetadata !(traverse resolvedLHS lhs)
!(traverse resolvedTy ns) !(traverse resolvedTy ns)
!(traverse resolvedTy tys) !(traverse resolvedTy tys)
Nothing Nothing
!(traverse resolvedHLHS hlhs) !(traverse resolvedHLHS hlhs)
(fromList !(traverse resolvedDecls (toList dlocs))) (fromList !(traverse resolvedDecls (toList dlocs)))
fname
semhl
semal
where where
resolvedLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm)) resolvedLHS : (NonEmptyFC, (Nat, ClosedTerm)) -> Core (NonEmptyFC, (Nat, ClosedTerm))
resolvedLHS (fc, (i, tm)) = pure (fc, (i, !(resolved gam tm))) resolvedLHS (fc, (i, tm)) = pure (fc, (i, !(resolved gam tm)))

View File

@ -2,6 +2,7 @@ module Core.Name
import Data.List import Data.List
import Data.Strings import Data.Strings
import Data.Maybe
import Decidable.Equality import Decidable.Equality
import Libraries.Text.PrettyPrint.Prettyprinter import Libraries.Text.PrettyPrint.Prettyprinter
import Libraries.Text.PrettyPrint.Prettyprinter.Util import Libraries.Text.PrettyPrint.Prettyprinter.Util
@ -75,6 +76,21 @@ isUserName (NS _ n) = isUserName n
isUserName (DN _ n) = isUserName n isUserName (DN _ n) = isUserName n
isUserName _ = True isUserName _ = True
||| True iff name can be traced back to a source name.
||| Used to determine whether it needs semantic highlighting.
export
isSourceName : Name -> Bool
isSourceName (NS _ n) = isSourceName n
isSourceName (UN _) = True
isSourceName (MN _ _) = False
isSourceName (PV n _) = isSourceName n
isSourceName (DN _ n) = isSourceName n
isSourceName (RF _) = True
isSourceName (Nested _ n) = isSourceName n
isSourceName (CaseBlock _ _) = False
isSourceName (WithBlock _ _) = False
isSourceName (Resolved _) = False
export export
isUN : Name -> Maybe String isUN : Name -> Maybe String
isUN (UN str) = Just str isUN (UN str) = Just str
@ -93,6 +109,19 @@ nameRoot (CaseBlock n _) = "$" ++ show n
nameRoot (WithBlock n _) = "$" ++ show n nameRoot (WithBlock n _) = "$" ++ show n
nameRoot (Resolved i) = "$" ++ show i nameRoot (Resolved i) = "$" ++ show i
export
displayName : Name -> (Maybe Namespace, String)
displayName (NS ns n) = mapFst (pure . maybe ns (ns <.>)) $ displayName n
displayName (UN n) = (Nothing, n)
displayName (MN n _) = (Nothing, n)
displayName (PV n _) = displayName n
displayName (DN n _) = (Nothing, n)
displayName (RF n) = (Nothing, n)
displayName (Nested _ n) = displayName n
displayName (CaseBlock outer _) = (Nothing, "case block in " ++ show outer)
displayName (WithBlock outer _) = (Nothing, "with block in " ++ show outer)
displayName (Resolved i) = (Nothing, "$resolved" ++ show i)
--- Drop a namespace from a name --- Drop a namespace from a name
export export
dropNS : Name -> Name dropNS : Name -> Name
@ -129,7 +158,8 @@ Pretty Name where
pretty (PV n d) = braces (pretty 'P' <+> colon <+> pretty n <+> colon <+> pretty d) pretty (PV n d) = braces (pretty 'P' <+> colon <+> pretty n <+> colon <+> pretty d)
pretty (DN str _) = pretty str pretty (DN str _) = pretty str
pretty (RF n) = "." <+> pretty n pretty (RF n) = "." <+> pretty n
pretty (Nested (outer, idx) inner) = pretty outer <+> colon <+> pretty idx <+> colon <+> pretty inner pretty (Nested (outer, idx) inner)
= pretty outer <+> colon <+> pretty idx <+> colon <+> pretty inner
pretty (CaseBlock outer _) = reflow "case block in" <++> pretty outer pretty (CaseBlock outer _) = reflow "case block in" <++> pretty outer
pretty (WithBlock outer _) = reflow "with block in" <++> pretty outer pretty (WithBlock outer _) = reflow "with block in" <++> pretty outer
pretty (Resolved x) = pretty "$resolved" <+> pretty x pretty (Resolved x) = pretty "$resolved" <+> pretty x

View File

@ -53,6 +53,7 @@ toString d@(MkDirs wdir sdir bdir ldir odir dfix edirs pdirs ldirs ddirs) =
public export public export
data CG = Chez data CG = Chez
| ChezSep
| Racket | Racket
| Gambit | Gambit
| Node | Node
@ -63,6 +64,7 @@ data CG = Chez
export export
Eq CG where Eq CG where
Chez == Chez = True Chez == Chez = True
ChezSep == ChezSep = True
Racket == Racket = True Racket == Racket = True
Gambit == Gambit = True Gambit == Gambit = True
Node == Node = True Node == Node = True
@ -74,6 +76,7 @@ Eq CG where
export export
Show CG where Show CG where
show Chez = "chez" show Chez = "chez"
show ChezSep = "chez-sep"
show Racket = "racket" show Racket = "racket"
show Gambit = "gambit" show Gambit = "gambit"
show Node = "node" show Node = "node"
@ -180,6 +183,7 @@ export
availableCGs : Options -> List (String, CG) availableCGs : Options -> List (String, CG)
availableCGs o availableCGs o
= [("chez", Chez), = [("chez", Chez),
("chez-sep", ChezSep),
("racket", Racket), ("racket", Racket),
("node", Node), ("node", Node),
("javascript", Javascript), ("javascript", Javascript),

View File

@ -177,11 +177,11 @@ Reify a => Reify (List a) where
export export
Reflect a => Reflect (List a) where Reflect a => Reflect (List a) where
reflect fc defs lhs env [] = appCon fc defs (preludetypes "Nil") [Erased fc False] reflect fc defs lhs env [] = appCon fc defs (basics "Nil") [Erased fc False]
reflect fc defs lhs env (x :: xs) reflect fc defs lhs env (x :: xs)
= do x' <- reflect fc defs lhs env x = do x' <- reflect fc defs lhs env x
xs' <- reflect fc defs lhs env xs xs' <- reflect fc defs lhs env xs
appCon fc defs (preludetypes "::") [Erased fc False, x', xs'] appCon fc defs (basics "::") [Erased fc False, x', xs']
export export
Reify a => Reify (List1 a) where Reify a => Reify (List1 a) where
@ -594,6 +594,11 @@ Reflect FC where
start' <- reflect fc defs lhs env start start' <- reflect fc defs lhs env start
end' <- reflect fc defs lhs env end end' <- reflect fc defs lhs env end
appCon fc defs (reflectiontt "MkFC") [fn', start', end'] appCon fc defs (reflectiontt "MkFC") [fn', start', end']
reflect fc defs lhs env (MkVirtualFC fn start end)
= do fn' <- reflect fc defs lhs env fn
start' <- reflect fc defs lhs env start
end' <- reflect fc defs lhs env end
appCon fc defs (reflectiontt "MkFC") [fn', start', end']
reflect fc defs lhs env EmptyFC = getCon fc defs (reflectiontt "EmptyFC") reflect fc defs lhs env EmptyFC = getCon fc defs (reflectiontt "EmptyFC")
{- {-

View File

@ -82,6 +82,38 @@ isConstantType (UN n) = case n of
_ => Nothing _ => Nothing
isConstantType _ = Nothing isConstantType _ = Nothing
export
isPrimType : Constant -> Bool
isPrimType (I x) = False
isPrimType (I8 x) = False
isPrimType (I16 x) = False
isPrimType (I32 x) = False
isPrimType (I64 x) = False
isPrimType (BI x) = False
isPrimType (B8 x) = False
isPrimType (B16 x) = False
isPrimType (B32 x) = False
isPrimType (B64 x) = False
isPrimType (Str x) = False
isPrimType (Ch x) = False
isPrimType (Db x) = False
isPrimType WorldVal = False
isPrimType Int8Type = True
isPrimType Int16Type = True
isPrimType Int32Type = True
isPrimType Int64Type = True
isPrimType IntType = True
isPrimType IntegerType = True
isPrimType Bits8Type = True
isPrimType Bits16Type = True
isPrimType Bits32Type = True
isPrimType Bits64Type = True
isPrimType StringType = True
isPrimType CharType = True
isPrimType DoubleType = True
isPrimType WorldType = True
export export
constantEq : (x, y : Constant) -> Maybe (x = y) constantEq : (x, y : Constant) -> Maybe (x = y)
constantEq (I x) (I y) = case decEq x y of constantEq (I x) (I y) = case decEq x y of

View File

@ -22,6 +22,8 @@ export
TTC FC where TTC FC where
toBuf b (MkFC file startPos endPos) toBuf b (MkFC file startPos endPos)
= do tag 0; toBuf b file; toBuf b startPos; toBuf b endPos = do tag 0; toBuf b file; toBuf b startPos; toBuf b endPos
toBuf b (MkVirtualFC file startPos endPos)
= do tag 2; toBuf b file; toBuf b startPos; toBuf b endPos
toBuf b EmptyFC = tag 1 toBuf b EmptyFC = tag 1
fromBuf b fromBuf b
@ -30,6 +32,9 @@ TTC FC where
s <- fromBuf b; e <- fromBuf b s <- fromBuf b; e <- fromBuf b
pure (MkFC f s e) pure (MkFC f s e)
1 => pure EmptyFC 1 => pure EmptyFC
2 => do f <- fromBuf b;
s <- fromBuf b; e <- fromBuf b
pure (MkVirtualFC f s e)
_ => corrupt "FC" _ => corrupt "FC"
export export
TTC Namespace where TTC Namespace where
@ -789,6 +794,7 @@ TTC CDef where
export export
TTC CG where TTC CG where
toBuf b Chez = tag 0 toBuf b Chez = tag 0
toBuf b ChezSep = tag 1
toBuf b Racket = tag 2 toBuf b Racket = tag 2
toBuf b Gambit = tag 3 toBuf b Gambit = tag 3
toBuf b (Other s) = do tag 4; toBuf b s toBuf b (Other s) = do tag 4; toBuf b s
@ -799,6 +805,7 @@ TTC CG where
fromBuf b fromBuf b
= case !getTag of = case !getTag of
0 => pure Chez 0 => pure Chez
1 => pure ChezSep
2 => pure Racket 2 => pure Racket
3 => pure Gambit 3 => pure Gambit
4 => do s <- fromBuf b 4 => do s <- fromBuf b
@ -999,14 +1006,14 @@ TTC GlobalDef where
= -- Only write full details for user specified names. The others will = -- Only write full details for user specified names. The others will
-- be holes where all we will ever need after loading is the definition -- be holes where all we will ever need after loading is the definition
do toBuf b (compexpr gdef) do toBuf b (compexpr gdef)
toBuf b (map toList (refersToRuntimeM gdef)) toBuf b (map NameMap.toList (refersToRuntimeM gdef))
toBuf b (location gdef) toBuf b (location gdef)
-- We don't need any of the rest for code generation, so if -- We don't need any of the rest for code generation, so if
-- we're decoding then, we can skip these (see Compiler.Common -- we're decoding then, we can skip these (see Compiler.Common
-- for how it's decoded minimally there) -- for how it's decoded minimally there)
toBuf b (multiplicity gdef) toBuf b (multiplicity gdef)
toBuf b (fullname gdef) toBuf b (fullname gdef)
toBuf b (map toList (refersToM gdef)) toBuf b (map NameMap.toList (refersToM gdef))
toBuf b (definition gdef) toBuf b (definition gdef)
when (isUserName (fullname gdef) || cwName (fullname gdef)) $ when (isUserName (fullname gdef) || cwName (fullname gdef)) $
do toBuf b (type gdef) do toBuf b (type gdef)

View File

@ -1228,28 +1228,6 @@ mutual
isDelay (NDelayed _ _ _) = True isDelay (NDelayed _ _ _) = True
isDelay _ = False isDelay _ = False
-- Try to get the type of the application inside the given term, to use in
-- eta expansion. If there's no application, return Nothing
getEtaType : {vars : _} ->
{auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} ->
Env Term vars -> Term vars ->
Core (Maybe (Term vars))
getEtaType env (Bind fc n b sc)
= do Just ty <- getEtaType (b :: env) sc
| Nothing => pure Nothing
pure (shrinkTerm ty (DropCons SubRefl))
getEtaType env (App fc f _)
= do fty <- getType env f
logGlue "unify.eta" 10 "Function type" env fty
case !(getNF fty) of
NBind _ _ (Pi _ _ _ ty) sc =>
do defs <- get Ctxt
empty <- clearDefs defs
pure (Just !(quote empty env ty))
_ => pure Nothing
getEtaType env _ = pure Nothing
isHoleApp : NF vars -> Bool isHoleApp : NF vars -> Bool
isHoleApp (NApp _ (NMeta _ _ _) _) = True isHoleApp (NApp _ (NMeta _ _ _) _) = True
isHoleApp _ = False isHoleApp _ = False
@ -1265,17 +1243,13 @@ mutual
if isHoleApp tmy if isHoleApp tmy
then unifyNoEta (lower mode) loc env tmx tmy then unifyNoEta (lower mode) loc env tmx tmy
else do empty <- clearDefs defs else do empty <- clearDefs defs
ety <- getEtaType env !(quote empty env tmx) domty <- quote empty env tx
case ety of etay <- nf defs env
Just argty => $ Bind xfc x (Lam fcx cx Explicit domty)
do etay <- nf defs env $ App xfc (weaken !(quote empty env tmy))
(Bind xfc x (Lam fcx cx Explicit argty) (Local xfc Nothing 0 First)
(App xfc logNF "unify" 10 "Expand" env etay
(weaken !(quote empty env tmy)) unify (lower mode) loc env tmx etay
(Local xfc Nothing 0 First)))
logNF "unify" 10 "Expand" env etay
unify mode loc env tmx etay
_ => unifyNoEta mode loc env tmx tmy
unifyD _ _ mode loc env tmx tmy@(NBind yfc y (Lam fcy cy iy ty) scy) unifyD _ _ mode loc env tmx tmy@(NBind yfc y (Lam fcy cy iy ty) scy)
= do defs <- get Ctxt = do defs <- get Ctxt
logNF "unify" 10 "EtaL" env tmx logNF "unify" 10 "EtaL" env tmx
@ -1283,17 +1257,13 @@ mutual
if isHoleApp tmx if isHoleApp tmx
then unifyNoEta (lower mode) loc env tmx tmy then unifyNoEta (lower mode) loc env tmx tmy
else do empty <- clearDefs defs else do empty <- clearDefs defs
ety <- getEtaType env !(quote empty env tmy) domty <- quote empty env ty
case ety of etax <- nf defs env
Just argty => $ Bind yfc y (Lam fcy cy Explicit domty)
do etax <- nf defs env $ App yfc (weaken !(quote empty env tmx))
(Bind yfc y (Lam fcy cy Explicit argty) (Local yfc Nothing 0 First)
(App yfc logNF "unify" 10 "Expand" env etax
(weaken !(quote empty env tmx)) unify (lower mode) loc env etax tmy
(Local yfc Nothing 0 First)))
logNF "unify" 10 "Expand" env etax
unify (lower mode) loc env etax tmy
_ => unifyNoEta (lower mode) loc env tmx tmy
unifyD _ _ mode loc env tmx tmy = unifyNoEta mode loc env tmx tmy unifyD _ _ mode loc env tmx tmy = unifyNoEta mode loc env tmx tmy
unifyWithLazyD _ _ mode loc env (NDelayed _ _ tmx) (NDelayed _ _ tmy) unifyWithLazyD _ _ mode loc env (NDelayed _ _ tmx) (NDelayed _ _ tmy)

View File

@ -87,7 +87,7 @@ mkPrec Prefix p = Prefix p
toTokList : {auto s : Ref Syn SyntaxInfo} -> toTokList : {auto s : Ref Syn SyntaxInfo} ->
PTerm -> Core (List (Tok OpStr PTerm)) PTerm -> Core (List (Tok OpStr PTerm))
toTokList (POp fc opn l r) toTokList (POp fc opFC opn l r)
= do syn <- get Syn = do syn <- get Syn
let op = nameRoot opn let op = nameRoot opn
case lookup op (infixes syn) of case lookup op (infixes syn) of
@ -95,16 +95,16 @@ toTokList (POp fc opn l r)
if any isOpChar (fastUnpack op) if any isOpChar (fastUnpack op)
then throw (GenericMsg fc $ "Unknown operator '" ++ op ++ "'") then throw (GenericMsg fc $ "Unknown operator '" ++ op ++ "'")
else do rtoks <- toTokList r else do rtoks <- toTokList r
pure (Expr l :: Op fc opn backtickPrec :: rtoks) pure (Expr l :: Op fc opFC opn backtickPrec :: rtoks)
Just (Prefix, _) => Just (Prefix, _) =>
throw (GenericMsg fc $ "'" ++ op ++ "' is a prefix operator") throw (GenericMsg fc $ "'" ++ op ++ "' is a prefix operator")
Just (fix, prec) => Just (fix, prec) =>
do rtoks <- toTokList r do rtoks <- toTokList r
pure (Expr l :: Op fc opn (mkPrec fix prec) :: rtoks) pure (Expr l :: Op fc opFC opn (mkPrec fix prec) :: rtoks)
where where
backtickPrec : OpPrec backtickPrec : OpPrec
backtickPrec = NonAssoc 1 backtickPrec = NonAssoc 1
toTokList (PPrefixOp fc opn arg) toTokList (PPrefixOp fc opFC opn arg)
= do syn <- get Syn = do syn <- get Syn
let op = nameRoot opn let op = nameRoot opn
case lookup op (prefixes syn) of case lookup op (prefixes syn) of
@ -112,7 +112,7 @@ toTokList (PPrefixOp fc opn arg)
throw (GenericMsg fc $ "'" ++ op ++ "' is not a prefix operator") throw (GenericMsg fc $ "'" ++ op ++ "' is not a prefix operator")
Just prec => Just prec =>
do rtoks <- toTokList arg do rtoks <- toTokList arg
pure (Op fc opn (Prefix prec) :: rtoks) pure (Op fc opFC opn (Prefix prec) :: rtoks)
toTokList t = pure [Expr t] toTokList t = pure [Expr t]
record BangData where record BangData where
@ -123,21 +123,39 @@ record BangData where
initBangs : BangData initBangs : BangData
initBangs = MkBangData 0 [] initBangs = MkBangData 0 []
addNS : Maybe Namespace -> Name -> Name
addNS (Just ns) n@(NS _ _) = n
addNS (Just ns) n = NS ns n
addNS _ n = n
bindFun : FC -> Maybe Namespace -> RawImp -> RawImp -> RawImp
bindFun fc ns ma f =
let fc = virtualiseFC fc in
IApp fc (IApp fc (IVar fc (addNS ns $ UN ">>=")) ma) f
seqFun : FC -> Maybe Namespace -> RawImp -> RawImp -> RawImp
seqFun fc ns ma mb =
let fc = virtualiseFC fc in
IApp fc (IApp fc (IVar fc (addNS ns (UN ">>"))) ma) mb
bindBangs : List (Name, FC, RawImp) -> RawImp -> RawImp bindBangs : List (Name, FC, RawImp) -> RawImp -> RawImp
bindBangs [] tm = tm bindBangs [] tm = tm
bindBangs ((n, fc, btm) :: bs) tm bindBangs ((n, fc, btm) :: bs) tm
= bindBangs bs $ IApp fc (IApp fc (IVar fc (UN ">>=")) btm) = bindBangs bs
(ILam EmptyFC top Explicit (Just n) $ bindFun fc Nothing btm
(Implicit fc False) tm) $ ILam EmptyFC top Explicit (Just n) (Implicit fc False) tm
idiomise : FC -> RawImp -> RawImp idiomise : FC -> RawImp -> RawImp
idiomise fc (IAlternative afc u alts) idiomise fc (IAlternative afc u alts)
= IAlternative afc (mapAltType (idiomise afc) u) (idiomise afc <$> alts) = IAlternative afc (mapAltType (idiomise afc) u) (idiomise afc <$> alts)
idiomise fc (IApp afc f a) idiomise fc (IApp afc f a)
= IApp fc (IApp fc (IVar fc (UN "<*>")) = let fc = virtualiseFC fc in
(idiomise afc f)) IApp fc (IApp fc (IVar fc (UN "<*>"))
a (idiomise afc f))
idiomise fc fn = IApp fc (IVar fc (UN "pure")) fn a
idiomise fc fn
= let fc = virtualiseFC fc in
IApp fc (IVar fc (UN "pure")) fn
pairname : Name pairname : Name
pairname = NS builtinNS (UN "Pair") pairname = NS builtinNS (UN "Pair")
@ -166,9 +184,11 @@ mutual
pure $ IPi fc rig !(traverse (desugar side ps') p) pure $ IPi fc rig !(traverse (desugar side ps') p)
mn !(desugarB side ps argTy) mn !(desugarB side ps argTy)
!(desugarB side ps' retTy) !(desugarB side ps' retTy)
desugarB side ps (PLam fc rig p pat@(PRef _ n@(UN nm)) argTy scope) desugarB side ps (PLam fc rig p pat@(PRef prefFC n@(UN nm)) argTy scope)
= if lowerFirst nm || nm == "_" = if lowerFirst nm || nm == "_"
then pure $ ILam fc rig !(traverse (desugar side ps) p) then do whenJust (isConcreteFC prefFC) \nfc
=> addSemanticDecorations [(nfc, Bound, Just n)]
pure $ ILam fc rig !(traverse (desugar side ps) p)
(Just n) !(desugarB side ps argTy) (Just n) !(desugarB side ps argTy)
!(desugar side (n :: ps) scope) !(desugar side (n :: ps) scope)
else pure $ ILam EmptyFC rig !(traverse (desugar side ps) p) else pure $ ILam EmptyFC rig !(traverse (desugar side ps) p)
@ -189,15 +209,17 @@ mutual
ICase fc (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False) ICase fc (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False)
[snd !(desugarClause ps True (MkPatClause fc pat scope []))] [snd !(desugarClause ps True (MkPatClause fc pat scope []))]
desugarB side ps (PLet fc rig (PRef prefFC n) nTy nVal scope []) desugarB side ps (PLet fc rig (PRef prefFC n) nTy nVal scope [])
= pure $ ILet fc prefFC rig n !(desugarB side ps nTy) !(desugarB side ps nVal) = do whenJust (isConcreteFC prefFC) \nfc =>
!(desugar side (n :: ps) scope) addSemanticDecorations [(nfc, Bound, Just n)]
pure $ ILet fc prefFC rig n !(desugarB side ps nTy) !(desugarB side ps nVal)
!(desugar side (n :: ps) scope)
desugarB side ps (PLet fc rig pat nTy nVal scope alts) desugarB side ps (PLet fc rig pat nTy nVal scope alts)
= pure $ ICase fc !(desugarB side ps nVal) !(desugarB side ps nTy) = pure $ ICase fc !(desugarB side ps nVal) !(desugarB side ps nTy)
!(traverse (map snd . desugarClause ps True) !(traverse (map snd . desugarClause ps True)
(MkPatClause fc pat scope [] :: alts)) (MkPatClause fc pat scope [] :: alts))
desugarB side ps (PCase fc x xs) desugarB side ps (PCase fc x xs)
= pure $ ICase fc !(desugarB side ps x) = pure $ ICase fc !(desugarB side ps x)
(Implicit fc False) (Implicit (virtualiseFC fc) False)
!(traverse (map snd . desugarClause ps True) xs) !(traverse (map snd . desugarClause ps True) xs)
desugarB side ps (PLocal fc xs scope) desugarB side ps (PLocal fc xs scope)
= let ps' = definedIn xs ++ ps in = let ps' = definedIn xs ++ ps in
@ -207,8 +229,10 @@ mutual
= pure $ IUpdate pfc !(traverse (desugarUpdate side ps) fs) = pure $ IUpdate pfc !(traverse (desugarUpdate side ps) fs)
!(desugarB side ps rec) !(desugarB side ps rec)
desugarB side ps (PUpdate fc fs) desugarB side ps (PUpdate fc fs)
= desugarB side ps (PLam fc top Explicit (PRef fc (MN "rec" 0)) (PImplicit fc) = desugarB side ps
(PApp fc (PUpdate fc fs) (PRef fc (MN "rec" 0)))) $ let vfc = virtualiseFC fc in
PLam vfc top Explicit (PRef vfc (MN "rec" 0)) (PImplicit vfc)
$ PApp vfc (PUpdate fc fs) (PRef vfc (MN "rec" 0))
desugarB side ps (PApp fc x y) desugarB side ps (PApp fc x y)
= pure $ IApp fc !(desugarB side ps x) !(desugarB side ps y) = pure $ IApp fc !(desugarB side ps x) !(desugarB side ps y)
desugarB side ps (PAutoApp fc x y) desugarB side ps (PAutoApp fc x y)
@ -230,24 +254,24 @@ mutual
[apply (IVar fc (UN "===")) [l', r'], [apply (IVar fc (UN "===")) [l', r'],
apply (IVar fc (UN "~=~")) [l', r']] apply (IVar fc (UN "~=~")) [l', r']]
desugarB side ps (PBracketed fc e) = desugarB side ps e desugarB side ps (PBracketed fc e) = desugarB side ps e
desugarB side ps (POp fc op l r) desugarB side ps (POp fc opFC op l r)
= do ts <- toTokList (POp fc op l r) = do ts <- toTokList (POp fc opFC op l r)
desugarTree side ps !(parseOps ts) desugarTree side ps !(parseOps ts)
desugarB side ps (PPrefixOp fc op arg) desugarB side ps (PPrefixOp fc opFC op arg)
= do ts <- toTokList (PPrefixOp fc op arg) = do ts <- toTokList (PPrefixOp fc opFC op arg)
desugarTree side ps !(parseOps ts) desugarTree side ps !(parseOps ts)
desugarB side ps (PSectionL fc op arg) desugarB side ps (PSectionL fc opFC op arg)
= do syn <- get Syn = do syn <- get Syn
-- It might actually be a prefix argument rather than a section -- It might actually be a prefix argument rather than a section
-- so check that first, otherwise desugar as a lambda -- so check that first, otherwise desugar as a lambda
case lookup (nameRoot op) (prefixes syn) of case lookup (nameRoot op) (prefixes syn) of
Nothing => Nothing =>
desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc)
(POp fc op (PRef fc (MN "arg" 0)) arg)) (POp fc opFC op (PRef fc (MN "arg" 0)) arg))
Just prec => desugarB side ps (PPrefixOp fc op arg) Just prec => desugarB side ps (PPrefixOp fc opFC op arg)
desugarB side ps (PSectionR fc arg op) desugarB side ps (PSectionR fc opFC arg op)
= desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc) = desugarB side ps (PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc)
(POp fc op arg (PRef fc (MN "arg" 0)))) (POp fc opFC op arg (PRef fc (MN "arg" 0))))
desugarB side ps (PSearch fc depth) = pure $ ISearch fc depth desugarB side ps (PSearch fc depth) = pure $ ISearch fc depth
desugarB side ps (PPrimVal fc (BI x)) desugarB side ps (PPrimVal fc (BI x))
= case !fromIntegerName of = case !fromIntegerName of
@ -255,20 +279,23 @@ mutual
pure $ IAlternative fc (UniqueDefault (IPrimVal fc (BI x))) pure $ IAlternative fc (UniqueDefault (IPrimVal fc (BI x)))
[IPrimVal fc (BI x), [IPrimVal fc (BI x),
IPrimVal fc (I (fromInteger x))] IPrimVal fc (I (fromInteger x))]
Just fi => pure $ IApp fc (IVar fc fi) Just fi =>
(IPrimVal fc (BI x)) let vfc = virtualiseFC fc in
pure $ IApp vfc (IVar vfc fi) (IPrimVal fc (BI x))
desugarB side ps (PPrimVal fc (Ch x)) desugarB side ps (PPrimVal fc (Ch x))
= case !fromCharName of = case !fromCharName of
Nothing => Nothing =>
pure $ IPrimVal fc (Ch x) pure $ IPrimVal fc (Ch x)
Just f => pure $ IApp fc (IVar fc f) Just f =>
(IPrimVal fc (Ch x)) let vfc = virtualiseFC fc in
pure $ IApp vfc (IVar vfc f) (IPrimVal fc (Ch x))
desugarB side ps (PPrimVal fc (Db x)) desugarB side ps (PPrimVal fc (Db x))
= case !fromDoubleName of = case !fromDoubleName of
Nothing => Nothing =>
pure $ IPrimVal fc (Db x) pure $ IPrimVal fc (Db x)
Just f => pure $ IApp fc (IVar fc f) Just f =>
(IPrimVal fc (Db x)) let vfc = virtualiseFC fc in
pure $ IApp vfc (IVar vfc f) (IPrimVal fc (Db x))
desugarB side ps (PPrimVal fc x) = pure $ IPrimVal fc x desugarB side ps (PPrimVal fc x) = pure $ IPrimVal fc x
desugarB side ps (PQuote fc tm) desugarB side ps (PQuote fc tm)
= pure $ IQuote fc !(desugarB side ps tm) = pure $ IQuote fc !(desugarB side ps tm)
@ -313,40 +340,41 @@ mutual
let val = idiomise fc itm let val = idiomise fc itm
logRaw "desugar.idiom" 10 "Desugared to" val logRaw "desugar.idiom" 10 "Desugared to" val
pure val pure val
desugarB side ps (PList fc args) desugarB side ps (PList fc nilFC args)
= expandList side ps fc args = expandList side ps nilFC args
desugarB side ps (PPair fc l r) desugarB side ps (PPair fc l r)
= do l' <- desugarB side ps l = do l' <- desugarB side ps l
r' <- desugarB side ps r r' <- desugarB side ps r
let pval = apply (IVar fc mkpairname) [l', r'] let pval = apply (IVar fc mkpairname) [l', r']
pure $ IAlternative fc (UniqueDefault pval) pure $ IAlternative fc (UniqueDefault pval)
[apply (IVar fc pairname) [l', r'], pval] [apply (IVar fc pairname) [l', r'], pval]
desugarB side ps (PDPair fc (PRef nfc (UN n)) (PImplicit _) r) desugarB side ps (PDPair fc opFC (PRef nfc (UN n)) (PImplicit _) r)
= do r' <- desugarB side ps r = do r' <- desugarB side ps r
let pval = apply (IVar fc mkdpairname) [IVar nfc (UN n), r'] let pval = apply (IVar opFC mkdpairname) [IVar nfc (UN n), r']
pure $ IAlternative fc (UniqueDefault pval) pure $ IAlternative fc (UniqueDefault pval)
[apply (IVar fc dpairname) [apply (IVar fc dpairname)
[Implicit nfc False, [Implicit nfc False,
ILam nfc top Explicit (Just (UN n)) (Implicit nfc False) r'], ILam nfc top Explicit (Just (UN n)) (Implicit nfc False) r'],
pval] pval]
desugarB side ps (PDPair fc (PRef nfc (UN n)) ty r) desugarB side ps (PDPair fc opFC (PRef nfc (UN n)) ty r)
= do ty' <- desugarB side ps ty = do ty' <- desugarB side ps ty
r' <- desugarB side ps r r' <- desugarB side ps r
pure $ apply (IVar fc dpairname) pure $ apply (IVar opFC dpairname)
[ty', [ty',
ILam nfc top Explicit (Just (UN n)) ty' r'] ILam nfc top Explicit (Just (UN n)) ty' r']
desugarB side ps (PDPair fc l (PImplicit _) r) desugarB side ps (PDPair fc opFC l (PImplicit _) r)
= do l' <- desugarB side ps l = do l' <- desugarB side ps l
r' <- desugarB side ps r r' <- desugarB side ps r
pure $ apply (IVar fc mkdpairname) [l', r'] pure $ apply (IVar opFC mkdpairname) [l', r']
desugarB side ps (PDPair fc l ty r) desugarB side ps (PDPair fc opFC l ty r)
= throw (GenericMsg fc "Invalid dependent pair type") = throw (GenericMsg fc "Invalid dependent pair type")
desugarB side ps (PUnit fc) desugarB side ps (PUnit fc)
= pure $ IAlternative fc (UniqueDefault (IVar fc (UN "MkUnit"))) = pure $ IAlternative fc (UniqueDefault (IVar fc (UN "MkUnit")))
[IVar fc (UN "Unit"), [IVar fc (UN "Unit"),
IVar fc (UN "MkUnit")] IVar fc (UN "MkUnit")]
desugarB side ps (PIfThenElse fc x t e) desugarB side ps (PIfThenElse fc x t e)
= pure $ ICase fc !(desugarB side ps x) (IVar fc (UN "Bool")) = let fc = virtualiseFC fc in
pure $ ICase fc !(desugarB side ps x) (IVar fc (UN "Bool"))
[PatClause fc (IVar fc (UN "True")) !(desugar side ps t), [PatClause fc (IVar fc (UN "True")) !(desugar side ps t),
PatClause fc (IVar fc (UN "False")) !(desugar side ps e)] PatClause fc (IVar fc (UN "False")) !(desugar side ps e)]
desugarB side ps (PComprehension fc ret conds) desugarB side ps (PComprehension fc ret conds)
@ -361,22 +389,15 @@ mutual
desugarB side ps (PRewrite fc rule tm) desugarB side ps (PRewrite fc rule tm)
= pure $ IRewrite fc !(desugarB side ps rule) !(desugarB side ps tm) = pure $ IRewrite fc !(desugarB side ps rule) !(desugarB side ps tm)
desugarB side ps (PRange fc start next end) desugarB side ps (PRange fc start next end)
= case next of = let fc = virtualiseFC fc in
Nothing => desugarB side ps $ case next of
desugarB side ps (PApp fc Nothing => papply fc (PRef fc (UN "rangeFromTo")) [start,end]
(PApp fc (PRef fc (UN "rangeFromTo")) Just n => papply fc (PRef fc (UN "rangeFromThenTo")) [start, n, end]
start) end)
Just n =>
desugarB side ps (PApp fc
(PApp fc
(PApp fc (PRef fc (UN "rangeFromThenTo"))
start) n) end)
desugarB side ps (PRangeStream fc start next) desugarB side ps (PRangeStream fc start next)
= case next of = let fc = virtualiseFC fc in
Nothing => desugarB side ps $ case next of
desugarB side ps (PApp fc (PRef fc (UN "rangeFrom")) start) Nothing => papply fc (PRef fc (UN "rangeFrom")) [start]
Just n => Just n => papply fc (PRef fc (UN "rangeFromThen")) [start, n]
desugarB side ps (PApp fc (PApp fc (PRef fc (UN "rangeFromThen")) start) n)
desugarB side ps (PUnifyLog fc lvl tm) desugarB side ps (PUnifyLog fc lvl tm)
= pure $ IUnifyLog fc lvl !(desugarB side ps tm) = pure $ IUnifyLog fc lvl !(desugarB side ps tm)
desugarB side ps (PPostfixApp fc rec projs) desugarB side ps (PPostfixApp fc rec projs)
@ -404,23 +425,21 @@ mutual
{auto c : Ref Ctxt Defs} -> {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
{auto m : Ref MD Metadata} -> {auto m : Ref MD Metadata} ->
Side -> List Name -> FC -> List PTerm -> Core RawImp Side -> List Name ->
expandList side ps fc [] = pure (IVar fc (UN "Nil")) (nilFC : FC) -> List (FC, PTerm) -> Core RawImp
expandList side ps fc (x :: xs) expandList side ps nilFC [] = pure (IVar nilFC (UN "Nil"))
= pure $ apply (IVar fc (UN "::")) expandList side ps nilFC ((consFC, x) :: xs)
[!(desugarB side ps x), !(expandList side ps fc xs)] = pure $ apply (IVar consFC (UN "::"))
[!(desugarB side ps x), !(expandList side ps nilFC xs)]
addNS : Maybe Namespace -> Name -> Name
addNS (Just ns) n@(NS _ _) = n
addNS (Just ns) n = NS ns n
addNS _ n = n
addFromString : {auto c : Ref Ctxt Defs} -> addFromString : {auto c : Ref Ctxt Defs} ->
FC -> RawImp -> Core RawImp FC -> RawImp -> Core RawImp
addFromString fc tm addFromString fc tm
= pure $ case !fromStringName of = pure $ case !fromStringName of
Nothing => tm Nothing => tm
Just f => IApp fc (IVar fc f) tm Just f =>
let fc = virtualiseFC fc in
IApp fc (IVar fc f) tm
expandString : {auto s : Ref Syn SyntaxInfo} -> expandString : {auto s : Ref Syn SyntaxInfo} ->
{auto b : Ref Bang BangData} -> {auto b : Ref Bang BangData} ->
@ -428,9 +447,11 @@ mutual
{auto m : Ref MD Metadata} -> {auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
Side -> List Name -> FC -> List PStr -> Core RawImp Side -> List Name -> FC -> List PStr -> Core RawImp
expandString side ps fc xs = pure $ case !(traverse toRawImp (filter notEmpty $ mergeStrLit xs)) of expandString side ps fc xs
[] => IPrimVal fc (Str "") = do xs <- traverse toRawImp (filter notEmpty $ mergeStrLit xs)
xs@(_::_) => foldr1 concatStr xs pure $ case xs of
[] => IPrimVal fc (Str "")
(_ :: _) => foldr1 concatStr xs
where where
toRawImp : PStr -> Core RawImp toRawImp : PStr -> Core RawImp
toRawImp (StrLiteral fc str) = pure $ IPrimVal fc (Str str) toRawImp (StrLiteral fc str) = pure $ IPrimVal fc (Str str)
@ -449,7 +470,10 @@ mutual
notEmpty (StrInterp _ _) = True notEmpty (StrInterp _ _) = True
concatStr : RawImp -> RawImp -> RawImp concatStr : RawImp -> RawImp -> RawImp
concatStr a b = IApp (getFC a) (IApp (getFC b) (IVar (getFC b) (UN "++")) a) b concatStr a b =
let aFC = virtualiseFC (getFC a)
bFC = virtualiseFC (getFC b)
in IApp aFC (IApp bFC (IVar bFC (UN "++")) a) b
trimMultiline : FC -> Nat -> List (List PStr) -> Core (List PStr) trimMultiline : FC -> Nat -> List (List PStr) -> Core (List PStr)
trimMultiline fc indent lines trimMultiline fc indent lines
@ -469,8 +493,7 @@ mutual
then throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters" then throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters"
else pure initLines else pure initLines
trimLast _ (initLines `snoc` xs) | Snoc xs initLines _ trimLast _ (initLines `snoc` xs) | Snoc xs initLines _
= let fc = fromMaybe fc $ findBy (\case StrInterp fc _ => Just fc; = let fc = fromMaybe fc $ findBy isStrInterp xs in
StrLiteral _ _ => Nothing) xs in
throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters" throw $ BadMultiline fc "Closing delimiter of multiline strings cannot be preceded by non-whitespace characters"
dropLastNL : List PStr -> List PStr dropLastNL : List PStr -> List PStr
@ -508,13 +531,14 @@ mutual
= do tm' <- desugar side ps tm = do tm' <- desugar side ps tm
rest' <- expandDo side ps topfc ns rest rest' <- expandDo side ps topfc ns rest
gam <- get Ctxt gam <- get Ctxt
pure $ IApp fc (IApp fc (IVar fc (addNS ns (UN ">>"))) tm') rest' pure $ seqFun fc ns tm' rest'
expandDo side ps topfc ns (DoBind fc nameFC n tm :: rest) expandDo side ps topfc ns (DoBind fc nameFC n tm :: rest)
= do tm' <- desugar side ps tm = do tm' <- desugar side ps tm
rest' <- expandDo side ps topfc ns rest rest' <- expandDo side ps topfc ns rest
pure $ IApp fc (IApp fc (IVar fc (addNS ns (UN ">>="))) tm') whenJust (isConcreteFC nameFC) \nfc => addSemanticDecorations [(nfc, Bound, Just n)]
(ILam nameFC top Explicit (Just n) pure $ bindFun fc ns tm'
(Implicit fc False) rest') $ ILam nameFC top Explicit (Just n)
(Implicit (virtualiseFC fc) False) rest'
expandDo side ps topfc ns (DoBindPat fc pat exp alts :: rest) expandDo side ps topfc ns (DoBindPat fc pat exp alts :: rest)
= do pat' <- desugar LHS ps pat = do pat' <- desugar LHS ps pat
(newps, bpat) <- bindNames False pat' (newps, bpat) <- bindNames False pat'
@ -522,19 +546,22 @@ mutual
alts' <- traverse (map snd . desugarClause ps True) alts alts' <- traverse (map snd . desugarClause ps True) alts
let ps' = newps ++ ps let ps' = newps ++ ps
rest' <- expandDo side ps' topfc ns rest rest' <- expandDo side ps' topfc ns rest
pure $ IApp fc (IApp fc (IVar fc (addNS ns (UN ">>="))) exp') let fcOriginal = fc
(ILam EmptyFC top Explicit (Just (MN "_" 0)) let fc = virtualiseFC fc
pure $ bindFun fc ns exp'
$ ILam EmptyFC top Explicit (Just (MN "_" 0))
(Implicit fc False) (Implicit fc False)
(ICase fc (IVar EmptyFC (MN "_" 0)) (ICase fc (IVar EmptyFC (MN "_" 0))
(Implicit fc False) (Implicit fc False)
(PatClause fc bpat rest' (PatClause fcOriginal bpat rest'
:: alts'))) :: alts'))
expandDo side ps topfc ns (DoLet fc lhsFC n rig ty tm :: rest) expandDo side ps topfc ns (DoLet fc lhsFC n rig ty tm :: rest)
= do b <- newRef Bang initBangs = do b <- newRef Bang initBangs
tm' <- desugarB side ps tm tm' <- desugarB side ps tm
ty' <- desugar side ps ty ty' <- desugar side ps ty
rest' <- expandDo side ps topfc ns rest rest' <- expandDo side ps topfc ns rest
let bind = ILet fc lhsFC rig n ty' tm' rest' whenJust (isConcreteFC lhsFC) \nfc => addSemanticDecorations [(nfc, Bound, Just n)]
let bind = ILet fc (virtualiseFC lhsFC) rig n ty' tm' rest'
bd <- get Bang bd <- get Bang
pure $ bindBangs (bangNames bd) bind pure $ bindBangs (bangNames bd) bind
expandDo side ps topfc ns (DoLetPat fc pat ty tm alts :: rest) expandDo side ps topfc ns (DoLetPat fc pat ty tm alts :: rest)
@ -547,6 +574,7 @@ mutual
let ps' = newps ++ ps let ps' = newps ++ ps
rest' <- expandDo side ps' topfc ns rest rest' <- expandDo side ps' topfc ns rest
bd <- get Bang bd <- get Bang
let fc = virtualiseFC fc
pure $ bindBangs (bangNames bd) $ pure $ bindBangs (bangNames bd) $
ICase fc tm' ty' ICase fc tm' ty'
(PatClause fc bpat rest' (PatClause fc bpat rest'
@ -566,20 +594,20 @@ mutual
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
{auto m : Ref MD Metadata} -> {auto m : Ref MD Metadata} ->
Side -> List Name -> Tree OpStr PTerm -> Core RawImp Side -> List Name -> Tree OpStr PTerm -> Core RawImp
desugarTree side ps (Infix loc (UN "=") l r) -- special case since '=' is special syntax desugarTree side ps (Infix loc eqFC (UN "=") l r) -- special case since '=' is special syntax
= do l' <- desugarTree side ps l = do l' <- desugarTree side ps l
r' <- desugarTree side ps r r' <- desugarTree side ps r
pure (IAlternative loc FirstSuccess pure (IAlternative loc FirstSuccess
[apply (IVar loc (UN "===")) [l', r'], [apply (IVar eqFC (UN "===")) [l', r'],
apply (IVar loc (UN "~=~")) [l', r']]) apply (IVar eqFC (UN "~=~")) [l', r']])
desugarTree side ps (Infix loc (UN "$") l r) -- special case since '$' is special syntax desugarTree side ps (Infix loc _ (UN "$") l r) -- special case since '$' is special syntax
= do l' <- desugarTree side ps l = do l' <- desugarTree side ps l
r' <- desugarTree side ps r r' <- desugarTree side ps r
pure (IApp loc l' r') pure (IApp loc l' r')
desugarTree side ps (Infix loc op l r) desugarTree side ps (Infix loc opFC op l r)
= do l' <- desugarTree side ps l = do l' <- desugarTree side ps l
r' <- desugarTree side ps r r' <- desugarTree side ps r
pure (IApp loc (IApp loc (IVar loc op) l') r') pure (IApp loc (IApp loc (IVar opFC op) l') r')
-- negation is a special case, since we can't have an operator with -- negation is a special case, since we can't have an operator with
-- two meanings otherwise -- two meanings otherwise
@ -587,7 +615,7 @@ mutual
-- Note: In case of negated signed integer literals, we apply the -- Note: In case of negated signed integer literals, we apply the
-- negation directly. Otherwise, the literal might be -- negation directly. Otherwise, the literal might be
-- truncated to 0 before being passed on to `negate`. -- truncated to 0 before being passed on to `negate`.
desugarTree side ps (Pre loc (UN "-") $ Leaf $ PPrimVal fc c) desugarTree side ps (Pre loc opFC (UN "-") $ Leaf $ PPrimVal fc c)
= let newFC = fromMaybe EmptyFC (mergeFC loc fc) = let newFC = fromMaybe EmptyFC (mergeFC loc fc)
continue = desugarTree side ps . Leaf . PPrimVal newFC continue = desugarTree side ps . Leaf . PPrimVal newFC
in case c of in case c of
@ -601,15 +629,15 @@ mutual
-- not a signed integer literal. proceed by desugaring -- not a signed integer literal. proceed by desugaring
-- and applying to `negate`. -- and applying to `negate`.
_ => do arg' <- desugarTree side ps (Leaf $ PPrimVal fc c) _ => do arg' <- desugarTree side ps (Leaf $ PPrimVal fc c)
pure (IApp loc (IVar loc (UN "negate")) arg') pure (IApp loc (IVar opFC (UN "negate")) arg')
desugarTree side ps (Pre loc (UN "-") arg) desugarTree side ps (Pre loc opFC (UN "-") arg)
= do arg' <- desugarTree side ps arg = do arg' <- desugarTree side ps arg
pure (IApp loc (IVar loc (UN "negate")) arg') pure (IApp loc (IVar opFC (UN "negate")) arg')
desugarTree side ps (Pre loc op arg) desugarTree side ps (Pre loc opFC op arg)
= do arg' <- desugarTree side ps arg = do arg' <- desugarTree side ps arg
pure (IApp loc (IVar loc op) arg') pure (IApp loc (IVar opFC op) arg')
desugarTree side ps (Leaf t) = desugarB side ps t desugarTree side ps (Leaf t) = desugarB side ps t
desugarType : {auto s : Ref Syn SyntaxInfo} -> desugarType : {auto s : Ref Syn SyntaxInfo} ->

View File

@ -286,9 +286,9 @@ getDocsForPTerm (PRef fc name) = pure $ [!(render styleAnn !(getDocsForName fc n
getDocsForPTerm (PPrimVal _ constant) = getDocsForPrimitive constant getDocsForPTerm (PPrimVal _ constant) = getDocsForPrimitive constant
getDocsForPTerm (PType _) = pure ["Type : Type\n\tThe type of all types is Type. The type of Type is Type."] getDocsForPTerm (PType _) = pure ["Type : Type\n\tThe type of all types is Type. The type of Type is Type."]
getDocsForPTerm (PString _ _) = pure ["String Literal\n\tDesugars to a fromString call"] getDocsForPTerm (PString _ _) = pure ["String Literal\n\tDesugars to a fromString call"]
getDocsForPTerm (PList _ _) = pure ["List Literal\n\tDesugars to (::) and Nil"] getDocsForPTerm (PList _ _ _) = pure ["List Literal\n\tDesugars to (::) and Nil"]
getDocsForPTerm (PPair _ _ _) = pure ["Pair Literal\n\tDesugars to MkPair or Pair"] getDocsForPTerm (PPair _ _ _) = pure ["Pair Literal\n\tDesugars to MkPair or Pair"]
getDocsForPTerm (PDPair _ _ _ _) = pure ["Dependant Pair Literal\n\tDesugars to MkDPair or DPair"] getDocsForPTerm (PDPair _ _ _ _ _) = pure ["Dependant Pair Literal\n\tDesugars to MkDPair or DPair"]
getDocsForPTerm (PUnit _) = pure ["Unit Literal\n\tDesugars to MkUnit or Unit"] getDocsForPTerm (PUnit _) = pure ["Unit Literal\n\tDesugars to MkUnit or Unit"]
getDocsForPTerm pterm = pure ["Docs not implemented for " ++ show pterm ++ " yet"] getDocsForPTerm pterm = pure ["Docs not implemented for " ++ show pterm ++ " yet"]

View File

@ -178,7 +178,7 @@ stMain cgs opts
when (checkVerbose opts) $ -- override Quiet if implicitly set when (checkVerbose opts) $ -- override Quiet if implicitly set
setOutput (REPL False) setOutput (REPL False)
u <- newRef UST initUState u <- newRef UST initUState
m <- newRef MD initMetadata m <- newRef MD (initMetadata $ fromMaybe "(interactive)" fname)
updateREPLOpts updateREPLOpts
session <- getSession session <- getSession
when (not $ nobanner session) $ do when (not $ nobanner session) $ do

View File

@ -124,7 +124,7 @@ elabImplementation : {vars : _} ->
Maybe (List ImpDecl) -> Maybe (List ImpDecl) ->
Core () Core ()
-- TODO: Refactor all these steps into separate functions -- TODO: Refactor all these steps into separate functions
elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named impName_in nusing mbody elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named impName_in nusing mbody
= do -- let impName_in = maybe (mkImplName fc iname ps) id impln = do -- let impName_in = maybe (mkImplName fc iname ps) id impln
-- If we're in a nested block, update the name -- If we're in a nested block, update the name
let impName_nest = case lookup impName_in (names nest) of let impName_nest = case lookup impName_in (names nest) of
@ -138,15 +138,15 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
inames <- lookupCtxtName iname (gamma defs) inames <- lookupCtxtName iname (gamma defs)
let [cndata] = concatMap (\n => lookupName n (ifaces syn)) let [cndata] = concatMap (\n => lookupName n (ifaces syn))
(map fst inames) (map fst inames)
| [] => undefinedName fc iname | [] => undefinedName vfc iname
| ns => throw (AmbiguousName fc (map fst ns)) | ns => throw (AmbiguousName vfc (map fst ns))
let cn : Name = fst cndata let cn : Name = fst cndata
let cdata : IFaceInfo = snd cndata let cdata : IFaceInfo = snd cndata
Just ity <- lookupTyExact cn (gamma defs) Just ity <- lookupTyExact cn (gamma defs)
| Nothing => undefinedName fc cn | Nothing => undefinedName vfc cn
Just conty <- lookupTyExact (iconstructor cdata) (gamma defs) Just conty <- lookupTyExact (iconstructor cdata) (gamma defs)
| Nothing => undefinedName fc (iconstructor cdata) | Nothing => undefinedName vfc (iconstructor cdata)
let impsp = nub (concatMap findIBinds ps ++ let impsp = nub (concatMap findIBinds ps ++
concatMap findIBinds (map snd cons)) concatMap findIBinds (map snd cons))
@ -171,14 +171,14 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
then [Inline] then [Inline]
else [Inline, Hint True] else [Inline, Hint True]
let initTy = bindImpls fc is $ bindConstraints fc AutoImplicit cons let initTy = bindImpls vfc is $ bindConstraints vfc AutoImplicit cons
(apply (IVar fc iname) ps) (apply (IVar vfc iname) ps)
let paramBinds = if !isUnboundImplicits let paramBinds = if !isUnboundImplicits
then findBindableNames True vars [] initTy then findBindableNames True vars [] initTy
else [] else []
let impTy = doBind paramBinds initTy let impTy = doBind paramBinds initTy
let impTyDecl = IClaim fc top vis opts (MkImpTy EmptyFC EmptyFC impName impTy) let impTyDecl = IClaim vfc top vis opts (MkImpTy EmptyFC EmptyFC impName impTy)
log "elab.implementation" 5 $ "Implementation type: " ++ show impTy log "elab.implementation" 5 $ "Implementation type: " ++ show impTy
@ -186,17 +186,17 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
-- If the body is empty, we're done for now (just declaring that -- If the body is empty, we're done for now (just declaring that
-- the implementation exists and define it later) -- the implementation exists and define it later)
when (defPass pass) $ maybe (pure ()) when (defPass pass) $
(\body_in => do whenJust mbody $ \body_in => do
defs <- get Ctxt defs <- get Ctxt
Just impTyc <- lookupTyExact impName (gamma defs) Just impTyc <- lookupTyExact impName (gamma defs)
| Nothing => throw (InternalError ("Can't happen, can't find type of " ++ show impName)) | Nothing => throw (InternalError ("Can't happen, can't find type of " ++ show impName))
methImps <- getMethImps [] impTyc methImps <- getMethImps [] impTyc
log "elab.implementation" 3 $ "Bind implicits to each method: " ++ show methImps log "elab.implementation" 3 $ "Bind implicits to each method: " ++ show methImps
-- 1.5. Lookup default definitions and add them to to body -- 1.5. Lookup default definitions and add them to the body
let (body, missing) let (body, missing)
= addDefaults fc impName = addDefaults vfc impName
(zip (params cdata) ps) (zip (params cdata) ps)
(map (dropNS . name) (methods cdata)) (map (dropNS . name) (methods cdata))
(defaults cdata) body_in (defaults cdata) body_in
@ -208,7 +208,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
defs <- get Ctxt defs <- get Ctxt
let hs = openHints defs -- snapshot open hint state let hs = openHints defs -- snapshot open hint state
log "elab.implementation" 10 $ "Open hints: " ++ (show (impName :: nusing)) log "elab.implementation" 10 $ "Open hints: " ++ (show (impName :: nusing))
traverse_ (\n => do n' <- checkUnambig fc n traverse_ (\n => do n' <- checkUnambig vfc n
addOpenHint n') nusing addOpenHint n') nusing
-- 2. Elaborate top level function types for this interface -- 2. Elaborate top level function types for this interface
@ -223,16 +223,16 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
let mtops = map (Builtin.fst . snd) fns let mtops = map (Builtin.fst . snd) fns
let con = iconstructor cdata let con = iconstructor cdata
let ilhs = impsApply (IVar EmptyFC impName) let ilhs = impsApply (IVar EmptyFC impName)
(map (\x => (x, IBindVar fc (show x))) (map (\x => (x, IBindVar vfc (show x)))
(map fst methImps)) (map fst methImps))
-- RHS is the constructor applied to a search for the necessary -- RHS is the constructor applied to a search for the necessary
-- parent constraints, then the method implementations -- parent constraints, then the method implementations
defs <- get Ctxt defs <- get Ctxt
let fldTys = getFieldArgs !(normaliseHoles defs [] conty) let fldTys = getFieldArgs !(normaliseHoles defs [] conty)
log "elab.implementation" 5 $ "Field types " ++ show fldTys log "elab.implementation" 5 $ "Field types " ++ show fldTys
let irhs = apply (autoImpsApply (IVar fc con) $ map (const (ISearch fc 500)) (parents cdata)) let irhs = apply (autoImpsApply (IVar vfc con) $ map (const (ISearch vfc 500)) (parents cdata))
(map (mkMethField methImps fldTys) fns) (map (mkMethField methImps fldTys) fns)
let impFn = IDef fc impName [PatClause fc ilhs irhs] let impFn = IDef vfc impName [PatClause vfc ilhs irhs]
log "elab.implementation" 5 $ "Implementation record: " ++ show impFn log "elab.implementation" 5 $ "Implementation record: " ++ show impFn
-- If it's a named implementation, add it as a global hint while -- If it's a named implementation, add it as a global hint while
@ -241,7 +241,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
-- Make sure we don't use this name to solve parent constraints -- Make sure we don't use this name to solve parent constraints
-- when elaborating the record, or we'll end up in a cycle! -- when elaborating the record, or we'll end up in a cycle!
setFlag fc impName BlockedHint setFlag vfc impName BlockedHint
-- Update nested names so we elaborate the body in the right -- Update nested names so we elaborate the body in the right
-- environment -- environment
@ -249,11 +249,11 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
let nest' = record { names $= (names' ++) } nest let nest' = record { names $= (names' ++) } nest
traverse_ (processDecl [] nest' env) [impFn] traverse_ (processDecl [] nest' env) [impFn]
unsetFlag fc impName BlockedHint unsetFlag vfc impName BlockedHint
setFlag fc impName TCInline setFlag vfc impName TCInline
-- it's the methods we're interested in, not the implementation -- it's the methods we're interested in, not the implementation
setFlag fc impName (SetTotal PartialOK) setFlag vfc impName (SetTotal PartialOK)
-- 4. (TODO: Order method bodies to be in declaration order, in -- 4. (TODO: Order method bodies to be in declaration order, in
-- case of dependencies) -- case of dependencies)
@ -270,19 +270,22 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
-- inline flag has done its job, and outside the interface -- inline flag has done its job, and outside the interface
-- it can hurt, so unset it now -- it can hurt, so unset it now
unsetFlag fc impName TCInline unsetFlag vfc impName TCInline
-- Reset the open hints (remove the named implementation) -- Reset the open hints (remove the named implementation)
setOpenHints hs setOpenHints hs
pure ()) mbody
where where
vfc : FC
vfc = virtualiseFC ifc
applyEnv : Name -> applyEnv : Name ->
Core (Name, (Maybe Name, List (Var vars), FC -> NameType -> Term vars)) Core (Name, (Maybe Name, List (Var vars), FC -> NameType -> Term vars))
applyEnv n applyEnv n
= do n' <- resolveName n = do n' <- resolveName n
pure (Resolved n', (Nothing, reverse (allVars env), pure (Resolved n', (Nothing, reverse (allVars env),
\fn, nt => applyToFull fc \fn, nt => applyToFull vfc
(Ref fc nt (Resolved n')) env)) (Ref vfc nt (Resolved n')) env))
-- For the method fields in the record, get the arguments we need to abstract -- For the method fields in the record, get the arguments we need to abstract
-- over -- over
@ -299,7 +302,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
impsApply : RawImp -> List (Name, RawImp) -> RawImp impsApply : RawImp -> List (Name, RawImp) -> RawImp
impsApply fn [] = fn impsApply fn [] = fn
impsApply fn ((n, arg) :: ns) impsApply fn ((n, arg) :: ns)
= impsApply (INamedApp fc fn n arg) ns = impsApply (INamedApp vfc fn n arg) ns
autoImpsApply : RawImp -> List RawImp -> RawImp autoImpsApply : RawImp -> List RawImp -> RawImp
autoImpsApply f [] = f autoImpsApply f [] = f
@ -308,7 +311,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
mkLam : List (Name, RigCount, PiInfo RawImp) -> RawImp -> RawImp mkLam : List (Name, RigCount, PiInfo RawImp) -> RawImp -> RawImp
mkLam [] tm = tm mkLam [] tm = tm
mkLam ((x, c, p) :: xs) tm mkLam ((x, c, p) :: xs) tm
= ILam EmptyFC c p (Just x) (Implicit fc False) (mkLam xs tm) = ILam EmptyFC c p (Just x) (Implicit vfc False) (mkLam xs tm)
applyTo : RawImp -> List (Name, RigCount, PiInfo RawImp) -> RawImp applyTo : RawImp -> List (Name, RigCount, PiInfo RawImp) -> RawImp
applyTo tm [] = tm applyTo tm [] = tm
@ -335,7 +338,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
mkLam argns mkLam argns
(impsApply (impsApply
(applyTo (IVar EmptyFC n) argns) (applyTo (IVar EmptyFC n) argns)
(map (\n => (n, IVar fc (UN (show n)))) imps)) (map (\n => (n, IVar vfc (UN (show n)))) imps))
where where
applyUpdate : (Name, RigCount, PiInfo RawImp) -> applyUpdate : (Name, RigCount, PiInfo RawImp) ->
(Name, RigCount, PiInfo RawImp) (Name, RigCount, PiInfo RawImp)
@ -354,12 +357,12 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
applyCon : Name -> Name -> Core (Name, RawImp) applyCon : Name -> Name -> Core (Name, RawImp)
applyCon impl n applyCon impl n
= do mn <- inCurrentNS (methName n) = do mn <- inCurrentNS (methName n)
pure (dropNS n, IVar fc mn) pure (dropNS n, IVar vfc mn)
bindImps : List (Name, RigCount, RawImp) -> RawImp -> RawImp bindImps : List (Name, RigCount, RawImp) -> RawImp -> RawImp
bindImps [] ty = ty bindImps [] ty = ty
bindImps ((n, c, t) :: ts) ty bindImps ((n, c, t) :: ts) ty
= IPi fc c Implicit (Just n) t (bindImps ts ty) = IPi vfc c Implicit (Just n) t (bindImps ts ty)
-- Return method name, specialised method name, implicit name updates, -- Return method name, specialised method name, implicit name updates,
-- and method type. Also return how the method name should be updated -- and method type. Also return how the method name should be updated
@ -390,7 +393,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
-- substitute in the explicit parameters. -- substitute in the explicit parameters.
let mty_iparams let mty_iparams
= substBindVars vars = substBindVars vars
(map (\n => (n, Implicit fc False)) imppnames) (map (\n => (n, Implicit vfc False)) imppnames)
mty_in mty_in
let mty_params let mty_params
= substNames vars (zip pnames ps) mty_iparams = substNames vars (zip pnames ps) mty_iparams
@ -400,7 +403,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
show mty_params show mty_params
let mbase = bindImps methImps $ let mbase = bindImps methImps $
bindConstraints fc AutoImplicit cons $ bindConstraints vfc AutoImplicit cons $
mty_params mty_params
let ibound = findImplicits mbase let ibound = findImplicits mbase
@ -417,8 +420,8 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
log "elab.implementation" 10 $ "Used names " ++ show ibound log "elab.implementation" 10 $ "Used names " ++ show ibound
let ibinds = map fst methImps let ibinds = map fst methImps
let methupds' = if isNil ibinds then [] let methupds' = if isNil ibinds then []
else [(n, impsApply (IVar fc n) else [(n, impsApply (IVar vfc n)
(map (\x => (x, IBindVar fc (show x))) ibinds))] (map (\x => (x, IBindVar vfc (show x))) ibinds))]
pure ((meth.name, n, upds, meth.count, meth.totalReq, mty), methupds') pure ((meth.name, n, upds, meth.count, meth.totalReq, mty), methupds')
@ -437,7 +440,7 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
mkTopMethDecl : (Name, Name, List (String, String), RigCount, Maybe TotalReq, RawImp) -> ImpDecl mkTopMethDecl : (Name, Name, List (String, String), RigCount, Maybe TotalReq, RawImp) -> ImpDecl
mkTopMethDecl (mn, n, upds, c, treq, mty) mkTopMethDecl (mn, n, upds, c, treq, mty)
= let opts = maybe opts_in (\t => Totality t :: opts_in) treq in = let opts = maybe opts_in (\t => Totality t :: opts_in) treq in
IClaim fc c vis opts (MkImpTy EmptyFC EmptyFC n mty) IClaim vfc c vis opts (MkImpTy EmptyFC EmptyFC n mty)
-- Given the method type (result of topMethType) return the mapping from -- Given the method type (result of topMethType) return the mapping from
-- top level method name to current implementation's method name -- top level method name to current implementation's method name
@ -489,9 +492,10 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
updateBody ns (IDef fc n cs) updateBody ns (IDef fc n cs)
= do cs' <- traverse (updateClause ns) cs = do cs' <- traverse (updateClause ns) cs
n' <- findMethName ns fc n n' <- findMethName ns fc n
log "ide-mode.highlight" 1 $ show (n, n', fc)
pure (IDef fc n' cs') pure (IDef fc n' cs')
updateBody ns _ updateBody ns e
= throw (GenericMsg fc = throw (GenericMsg (getFC e)
"Implementation body can only contain definitions") "Implementation body can only contain definitions")
addTransform : Name -> List (Name, Name) -> addTransform : Name -> List (Name, Name) ->
@ -501,16 +505,16 @@ elabImplementation {vars} fc vis opts_in pass env nest is cons iname ps named im
= do log "elab.implementation" 3 $ = do log "elab.implementation" 3 $
"Adding transform for " ++ show meth.name ++ " : " ++ show meth.type ++ "Adding transform for " ++ show meth.name ++ " : " ++ show meth.type ++
"\n\tfor " ++ show iname ++ " in " ++ show ns "\n\tfor " ++ show iname ++ " in " ++ show ns
let lhs = INamedApp fc (IVar fc meth.name) let lhs = INamedApp vfc (IVar vfc meth.name)
(UN "__con") (UN "__con")
(IVar fc iname) (IVar vfc iname)
let Just mname = lookup (dropNS meth.name) ns let Just mname = lookup (dropNS meth.name) ns
| Nothing => pure () | Nothing => pure ()
let rhs = IVar fc mname let rhs = IVar vfc mname
log "elab.implementation" 5 $ show lhs ++ " ==> " ++ show rhs log "elab.implementation" 5 $ show lhs ++ " ==> " ++ show rhs
handleUnify handleUnify
(processDecl [] nest env (processDecl [] nest env
(ITransform fc (UN (show meth.name ++ " " ++ show iname)) lhs rhs)) (ITransform vfc (UN (show meth.name ++ " " ++ show iname)) lhs rhs))
(\err => (\err =>
log "elab.implementation" 5 $ "Can't add transform " ++ log "elab.implementation" 5 $ "Can't add transform " ++
show lhs ++ " ==> " ++ show rhs ++ show lhs ++ " ==> " ++ show rhs ++

View File

@ -4,6 +4,7 @@ import Core.Binary
import Core.Context import Core.Context
import Core.Context.Log import Core.Context.Log
import Core.Core import Core.Core
import Core.Name
import Core.Env import Core.Env
import Core.Metadata import Core.Metadata
import Core.TT import Core.TT
@ -111,20 +112,24 @@ mkIfaceData : {vars : _} ->
List (Maybe Name, RigCount, RawImp) -> List (Maybe Name, RigCount, RawImp) ->
Name -> Name -> List (Name, (RigCount, RawImp)) -> Name -> Name -> List (Name, (RigCount, RawImp)) ->
List Name -> List (Name, RigCount, RawImp) -> Core ImpDecl List Name -> List (Name, RigCount, RawImp) -> Core ImpDecl
mkIfaceData {vars} fc vis env constraints n conName ps dets meths mkIfaceData {vars} ifc vis env constraints n conName ps dets meths
= let opts = if isNil dets = let opts = if isNil dets
then [NoHints, UniqueSearch] then [NoHints, UniqueSearch]
else [NoHints, UniqueSearch, SearchBy dets] else [NoHints, UniqueSearch, SearchBy dets]
pNames = map fst ps pNames = map fst ps
retty = apply (IVar fc n) (map (IVar EmptyFC) pNames) retty = apply (IVar vfc n) (map (IVar EmptyFC) pNames)
conty = mkTy Implicit (map jname ps) $ conty = mkTy Implicit (map jname ps) $
mkTy AutoImplicit (map bhere constraints) (mkTy Explicit (map bname meths) retty) mkTy AutoImplicit (map bhere constraints) (mkTy Explicit (map bname meths) retty)
con = MkImpTy EmptyFC EmptyFC conName !(bindTypeNames [] (pNames ++ map fst meths ++ vars) conty) in con = MkImpTy EmptyFC EmptyFC conName !(bindTypeNames [] (pNames ++ map fst meths ++ vars) conty) in
pure $ IData fc vis (MkImpData fc n pure $ IData vfc vis (MkImpData vfc n
!(bindTypeNames [] (pNames ++ map fst meths ++ vars) !(bindTypeNames [] (pNames ++ map fst meths ++ vars)
(mkDataTy fc ps)) (mkDataTy vfc ps))
opts [con]) opts [con])
where where
vfc : FC
vfc = virtualiseFC ifc
jname : (Name, (RigCount, RawImp)) -> (Maybe Name, RigCount, RawImp) jname : (Name, (RigCount, RawImp)) -> (Maybe Name, RigCount, RawImp)
jname (n, rig, t) = (Just n, rig, t) jname (n, rig, t) = (Just n, rig, t)
@ -138,7 +143,7 @@ mkIfaceData {vars} fc vis env constraints n conName ps dets meths
List (Maybe Name, RigCount, RawImp) -> RawImp -> RawImp List (Maybe Name, RigCount, RawImp) -> RawImp -> RawImp
mkTy imp [] ret = ret mkTy imp [] ret = ret
mkTy imp ((n, c, argty) :: args) ret mkTy imp ((n, c, argty) :: args) ret
= IPi fc c imp n argty (mkTy imp args ret) = IPi vfc c imp n argty (mkTy imp args ret)
-- Get the implicit arguments for a method declaration or constraint hint -- Get the implicit arguments for a method declaration or constraint hint
-- to allow us to build the data declaration -- to allow us to build the data declaration
@ -181,31 +186,31 @@ getMethToplevel : {vars : _} ->
Core (List ImpDecl) Core (List ImpDecl)
getMethToplevel {vars} env vis iname cname constraints allmeths params sig getMethToplevel {vars} env vis iname cname constraints allmeths params sig
= do let paramNames = map fst params = do let paramNames = map fst params
let ity = apply (IVar fc iname) (map (IVar EmptyFC) paramNames) let ity = apply (IVar vfc iname) (map (IVar EmptyFC) paramNames)
-- Make the constraint application explicit for any method names -- Make the constraint application explicit for any method names
-- which appear in other method types -- which appear in other method types
let ty_constr = let ty_constr =
substNames vars (map applyCon allmeths) sig.type substNames vars (map applyCon allmeths) sig.type
ty_imp <- bindTypeNames [] vars (bindPs params $ bindIFace fc ity ty_constr) ty_imp <- bindTypeNames [] vars (bindPs params $ bindIFace vfc ity ty_constr)
cn <- inCurrentNS sig.name cn <- inCurrentNS sig.name
let tydecl = IClaim fc sig.count vis (if sig.isData then [Inline, Invertible] let tydecl = IClaim vfc sig.count vis (if sig.isData then [Inline, Invertible]
else [Inline]) else [Inline])
(MkImpTy fc sig.nameLoc cn ty_imp) (MkImpTy vfc sig.nameLoc cn ty_imp)
let conapp = apply (IVar fc cname) let conapp = apply (IVar vfc cname)
(map (IBindVar EmptyFC) (map bindName allmeths)) (map (IBindVar EmptyFC) (map bindName allmeths))
let argns = getExplicitArgs 0 sig.type let argns = getExplicitArgs 0 sig.type
-- eta expand the RHS so that we put implicits in the right place -- eta expand the RHS so that we put implicits in the right place
let fnclause = PatClause fc (INamedApp fc (IVar fc cn) let fnclause = PatClause vfc (INamedApp vfc (IVar sig.location cn)
(UN "__con") (UN "__con")
conapp) conapp)
(mkLam argns (mkLam argns
(apply (IVar EmptyFC (methName sig.name)) (apply (IVar EmptyFC (methName sig.name))
(map (IVar EmptyFC) argns))) (map (IVar EmptyFC) argns)))
let fndef = IDef fc cn [fnclause] let fndef = IDef vfc cn [fnclause]
pure [tydecl, fndef] pure [tydecl, fndef]
where where
fc : FC vfc : FC
fc = sig.location vfc = virtualiseFC sig.location
-- Bind the type parameters given explicitly - there might be information -- Bind the type parameters given explicitly - there might be information
-- in there that we can't infer after all -- in there that we can't infer after all
@ -216,7 +221,7 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig
applyCon : Name -> (Name, RawImp) applyCon : Name -> (Name, RawImp)
applyCon n = let name = UN "__con" in applyCon n = let name = UN "__con" in
(n, INamedApp fc (IVar fc n) name (IVar fc name)) (n, INamedApp vfc (IVar vfc n) name (IVar vfc name))
getExplicitArgs : Int -> RawImp -> List Name getExplicitArgs : Int -> RawImp -> List Name
getExplicitArgs i (IPi _ _ Explicit n _ sc) getExplicitArgs i (IPi _ _ Explicit n _ sc)
@ -227,7 +232,7 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig
mkLam : List Name -> RawImp -> RawImp mkLam : List Name -> RawImp -> RawImp
mkLam [] tm = tm mkLam [] tm = tm
mkLam (x :: xs) tm mkLam (x :: xs) tm
= ILam EmptyFC top Explicit (Just x) (Implicit fc False) (mkLam xs tm) = ILam EmptyFC top Explicit (Just x) (Implicit vfc False) (mkLam xs tm)
bindName : Name -> String bindName : Name -> String
bindName (UN n) = "__bind_" ++ n bindName (UN n) = "__bind_" ++ n
@ -337,10 +342,10 @@ elabInterface : {vars : _} ->
(conName : Maybe Name) -> (conName : Maybe Name) ->
List ImpDecl -> List ImpDecl ->
Core () Core ()
elabInterface {vars} fc vis env nest constraints iname params dets mcon body elabInterface {vars} ifc vis env nest constraints iname params dets mcon body
= do fullIName <- getFullName iname = do fullIName <- getFullName iname
ns_iname <- inCurrentNS fullIName ns_iname <- inCurrentNS fullIName
let conName_in = maybe (mkCon fc fullIName) id mcon let conName_in = maybe (mkCon vfc fullIName) id mcon
-- Machine generated names need to be qualified when looking them up -- Machine generated names need to be qualified when looking them up
conName <- inCurrentNS conName_in conName <- inCurrentNS conName_in
let meth_sigs = mapMaybe getSig body let meth_sigs = mapMaybe getSig body
@ -357,13 +362,16 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
pure (record { name = n } mt)) meth_decls pure (record { name = n } mt)) meth_decls
defs <- get Ctxt defs <- get Ctxt
Just ty <- lookupTyExact ns_iname (gamma defs) Just ty <- lookupTyExact ns_iname (gamma defs)
| Nothing => undefinedName fc iname | Nothing => undefinedName ifc iname
let implParams = getImplParams ty let implParams = getImplParams ty
updateIfaceSyn ns_iname conName updateIfaceSyn ns_iname conName
implParams paramNames (map snd constraints) implParams paramNames (map snd constraints)
ns_meths ds ns_meths ds
where where
vfc : FC
vfc = virtualiseFC ifc
paramNames : List Name paramNames : List Name
paramNames = map fst params paramNames = map fst params
@ -387,7 +395,7 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
consts <- traverse (getMethDecl env nest params meth_names . (top,)) constraints consts <- traverse (getMethDecl env nest params meth_names . (top,)) constraints
log "elab.interface" 5 $ "Constraints: " ++ show consts log "elab.interface" 5 $ "Constraints: " ++ show consts
dt <- mkIfaceData fc vis env consts iname conName params dt <- mkIfaceData vfc vis env consts iname conName params
dets meths dets meths
log "elab.interface" 10 $ "Methods: " ++ show meths log "elab.interface" 10 $ "Methods: " ++ show meths
log "elab.interface" 5 $ "Making interface data type " ++ show dt log "elab.interface" 5 $ "Making interface data type " ++ show dt
@ -406,9 +414,9 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
log "elab.interface" 5 $ "Top level methods: " ++ show fns log "elab.interface" 5 $ "Top level methods: " ++ show fns
traverse_ (processDecl [] nest env) fns traverse_ (processDecl [] nest env) fns
traverse_ (\n => do mn <- inCurrentNS n traverse_ (\n => do mn <- inCurrentNS n
setFlag fc mn Inline setFlag vfc mn Inline
setFlag fc mn TCInline setFlag vfc mn TCInline
setFlag fc mn Overloadable) meth_names setFlag vfc mn Overloadable) meth_names
-- Check that a default definition is correct. We just discard it here once -- Check that a default definition is correct. We just discard it here once
-- we know it's okay, since we'll need to re-elaborate it for each -- we know it's okay, since we'll need to re-elaborate it for each
@ -416,8 +424,9 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
elabDefault : List Declaration -> elabDefault : List Declaration ->
(FC, List FnOpt, Name, List ImpClause) -> (FC, List FnOpt, Name, List ImpClause) ->
Core (Name, List ImpClause) Core (Name, List ImpClause)
elabDefault tydecls (fc, opts, n, cs) elabDefault tydecls (dfc, opts, n, cs)
= do -- orig <- branch = do -- orig <- branch
let dn_in = MN ("Default implementation of " ++ show n) 0 let dn_in = MN ("Default implementation of " ++ show n) 0
dn <- inCurrentNS dn_in dn <- inCurrentNS dn_in
@ -425,37 +434,40 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
the (Core (RigCount, RawImp)) $ the (Core (RigCount, RawImp)) $
case findBy (\ d => d <$ guard (n == d.name)) tydecls of case findBy (\ d => d <$ guard (n == d.name)) tydecls of
Just d => pure (d.count, d.type) Just d => pure (d.count, d.type)
Nothing => throw (GenericMsg fc ("No method named " ++ show n ++ " in interface " ++ show iname)) Nothing => throw (GenericMsg dfc ("No method named " ++ show n ++ " in interface " ++ show iname))
let ity = apply (IVar fc iname) (map (IVar fc) paramNames) let ity = apply (IVar vdfc iname) (map (IVar vdfc) paramNames)
-- Substitute the method names with their top level function -- Substitute the method names with their top level function
-- name, so they don't get implicitly bound in the name -- name, so they don't get implicitly bound in the name
methNameMap <- traverse (\d => methNameMap <- traverse (\d =>
do let n = d.name do let n = d.name
cn <- inCurrentNS n cn <- inCurrentNS n
pure (n, applyParams (IVar fc cn) paramNames)) pure (n, applyParams (IVar vdfc cn) paramNames))
tydecls tydecls
let dty = bindPs params -- bind parameters let dty = bindPs params -- bind parameters
$ bindIFace fc ity -- bind interface (?!) $ bindIFace vdfc ity -- bind interface (?!)
$ substNames vars methNameMap dty $ substNames vars methNameMap dty
dty_imp <- bindTypeNames [] (map name tydecls ++ vars) dty dty_imp <- bindTypeNames [] (map name tydecls ++ vars) dty
log "elab.interface.default" 5 $ "Default method " ++ show dn ++ " : " ++ show dty_imp log "elab.interface.default" 5 $ "Default method " ++ show dn ++ " : " ++ show dty_imp
let dtydecl = IClaim fc rig vis [] (MkImpTy EmptyFC EmptyFC dn dty_imp) let dtydecl = IClaim vdfc rig vis [] (MkImpTy EmptyFC EmptyFC dn dty_imp)
processDecl [] nest env dtydecl processDecl [] nest env dtydecl
let cs' = map (changeName dn) cs cs' <- traverse (changeName dn) cs
log "elab.interface.default" 5 $ "Default method body " ++ show cs' log "elab.interface.default" 5 $ "Default method body " ++ show cs'
processDecl [] nest env (IDef fc dn cs') processDecl [] nest env (IDef vdfc dn cs')
-- Reset the original context, we don't need to keep the definition -- Reset the original context, we don't need to keep the definition
-- Actually we do for the metadata and name map! -- Actually we do for the metadata and name map!
-- put Ctxt orig -- put Ctxt orig
pure (n, cs) pure (n, cs)
where where
vdfc : FC
vdfc = virtualiseFC dfc
-- Bind the type parameters given explicitly - there might be information -- Bind the type parameters given explicitly - there might be information
-- in there that we can't infer after all -- in there that we can't infer after all
bindPs : List (Name, (RigCount, RawImp)) -> RawImp -> RawImp bindPs : List (Name, (RigCount, RawImp)) -> RawImp -> RawImp
@ -466,38 +478,48 @@ elabInterface {vars} fc vis env nest constraints iname params dets mcon body
applyParams : RawImp -> List Name -> RawImp applyParams : RawImp -> List Name -> RawImp
applyParams tm [] = tm applyParams tm [] = tm
applyParams tm (UN n :: ns) applyParams tm (UN n :: ns)
= applyParams (INamedApp fc tm (UN n) (IBindVar fc n)) ns = applyParams (INamedApp vdfc tm (UN n) (IBindVar vdfc n)) ns
applyParams tm (_ :: ns) = applyParams tm ns applyParams tm (_ :: ns) = applyParams tm ns
changeNameTerm : Name -> RawImp -> RawImp changeNameTerm : Name -> RawImp -> Core RawImp
changeNameTerm dn (IVar fc n') changeNameTerm dn (IVar fc n')
= if n == n' then IVar fc dn else IVar fc n' = do if n /= n' then pure (IVar fc n') else do
log "ide-mode.highlight" 7 $
"elabDefault is trying to add Function: " ++ show n ++ " (" ++ show fc ++")"
whenJust (isConcreteFC fc) \nfc => do
log "ide-mode.highlight" 7 $ "elabDefault is adding Function: " ++ show n
addSemanticDecorations [(nfc, Function, Just n)]
pure (IVar fc dn)
changeNameTerm dn (IApp fc f arg) changeNameTerm dn (IApp fc f arg)
= IApp fc (changeNameTerm dn f) arg = IApp fc <$> changeNameTerm dn f <*> pure arg
changeNameTerm dn (IAutoApp fc f arg) changeNameTerm dn (IAutoApp fc f arg)
= IAutoApp fc (changeNameTerm dn f) arg = IAutoApp fc <$> changeNameTerm dn f <*> pure arg
changeNameTerm dn (INamedApp fc f x arg) changeNameTerm dn (INamedApp fc f x arg)
= INamedApp fc (changeNameTerm dn f) x arg = INamedApp fc <$> changeNameTerm dn f <*> pure x <*> pure arg
changeNameTerm dn tm = tm changeNameTerm dn tm = pure tm
changeName : Name -> ImpClause -> ImpClause changeName : Name -> ImpClause -> Core ImpClause
changeName dn (PatClause fc lhs rhs) changeName dn (PatClause fc lhs rhs)
= PatClause fc (changeNameTerm dn lhs) rhs = PatClause fc <$> changeNameTerm dn lhs <*> pure rhs
changeName dn (WithClause fc lhs wval prf flags cs) changeName dn (WithClause fc lhs wval prf flags cs)
= WithClause fc (changeNameTerm dn lhs) wval = WithClause fc
prf flags (map (changeName dn) cs) <$> changeNameTerm dn lhs
<*> pure wval
<*> pure prf
<*> pure flags
<*> traverse (changeName dn) cs
changeName dn (ImpossibleClause fc lhs) changeName dn (ImpossibleClause fc lhs)
= ImpossibleClause fc (changeNameTerm dn lhs) = ImpossibleClause fc <$> changeNameTerm dn lhs
elabConstraintHints : (conName : Name) -> List Name -> elabConstraintHints : (conName : Name) -> List Name ->
Core () Core ()
elabConstraintHints conName meth_names elabConstraintHints conName meth_names
= do let nconstraints = nameCons 0 constraints = do let nconstraints = nameCons 0 constraints
chints <- traverse (getConstraintHint fc env vis iname conName chints <- traverse (getConstraintHint vfc env vis iname conName
(map fst nconstraints) (map fst nconstraints)
meth_names meth_names
paramNames) nconstraints paramNames) nconstraints
log "elab.interface" 5 $ "Constraint hints from " ++ show constraints ++ ": " ++ show chints log "elab.interface" 5 $ "Constraint hints from " ++ show constraints ++ ": " ++ show chints
traverse_ (processDecl [] nest env) (concatMap snd chints) traverse_ (processDecl [] nest env) (concatMap snd chints)
traverse_ (\n => do mn <- inCurrentNS n traverse_ (\n => do mn <- inCurrentNS n
setFlag fc mn TCInline) (map fst chints) setFlag vfc mn TCInline) (map fst chints)

View File

@ -66,8 +66,9 @@ pshowNoNorm env tm
ploc : {auto o : Ref ROpts REPLOpts} -> ploc : {auto o : Ref ROpts REPLOpts} ->
FC -> Core (Doc IdrisAnn) FC -> Core (Doc IdrisAnn)
ploc EmptyFC = pure emptyDoc ploc fc = do
ploc fc@(MkFC fn s e) = do let Just (fn, s, e) = isNonEmptyFC fc
| Nothing => pure emptyDoc
let (sr, sc) = mapHom (fromInteger . cast) s let (sr, sc) = mapHom (fromInteger . cast) s
let (er, ec) = mapHom (fromInteger . cast) e let (er, ec) = mapHom (fromInteger . cast) e
let nsize = length $ show (er + 1) let nsize = length $ show (er + 1)
@ -91,10 +92,12 @@ ploc fc@(MkFC fn s e) = do
-- Assumes the two FCs are sorted -- Assumes the two FCs are sorted
ploc2 : {auto o : Ref ROpts REPLOpts} -> ploc2 : {auto o : Ref ROpts REPLOpts} ->
FC -> FC -> Core (Doc IdrisAnn) FC -> FC -> Core (Doc IdrisAnn)
ploc2 fc EmptyFC = ploc fc ploc2 fc1 fc2 =
ploc2 EmptyFC fc = ploc fc do let Just (fn1, s1, e1) = isNonEmptyFC fc1
ploc2 (MkFC fn1 s1 e1) (MkFC fn2 s2 e2) = | Nothing => ploc fc2
do let (sr1, sc1) = mapHom (fromInteger . cast) s1 let Just (fn2, s2, e2) = isNonEmptyFC fc2
| Nothing => ploc fc1
let (sr1, sc1) = mapHom (fromInteger . cast) s1
let (sr2, sc2) = mapHom (fromInteger . cast) s2 let (sr2, sc2) = mapHom (fromInteger . cast) s2
let (er1, ec1) = mapHom (fromInteger . cast) e1 let (er1, ec1) = mapHom (fromInteger . cast) e1
let (er2, ec2) = mapHom (fromInteger . cast) e2 let (er2, ec2) = mapHom (fromInteger . cast) e2
@ -179,10 +182,12 @@ perror (PatternVariableUnifies fc env n tm)
prettyVar (PV n _) = prettyVar n prettyVar (PV n _) = prettyVar n
prettyVar n = pretty n prettyVar n = pretty n
order : FC -> FC -> (FC, FC) order : FC -> FC -> (FC, FC)
order EmptyFC fc2 = (EmptyFC, fc2) order fc1 fc2 =
order fc1 EmptyFC = (fc1, EmptyFC) let Just (_, sr1, sc1) = isNonEmptyFC fc1
order fc1@(MkFC _ sr1 sc1) fc2@(MkFC _ sr2 sc2) = | Nothing => (EmptyFC, fc2)
if sr1 < sr2 then (fc1, fc2) else if sr1 == sr2 && sc1 < sc2 then (fc1, fc2) else (fc2, fc1) Just (_, sr2, sc2) = isNonEmptyFC fc2
| Nothing => (fc1, EmptyFC)
in if sr1 < sr2 then (fc1, fc2) else if sr1 == sr2 && sc1 < sc2 then (fc1, fc2) else (fc2, fc1)
perror (CyclicMeta fc env n tm) perror (CyclicMeta fc env n tm)
= pure $ errorDesc (reflow "Cycle detected in solution of metavariable" <++> meta (pretty !(prettyName n)) <++> equals = pure $ errorDesc (reflow "Cycle detected in solution of metavariable" <++> meta (pretty !(prettyName n)) <++> equals
<++> code !(pshow env tm)) <+> line <+> !(ploc fc) <++> code !(pshow env tm)) <+> line <+> !(ploc fc)

View File

@ -48,7 +48,7 @@ toStrUpdate (UN n, term)
where where
bracket : PTerm -> PTerm bracket : PTerm -> PTerm
bracket tm@(PRef _ _) = tm bracket tm@(PRef _ _) = tm
bracket tm@(PList _ _) = tm bracket tm@(PList _ _ _) = tm
bracket tm@(PPair _ _ _) = tm bracket tm@(PPair _ _ _) = tm
bracket tm@(PUnit _) = tm bracket tm@(PUnit _) = tm
bracket tm@(PComprehension _ _ _) = tm bracket tm@(PComprehension _ _ _) = tm

View File

@ -1,6 +1,8 @@
module Idris.IDEMode.Commands module Idris.IDEMode.Commands
import Core.Core import Core.Core
import Core.Context
import Core.Context.Log
import Core.Name import Core.Name
import public Idris.REPL.Opts import public Idris.REPL.Opts
import Libraries.Utils.Hex import Libraries.Utils.Hex
@ -281,9 +283,10 @@ sendStr f st =
map (const ()) (fPutStr f st) map (const ()) (fPutStr f st)
export export
send : SExpable a => File -> a -> Core () send : {auto c : Ref Ctxt Defs} -> SExpable a => File -> a -> Core ()
send f resp send f resp
= do let r = show (toSExp resp) ++ "\n" = do let r = show (toSExp resp) ++ "\n"
log "ide-mode.send" 20 r
coreLift $ sendStr f $ leftPad '0' 6 (asHex (cast (length r))) coreLift $ sendStr f $ leftPad '0' 6 (asHex (cast (length r)))
coreLift $ sendStr f r coreLift $ sendStr f r
coreLift $ fflush f coreLift $ fflush f

View File

@ -5,6 +5,9 @@ module Idris.IDEMode.Parser
import Idris.IDEMode.Commands import Idris.IDEMode.Commands
import Core.Core import Core.Core
import Core.Name
import Core.Metadata
import Core.FC
import Data.Maybe import Data.Maybe
import Data.List import Data.List
@ -76,11 +79,11 @@ sexp
pure (SExpList xs) pure (SExpList xs)
ideParser : {e : _} -> ideParser : {e : _} ->
(fname : String) -> String -> Grammar Token e ty -> Either Error ty (fname : String) -> String -> Grammar SemanticDecorations Token e ty -> Either Error ty
ideParser fname str p ideParser fname str p
= do toks <- mapError (fromLexError fname) $ idelex str = do toks <- mapError (fromLexError fname) $ idelex str
parsed <- mapError (fromParsingError fname) $ parse p toks (decor, (parsed, _)) <- mapError (fromParsingError fname) $ parseWith p toks
Right (fst parsed) Right parsed
export export

View File

@ -133,7 +133,8 @@ getInput f
pure (pack inp) pure (pack inp)
||| Do nothing and tell the user to wait for us to implmement this (or join the effort!) ||| Do nothing and tell the user to wait for us to implmement this (or join the effort!)
todoCmd : {auto o : Ref ROpts REPLOpts} -> todoCmd : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} ->
String -> Core () String -> Core ()
todoCmd cmdName = iputStrLn $ reflow $ cmdName ++ ": command not yet implemented. Hopefully soon!" todoCmd cmdName = iputStrLn $ reflow $ cmdName ++ ": command not yet implemented. Hopefully soon!"
@ -260,24 +261,24 @@ processCatch cmd
msg <- perror err msg <- perror err
pure $ REPL $ REPLError msg) pure $ REPL $ REPLError msg)
idePutStrLn : File -> Integer -> String -> Core () idePutStrLn : {auto c : Ref Ctxt Defs} -> File -> Integer -> String -> Core ()
idePutStrLn outf i msg idePutStrLn outf i msg
= send outf (SExpList [SymbolAtom "write-string", = send outf (SExpList [SymbolAtom "write-string",
toSExp msg, toSExp i]) toSExp msg, toSExp i])
returnFromIDE : File -> Integer -> SExp -> Core () returnFromIDE : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core ()
returnFromIDE outf i msg returnFromIDE outf i msg
= do send outf (SExpList [SymbolAtom "return", msg, toSExp i]) = do send outf (SExpList [SymbolAtom "return", msg, toSExp i])
printIDEResult : File -> Integer -> SExp -> Core () printIDEResult : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core ()
printIDEResult outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "ok", toSExp msg]) printIDEResult outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "ok", toSExp msg])
printIDEResultWithHighlight : File -> Integer -> SExp -> Core () printIDEResultWithHighlight : {auto c : Ref Ctxt Defs} -> File -> Integer -> SExp -> Core ()
printIDEResultWithHighlight outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "ok", toSExp msg printIDEResultWithHighlight outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "ok", toSExp msg
-- TODO return syntax highlighted result -- TODO return syntax highlighted result
, SExpList []]) , SExpList []])
printIDEError : Ref ROpts REPLOpts => File -> Integer -> Doc IdrisAnn -> Core () printIDEError : Ref ROpts REPLOpts => {auto c : Ref Ctxt Defs} -> File -> Integer -> Doc IdrisAnn -> Core ()
printIDEError outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "error", toSExp !(renderWithoutColor msg) ]) printIDEError outf i msg = returnFromIDE outf i (SExpList [SymbolAtom "error", toSExp !(renderWithoutColor msg) ])
SExpable REPLEval where SExpable REPLEval where

View File

@ -1,35 +1,34 @@
module Idris.IDEMode.SyntaxHighlight module Idris.IDEMode.SyntaxHighlight
import Core.Context import Core.Context
import Core.Context.Log
import Core.InitPrimitives import Core.InitPrimitives
import Core.Metadata import Core.Metadata
import Core.TT import Core.TT
import Idris.REPL import Idris.REPL
import Idris.Syntax
import Idris.DocString
import Idris.IDEMode.Commands import Idris.IDEMode.Commands
import Data.List import Data.List
import Data.Maybe
import Libraries.Data.PosMap
%default covering %default covering
data Decoration : Type where
Typ : Decoration
Function : Decoration
Data : Decoration
Keyword : Decoration
Bound : Decoration
SExpable Decoration where SExpable Decoration where
toSExp Typ = SymbolAtom "type" toSExp Typ = SExpList [ SymbolAtom "decor", SymbolAtom "type"]
toSExp Function = SymbolAtom "function" toSExp Function = SExpList [ SymbolAtom "decor", SymbolAtom "function"]
toSExp Data = SymbolAtom "data" toSExp Data = SExpList [ SymbolAtom "decor", SymbolAtom "data"]
toSExp Keyword = SymbolAtom "keyword" toSExp Keyword = SExpList [ SymbolAtom "decor", SymbolAtom "keyword"]
toSExp Bound = SymbolAtom "bound" toSExp Bound = SExpList [ SymbolAtom "decor", SymbolAtom "bound"]
record Highlight where record Highlight where
constructor MkHighlight constructor MkHighlight
location : NonEmptyFC location : NonEmptyFC
name : Name name : String
isImplicit : Bool isImplicit : Bool
key : String key : String
decor : Decoration decor : Decoration
@ -38,19 +37,26 @@ record Highlight where
ns : String ns : String
SExpable FC where SExpable FC where
toSExp (MkFC fname (startLine, startCol) (endLine, endCol)) toSExp fc = case isNonEmptyFC fc of
= SExpList [ SExpList [ SymbolAtom "filename", StringAtom fname ] Just (fname , (startLine, startCol), (endLine, endCol)) =>
, SExpList [ SymbolAtom "start", IntegerAtom (cast startLine + 1), IntegerAtom (cast startCol + 1) ] SExpList [ SExpList [ SymbolAtom "filename", StringAtom fname ]
, SExpList [ SymbolAtom "end", IntegerAtom (cast endLine + 1), IntegerAtom (cast endCol + 1) ] , SExpList [ SymbolAtom "start"
, IntegerAtom (cast startLine + 1)
, IntegerAtom (cast startCol + 1)
]
, SExpList [ SymbolAtom "end"
, IntegerAtom (cast endLine + 1)
, IntegerAtom (cast endCol)
]
] ]
toSExp EmptyFC = SExpList [] Nothing => SExpList []
SExpable Highlight where SExpable Highlight where
toSExp (MkHighlight loc nam impl k dec doc t ns) toSExp (MkHighlight loc nam impl k dec doc t ns)
= SExpList [ toSExp $ justFC loc = SExpList [ toSExp $ justFC loc
, SExpList [ SExpList [ SymbolAtom "name", StringAtom (show nam) ] , SExpList [ SExpList [ SymbolAtom "name", StringAtom nam ]
, SExpList [ SymbolAtom "namespace", StringAtom ns ] , SExpList [ SymbolAtom "namespace", StringAtom ns ]
, SExpList [ SymbolAtom "decor", toSExp dec ] , toSExp dec
, SExpList [ SymbolAtom "implicit", toSExp impl ] , SExpList [ SymbolAtom "implicit", toSExp impl ]
, SExpList [ SymbolAtom "key", StringAtom k ] , SExpList [ SymbolAtom "key", StringAtom k ]
, SExpList [ SymbolAtom "doc-overview", StringAtom doc ] , SExpList [ SymbolAtom "doc-overview", StringAtom doc ]
@ -58,12 +64,13 @@ SExpable Highlight where
] ]
] ]
inFile : String -> (NonEmptyFC, (Name, Nat, ClosedTerm)) -> Bool inFile : (s : String) -> (NonEmptyFC, a) -> Bool
inFile fname ((file, _, _), _) = file == fname inFile fname ((file, _, _), _) = file == fname
||| Output some data using current dialog index ||| Output some data using current dialog index
export export
printOutput : {auto o : Ref ROpts REPLOpts} -> printOutput : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} ->
SExp -> Core () SExp -> Core ()
printOutput msg printOutput msg
= do opts <- get ROpts = do opts <- get ROpts
@ -74,7 +81,8 @@ printOutput msg
msg, toSExp i]) msg, toSExp i])
outputHighlight : {auto opts : Ref ROpts REPLOpts} -> outputHighlight : {auto c : Ref Ctxt Defs} ->
{auto opts : Ref ROpts REPLOpts} ->
Highlight -> Core () Highlight -> Core ()
outputHighlight h = outputHighlight h =
printOutput $ SExpList [ SymbolAtom "ok" printOutput $ SExpList [ SymbolAtom "ok"
@ -86,37 +94,46 @@ outputHighlight h =
hlt : List Highlight hlt : List Highlight
hlt = [h] hlt = [h]
outputNameSyntax : {auto opts : Ref ROpts REPLOpts} -> lwOutputHighlight :
(NonEmptyFC, (Name, Nat, ClosedTerm)) -> Core () {auto c : Ref Ctxt Defs} ->
outputNameSyntax (fc, (name, _, term)) = {auto opts : Ref ROpts REPLOpts} ->
let dec = case term of (NonEmptyFC, Decoration) -> Core ()
(Local fc x idx y) => Just Bound lwOutputHighlight (nfc,decor) =
printOutput $ SExpList [ SymbolAtom "ok"
, SExpList [ SymbolAtom "highlight-source"
, toSExp $ the (List _) [
SExpList [ toSExp $ justFC nfc
, SExpList [ toSExp decor]
]]]]
-- See definition of NameType in Core.TT for possible values of Ref's nametype field
-- data NameType : Type where
-- Bound : NameType outputNameSyntax : {auto c : Ref Ctxt Defs} ->
-- Func : NameType {auto s : Ref Syn SyntaxInfo} ->
-- DataCon : (tag : Int) -> (arity : Nat) -> NameType {auto opts : Ref ROpts REPLOpts} ->
-- TyCon : (tag : Int) -> (arity : Nat) -> NameType (NonEmptyFC, Decoration, Name) -> Core ()
(Ref fc Bound name) => Just Bound outputNameSyntax (nfc, decor, nm) = do
(Ref fc Func name) => Just Function defs <- get Ctxt
(Ref fc (DataCon tag arity) name) => Just Data log "ide-mode.highlight" 20 $ "highlighting at " ++ show nfc
(Ref fc (TyCon tag arity) name) => Just Typ ++ ": " ++ show nm
(Meta fc x y xs) => Just Bound ++ "\nAs: " ++ show decor
(Bind fc x b scope) => Just Bound let fc = justFC nfc
(App fc fn arg) => Just Bound let (mns, name) = displayName nm
(As fc x as pat) => Just Bound outputHighlight $ MkHighlight
(TDelayed fc x y) => Nothing { location = nfc
(TDelay fc x ty arg) => Nothing , name
(TForce fc x y) => Nothing , isImplicit = False
(PrimVal fc c) => Just Typ , key = ""
(Erased fc imp) => Just Bound , decor
(TType fc) => Just Typ , docOverview = "" --!(getDocsForName fc nm)
hilite = Prelude.map (\ d => MkHighlight fc name False "" d "" (show term) "") dec , typ = "" -- TODO: extract type maybe "" show !(lookupTyExact nm (gamma defs))
in maybe (pure ()) outputHighlight hilite , ns = maybe "" show mns
}
export export
outputSyntaxHighlighting : {auto m : Ref MD Metadata} -> outputSyntaxHighlighting : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto s : Ref Syn SyntaxInfo} ->
{auto opts : Ref ROpts REPLOpts} -> {auto opts : Ref ROpts REPLOpts} ->
String -> String ->
REPLResult -> REPLResult ->
@ -124,9 +141,28 @@ outputSyntaxHighlighting : {auto m : Ref MD Metadata} ->
outputSyntaxHighlighting fname loadResult = do outputSyntaxHighlighting fname loadResult = do
opts <- get ROpts opts <- get ROpts
when (opts.synHighlightOn) $ do when (opts.synHighlightOn) $ do
allNames <- the (Core ?) $ filter (inFile fname) . names <$> get MD meta <- get MD
-- decls <- filter (inFile fname) . tydecls <$> get MD let allNames = filter (inFile fname) $ toList meta.nameLocMap
_ <- traverse outputNameSyntax allNames -- ++ decls) --decls <- filter (inFile fname) . tydecls <$> get MD
pure () --_ <- traverse outputNameSyntax allNames -- ++ decls)
let semHigh = meta.semanticHighlighting
log "ide-mode.highlight" 19 $
"Semantic metadata is: " ++ show semHigh
let aliases
: List ASemanticDecoration
= flip foldMap meta.semanticAliases $ \ (from, to) =>
let decors = uncurry exactRange (snd to) semHigh in
map (\ ((fnm, loc), rest) => ((fnm, snd from), rest)) decors
log "ide-mode.highlight.alias" 19 $
"Semantic metadata from aliases is: " ++ show aliases
traverse_ {b = Unit}
(\(nfc, decor, mn) =>
case mn of
Nothing => lwOutputHighlight (nfc, decor)
Just n => outputNameSyntax (nfc, decor, n))
(aliases ++ toList semHigh)
pure loadResult pure loadResult

View File

@ -178,7 +178,7 @@ buildMod loc num len mod
Nothing => True Nothing => True
Just t => any (\x => x > t) (srcTime :: map snd depTimes) Just t => any (\x => x > t) (srcTime :: map snd depTimes)
u <- newRef UST initUState u <- newRef UST initUState
m <- newRef MD initMetadata m <- newRef MD (initMetadata src)
put Syn initSyntax put Syn initSyntax
if needsBuilding if needsBuilding
@ -220,7 +220,7 @@ buildDeps fname
case ok of case ok of
[] => do -- On success, reload the main ttc in a clean context [] => do -- On success, reload the main ttc in a clean context
clearCtxt; addPrimitives clearCtxt; addPrimitives
put MD initMetadata put MD (initMetadata fname)
mainttc <- getTTCFileName fname "ttc" mainttc <- getTTCFileName fname "ttc"
log "import" 10 $ "Reloading " ++ show mainttc ++ " from " ++ fname log "import" 10 $ "Reloading " ++ show mainttc ++ " from " ++ fname
readAsMain mainttc readAsMain mainttc

View File

@ -159,7 +159,7 @@ field fname
pure [LT (MkPkgVersion (fromInteger <$> vs)) True, pure [LT (MkPkgVersion (fromInteger <$> vs)) True,
GT (MkPkgVersion (fromInteger <$> vs)) True] GT (MkPkgVersion (fromInteger <$> vs)) True]
mkBound : List Bound -> PkgVersionBounds -> PackageEmptyRule PkgVersionBounds mkBound : List Bound -> PkgVersionBounds -> EmptyRule PkgVersionBounds
mkBound (LT b i :: bs) pkgbs mkBound (LT b i :: bs) pkgbs
= maybe (mkBound bs (record { upperBound = Just b, = maybe (mkBound bs (record { upperBound = Just b,
upperInclusive = i } pkgbs)) upperInclusive = i } pkgbs))
@ -285,7 +285,7 @@ compileMain : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} -> {auto o : Ref ROpts REPLOpts} ->
Name -> String -> String -> Core () Name -> String -> String -> Core ()
compileMain mainn mmod exec compileMain mainn mmod exec
= do m <- newRef MD initMetadata = do m <- newRef MD (initMetadata mmod)
u <- newRef UST initUState u <- newRef UST initUState
ignore $ loadMainFile mmod ignore $ loadMainFile mmod
ignore $ compileExp (PRef replFC mainn) exec ignore $ compileExp (PRef replFC mainn) exec
@ -559,7 +559,7 @@ runRepl : {auto c : Ref Ctxt Defs} ->
Core () Core ()
runRepl fname = do runRepl fname = do
u <- newRef UST initUState u <- newRef UST initUState
m <- newRef MD initMetadata m <- newRef MD (initMetadata $ fromMaybe "(interactive)" fname)
the (Core ()) $ the (Core ()) $
case fname of case fname of
Nothing => pure () Nothing => pure ()

File diff suppressed because it is too large Load Diff

View File

@ -56,7 +56,7 @@ mkLets : FileName ->
List1 (WithBounds (Either LetBinder LetDecl)) -> List1 (WithBounds (Either LetBinder LetDecl)) ->
PTerm -> PTerm PTerm -> PTerm
mkLets fname = letFactory buildLets mkLets fname = letFactory buildLets
(\ decls, scope => PLocal (boundToFC fname decls) decls.val scope) (\ decls, scope => PLocal (virtualiseFC $ boundToFC fname decls) decls.val scope)
where where

View File

@ -294,10 +294,10 @@ mutual
go d (PDotted _ p) = dot <+> go d p go d (PDotted _ p) = dot <+> go d p
go d (PImplicit _) = "_" go d (PImplicit _) = "_"
go d (PInfer _) = "?" go d (PInfer _) = "?"
go d (POp _ op x y) = parenthesise (d > appPrec) $ group $ go startPrec x <++> prettyOp op <++> go startPrec y go d (POp _ _ op x y) = parenthesise (d > appPrec) $ group $ go startPrec x <++> prettyOp op <++> go startPrec y
go d (PPrefixOp _ op x) = parenthesise (d > appPrec) $ pretty op <+> go startPrec x go d (PPrefixOp _ _ op x) = parenthesise (d > appPrec) $ pretty op <+> go startPrec x
go d (PSectionL _ op x) = parens (prettyOp op <++> go startPrec x) go d (PSectionL _ _ op x) = parens (prettyOp op <++> go startPrec x)
go d (PSectionR _ x op) = parens (go startPrec x <++> prettyOp op) go d (PSectionR _ _ x op) = parens (go startPrec x <++> prettyOp op)
go d (PEq fc l r) = parenthesise (d > appPrec) $ go startPrec l <++> equals <++> go startPrec r go d (PEq fc l r) = parenthesise (d > appPrec) $ go startPrec l <++> equals <++> go startPrec r
go d (PBracketed _ tm) = parens (go startPrec tm) go d (PBracketed _ tm) = parens (go startPrec tm)
go d (PString _ xs) = parenthesise (d > appPrec) $ hsep $ punctuate "++" (prettyString <$> xs) go d (PString _ xs) = parenthesise (d > appPrec) $ hsep $ punctuate "++" (prettyString <$> xs)
@ -305,10 +305,10 @@ mutual
go d (PDoBlock _ ns ds) = parenthesise (d > appPrec) $ group $ align $ hang 2 $ do_ <++> (vsep $ punctuate semi (prettyDo <$> ds)) go d (PDoBlock _ ns ds) = parenthesise (d > appPrec) $ group $ align $ hang 2 $ do_ <++> (vsep $ punctuate semi (prettyDo <$> ds))
go d (PBang _ tm) = "!" <+> go d tm go d (PBang _ tm) = "!" <+> go d tm
go d (PIdiom _ tm) = enclose (pretty "[|") (pretty "|]") (go startPrec tm) go d (PIdiom _ tm) = enclose (pretty "[|") (pretty "|]") (go startPrec tm)
go d (PList _ xs) = brackets (group $ align $ vsep $ punctuate comma (go startPrec <$> xs)) go d (PList _ _ xs) = brackets (group $ align $ vsep $ punctuate comma (go startPrec . snd <$> xs))
go d (PPair _ l r) = group $ parens (go startPrec l <+> comma <+> line <+> go startPrec r) go d (PPair _ l r) = group $ parens (go startPrec l <+> comma <+> line <+> go startPrec r)
go d (PDPair _ l (PImplicit _) r) = group $ parens (go startPrec l <++> pretty "**" <+> line <+> go startPrec r) go d (PDPair _ _ l (PImplicit _) r) = group $ parens (go startPrec l <++> pretty "**" <+> line <+> go startPrec r)
go d (PDPair _ l ty r) = group $ parens (go startPrec l <++> colon <++> go startPrec ty <++> pretty "**" <+> line <+> go startPrec r) go d (PDPair _ _ l ty r) = group $ parens (go startPrec l <++> colon <++> go startPrec ty <++> pretty "**" <+> line <+> go startPrec r)
go d (PUnit _) = "()" go d (PUnit _) = "()"
go d (PIfThenElse _ x t e) = go d (PIfThenElse _ x t e) =
parenthesise (d > appPrec) $ group $ align $ hang 2 $ vsep parenthesise (d > appPrec) $ group $ align $ hang 2 $ vsep

View File

@ -197,7 +197,7 @@ readHeader path
-- Stop at the first :, that's definitely not part of the header, to -- Stop at the first :, that's definitely not part of the header, to
-- save lexing the whole file unnecessarily -- save lexing the whole file unnecessarily
setCurrentElabSource res -- for error printing purposes setCurrentElabSource res -- for error printing purposes
let Right mod = runParserTo path (isLitFile path) (is ':') res (progHdr path) let Right (decor, mod) = runParserTo path (isLitFile path) (is ':') res (progHdr path)
| Left err => throw err | Left err => throw err
pure mod pure mod
@ -259,15 +259,16 @@ processMod srcf ttcf msg sourcecode
pure Nothing pure Nothing
else -- needs rebuilding else -- needs rebuilding
do iputStrLn msg do iputStrLn msg
Right mod <- logTime ("++ Parsing " ++ srcf) $ Right (decor, mod) <- logTime ("++ Parsing " ++ srcf) $
pure (runParser srcf (isLitFile srcf) sourcecode (do p <- prog srcf; eoi; pure p)) pure (runParser srcf (isLitFile srcf) sourcecode (do p <- prog srcf; eoi; pure p))
| Left err => pure (Just [err]) | Left err => pure (Just [err])
addSemanticDecorations decor
initHash initHash
traverse_ addPublicHash (sort hs) traverse_ addPublicHash (sort hs)
resetNextVar resetNextVar
when (ns /= nsAsModuleIdent mainNS) $ when (ns /= nsAsModuleIdent mainNS) $
do let MkFC fname _ _ = headerloc mod do let Just fname = map file (isNonEmptyFC $ headerloc mod)
| EmptyFC => throw (InternalError "No file name") | Nothing => throw (InternalError "No file name")
d <- getDirs d <- getDirs
ns' <- pathToNS (working_dir d) (source_dir d) fname ns' <- pathToNS (working_dir d) (source_dir d) fname
when (ns /= ns') $ when (ns /= ns') $

View File

@ -1,6 +1,7 @@
module Idris.REPL module Idris.REPL
import Compiler.Scheme.Chez import Compiler.Scheme.Chez
import Compiler.Scheme.ChezSep
import Compiler.Scheme.Racket import Compiler.Scheme.Racket
import Compiler.Scheme.Gambit import Compiler.Scheme.Gambit
import Compiler.ES.Node import Compiler.ES.Node
@ -18,6 +19,7 @@ import Core.Env
import Core.InitPrimitives import Core.InitPrimitives
import Core.LinearCheck import Core.LinearCheck
import Core.Metadata import Core.Metadata
import Core.FC
import Core.Normalise import Core.Normalise
import Core.Options import Core.Options
import Core.Termination import Core.Termination
@ -210,6 +212,7 @@ findCG
= do defs <- get Ctxt = do defs <- get Ctxt
case codegen (session (options defs)) of case codegen (session (options defs)) of
Chez => pure codegenChez Chez => pure codegenChez
ChezSep => pure codegenChezSep
Racket => pure codegenRacket Racket => pure codegenRacket
Gambit => pure codegenGambit Gambit => pure codegenGambit
Node => pure codegenNode Node => pure codegenNode
@ -638,7 +641,7 @@ loadMainFile : {auto c : Ref Ctxt Defs} ->
loadMainFile f loadMainFile f
= do opts <- get ROpts = do opts <- get ROpts
put ROpts (record { evalResultName = Nothing } opts) put ROpts (record { evalResultName = Nothing } opts)
resetContext resetContext f
Right res <- coreLift (readFile f) Right res <- coreLift (readFile f)
| Left err => do setSource "" | Left err => do setSource ""
pure (ErrorLoadingFile f err) pure (ErrorLoadingFile f err)
@ -970,16 +973,18 @@ processCatch cmd
pure $ REPLError msg pure $ REPLError msg
) )
parseEmptyCmd : SourceEmptyRule (Maybe REPLCmd) parseEmptyCmd : EmptyRule (Maybe REPLCmd)
parseEmptyCmd = eoi *> (pure Nothing) parseEmptyCmd = eoi *> (pure Nothing)
parseCmd : SourceEmptyRule (Maybe REPLCmd) parseCmd : EmptyRule (Maybe REPLCmd)
parseCmd = do c <- command; eoi; pure $ Just c parseCmd = do c <- command; eoi; pure $ Just c
export export
parseRepl : String -> Either Error (Maybe REPLCmd) parseRepl : String -> Either Error (Maybe REPLCmd)
parseRepl inp parseRepl inp
= runParser "(interactive)" Nothing inp (parseEmptyCmd <|> parseCmd) = case runParser "(interactive)" Nothing inp (parseEmptyCmd <|> parseCmd) of
Left err => Left err
Right (decor, result) => Right result
export export
interpret : {auto c : Ref Ctxt Defs} -> interpret : {auto c : Ref Ctxt Defs} ->

View File

@ -29,7 +29,8 @@ import System.File
-- Output informational messages, unless quiet flag is set -- Output informational messages, unless quiet flag is set
export export
iputStrLn : {auto o : Ref ROpts REPLOpts} -> iputStrLn : {auto c : Ref Ctxt Defs} ->
{auto o : Ref ROpts REPLOpts} ->
Doc IdrisAnn -> Core () Doc IdrisAnn -> Core ()
iputStrLn msg iputStrLn msg
= do opts <- get ROpts = do opts <- get ROpts
@ -118,8 +119,7 @@ emitWarnings
put Ctxt (record { warnings = [] } defs) put Ctxt (record { warnings = [] } defs)
getFCLine : FC -> Maybe Int getFCLine : FC -> Maybe Int
getFCLine (MkFC _ (line, _) _) = Just line getFCLine = map startLine . isNonEmptyFC
getFCLine EmptyFC = Nothing
export export
updateErrorLine : {auto o : Ref ROpts REPLOpts} -> updateErrorLine : {auto o : Ref ROpts REPLOpts} ->
@ -136,14 +136,15 @@ resetContext : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
{auto s : Ref Syn SyntaxInfo} -> {auto s : Ref Syn SyntaxInfo} ->
{auto m : Ref MD Metadata} -> {auto m : Ref MD Metadata} ->
(source : String) ->
Core () Core ()
resetContext resetContext fname
= do defs <- get Ctxt = do defs <- get Ctxt
put Ctxt (record { options = clearNames (options defs) } !initDefs) put Ctxt (record { options = clearNames (options defs) } !initDefs)
addPrimitives addPrimitives
put UST initUState put UST initUState
put Syn initSyntax put Syn initSyntax
put MD initMetadata put MD (initMetadata fname)
public export public export
data EditResult : Type where data EditResult : Type where

View File

@ -29,11 +29,11 @@ unbracketApp tm = tm
-- TODO: Deal with precedences -- TODO: Deal with precedences
mkOp : {auto s : Ref Syn SyntaxInfo} -> mkOp : {auto s : Ref Syn SyntaxInfo} ->
PTerm -> Core PTerm PTerm -> Core PTerm
mkOp tm@(PApp fc (PApp _ (PRef _ n) x) y) mkOp tm@(PApp fc (PApp _ (PRef opFC n) x) y)
= do syn <- get Syn = do syn <- get Syn
case StringMap.lookup (nameRoot n) (infixes syn) of case StringMap.lookup (nameRoot n) (infixes syn) of
Nothing => pure tm Nothing => pure tm
Just _ => pure (POp fc n (unbracketApp x) (unbracketApp y)) Just _ => pure (POp fc opFC n (unbracketApp x) (unbracketApp y))
mkOp tm = pure tm mkOp tm = pure tm
export export
@ -44,10 +44,10 @@ addBracket fc tm = if needed tm then PBracketed fc tm else tm
needed (PBracketed _ _) = False needed (PBracketed _ _) = False
needed (PRef _ _) = False needed (PRef _ _) = False
needed (PPair _ _ _) = False needed (PPair _ _ _) = False
needed (PDPair _ _ _ _) = False needed (PDPair _ _ _ _ _) = False
needed (PUnit _) = False needed (PUnit _) = False
needed (PComprehension _ _ _) = False needed (PComprehension _ _ _) = False
needed (PList _ _) = False needed (PList _ _ _) = False
needed (PPrimVal _ _) = False needed (PPrimVal _ _) = False
needed tm = True needed tm = True
@ -113,13 +113,13 @@ mutual
||| Put the special names (Nil, ::, Pair, Z, S, etc) back as syntax ||| Put the special names (Nil, ::, Pair, Z, S, etc) back as syntax
||| Returns `Nothing` in case there was nothing to resugar. ||| Returns `Nothing` in case there was nothing to resugar.
sugarAppM : PTerm -> Maybe PTerm sugarAppM : PTerm -> Maybe PTerm
sugarAppM (PApp fc (PApp _ (PRef _ (NS ns nm)) l) r) = sugarAppM (PApp fc (PApp _ (PRef opFC (NS ns nm)) l) r) =
if builtinNS == ns if builtinNS == ns
then case nameRoot nm of then case nameRoot nm of
"Pair" => pure $ PPair fc (unbracket l) (unbracket r) "Pair" => pure $ PPair fc (unbracket l) (unbracket r)
"MkPair" => pure $ PPair fc (unbracket l) (unbracket r) "MkPair" => pure $ PPair fc (unbracket l) (unbracket r)
"DPair" => case unbracket r of "DPair" => case unbracket r of
PLam _ _ _ n _ r' => pure $ PDPair fc n (unbracket l) (unbracket r') PLam _ _ _ n _ r' => pure $ PDPair fc opFC n (unbracket l) (unbracket r')
_ => Nothing _ => Nothing
"Equal" => pure $ PEq fc (unbracket l) (unbracket r) "Equal" => pure $ PEq fc (unbracket l) (unbracket r)
"===" => pure $ PEq fc (unbracket l) (unbracket r) "===" => pure $ PEq fc (unbracket l) (unbracket r)
@ -127,8 +127,8 @@ mutual
_ => Nothing _ => Nothing
else if nameRoot nm == "::" else if nameRoot nm == "::"
then case sugarApp (unbracket r) of then case sugarApp (unbracket r) of
PList fc xs => pure $ PList fc (unbracketApp l :: xs) PList fc nilFC xs => pure $ PList fc nilFC ((opFC, unbracketApp l) :: xs)
_ => Nothing _ => Nothing
else Nothing else Nothing
sugarAppM tm = sugarAppM tm =
-- refolding natural numbers if the expression is a constant -- refolding natural numbers if the expression is a constant
@ -142,7 +142,7 @@ mutual
"MkUnit" => pure $ PUnit fc "MkUnit" => pure $ PUnit fc
_ => Nothing _ => Nothing
else if nameRoot nm == "Nil" else if nameRoot nm == "Nil"
then pure $ PList fc [] then pure $ PList fc fc []
else Nothing else Nothing
_ => Nothing _ => Nothing
@ -464,14 +464,14 @@ cleanPTerm ptm
cleanNode : PTerm -> Core PTerm cleanNode : PTerm -> Core PTerm
cleanNode (PRef fc nm) = cleanNode (PRef fc nm) =
PRef fc <$> cleanName nm PRef fc <$> cleanName nm
cleanNode (POp fc op x y) = cleanNode (POp fc opFC op x y) =
(\ op => POp fc op x y) <$> cleanName op (\ op => POp fc opFC op x y) <$> cleanName op
cleanNode (PPrefixOp fc op x) = cleanNode (PPrefixOp fc opFC op x) =
(\ op => PPrefixOp fc op x) <$> cleanName op (\ op => PPrefixOp fc opFC op x) <$> cleanName op
cleanNode (PSectionL fc op x) = cleanNode (PSectionL fc opFC op x) =
(\ op => PSectionL fc op x) <$> cleanName op (\ op => PSectionL fc opFC op x) <$> cleanName op
cleanNode (PSectionR fc x op) = cleanNode (PSectionR fc opFC x op) =
PSectionR fc x <$> cleanName op PSectionR fc opFC x <$> cleanName op
cleanNode tm = pure tm cleanNode tm = pure tm
toCleanPTerm : {auto c : Ref Ctxt Defs} -> toCleanPTerm : {auto c : Ref Ctxt Defs} ->

View File

@ -82,10 +82,10 @@ mutual
-- Operators -- Operators
POp : FC -> OpStr -> PTerm -> PTerm -> PTerm POp : (full, opFC : FC) -> OpStr -> PTerm -> PTerm -> PTerm
PPrefixOp : FC -> OpStr -> PTerm -> PTerm PPrefixOp : (full, opFC : FC) -> OpStr -> PTerm -> PTerm
PSectionL : FC -> OpStr -> PTerm -> PTerm PSectionL : (full, opFC : FC) -> OpStr -> PTerm -> PTerm
PSectionR : FC -> PTerm -> OpStr -> PTerm PSectionR : (full, opFC : FC) -> PTerm -> OpStr -> PTerm
PEq : FC -> PTerm -> PTerm -> PTerm PEq : FC -> PTerm -> PTerm -> PTerm
PBracketed : FC -> PTerm -> PTerm PBracketed : FC -> PTerm -> PTerm
@ -95,9 +95,10 @@ mutual
PDoBlock : FC -> Maybe Namespace -> List PDo -> PTerm PDoBlock : FC -> Maybe Namespace -> List PDo -> PTerm
PBang : FC -> PTerm -> PTerm PBang : FC -> PTerm -> PTerm
PIdiom : FC -> PTerm -> PTerm PIdiom : FC -> PTerm -> PTerm
PList : FC -> List PTerm -> PTerm PList : (full, nilFC : FC) -> List (FC, PTerm) -> PTerm
-- ^ location of the conses
PPair : FC -> PTerm -> PTerm -> PTerm PPair : FC -> PTerm -> PTerm -> PTerm
PDPair : FC -> PTerm -> PTerm -> PTerm -> PTerm PDPair : (full, opFC : FC) -> PTerm -> PTerm -> PTerm -> PTerm
PUnit : FC -> PTerm PUnit : FC -> PTerm
PIfThenElse : FC -> PTerm -> PTerm -> PTerm -> PTerm PIfThenElse : FC -> PTerm -> PTerm -> PTerm -> PTerm
PComprehension : FC -> PTerm -> List PDo -> PTerm PComprehension : FC -> PTerm -> List PDo -> PTerm
@ -146,10 +147,10 @@ mutual
getPTermLoc (PDotted fc _) = fc getPTermLoc (PDotted fc _) = fc
getPTermLoc (PImplicit fc) = fc getPTermLoc (PImplicit fc) = fc
getPTermLoc (PInfer fc) = fc getPTermLoc (PInfer fc) = fc
getPTermLoc (POp fc _ _ _) = fc getPTermLoc (POp fc _ _ _ _) = fc
getPTermLoc (PPrefixOp fc _ _) = fc getPTermLoc (PPrefixOp fc _ _ _) = fc
getPTermLoc (PSectionL fc _ _) = fc getPTermLoc (PSectionL fc _ _ _) = fc
getPTermLoc (PSectionR fc _ _) = fc getPTermLoc (PSectionR fc _ _ _) = fc
getPTermLoc (PEq fc _ _) = fc getPTermLoc (PEq fc _ _) = fc
getPTermLoc (PBracketed fc _) = fc getPTermLoc (PBracketed fc _) = fc
getPTermLoc (PString fc _) = fc getPTermLoc (PString fc _) = fc
@ -157,9 +158,9 @@ mutual
getPTermLoc (PDoBlock fc _ _) = fc getPTermLoc (PDoBlock fc _ _) = fc
getPTermLoc (PBang fc _) = fc getPTermLoc (PBang fc _) = fc
getPTermLoc (PIdiom fc _) = fc getPTermLoc (PIdiom fc _) = fc
getPTermLoc (PList fc _) = fc getPTermLoc (PList fc _ _) = fc
getPTermLoc (PPair fc _ _) = fc getPTermLoc (PPair fc _ _) = fc
getPTermLoc (PDPair fc _ _ _) = fc getPTermLoc (PDPair fc _ _ _ _) = fc
getPTermLoc (PUnit fc) = fc getPTermLoc (PUnit fc) = fc
getPTermLoc (PIfThenElse fc _ _ _) = fc getPTermLoc (PIfThenElse fc _ _ _) = fc
getPTermLoc (PComprehension fc _ _) = fc getPTermLoc (PComprehension fc _ _) = fc
@ -360,10 +361,15 @@ mutual
getPDeclLoc (PDirective fc _) = fc getPDeclLoc (PDirective fc _) = fc
getPDeclLoc (PBuiltin fc _ _) = fc getPDeclLoc (PBuiltin fc _ _) = fc
export export
isPDef : PDecl -> Maybe (FC, List PClause) isStrInterp : PStr -> Maybe FC
isPDef (PDef annot cs) = Just (annot, cs) isStrInterp (StrInterp fc _) = Just fc
isPDef _ = Nothing isStrInterp (StrLiteral _ _) = Nothing
export
isPDef : PDecl -> Maybe (FC, List PClause)
isPDef (PDef annot cs) = Just (annot, cs)
isPDef _ = Nothing
definedInData : PDataDecl -> List Name definedInData : PDataDecl -> List Name
@ -604,10 +610,10 @@ mutual
showPrec d (PDotted _ p) = "." ++ showPrec d p showPrec d (PDotted _ p) = "." ++ showPrec d p
showPrec _ (PImplicit _) = "_" showPrec _ (PImplicit _) = "_"
showPrec _ (PInfer _) = "?" showPrec _ (PInfer _) = "?"
showPrec d (POp _ op x y) = showPrec d x ++ " " ++ showPrecOp d op ++ " " ++ showPrec d y showPrec d (POp _ _ op x y) = showPrec d x ++ " " ++ showPrecOp d op ++ " " ++ showPrec d y
showPrec d (PPrefixOp _ op x) = showPrec d op ++ showPrec d x showPrec d (PPrefixOp _ _ op x) = showPrec d op ++ showPrec d x
showPrec d (PSectionL _ op x) = "(" ++ showPrecOp d op ++ " " ++ showPrec d x ++ ")" showPrec d (PSectionL _ _ op x) = "(" ++ showPrecOp d op ++ " " ++ showPrec d x ++ ")"
showPrec d (PSectionR _ x op) = "(" ++ showPrec d x ++ " " ++ showPrecOp d op ++ ")" showPrec d (PSectionR _ _ x op) = "(" ++ showPrec d x ++ " " ++ showPrecOp d op ++ ")"
showPrec d (PEq fc l r) = showPrec d l ++ " = " ++ showPrec d r showPrec d (PEq fc l r) = showPrec d l ++ " = " ++ showPrec d r
showPrec d (PBracketed _ tm) = "(" ++ showPrec d tm ++ ")" showPrec d (PBracketed _ tm) = "(" ++ showPrec d tm ++ ")"
showPrec d (PString _ xs) = join " ++ " $ show <$> xs showPrec d (PString _ xs) = join " ++ " $ show <$> xs
@ -616,11 +622,11 @@ mutual
= "do " ++ showSep " ; " (map showDo ds) = "do " ++ showSep " ; " (map showDo ds)
showPrec d (PBang _ tm) = "!" ++ showPrec d tm showPrec d (PBang _ tm) = "!" ++ showPrec d tm
showPrec d (PIdiom _ tm) = "[|" ++ showPrec d tm ++ "|]" showPrec d (PIdiom _ tm) = "[|" ++ showPrec d tm ++ "|]"
showPrec d (PList _ xs) showPrec d (PList _ _ xs)
= "[" ++ showSep ", " (map (showPrec d) xs) ++ "]" = "[" ++ showSep ", " (map (showPrec d . snd) xs) ++ "]"
showPrec d (PPair _ l r) = "(" ++ showPrec d l ++ ", " ++ showPrec d r ++ ")" showPrec d (PPair _ l r) = "(" ++ showPrec d l ++ ", " ++ showPrec d r ++ ")"
showPrec d (PDPair _ l (PImplicit _) r) = "(" ++ showPrec d l ++ " ** " ++ showPrec d r ++ ")" showPrec d (PDPair _ _ l (PImplicit _) r) = "(" ++ showPrec d l ++ " ** " ++ showPrec d r ++ ")"
showPrec d (PDPair _ l ty r) = "(" ++ showPrec d l ++ " : " ++ showPrec d ty ++ showPrec d (PDPair _ _ l ty r) = "(" ++ showPrec d l ++ " : " ++ showPrec d ty ++
" ** " ++ showPrec d r ++ ")" " ** " ++ showPrec d r ++ ")"
showPrec _ (PUnit _) = "()" showPrec _ (PUnit _) = "()"
showPrec d (PIfThenElse _ x t e) = "if " ++ showPrec d x ++ " then " ++ showPrec d t ++ showPrec d (PIfThenElse _ x t e) = "if " ++ showPrec d x ++ " then " ++ showPrec d t ++
@ -941,18 +947,18 @@ mapPTermM f = goPTerm where
>>= f >>= f
goPTerm t@(PImplicit _) = f t goPTerm t@(PImplicit _) = f t
goPTerm t@(PInfer _) = f t goPTerm t@(PInfer _) = f t
goPTerm (POp fc x y z) = goPTerm (POp fc opFC x y z) =
POp fc x <$> goPTerm y POp fc opFC x <$> goPTerm y
<*> goPTerm z <*> goPTerm z
>>= f >>= f
goPTerm (PPrefixOp fc x y) = goPTerm (PPrefixOp fc opFC x y) =
PPrefixOp fc x <$> goPTerm y PPrefixOp fc opFC x <$> goPTerm y
>>= f >>= f
goPTerm (PSectionL fc x y) = goPTerm (PSectionL fc opFC x y) =
PSectionL fc x <$> goPTerm y PSectionL fc opFC x <$> goPTerm y
>>= f >>= f
goPTerm (PSectionR fc x y) = goPTerm (PSectionR fc opFC x y) =
PSectionR fc <$> goPTerm x PSectionR fc opFC <$> goPTerm x
<*> pure y <*> pure y
>>= f >>= f
goPTerm (PEq fc x y) = goPTerm (PEq fc x y) =
@ -977,15 +983,16 @@ mapPTermM f = goPTerm where
goPTerm (PIdiom fc x) = goPTerm (PIdiom fc x) =
PIdiom fc <$> goPTerm x PIdiom fc <$> goPTerm x
>>= f >>= f
goPTerm (PList fc xs) = goPTerm (PList fc nilFC xs) =
PList fc <$> goPTerms xs PList fc nilFC <$> goPairedPTerms xs
>>= f >>= f
goPTerm (PPair fc x y) = goPTerm (PPair fc x y) =
PPair fc <$> goPTerm x PPair fc <$> goPTerm x
<*> goPTerm y <*> goPTerm y
>>= f >>= f
goPTerm (PDPair fc x y z) = goPTerm (PDPair fc opFC x y z) =
PDPair fc <$> goPTerm x PDPair fc opFC
<$> goPTerm x
<*> goPTerm y <*> goPTerm y
<*> goPTerm z <*> goPTerm z
>>= f >>= f

View File

@ -227,6 +227,11 @@ export
singleton : Name -> v -> NameMap v singleton : Name -> v -> NameMap v
singleton n v = M Z $ Leaf n v singleton n v = M Z $ Leaf n v
export
null : NameMap v -> Bool
null Empty = True
null _ = False
export export
lookup : Name -> NameMap v -> Maybe v lookup : Name -> NameMap v -> Maybe v
lookup _ Empty = Nothing lookup _ Empty = Nothing

View File

@ -577,18 +577,27 @@ greater : FilePos -> Interval -> Bool
greater k (MkInterval (MkRange low _)) = fst low > k greater k (MkInterval (MkRange low _)) = fst low > k
greater k NoInterval = False greater k NoInterval = False
-- Finds all the interval that overlaps with the given interval. ||| Finds all the intervals that overlap with the given interval.
-- takeUntil selects all the intervals within the given upper bound, export
-- however the remaining interval are not necessarily adjacent in
-- the sequence, thus it drops elements until the next intersecting
-- interval with dropUntil.
inRange : MeasureRM a => FilePos -> FilePos -> PosMap a -> List a inRange : MeasureRM a => FilePos -> FilePos -> PosMap a -> List a
inRange low high t = matches (takeUntil (greater high) t) inRange low high t = matches (takeUntil (greater high) t)
-- takeUntil selects all the intervals within the given upper bound,
-- however the remaining interval are not necessarily adjacent in
-- the sequence, thus it drops elements until the next intersecting
-- interval with dropUntil.
where matches : PosMap a -> List a where matches : PosMap a -> List a
matches xs = case viewl (dropUntil (atleast low) xs) of matches xs = case viewl (dropUntil (atleast low) xs) of
EmptyL => [] EmptyL => []
x <: xs' => x :: assert_total (matches xs') x <: xs' => x :: assert_total (matches xs')
||| Finds the values matching the exact interval input
export
exactRange : MeasureRM a => FilePos -> FilePos -> PosMap a -> List a
exactRange low high t = flip mapMaybe (inRange low high t) $ \ a =>
do let (MkRange rng _) = measureRM a
guard (rng == (low, high))
pure a
||| Returns all the interval that contains the given point. ||| Returns all the interval that contains the given point.
export export
searchPos : MeasureRM a => FilePos -> PosMap a -> List a searchPos : MeasureRM a => FilePos -> PosMap a -> List a

View File

@ -247,12 +247,12 @@ toList (M _ t) = treeToList t
||| Gets the keys of the map. ||| Gets the keys of the map.
export export
keys : SortedMap k v -> List k keys : SortedMap k v -> List k
keys = map fst . toList keys = map fst . SortedMap.toList
||| Gets the values of the map. Could contain duplicates. ||| Gets the values of the map. Could contain duplicates.
export export
values : SortedMap k v -> List v values : SortedMap k v -> List v
values = map snd . toList values = map snd . SortedMap.toList
treeMap : (a -> b) -> Tree n k a o -> Tree n k b o treeMap : (a -> b) -> Tree n k a o -> Tree n k b o
treeMap f (Leaf k v) = Leaf k (f v) treeMap f (Leaf k v) = Leaf k (f v)
@ -299,7 +299,7 @@ mergeWith : (v -> v -> v) -> SortedMap k v -> SortedMap k v -> SortedMap k v
mergeWith f x y = insertFrom inserted x where mergeWith f x y = insertFrom inserted x where
inserted : List (k, v) inserted : List (k, v)
inserted = do inserted = do
(k, v) <- toList y (k, v) <- SortedMap.toList y
let v' = (maybe id f $ lookup k x) v let v' = (maybe id f $ lookup k x) v
pure (k, v') pure (k, v')
@ -314,9 +314,16 @@ export
mergeLeft : SortedMap k v -> SortedMap k v -> SortedMap k v mergeLeft : SortedMap k v -> SortedMap k v -> SortedMap k v
mergeLeft = mergeWith const mergeLeft = mergeWith const
export
adjust : k -> (v -> v) -> SortedMap k v -> SortedMap k v
adjust k f m =
case lookup k m of
Nothing => m
Just v => insert k (f v) m
export export
(Show k, Show v) => Show (SortedMap k v) where (Show k, Show v) => Show (SortedMap k v) where
show m = "fromList " ++ (show $ toList m) show m = "fromList " ++ (show $ SortedMap.toList m)
-- TODO: is this the right variant of merge to use for this? I think it is, but -- 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 -- I could also see the advantages of using `mergeLeft`. The current approach is

View File

@ -261,12 +261,12 @@ toList (M _ t) = treeToList t
||| Gets the Keys of the map. ||| Gets the Keys of the map.
export export
keys : StringMap v -> List String keys : StringMap v -> List String
keys = map fst . toList keys = map fst . StringMap.toList
||| Gets the values of the map. Could contain duplicates. ||| Gets the values of the map. Could contain duplicates.
export export
values : StringMap v -> List v values : StringMap v -> List v
values = map snd . toList values = map snd . StringMap.toList
treeMap : (a -> b) -> Tree n a -> Tree n b treeMap : (a -> b) -> Tree n a -> Tree n b
treeMap f (Leaf k v) = Leaf k (f v) treeMap f (Leaf k v) = Leaf k (f v)
@ -286,7 +286,7 @@ mergeWith : (v -> v -> v) -> StringMap v -> StringMap v -> StringMap v
mergeWith f x y = insertFrom inserted x where mergeWith f x y = insertFrom inserted x where
inserted : List (Key, v) inserted : List (Key, v)
inserted = do inserted = do
(k, v) <- toList y (k, v) <- StringMap.toList y
let v' = (maybe id f $ lookup k x) v let v' = (maybe id f $ lookup k x) v
pure (k, v') pure (k, v')
@ -310,7 +310,7 @@ adjust k f m =
export export
Show v => Show (StringMap v) where Show v => Show (StringMap v) where
show m = show $ map {b=String} (\(k,v) => k ++ "->" ++ show v) $ toList m show m = show $ map {b=String} (\(k,v) => k ++ "->" ++ show v) $ StringMap.toList m
-- TODO: is this the right variant of merge to use for this? I think it is, but -- 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 -- I could also see the advantages of using `mergeLeft`. The current approach is

View File

@ -62,5 +62,5 @@ foldWithKeysM {a} {m} {b} fk fv = go []
z <- fk ks' z <- fk ks'
pure $ x <+> y <+> z) pure $ x <+> y <+> z)
neutral neutral
(toList sm)) (StringMap.toList sm))
nd nd

View File

@ -15,7 +15,7 @@ import public Libraries.Text.Token
export export
match : (Eq k, TokenKind k) => match : (Eq k, TokenKind k) =>
(kind : k) -> (kind : k) ->
Grammar (Token k) True (TokType kind) Grammar state (Token k) True (TokType kind)
match k = terminal "Unrecognised input" $ match k = terminal "Unrecognised input" $
\t => if t.kind == k \t => if t.kind == k
then Just $ tokValue k t.text then Just $ tokValue k t.text
@ -25,8 +25,8 @@ match k = terminal "Unrecognised input" $
||| match. May match the empty input. ||| match. May match the empty input.
export export
option : {c : Bool} -> option : {c : Bool} ->
(def : a) -> (p : Grammar tok c a) -> (def : a) -> (p : Grammar state tok c a) ->
Grammar tok False a Grammar state tok False a
option {c = False} def p = p <|> pure def option {c = False} def p = p <|> pure def
option {c = True} def p = p <|> pure def option {c = True} def p = p <|> pure def
@ -34,8 +34,8 @@ option {c = True} def p = p <|> pure def
||| To provide a default value, use `option` instead. ||| To provide a default value, use `option` instead.
export export
optional : {c : Bool} -> optional : {c : Bool} ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok False (Maybe a) Grammar state tok False (Maybe a)
optional p = option Nothing (map Just p) optional p = option Nothing (map Just p)
||| Try to parse one thing or the other, producing a value that indicates ||| Try to parse one thing or the other, producing a value that indicates
@ -43,9 +43,9 @@ optional p = option Nothing (map Just p)
||| takes priority. ||| takes priority.
export export
choose : {c1, c2 : Bool} -> choose : {c1, c2 : Bool} ->
(l : Grammar tok c1 a) -> (l : Grammar state tok c1 a) ->
(r : Grammar tok c2 b) -> (r : Grammar state tok c2 b) ->
Grammar tok (c1 && c2) (Either a b) Grammar state tok (c1 && c2) (Either a b)
choose l r = map Left l <|> map Right r choose l r = map Left l <|> map Right r
||| Produce a grammar by applying a function to each element of a container, ||| Produce a grammar by applying a function to each element of a container,
@ -53,9 +53,9 @@ choose l r = map Left l <|> map Right r
||| container is empty. ||| container is empty.
export export
choiceMap : {c : Bool} -> choiceMap : {c : Bool} ->
(a -> Grammar tok c b) -> (a -> Grammar state tok c b) ->
Foldable t => t a -> Foldable t => t a ->
Grammar tok c b Grammar state tok c b
choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in
f x <|> acc) f x <|> acc)
(fail "No more options") xs (fail "No more options") xs
@ -67,28 +67,28 @@ choiceMap {c} f xs = foldr (\x, acc => rewrite sym (andSameNeutral c) in
export export
choice : Foldable t => choice : Foldable t =>
{c : Bool} -> {c : Bool} ->
t (Grammar tok c a) -> t (Grammar state tok c a) ->
Grammar tok c a Grammar state tok c a
choice = choiceMap id choice = choiceMap id
mutual mutual
||| Parse one or more things ||| Parse one or more things
export export
some : Grammar tok True a -> some : Grammar state tok True a ->
Grammar tok True (List1 a) Grammar state tok True (List1 a)
some p = pure (!p ::: !(many p)) some p = pure (!p ::: !(many p))
||| Parse zero or more things (may match the empty input) ||| Parse zero or more things (may match the empty input)
export export
many : Grammar tok True a -> many : Grammar state tok True a ->
Grammar tok False (List a) Grammar state tok False (List a)
many p = option [] (forget <$> some p) many p = option [] (forget <$> some p)
mutual mutual
private private
count1 : (q : Quantity) -> count1 : (q : Quantity) ->
(p : Grammar tok True a) -> (p : Grammar state tok True a) ->
Grammar tok True (List a) Grammar state tok True (List a)
count1 q p = do x <- p count1 q p = do x <- p
seq (count q p) seq (count q p)
(\xs => pure (x :: xs)) (\xs => pure (x :: xs))
@ -96,8 +96,8 @@ mutual
||| Parse `p`, repeated as specified by `q`, returning the list of values. ||| Parse `p`, repeated as specified by `q`, returning the list of values.
export export
count : (q : Quantity) -> count : (q : Quantity) ->
(p : Grammar tok True a) -> (p : Grammar state tok True a) ->
Grammar tok (isSucc (min q)) (List a) Grammar state tok (isSucc (min q)) (List a)
count (Qty Z Nothing) p = many p count (Qty Z Nothing) p = many p
count (Qty Z (Just Z)) _ = pure [] count (Qty Z (Just Z)) _ = pure []
count (Qty Z (Just (S max))) p = option [] $ count1 (atMost max) p count (Qty Z (Just (S max))) p = option [] $ count1 (atMost max) p
@ -110,9 +110,9 @@ mutual
||| list of values from `p`. Guaranteed to consume input. ||| list of values from `p`. Guaranteed to consume input.
export export
someTill : {c : Bool} -> someTill : {c : Bool} ->
(end : Grammar tok c e) -> (end : Grammar state tok c e) ->
(p : Grammar tok True a) -> (p : Grammar state tok True a) ->
Grammar tok True (List1 a) Grammar state tok True (List1 a)
someTill {c} end p = do x <- p someTill {c} end p = do x <- p
seq (manyTill end p) seq (manyTill end p)
(\xs => pure (x ::: xs)) (\xs => pure (x ::: xs))
@ -121,9 +121,9 @@ mutual
||| list of values from `p`. Guaranteed to consume input if `end` consumes. ||| list of values from `p`. Guaranteed to consume input if `end` consumes.
export export
manyTill : {c : Bool} -> manyTill : {c : Bool} ->
(end : Grammar tok c e) -> (end : Grammar state tok c e) ->
(p : Grammar tok True a) -> (p : Grammar state tok True a) ->
Grammar tok c (List a) Grammar state tok c (List a)
manyTill {c} end p = rewrite sym (andTrueNeutral c) in manyTill {c} end p = rewrite sym (andTrueNeutral c) in
map (const []) end <|> (forget <$> someTill end p) map (const []) end <|> (forget <$> someTill end p)
@ -132,9 +132,9 @@ mutual
||| returning its value. ||| returning its value.
export export
afterSome : {c : Bool} -> afterSome : {c : Bool} ->
(skip : Grammar tok True s) -> (skip : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok True a Grammar state tok True a
afterSome skip p = do ignore $ skip afterSome skip p = do ignore $ skip
afterMany skip p afterMany skip p
@ -142,18 +142,18 @@ mutual
||| returning its value. ||| returning its value.
export export
afterMany : {c : Bool} -> afterMany : {c : Bool} ->
(skip : Grammar tok True s) -> (skip : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok c a Grammar state tok c a
afterMany {c} skip p = rewrite sym (andTrueNeutral c) in afterMany {c} skip p = rewrite sym (andTrueNeutral c) in
p <|> afterSome skip p p <|> afterSome skip p
||| Parse one or more things, each separated by another thing. ||| Parse one or more things, each separated by another thing.
export export
sepBy1 : {c : Bool} -> sepBy1 : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok c (List1 a) Grammar state tok c (List1 a)
sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
[| p ::: many (sep *> p) |] [| p ::: many (sep *> p) |]
@ -161,18 +161,18 @@ sepBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
||| match the empty input. ||| match the empty input.
export export
sepBy : {c : Bool} -> sepBy : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok False (List a) Grammar state tok False (List a)
sepBy sep p = option [] $ forget <$> sepBy1 sep p sepBy sep p = option [] $ forget <$> sepBy1 sep p
||| Parse one or more instances of `p` separated by and optionally terminated by ||| Parse one or more instances of `p` separated by and optionally terminated by
||| `sep`. ||| `sep`.
export export
sepEndBy1 : {c : Bool} -> sepEndBy1 : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok c (List1 a) Grammar state tok c (List1 a)
sepEndBy1 {c} sep p = rewrite sym (orFalseNeutral c) in sepEndBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
sepBy1 sep p <* optional sep sepBy1 sep p <* optional sep
@ -180,32 +180,49 @@ sepEndBy1 {c} sep p = rewrite sym (orFalseNeutral c) in
||| by `sep`. Will not match a separator by itself. ||| by `sep`. Will not match a separator by itself.
export export
sepEndBy : {c : Bool} -> sepEndBy : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok False (List a) Grammar state tok False (List a)
sepEndBy sep p = option [] $ forget <$> sepEndBy1 sep p sepEndBy sep p = option [] $ forget <$> sepEndBy1 sep p
||| Parse one or more instances of `p`, separated and terminated by `sep`. ||| Parse one or more instances of `p`, separated and terminated by `sep`.
export export
endBy1 : {c : Bool} -> endBy1 : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok True (List1 a) Grammar state tok True (List1 a)
endBy1 {c} sep p = some $ rewrite sym (orTrueTrue c) in endBy1 {c} sep p = some $ rewrite sym (orTrueTrue c) in
p <* sep p <* sep
export export
endBy : {c : Bool} -> endBy : {c : Bool} ->
(sep : Grammar tok True s) -> (sep : Grammar state tok True s) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok False (List a) Grammar state tok False (List a)
endBy sep p = option [] $ forget <$> endBy1 sep p endBy sep p = option [] $ forget <$> endBy1 sep p
||| Parse an instance of `p` that is between `left` and `right`. ||| Parse an instance of `p` that is between `left` and `right`.
export export
between : {c : Bool} -> between : {c : Bool} ->
(left : Grammar tok True l) -> (left : Grammar state tok True l) ->
(right : Grammar tok True r) -> (right : Grammar state tok True r) ->
(p : Grammar tok c a) -> (p : Grammar state tok c a) ->
Grammar tok True a Grammar state tok True a
between left right contents = left *> contents <* right between left right contents = left *> contents <* right
export
location : Grammar state token False (Int, Int)
location = startBounds <$> position
export
endLocation : Grammar state token False (Int, Int)
endLocation = endBounds <$> position
export
column : Grammar state token False Int
column = snd <$> location
public export
when : Bool -> Lazy (Grammar state token False ()) -> Grammar state token False ()
when True y = y
when False y = pure ()

View File

@ -16,37 +16,39 @@ import public Libraries.Text.Bounded
||| to be non-empty - that is, successfully parsing the language is guaranteed ||| to be non-empty - that is, successfully parsing the language is guaranteed
||| to consume some input. ||| to consume some input.
export export
data Grammar : (tok : Type) -> (consumes : Bool) -> Type -> Type where data Grammar : (state : Type) -> (tok : Type) -> (consumes : Bool) -> Type -> Type where
Empty : (val : ty) -> Grammar tok False ty Empty : (val : ty) -> Grammar state tok False ty
Terminal : String -> (tok -> Maybe a) -> Grammar tok True a Terminal : String -> (tok -> Maybe a) -> Grammar state tok True a
NextIs : String -> (tok -> Bool) -> Grammar tok False tok NextIs : String -> (tok -> Bool) -> Grammar state tok False tok
EOF : Grammar tok False () EOF : Grammar state tok False ()
Fail : (location : Maybe Bounds) -> Bool -> String -> Grammar tok c ty Fail : (location : Maybe Bounds) -> Bool -> String -> Grammar state tok c ty
Try : Grammar tok c ty -> Grammar tok c ty Try : Grammar state tok c ty -> Grammar state tok c ty
Commit : Grammar tok False () Commit : Grammar state tok False ()
MustWork : Grammar tok c a -> Grammar tok c a MustWork : Grammar state tok c a -> Grammar state tok c a
SeqEat : {c2 : Bool} -> SeqEat : {c2 : Bool} ->
Grammar tok True a -> Inf (a -> Grammar tok c2 b) -> Grammar state tok True a -> Inf (a -> Grammar state tok c2 b) ->
Grammar tok True b Grammar state tok True b
SeqEmpty : {c1, c2 : Bool} -> SeqEmpty : {c1, c2 : Bool} ->
Grammar tok c1 a -> (a -> Grammar tok c2 b) -> Grammar state tok c1 a -> (a -> Grammar state tok c2 b) ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
ThenEat : {c2 : Bool} -> ThenEat : {c2 : Bool} ->
Grammar tok True () -> Inf (Grammar tok c2 a) -> Grammar state tok True () -> Inf (Grammar state tok c2 a) ->
Grammar tok True a Grammar state tok True a
ThenEmpty : {c1, c2 : Bool} -> ThenEmpty : {c1, c2 : Bool} ->
Grammar tok c1 () -> Grammar tok c2 a -> Grammar state tok c1 () -> Grammar state tok c2 a ->
Grammar tok (c1 || c2) a Grammar state tok (c1 || c2) a
Alt : {c1, c2 : Bool} -> Alt : {c1, c2 : Bool} ->
Grammar tok c1 ty -> Lazy (Grammar tok c2 ty) -> Grammar state tok c1 ty -> Lazy (Grammar state tok c2 ty) ->
Grammar tok (c1 && c2) ty Grammar state tok (c1 && c2) ty
Bounds : Grammar tok c ty -> Grammar tok c (WithBounds ty) Bounds : Grammar state tok c ty -> Grammar state tok c (WithBounds ty)
Position : Grammar tok False Bounds Position : Grammar state tok False Bounds
Act : state -> Grammar state tok False ()
||| Sequence two grammars. If either consumes some input, the sequence is ||| Sequence two grammars. If either consumes some input, the sequence is
||| guaranteed to consume some input. If the first one consumes input, the ||| guaranteed to consume some input. If the first one consumes input, the
@ -54,9 +56,9 @@ data Grammar : (tok : Type) -> (consumes : Bool) -> Type -> Type where
||| consumed and therefore the input is smaller) ||| consumed and therefore the input is smaller)
export %inline export %inline
(>>=) : {c1, c2 : Bool} -> (>>=) : {c1, c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
inf c1 (a -> Grammar tok c2 b) -> inf c1 (a -> Grammar state tok c2 b) ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
(>>=) {c1 = False} = SeqEmpty (>>=) {c1 = False} = SeqEmpty
(>>=) {c1 = True} = SeqEat (>>=) {c1 = True} = SeqEat
@ -66,9 +68,9 @@ export %inline
||| consumed and therefore the input is smaller) ||| consumed and therefore the input is smaller)
public export %inline %tcinline public export %inline %tcinline
(>>) : {c1, c2 : Bool} -> (>>) : {c1, c2 : Bool} ->
Grammar tok c1 () -> Grammar state tok c1 () ->
inf c1 (Grammar tok c2 a) -> inf c1 (Grammar state tok c2 a) ->
Grammar tok (c1 || c2) a Grammar state tok (c1 || c2) a
(>>) {c1 = False} = ThenEmpty (>>) {c1 = False} = ThenEmpty
(>>) {c1 = True} = ThenEat (>>) {c1 = True} = ThenEat
@ -77,23 +79,23 @@ public export %inline %tcinline
||| of `>>=`. ||| of `>>=`.
export %inline export %inline
seq : {c1,c2 : Bool} -> seq : {c1,c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
(a -> Grammar tok c2 b) -> (a -> Grammar state tok c2 b) ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
seq = SeqEmpty seq = SeqEmpty
||| Sequence a grammar followed by the grammar it returns. ||| Sequence a grammar followed by the grammar it returns.
export %inline export %inline
join : {c1,c2 : Bool} -> join : {c1,c2 : Bool} ->
Grammar tok c1 (Grammar tok c2 a) -> Grammar state tok c1 (Grammar state tok c2 a) ->
Grammar tok (c1 || c2) a Grammar state tok (c1 || c2) a
join {c1 = False} p = SeqEmpty p id join {c1 = False} p = SeqEmpty p id
join {c1 = True} p = SeqEat p id join {c1 = True} p = SeqEat p id
||| Allows the result of a grammar to be mapped to a different value. ||| Allows the result of a grammar to be mapped to a different value.
export export
{c : _} -> {c : _} ->
Functor (Grammar tok c) where Functor (Grammar state tok c) where
map f (Empty val) = Empty (f val) map f (Empty val) = Empty (f val)
map f (Fail bd fatal msg) = Fail bd fatal msg map f (Fail bd fatal msg) = Fail bd fatal msg
map f (Try g) = Try (map f g) map f (Try g) = Try (map f g)
@ -119,9 +121,9 @@ Functor (Grammar tok c) where
||| guaranteed to consume. ||| guaranteed to consume.
export %inline export %inline
(<|>) : {c1,c2 : Bool} -> (<|>) : {c1,c2 : Bool} ->
Grammar tok c1 ty -> Grammar state tok c1 ty ->
Lazy (Grammar tok c2 ty) -> Lazy (Grammar state tok c2 ty) ->
Grammar tok (c1 && c2) ty Grammar state tok (c1 && c2) ty
(<|>) = Alt (<|>) = Alt
infixr 2 <||> infixr 2 <||>
@ -129,9 +131,9 @@ infixr 2 <||>
||| combination is guaranteed to consume. ||| combination is guaranteed to consume.
export export
(<||>) : {c1,c2 : Bool} -> (<||>) : {c1,c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
Lazy (Grammar tok c2 b) -> Lazy (Grammar state tok c2 b) ->
Grammar tok (c1 && c2) (Either a b) Grammar state tok (c1 && c2) (Either a b)
(<||>) p q = (Left <$> p) <|> (Right <$> q) (<||>) p q = (Left <$> p) <|> (Right <$> q)
||| Sequence a grammar with value type `a -> b` and a grammar ||| Sequence a grammar with value type `a -> b` and a grammar
@ -140,33 +142,37 @@ export
||| Guaranteed to consume if either grammar consumes. ||| Guaranteed to consume if either grammar consumes.
export %inline export %inline
(<*>) : {c1, c2 : Bool} -> (<*>) : {c1, c2 : Bool} ->
Grammar tok c1 (a -> b) -> Grammar state tok c1 (a -> b) ->
Grammar tok c2 a -> Grammar state tok c2 a ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
(<*>) x y = SeqEmpty x (\f => map f y) (<*>) x y = SeqEmpty x (\f => map f y)
||| Sequence two grammars. If both succeed, use the value of the first one. ||| Sequence two grammars. If both succeed, use the value of the first one.
||| Guaranteed to consume if either grammar consumes. ||| Guaranteed to consume if either grammar consumes.
export %inline export %inline
(<*) : {c1,c2 : Bool} -> (<*) : {c1,c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
Grammar tok c2 b -> Grammar state tok c2 b ->
Grammar tok (c1 || c2) a Grammar state tok (c1 || c2) a
(<*) x y = map const x <*> y (<*) x y = map const x <*> y
||| Sequence two grammars. If both succeed, use the value of the second one. ||| Sequence two grammars. If both succeed, use the value of the second one.
||| Guaranteed to consume if either grammar consumes. ||| Guaranteed to consume if either grammar consumes.
export %inline export %inline
(*>) : {c1,c2 : Bool} -> (*>) : {c1,c2 : Bool} ->
Grammar tok c1 a -> Grammar state tok c1 a ->
Grammar tok c2 b -> Grammar state tok c2 b ->
Grammar tok (c1 || c2) b Grammar state tok (c1 || c2) b
(*>) x y = map (const id) x <*> y (*>) x y = map (const id) x <*> y
export %inline
act : state -> Grammar state tok False ()
act = Act
||| Produce a grammar that can parse a different type of token by providing a ||| Produce a grammar that can parse a different type of token by providing a
||| function converting the new token type into the original one. ||| function converting the new token type into the original one.
export export
mapToken : (a -> b) -> Grammar b c ty -> Grammar a c ty mapToken : (a -> b) -> Grammar state b c ty -> Grammar state a c ty
mapToken f (Empty val) = Empty val mapToken f (Empty val) = Empty val
mapToken f (Terminal msg g) = Terminal msg (g . f) mapToken f (Terminal msg g) = Terminal msg (g . f)
mapToken f (NextIs msg g) = SeqEmpty (NextIs msg (g . f)) (Empty . f) mapToken f (NextIs msg g) = SeqEmpty (NextIs msg (g . f)) (Empty . f)
@ -186,149 +192,151 @@ mapToken f (ThenEmpty act next)
mapToken f (Alt x y) = Alt (mapToken f x) (mapToken f y) mapToken f (Alt x y) = Alt (mapToken f x) (mapToken f y)
mapToken f (Bounds act) = Bounds (mapToken f act) mapToken f (Bounds act) = Bounds (mapToken f act)
mapToken f Position = Position mapToken f Position = Position
mapToken f (Act action) = Act action
||| Always succeed with the given value. ||| Always succeed with the given value.
export %inline export %inline
pure : (val : ty) -> Grammar tok False ty pure : (val : ty) -> Grammar state tok False ty
pure = Empty pure = Empty
||| Check whether the next token satisfies a predicate ||| Check whether the next token satisfies a predicate
export %inline export %inline
nextIs : String -> (tok -> Bool) -> Grammar tok False tok nextIs : String -> (tok -> Bool) -> Grammar state tok False tok
nextIs = NextIs nextIs = NextIs
||| Look at the next token in the input ||| Look at the next token in the input
export %inline export %inline
peek : Grammar tok False tok peek : Grammar state tok False tok
peek = nextIs "Unrecognised token" (const True) peek = nextIs "Unrecognised token" (const True)
||| Succeeds if running the predicate on the next token returns Just x, ||| Succeeds if running the predicate on the next token returns Just x,
||| returning x. Otherwise fails. ||| returning x. Otherwise fails.
export %inline export %inline
terminal : String -> (tok -> Maybe a) -> Grammar tok True a terminal : String -> (tok -> Maybe a) -> Grammar state tok True a
terminal = Terminal terminal = Terminal
||| Always fail with a message ||| Always fail with a message
export %inline export %inline
fail : String -> Grammar tok c ty fail : String -> Grammar state tok c ty
fail = Fail Nothing False fail = Fail Nothing False
||| Always fail with a message and a location ||| Always fail with a message and a location
export %inline export %inline
failLoc : Bounds -> String -> Grammar tok c ty failLoc : Bounds -> String -> Grammar state tok c ty
failLoc b = Fail (Just b) False failLoc b = Fail (Just b) False
export %inline export %inline
fatalError : String -> Grammar tok c ty fatalError : String -> Grammar state tok c ty
fatalError = Fail Nothing True fatalError = Fail Nothing True
export %inline export %inline
fatalLoc : Bounds -> String -> Grammar tok c ty fatalLoc : Bounds -> String -> Grammar state tok c ty
fatalLoc b = Fail (Just b) True fatalLoc b = Fail (Just b) True
||| Catch a fatal error ||| Catch a fatal error
export %inline export %inline
try : Grammar tok c ty -> Grammar tok c ty try : Grammar state tok c ty -> Grammar state tok c ty
try = Try try = Try
||| Succeed if the input is empty ||| Succeed if the input is empty
export %inline export %inline
eof : Grammar tok False () eof : Grammar state tok False ()
eof = EOF eof = EOF
||| Commit to an alternative; if the current branch of an alternative ||| Commit to an alternative; if the current branch of an alternative
||| fails to parse, no more branches will be tried ||| fails to parse, no more branches will be tried
export %inline export %inline
commit : Grammar tok False () commit : Grammar state tok False ()
commit = Commit commit = Commit
||| If the parser fails, treat it as a fatal error ||| If the parser fails, treat it as a fatal error
export %inline export %inline
mustWork : {c : Bool} -> Grammar tok c ty -> Grammar tok c ty mustWork : {c : Bool} -> Grammar state tok c ty -> Grammar state tok c ty
mustWork = MustWork mustWork = MustWork
export %inline export %inline
bounds : Grammar tok c ty -> Grammar tok c (WithBounds ty) bounds : Grammar state tok c ty -> Grammar state tok c (WithBounds ty)
bounds = Bounds bounds = Bounds
export %inline export %inline
position : Grammar tok False Bounds position : Grammar state tok False Bounds
position = Position position = Position
data ParseResult : Type -> Type -> Type where data ParseResult : Type -> Type -> Type -> Type where
Failure : (committed : Bool) -> (fatal : Bool) -> Failure : (committed : Bool) -> (fatal : Bool) ->
(err : String) -> (location : Maybe Bounds) -> ParseResult tok ty (err : String) -> (location : Maybe Bounds) -> ParseResult state tok ty
Res : (committed : Bool) -> Res : state -> (committed : Bool) ->
(val : WithBounds ty) -> (more : List (WithBounds tok)) -> ParseResult tok ty (val : WithBounds ty) -> (more : List (WithBounds tok)) -> ParseResult state tok ty
mergeWith : WithBounds ty -> ParseResult tok sy -> ParseResult tok sy mergeWith : WithBounds ty -> ParseResult state tok sy -> ParseResult state tok sy
mergeWith x (Res committed val more) = Res committed (mergeBounds x val) more mergeWith x (Res s committed val more) = Res s committed (mergeBounds x val) more
mergeWith x v = v mergeWith x v = v
doParse : (commit : Bool) -> doParse : Semigroup state => state -> (commit : Bool) ->
(act : Grammar tok c ty) -> (act : Grammar state tok c ty) ->
(xs : List (WithBounds tok)) -> (xs : List (WithBounds tok)) ->
ParseResult tok ty ParseResult state tok ty
doParse com (Empty val) xs = Res com (irrelevantBounds val) xs doParse s com (Empty val) xs = Res s com (irrelevantBounds val) xs
doParse com (Fail location fatal str) xs = Failure com fatal str (maybe (bounds <$> head' xs) Just location) doParse s com (Fail location fatal str) xs = Failure com fatal str (maybe (bounds <$> head' xs) Just location)
doParse com (Try g) xs = case doParse com g xs of doParse s com (Try g) xs = case doParse s com g xs of
-- recover from fatal match but still propagate the 'commit' -- recover from fatal match but still propagate the 'commit'
Failure com _ msg ts => Failure com False msg ts Failure com _ msg ts => Failure com False msg ts
res => res res => res
doParse com Commit xs = Res True (irrelevantBounds ()) xs doParse s com Commit xs = Res s True (irrelevantBounds ()) xs
doParse com (MustWork g) xs = doParse s com (MustWork g) xs =
case assert_total (doParse com g xs) of case assert_total (doParse s com g xs) of
Failure com' _ msg ts => Failure com' True msg ts Failure com' _ msg ts => Failure com' True msg ts
res => res res => res
doParse com (Terminal err f) [] = Failure com False "End of input" Nothing doParse s com (Terminal err f) [] = Failure com False "End of input" Nothing
doParse com (Terminal err f) (x :: xs) = doParse s com (Terminal err f) (x :: xs) =
case f x.val of case f x.val of
Nothing => Failure com False err (Just x.bounds) Nothing => Failure com False err (Just x.bounds)
Just a => Res com (const a <$> x) xs Just a => Res s com (const a <$> x) xs
doParse com EOF [] = Res com (irrelevantBounds ()) [] doParse s com EOF [] = Res s com (irrelevantBounds ()) []
doParse com EOF (x :: xs) = Failure com False "Expected end of input" (Just x.bounds) doParse s com EOF (x :: xs) = Failure com False "Expected end of input" (Just x.bounds)
doParse com (NextIs err f) [] = Failure com False "End of input" Nothing doParse s com (NextIs err f) [] = Failure com False "End of input" Nothing
doParse com (NextIs err f) (x :: xs) doParse s com (NextIs err f) (x :: xs)
= if f x.val = if f x.val
then Res com (removeIrrelevance x) (x :: xs) then Res s com (removeIrrelevance x) (x :: xs)
else Failure com False err (Just x.bounds) else Failure com False err (Just x.bounds)
doParse com (Alt {c1} {c2} x y) xs doParse s com (Alt {c1} {c2} x y) xs
= case doParse False x xs of = case doParse s False x xs of
Failure com' fatal msg ts Failure com' fatal msg ts
=> if com' || fatal => if com' || fatal
-- If the alternative had committed, don't try the -- If the alternative had committed, don't try the
-- other branch (and reset commit flag) -- other branch (and reset commit flag)
then Failure com fatal msg ts then Failure com fatal msg ts
else assert_total (doParse False y xs) else assert_total (doParse s False y xs)
-- Successfully parsed the first option, so use the outer commit flag -- Successfully parsed the first option, so use the outer commit flag
Res _ val xs => Res com val xs Res s _ val xs => Res s com val xs
doParse com (SeqEmpty act next) xs doParse s com (SeqEmpty act next) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res s com v xs =>
mergeWith v (assert_total $ doParse com (next v.val) xs) mergeWith v (assert_total $ doParse s com (next v.val) xs)
doParse com (SeqEat act next) xs doParse s com (SeqEat act next) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res s com v xs =>
mergeWith v (assert_total $ doParse com (next v.val) xs) mergeWith v (assert_total $ doParse s com (next v.val) xs)
doParse com (ThenEmpty act next) xs doParse s com (ThenEmpty act next) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res s com v xs =>
mergeWith v (assert_total $ doParse com next xs) mergeWith v (assert_total $ doParse s com next xs)
doParse com (ThenEat act next) xs doParse s com (ThenEat act next) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res s com v xs =>
mergeWith v (assert_total $ doParse com next xs) mergeWith v (assert_total $ doParse s com next xs)
doParse com (Bounds act) xs doParse s com (Bounds act) xs
= case assert_total (doParse com act xs) of = case assert_total (doParse s com act xs) of
Failure com fatal msg ts => Failure com fatal msg ts Failure com fatal msg ts => Failure com fatal msg ts
Res com v xs => Res com (const v <$> v) xs Res s com v xs => Res s com (const v <$> v) xs
doParse com Position [] = Failure com False "End of input" Nothing doParse s com Position [] = Failure com False "End of input" Nothing
doParse com Position (x :: xs) doParse s com Position (x :: xs)
= Res com (irrelevantBounds x.bounds) (x :: xs) = Res s com (irrelevantBounds x.bounds) (x :: xs)
doParse s com (Act action) xs = Res (s <+> action) com (irrelevantBounds ()) xs
public export public export
data ParsingError tok = Error String (Maybe Bounds) data ParsingError tok = Error String (Maybe Bounds)
@ -337,9 +345,17 @@ data ParsingError tok = Error String (Maybe Bounds)
||| returns a pair of the parse result and the unparsed tokens (the remaining ||| returns a pair of the parse result and the unparsed tokens (the remaining
||| input). ||| input).
export export
parse : {c : Bool} -> (act : Grammar tok c ty) -> (xs : List (WithBounds tok)) -> parse : {c : Bool} -> (act : Grammar () tok c ty) -> (xs : List (WithBounds tok)) ->
Either (ParsingError tok) (ty, List (WithBounds tok)) Either (ParsingError tok) (ty, List (WithBounds tok))
parse act xs parse act xs
= case doParse False act xs of = case doParse neutral False act xs of
Failure _ _ msg ts => Left (Error msg ts) Failure _ _ msg ts => Left (Error msg ts)
Res _ v rest => Right (v.val, rest) Res _ _ v rest => Right (v.val, rest)
export
parseWith : Monoid state => {c : Bool} -> (act : Grammar state tok c ty) -> (xs : List (WithBounds tok)) ->
Either (ParsingError tok) (state, ty, List (WithBounds tok))
parseWith act xs
= case doParse neutral False act xs of
Failure _ _ msg ts => Left (Error msg ts)
Res s _ v rest => Right (s, v.val, rest)

View File

@ -136,6 +136,9 @@ Eq PathTokenKind where
PathToken : Type PathToken : Type
PathToken = Token PathTokenKind PathToken = Token PathTokenKind
PathGrammar : Bool -> Type -> Type
PathGrammar = Grammar () PathToken
TokenKind PathTokenKind where TokenKind PathTokenKind where
TokType PTText = String TokType PTText = String
TokType (PTPunct _) = () TokType (PTPunct _) = ()
@ -156,7 +159,7 @@ lexPath : String -> List (WithBounds PathToken)
lexPath str = let (tokens, _, _, _) = lex pathTokenMap str in tokens lexPath str = let (tokens, _, _, _) = lex pathTokenMap str in tokens
-- match both '/' and '\\' regardless of the platform. -- match both '/' and '\\' regardless of the platform.
bodySeparator : Grammar PathToken True () bodySeparator : PathGrammar True ()
bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/') bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/')
-- Windows will automatically translate '/' to '\\'. And the verbatim prefix, -- Windows will automatically translate '/' to '\\'. And the verbatim prefix,
@ -164,7 +167,7 @@ bodySeparator = (match $ PTPunct '\\') <|> (match $ PTPunct '/')
-- However, we just parse it and ignore it. -- However, we just parse it and ignore it.
-- --
-- Example: \\?\ -- Example: \\?\
verbatim : Grammar PathToken True () verbatim : PathGrammar True ()
verbatim = verbatim =
do do
ignore $ count (exactly 2) $ match $ PTPunct '\\' ignore $ count (exactly 2) $ match $ PTPunct '\\'
@ -173,7 +176,7 @@ verbatim =
pure () pure ()
-- Example: \\server\share -- Example: \\server\share
unc : Grammar PathToken True Volume unc : PathGrammar True Volume
unc = unc =
do do
ignore $ count (exactly 2) $ match $ PTPunct '\\' ignore $ count (exactly 2) $ match $ PTPunct '\\'
@ -183,7 +186,7 @@ unc =
pure $ UNC server share pure $ UNC server share
-- Example: \\?\server\share -- Example: \\?\server\share
verbatimUnc : Grammar PathToken True Volume verbatimUnc : PathGrammar True Volume
verbatimUnc = verbatimUnc =
do do
verbatim verbatim
@ -193,7 +196,7 @@ verbatimUnc =
pure $ UNC server share pure $ UNC server share
-- Example: C: -- Example: C:
disk : Grammar PathToken True Volume disk : PathGrammar True Volume
disk = disk =
do do
text <- match PTText text <- match PTText
@ -204,31 +207,31 @@ disk =
pure $ Disk (toUpper disk) pure $ Disk (toUpper disk)
-- Example: \\?\C: -- Example: \\?\C:
verbatimDisk : Grammar PathToken True Volume verbatimDisk : PathGrammar True Volume
verbatimDisk = verbatimDisk =
do do
verbatim verbatim
disk <- disk disk <- disk
pure disk pure disk
parseVolume : Grammar PathToken True Volume parseVolume : PathGrammar True Volume
parseVolume = parseVolume =
verbatimUnc verbatimUnc
<|> verbatimDisk <|> verbatimDisk
<|> unc <|> unc
<|> disk <|> disk
parseBody : Grammar PathToken True Body parseBody : PathGrammar True Body
parseBody = parseBody =
do do
text <- match PTText text <- match PTText
the (Grammar _ False _) $ the (PathGrammar False _) $
case text of case text of
".." => pure ParentDir ".." => pure ParentDir
"." => pure CurDir "." => pure CurDir
normal => pure (Normal normal) normal => pure (Normal normal)
parsePath : Grammar PathToken False Path parsePath : PathGrammar False Path
parsePath = parsePath =
do do
vol <- optional parseVolume vol <- optional parseVolume

View File

@ -21,8 +21,10 @@ data OpPrec
-- Tokens are either operators or already parsed expressions in some -- Tokens are either operators or already parsed expressions in some
-- higher level language -- higher level language
public export public export
data Tok op a = Op FC op OpPrec data Tok op a
| Expr a = ||| The second FC is for the operator alone
Op FC FC op OpPrec
| Expr a
-- The result of shunting is a parse tree with the precedences made explicit -- The result of shunting is a parse tree with the precedences made explicit
-- in the tree. -- in the tree.
@ -34,14 +36,14 @@ data Tok op a = Op FC op OpPrec
-- there's a better way though! -- there's a better way though!
public export public export
data Tree op a = Infix FC op (Tree op a) (Tree op a) data Tree op a = Infix FC FC op (Tree op a) (Tree op a)
| Pre FC op (Tree op a) | Pre FC FC op (Tree op a)
| Leaf a | Leaf a
export export
(Show op, Show a) => Show (Tree op a) where (Show op, Show a) => Show (Tree op a) where
show (Infix _ op l r) = "(" ++ show op ++ " " ++ show l ++ " " ++ show r ++ ")" show (Infix _ _ op l r) = "(" ++ show op ++ " " ++ show l ++ " " ++ show r ++ ")"
show (Pre _ op l) = "(" ++ show op ++ " " ++ show l ++ ")" show (Pre _ _ op l) = "(" ++ show op ++ " " ++ show l ++ ")"
show (Leaf val) = show val show (Leaf val) = show val
Show OpPrec where Show OpPrec where
@ -52,7 +54,7 @@ Show OpPrec where
export export
(Show op, Show a) => Show (Tok op a) where (Show op, Show a) => Show (Tok op a) where
show (Op _ op p) = show op ++ " " ++ show p show (Op _ _ op p) = show op ++ " " ++ show p
show (Expr val) = show val show (Expr val) = show val
-- Label for the output queue state -- Label for the output queue state
@ -60,9 +62,9 @@ data Out : Type where
output : List (Tree op a) -> Tok op a -> output : List (Tree op a) -> Tok op a ->
Core (List (Tree op a)) Core (List (Tree op a))
output [] (Op _ _ _) = throw (InternalError "Invalid input to shunting") output [] (Op _ _ _ _) = throw (InternalError "Invalid input to shunting")
output (x :: stk) (Op loc str (Prefix _)) = pure $ Pre loc str x :: stk output (x :: stk) (Op loc opFC str (Prefix _)) = pure $ Pre loc opFC str x :: stk
output (x :: y :: stk) (Op loc str _) = pure $ Infix loc str y x :: stk output (x :: y :: stk) (Op loc opFC str _) = pure $ Infix loc opFC str y x :: stk
output stk (Expr a) = pure $ Leaf a :: stk output stk (Expr a) = pure $ Leaf a :: stk
output _ _ = throw (InternalError "Invalid input to shunting") output _ _ = throw (InternalError "Invalid input to shunting")
@ -101,36 +103,32 @@ higher loc opl l opr r
((getPrec l == getPrec r) && isLAssoc l) ((getPrec l == getPrec r) && isLAssoc l)
processStack : Show op => {auto o : Ref Out (List (Tree op a))} -> processStack : Show op => {auto o : Ref Out (List (Tree op a))} ->
List (FC, op, OpPrec) -> op -> OpPrec -> List (FC, FC, op, OpPrec) -> op -> OpPrec ->
Core (List (FC, op, OpPrec)) Core (List (FC, FC, op, OpPrec))
processStack [] op prec = pure [] processStack [] op prec = pure []
processStack ((loc, opx, sprec) :: xs) opy prec processStack (x@(loc, opFC, opx, sprec) :: xs) opy prec
= if !(higher loc opx sprec opy prec) = if !(higher loc opx sprec opy prec)
then do emit (Op loc opx sprec) then do emit (Op loc opFC opx sprec)
processStack xs opy prec processStack xs opy prec
else pure ((loc, opx, sprec) :: xs) else pure (x :: xs)
shunt : Show op => {auto o : Ref Out (List (Tree op a))} -> shunt : Show op => {auto o : Ref Out (List (Tree op a))} ->
(opstk : List (FC, op, OpPrec)) -> (opstk : List (FC, FC, op, OpPrec)) ->
List (Tok op a) -> Core (Tree op a) List (Tok op a) -> Core (Tree op a)
shunt stk (Expr x :: rest) shunt stk (Expr x :: rest)
= do emit (Expr x) = do emit (Expr x)
shunt stk rest shunt stk rest
shunt stk (Op loc op prec :: rest) shunt stk (Op loc opFC op prec :: rest)
= do stk' <- processStack stk op prec = do stk' <- processStack stk op prec
shunt ((loc, op, prec) :: stk') rest shunt ((loc, opFC, op, prec) :: stk') rest
shunt stk [] shunt stk []
= do traverse_ (\s => emit (Op (sloc s) (sop s) (sprec s))) stk = do traverse_ (emit . mkOp) stk
[out] <- get Out [out] <- get Out
| out => throw (InternalError "Invalid input to shunting") | out => throw (InternalError "Invalid input to shunting")
pure out pure out
where where
sloc : (annot, b, c) -> annot mkOp : (FC, FC, op, OpPrec) -> Tok op a
sloc (x, y, z) = x mkOp (loc, opFC, op, prec) = Op loc opFC op prec
sop : (annot, b, c) -> b
sop (x, y, z) = y
sprec : (annot, b, c) -> c
sprec (x, y, z) = z
export export
parseOps : Show op => List (Tok op a) -> Core (Tree op a) parseOps : Show op => List (Tok op a) -> Core (Tree op a)

View File

@ -1,26 +0,0 @@
module Parser.Rule.Common
import public Libraries.Text.Lexer
import public Libraries.Text.Parser
%default total
public export
Rule : Type -> Type -> Type
Rule token ty = Grammar token True ty
public export
EmptyRule : Type -> Type -> Type
EmptyRule token ty = Grammar token False ty
export
location : {token : _} -> EmptyRule token (Int, Int)
location = startBounds <$> position
export
endLocation : {token : _} -> EmptyRule token (Int, Int)
endLocation = endBounds <$> position
export
column : {token : _ } -> EmptyRule token Int
column = snd <$> location

View File

@ -1,7 +1,6 @@
module Parser.Rule.Package module Parser.Rule.Package
import public Parser.Lexer.Package import public Parser.Lexer.Package
import public Parser.Rule.Common
import Data.List import Data.List
import Data.List1 import Data.List1
@ -12,11 +11,11 @@ import Core.Name.Namespace
public export public export
Rule : Type -> Type Rule : Type -> Type
Rule = Rule Token Rule = Grammar () Token True
public export public export
PackageEmptyRule : Type -> Type EmptyRule : Type -> Type
PackageEmptyRule = EmptyRule Token EmptyRule = Grammar () Token False
export export
equals : Rule () equals : Rule ()

View File

@ -1,10 +1,11 @@
module Parser.Rule.Source module Parser.Rule.Source
import public Parser.Lexer.Source import public Parser.Lexer.Source
import public Parser.Rule.Common
import public Parser.Support import public Parser.Support
import Core.Context
import Core.TT import Core.TT
import Core.Metadata
import Data.List1 import Data.List1
import Data.Strings import Data.Strings
import Libraries.Data.List.Extra import Libraries.Data.List.Extra
@ -19,14 +20,14 @@ import Libraries.Data.String.Extra
public export public export
Rule : Type -> Type Rule : Type -> Type
Rule = Rule Token Rule ty = Grammar SemanticDecorations Token True ty
public export public export
SourceEmptyRule : Type -> Type EmptyRule : Type -> Type
SourceEmptyRule = EmptyRule Token EmptyRule ty = Grammar SemanticDecorations Token False ty
export export
eoi : SourceEmptyRule () eoi : EmptyRule ()
eoi = ignore $ nextIs "Expected end of input" isEOI eoi = ignore $ nextIs "Expected end of input" isEOI
where where
isEOI : Token -> Bool isEOI : Token -> Bool
@ -181,6 +182,13 @@ pragma n =
else Nothing else Nothing
_ => Nothing) _ => Nothing)
export
builtinType : Rule BuiltinType
builtinType =
BuiltinNatural <$ exactIdent "Natural"
<|> NaturalToInteger <$ exactIdent "NaturalToInteger"
<|> IntegerToNatural <$ exactIdent "IntegerToNatural"
export export
operator : Rule Name operator : Rule Name
operator operator
@ -208,11 +216,11 @@ namespacedIdent
Ident i => Just (Nothing, i) Ident i => Just (Nothing, i)
_ => Nothing) _ => Nothing)
isCapitalisedIdent : WithBounds String -> SourceEmptyRule () isCapitalisedIdent : WithBounds String -> EmptyRule ()
isCapitalisedIdent str = isCapitalisedIdent str =
let val = str.val let val = str.val
loc = str.bounds loc = str.bounds
err : SourceEmptyRule () err : EmptyRule ()
= failLoc loc ("Expected a capitalised identifier, got: " ++ val) = failLoc loc ("Expected a capitalised identifier, got: " ++ val)
in case strM val of in case strM val of
StrNil => err StrNil => err
@ -249,7 +257,7 @@ reservedNames
, "String", "Char", "Double", "Lazy", "Inf", "Force", "Delay" , "String", "Char", "Double", "Lazy", "Inf", "Force", "Delay"
] ]
isNotReservedIdent : WithBounds String -> SourceEmptyRule () isNotReservedIdent : WithBounds String -> EmptyRule ()
isNotReservedIdent x isNotReservedIdent x
= if x.val `elem` reservedNames = if x.val `elem` reservedNames
then failLoc x.bounds $ "can't use reserved name " ++ x.val then failLoc x.bounds $ "can't use reserved name " ++ x.val
@ -260,7 +268,7 @@ opNonNS : Rule Name
opNonNS = symbol "(" *> (operator <|> postfixProj) <* symbol ")" opNonNS = symbol "(" *> (operator <|> postfixProj) <* symbol ")"
identWithCapital : (capitalised : Bool) -> WithBounds String -> identWithCapital : (capitalised : Bool) -> WithBounds String ->
SourceEmptyRule () EmptyRule ()
identWithCapital b x = if b then isCapitalisedIdent x else pure () identWithCapital b x = if b then isCapitalisedIdent x else pure ()
nameWithCapital : (capitalised : Bool) -> Rule Name nameWithCapital : (capitalised : Bool) -> Rule Name
@ -269,7 +277,7 @@ nameWithCapital b = opNonNS <|> do
opNS nsx <|> nameNS nsx opNS nsx <|> nameNS nsx
where where
nameNS : WithBounds (Maybe Namespace, String) -> SourceEmptyRule Name nameNS : WithBounds (Maybe Namespace, String) -> EmptyRule Name
nameNS nsx = do nameNS nsx = do
let id = snd <$> nsx let id = snd <$> nsx
identWithCapital b id identWithCapital b id
@ -317,23 +325,22 @@ export
init : IndentInfo init : IndentInfo
init = 0 init = 0
continueF : SourceEmptyRule () -> (indent : IndentInfo) -> SourceEmptyRule () continueF : EmptyRule () -> (indent : IndentInfo) -> EmptyRule ()
continueF err indent continueF err indent
= do eoi; err = do eoi; err
<|> do keyword "where"; err <|> do keyword "where"; err
<|> do col <- Common.column <|> do col <- column
if col <= indent when (col <= indent)
then err err
else pure ()
||| Fail if this is the end of a block entry or end of file ||| Fail if this is the end of a block entry or end of file
export export
continue : (indent : IndentInfo) -> SourceEmptyRule () continue : (indent : IndentInfo) -> EmptyRule ()
continue = continueF (fail "Unexpected end of expression") continue = continueF (fail "Unexpected end of expression")
||| As 'continue' but failing is fatal (i.e. entire parse fails) ||| As 'continue' but failing is fatal (i.e. entire parse fails)
export export
mustContinue : (indent : IndentInfo) -> Maybe String -> SourceEmptyRule () mustContinue : (indent : IndentInfo) -> Maybe String -> EmptyRule ()
mustContinue indent Nothing mustContinue indent Nothing
= continueF (fatalError "Unexpected end of expression") indent = continueF (fatalError "Unexpected end of expression") indent
mustContinue indent (Just req) mustContinue indent (Just req)
@ -355,7 +362,7 @@ Show ValidIndent where
show (AfterPos i) = "[after " ++ show i ++ "]" show (AfterPos i) = "[after " ++ show i ++ "]"
show EndOfBlock = "[EOB]" show EndOfBlock = "[EOB]"
checkValid : ValidIndent -> Int -> SourceEmptyRule () checkValid : ValidIndent -> Int -> EmptyRule ()
checkValid AnyIndent c = pure () checkValid AnyIndent c = pure ()
checkValid (AtPos x) c = if c == x checkValid (AtPos x) c = if c == x
then pure () then pure ()
@ -386,29 +393,27 @@ isTerminator _ = False
||| It's the end if we have a terminating token, or the next token starts ||| It's the end if we have a terminating token, or the next token starts
||| in or before indent. Works by looking ahead but not consuming. ||| in or before indent. Works by looking ahead but not consuming.
export export
atEnd : (indent : IndentInfo) -> SourceEmptyRule () atEnd : (indent : IndentInfo) -> EmptyRule ()
atEnd indent atEnd indent
= eoi = eoi
<|> do ignore $ nextIs "Expected end of block" isTerminator <|> do ignore $ nextIs "Expected end of block" isTerminator
<|> do col <- Common.column <|> do col <- column
if (col <= indent) when (not (col <= indent))
then pure () $ fail "Not the end of a block entry"
else fail "Not the end of a block entry"
-- Check we're at the end, but only by looking at indentation -- Check we're at the end, but only by looking at indentation
export export
atEndIndent : (indent : IndentInfo) -> SourceEmptyRule () atEndIndent : (indent : IndentInfo) -> EmptyRule ()
atEndIndent indent atEndIndent indent
= eoi = eoi
<|> do col <- Common.column <|> do col <- column
if col <= indent when (not (col <= indent))
then pure () $ fail "Not the end of a block entry"
else fail "Not the end of a block entry"
-- Parse a terminator, return where the next block entry -- Parse a terminator, return where the next block entry
-- must start, given where the current block entry started -- must start, given where the current block entry started
terminator : ValidIndent -> Int -> SourceEmptyRule ValidIndent terminator : ValidIndent -> Int -> EmptyRule ValidIndent
terminator valid laststart terminator valid laststart
= do eoi = do eoi
pure EndOfBlock pure EndOfBlock
@ -430,7 +435,7 @@ terminator valid laststart
-- Expected indentation for the next token can either be anything (if -- Expected indentation for the next token can either be anything (if
-- we're inside a brace delimited block) or in exactly the initial column -- we're inside a brace delimited block) or in exactly the initial column
-- (if we're inside an indentation delimited block) -- (if we're inside an indentation delimited block)
afterDedent : ValidIndent -> Int -> SourceEmptyRule ValidIndent afterDedent : ValidIndent -> Int -> EmptyRule ValidIndent
afterDedent AnyIndent col afterDedent AnyIndent col
= if col <= laststart = if col <= laststart
then pure AnyIndent then pure AnyIndent
@ -456,7 +461,7 @@ blockEntry valid rule
pure (p, valid') pure (p, valid')
blockEntries : ValidIndent -> (IndentInfo -> Rule ty) -> blockEntries : ValidIndent -> (IndentInfo -> Rule ty) ->
SourceEmptyRule (List ty) EmptyRule (List ty)
blockEntries valid rule blockEntries valid rule
= do eoi; pure [] = do eoi; pure []
<|> do res <- blockEntry valid rule <|> do res <- blockEntry valid rule
@ -465,7 +470,7 @@ blockEntries valid rule
<|> pure [] <|> pure []
export export
block : (IndentInfo -> Rule ty) -> SourceEmptyRule (List ty) block : (IndentInfo -> Rule ty) -> EmptyRule (List ty)
block item block item
= do symbol "{" = do symbol "{"
commit commit
@ -481,35 +486,35 @@ block item
||| by curly braces). `rule` is a function of the actual indentation ||| by curly braces). `rule` is a function of the actual indentation
||| level. ||| level.
export export
blockAfter : Int -> (IndentInfo -> Rule ty) -> SourceEmptyRule (List ty) blockAfter : Int -> (IndentInfo -> Rule ty) -> EmptyRule (List ty)
blockAfter mincol item blockAfter mincol item
= do symbol "{" = do symbol "{"
commit commit
ps <- blockEntries AnyIndent item ps <- blockEntries AnyIndent item
symbol "}" symbol "}"
pure ps pure ps
<|> do col <- Common.column <|> do col <- column
if col <= mincol ifThenElse (col <= mincol)
then pure [] (pure [])
else blockEntries (AtPos col) item $ blockEntries (AtPos col) item
export export
blockWithOptHeaderAfter : blockWithOptHeaderAfter :
(column : Int) -> (column : Int) ->
(header : IndentInfo -> Rule hd) -> (header : IndentInfo -> Rule hd) ->
(item : IndentInfo -> Rule ty) -> (item : IndentInfo -> Rule ty) ->
SourceEmptyRule (Maybe hd, List ty) EmptyRule (Maybe hd, List ty)
blockWithOptHeaderAfter {ty} mincol header item blockWithOptHeaderAfter {ty} mincol header item
= do symbol "{" = do symbol "{"
commit commit
hidt <- optional $ blockEntry AnyIndent header hidt <- optional $ blockEntry AnyIndent header
restOfBlock hidt restOfBlock hidt
<|> do col <- Common.column <|> do col <- column
if col <= mincol ifThenElse (col <= mincol)
then pure (Nothing, []) (pure (Nothing, []))
else do hidt <- optional $ blockEntry (AtPos col) header $ do hidt <- optional $ blockEntry (AtPos col) header
ps <- blockEntries (AtPos col) item ps <- blockEntries (AtPos col) item
pure (map fst hidt, ps) pure (map fst hidt, ps)
where where
restOfBlock : Maybe (hd, ValidIndent) -> Rule (Maybe hd, List ty) restOfBlock : Maybe (hd, ValidIndent) -> Rule (Maybe hd, List ty)
restOfBlock (Just (h, idt)) = do ps <- blockEntries idt item restOfBlock (Just (h, idt)) = do ps <- blockEntries idt item

View File

@ -5,6 +5,9 @@ import public Parser.Rule.Source
import public Parser.Unlit import public Parser.Unlit
import Core.Core import Core.Core
import Core.Name
import Core.Metadata
import Core.FC
import System.File import System.File
import Libraries.Utils.Either import Libraries.Utils.Either
@ -14,21 +17,21 @@ export
runParserTo : {e : _} -> runParserTo : {e : _} ->
(fname : String) -> (fname : String) ->
Maybe LiterateStyle -> Lexer -> Maybe LiterateStyle -> Lexer ->
String -> Grammar Token e ty -> Either Error ty String -> Grammar SemanticDecorations Token e ty -> Either Error (SemanticDecorations, ty)
runParserTo fname lit reject str p runParserTo fname lit reject str p
= do str <- mapError (fromLitError fname) $ unlit lit str = do str <- mapError (fromLitError fname) $ unlit lit str
toks <- mapError (fromLexError fname) $ lexTo reject str toks <- mapError (fromLexError fname) $ lexTo reject str
parsed <- mapError (fromParsingError fname) $ parse p toks (decs, (parsed, _)) <- mapError (fromParsingError fname) $ parseWith p toks
Right (fst parsed) Right (decs, parsed)
export export
runParser : {e : _} -> runParser : {e : _} ->
(fname : String) -> Maybe LiterateStyle -> String -> (fname : String) -> Maybe LiterateStyle -> String ->
Grammar Token e ty -> Either Error ty Grammar SemanticDecorations Token e ty -> Either Error (SemanticDecorations, ty)
runParser fname lit = runParserTo fname lit (pred $ const False) runParser fname lit = runParserTo fname lit (pred $ const False)
export covering export covering
parseFile : (fname : String) -> Rule ty -> IO (Either Error ty) parseFile : (fname : String) -> Rule ty -> IO (Either Error (SemanticDecorations, ty))
parseFile fname p parseFile fname p
= do Right str <- readFile fname = do Right str <- readFile fname
| Left err => pure (Left (FileErr fname err)) | Left err => pure (Left (FileErr fname err))

View File

@ -55,6 +55,14 @@ getNameType rigc env fc x
do est <- get EST do est <- get EST
put EST put EST
(record { linearUsed $= ((MkVar lv) :: ) } est) (record { linearUsed $= ((MkVar lv) :: ) } est)
log "ide-mode.highlight" 8
$ "getNameType is trying to add Bound: "
++ show x ++ " (" ++ show fc ++ ")"
when (isSourceName x) $
whenJust (isConcreteFC fc) \nfc => do
log "ide-mode.highlight" 7 $ "getNameType is adding Bound: " ++ show x
addSemanticDecorations [(nfc, Bound, Just x)]
pure (Local fc (Just (isLet binder)) _ lv, gnf env bty) pure (Local fc (Just (isLet binder)) _ lv, gnf env bty)
Nothing => Nothing =>
do defs <- get Ctxt do defs <- get Ctxt
@ -68,6 +76,18 @@ getNameType rigc env fc x
DCon t a _ => DataCon t a DCon t a _ => DataCon t a
TCon t a _ _ _ _ _ _ => TyCon t a TCon t a _ _ _ _ _ _ => TyCon t a
_ => Func _ => Func
log "ide-mode.highlight" 8
$ "getNameType is trying to add something for: "
++ show def.fullname ++ " (" ++ show fc ++ ")"
when (isSourceName def.fullname) $
whenJust (isConcreteFC fc) \nfc => do
let decor = nameTypeDecoration nt
log "ide-mode.highlight" 7
$ "getNameType is adding " ++ show decor ++ ": " ++ show def.fullname
addSemanticDecorations [(nfc, decor, Just def.fullname)]
pure (Ref fc nt (Resolved i), gnf env (embed (type def))) pure (Ref fc nt (Resolved i), gnf env (embed (type def)))
where where
rigSafe : RigCount -> RigCount -> Core () rigSafe : RigCount -> RigCount -> Core ()
@ -89,7 +109,7 @@ getVarType rigc nest env fc x
Just (nestn, argns, tmf) => Just (nestn, argns, tmf) =>
do defs <- get Ctxt do defs <- get Ctxt
let arglen = length argns let arglen = length argns
let n' = maybe x id nestn let n' = fromMaybe x nestn
case !(lookupCtxtExact n' (gamma defs)) of case !(lookupCtxtExact n' (gamma defs)) of
Nothing => undefinedName fc n' Nothing => undefinedName fc n'
Just ndef => Just ndef =>
@ -110,6 +130,14 @@ getVarType rigc nest env fc x
log "metadata.names" 7 $ "getVarType is adding ↓" log "metadata.names" 7 $ "getVarType is adding ↓"
addNameType fc x env tyenv addNameType fc x env tyenv
when (isSourceName ndef.fullname) $
whenJust (isConcreteFC fc) \nfc => do
let decor = nameTypeDecoration nt
log "ide-mode.highlight" 7
$ "getNameType is adding "++ show decor ++": "
++ show ndef.fullname
addSemanticDecorations [(nfc, decor, Just ndef.fullname)]
pure (tm, arglen, gnf env tyenv) pure (tm, arglen, gnf env tyenv)
where where
useVars : {vars : _} -> useVars : {vars : _} ->

View File

@ -515,7 +515,7 @@ successful : {vars : _} ->
Bool -> -- constraints allowed Bool -> -- constraints allowed
List (Maybe Name, Core a) -> List (Maybe Name, Core a) ->
Core (List (Either (Maybe Name, Error) Core (List (Either (Maybe Name, Error)
(Nat, a, Defs, UState, EState vars))) (Nat, a, Defs, UState, EState vars, Metadata)))
successful allowCons [] = pure [] successful allowCons [] = pure []
successful allowCons ((tm, elab) :: elabs) successful allowCons ((tm, elab) :: elabs)
= do ust <- get UST = do ust <- get UST
@ -555,7 +555,7 @@ successful allowCons ((tm, elab) :: elabs)
elabs' <- successful allowCons elabs elabs' <- successful allowCons elabs
-- Record success, and the state we ended at -- Record success, and the state we ended at
pure (Right (minus ncons' ncons, pure (Right (minus ncons' ncons,
res, defs', ust', est') :: elabs')) res, defs', ust', est', md') :: elabs'))
(\err => do put UST ust (\err => do put UST ust
put EST est put EST est
put MD md put MD md
@ -576,9 +576,10 @@ exactlyOne' allowCons fc env [(tm, elab)] = elab
exactlyOne' {vars} allowCons fc env all exactlyOne' {vars} allowCons fc env all
= do elabs <- successful allowCons all = do elabs <- successful allowCons all
case getRight elabs of case getRight elabs of
Right (res, defs, ust, est) => Right (res, defs, ust, est, md) =>
do put UST ust do put UST ust
put EST est put EST est
put MD md
put Ctxt defs put Ctxt defs
commit commit
pure res pure res

View File

@ -416,6 +416,12 @@ checkBindVar rig elabinfo nest env fc str topexp
noteLHSPatVar elabmode (UN str) noteLHSPatVar elabmode (UN str)
notePatVar n notePatVar n
est <- get EST est <- get EST
whenJust (isConcreteFC fc) \nfc => do
log "ide-mode.highlight" 7 $ "getNameType is adding Bound: " ++ show n
addSemanticDecorations [(nfc, Bound, Just n)]
case lookup n (boundNames est) of case lookup n (boundNames est) of
Nothing => Nothing =>
do (tm, exp, bty) <- mkPatternHole fc rig n env do (tm, exp, bty) <- mkPatternHole fc rig n env

View File

@ -245,7 +245,7 @@ checkQuoteDecl rig elabinfo nest env fc ds exp
qds <- reflect fc defs (onLHS (elabMode elabinfo)) env ds' qds <- reflect fc defs (onLHS (elabMode elabinfo)) env ds'
unqs <- get Unq unqs <- get Unq
qd <- getCon fc defs (reflectionttimp "Decl") qd <- getCon fc defs (reflectionttimp "Decl")
qty <- appCon fc defs (preludetypes "List") [qd] qty <- appCon fc defs (basics "List") [qd]
checkExp rig elabinfo env fc checkExp rig elabinfo env fc
!(bindUnqs unqs rig elabinfo nest env qds) !(bindUnqs unqs rig elabinfo nest env qds)
(gnf env qty) exp (gnf env qty) exp

View File

@ -177,18 +177,21 @@ recUpdate : {vars : _} ->
List IFieldUpdate -> List IFieldUpdate ->
(rec : RawImp) -> (grecty : Glued vars) -> (rec : RawImp) -> (grecty : Glued vars) ->
Core RawImp Core RawImp
recUpdate rigc elabinfo loc nest env flds rec grecty recUpdate rigc elabinfo iloc nest env flds rec grecty
= do defs <- get Ctxt = do defs <- get Ctxt
rectynf <- getNF grecty rectynf <- getNF grecty
let Just rectyn = getRecordType env rectynf let Just rectyn = getRecordType env rectynf
| Nothing => throw (RecordTypeNeeded loc env) | Nothing => throw (RecordTypeNeeded iloc env)
fldn <- genFieldName "__fld" fldn <- genFieldName "__fld"
sides <- getAllSides loc flds rectyn rec sides <- getAllSides iloc flds rectyn rec
(Field Nothing fldn (IVar loc (UN fldn))) (Field Nothing fldn (IVar vloc (UN fldn)))
pure $ ICase loc rec (Implicit loc False) [mkClause sides] pure $ ICase vloc rec (Implicit vloc False) [mkClause sides]
where where
vloc : FC
vloc = virtualiseFC iloc
mkClause : Rec -> ImpClause mkClause : Rec -> ImpClause
mkClause rec = PatClause loc (toLHS loc rec) (toRHS loc rec) mkClause rec = PatClause vloc (toLHS vloc rec) (toRHS vloc rec)
needType : Error -> Bool needType : Error -> Bool
needType (RecordTypeNeeded _ _) = True needType (RecordTypeNeeded _ _) = True

View File

@ -107,19 +107,20 @@ checkRewrite : {vars : _} ->
Core (Term vars, Glued vars) Core (Term vars, Glued vars)
checkRewrite rigc elabinfo nest env fc rule tm Nothing checkRewrite rigc elabinfo nest env fc rule tm Nothing
= throw (GenericMsg fc "Can't infer a type for rewrite") = throw (GenericMsg fc "Can't infer a type for rewrite")
checkRewrite {vars} rigc elabinfo nest env fc rule tm (Just expected) checkRewrite {vars} rigc elabinfo nest env ifc rule tm (Just expected)
= delayOnFailure fc rigc env expected rewriteErr 10 (\delayed => = delayOnFailure ifc rigc env expected rewriteErr 10 $ \delayed =>
do (rulev, grulet) <- check erased elabinfo nest env rule Nothing do let vfc = virtualiseFC ifc
(rulev, grulet) <- check erased elabinfo nest env rule Nothing
rulet <- getTerm grulet rulet <- getTerm grulet
expTy <- getTerm expected expTy <- getTerm expected
when delayed $ log "elab.rewrite" 5 "Retrying rewrite" when delayed $ log "elab.rewrite" 5 "Retrying rewrite"
(lemma, pred, predty) <- elabRewrite fc env expTy rulet (lemma, pred, predty) <- elabRewrite vfc env expTy rulet
rname <- genVarName "_" rname <- genVarName "_"
pname <- genVarName "_" pname <- genVarName "_"
let pbind = Let fc erased pred predty let pbind = Let vfc erased pred predty
let rbind = Let fc erased (weaken rulev) (weaken rulet) let rbind = Let vfc erased (weaken rulev) (weaken rulet)
let env' = rbind :: pbind :: env let env' = rbind :: pbind :: env
@ -128,16 +129,15 @@ checkRewrite {vars} rigc elabinfo nest env fc rule tm (Just expected)
-- implicits for the rewriting lemma are in the right place. But, -- implicits for the rewriting lemma are in the right place. But,
-- we still need the right type for the EState, so weaken it once -- we still need the right type for the EState, so weaken it once
-- for each of the let bindings above. -- for each of the let bindings above.
(rwtm, grwty) <- inScope fc (pbind :: env) (rwtm, grwty) <-
(\e' => inScope {e=e'} fc env' inScope vfc (pbind :: env) $ \e' =>
(\e'' => check {e = e''} {vars = rname :: pname :: vars} inScope {e=e'} vfc env' $ \e'' =>
rigc elabinfo (weaken (weaken nest)) env' let offset = mkSizeOf [rname, pname] in
(apply (IVar fc lemma) [IVar fc pname, check {e = e''} rigc elabinfo (weakenNs offset nest) env'
IVar fc rname, (apply (IVar vfc lemma) [IVar vfc pname,
tm]) IVar vfc rname,
(Just (gnf env' tm])
(weakenNs (mkSizeOf [rname, pname]) expTy))) (Just (gnf env' (weakenNs offset expTy)))
))
rwty <- getTerm grwty rwty <- getTerm grwty
pure (Bind fc pname pbind (Bind fc rname rbind rwtm), let binding = Bind vfc pname pbind . Bind vfc rname rbind
gnf env (Bind fc pname pbind (Bind fc rname rbind rwty)))) pure (binding rwtm, gnf env (binding rwty))

View File

@ -330,7 +330,7 @@ mkCase : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
Int -> RawImp -> RawImp -> Core ClauseUpdate Int -> RawImp -> RawImp -> Core ClauseUpdate
mkCase {c} {u} fn orig lhs_raw mkCase {c} {u} fn orig lhs_raw
= do m <- newRef MD initMetadata = do m <- newRef MD (initMetadata "(interactive)")
defs <- get Ctxt defs <- get Ctxt
ust <- get UST ust <- get UST
catch catch

View File

@ -2,6 +2,7 @@ module TTImp.Parser
import Core.Context import Core.Context
import Core.Core import Core.Core
import Core.Metadata
import Core.Env import Core.Env
import Core.TT import Core.TT
import Parser.Source import Parser.Source
@ -81,7 +82,7 @@ visOption
<|> do keyword "private" <|> do keyword "private"
pure Private pure Private
visibility : SourceEmptyRule Visibility visibility : EmptyRule Visibility
visibility visibility
= visOption = visOption
<|> pure Private <|> pure Private
@ -124,7 +125,7 @@ visOpt
pure (Right opt) pure (Right opt)
getVisibility : Maybe Visibility -> List (Either Visibility FnOpt) -> getVisibility : Maybe Visibility -> List (Either Visibility FnOpt) ->
SourceEmptyRule Visibility EmptyRule Visibility
getVisibility Nothing [] = pure Private getVisibility Nothing [] = pure Private
getVisibility (Just vis) [] = pure vis getVisibility (Just vis) [] = pure vis
getVisibility Nothing (Left x :: xs) = getVisibility (Just x) xs getVisibility Nothing (Left x :: xs) = getVisibility (Just x) xs
@ -216,13 +217,13 @@ mutual
symbol ")" symbol ")"
pure e pure e
multiplicity : SourceEmptyRule (Maybe Integer) multiplicity : EmptyRule (Maybe Integer)
multiplicity multiplicity
= do c <- intLit = do c <- intLit
pure (Just c) pure (Just c)
<|> pure Nothing <|> pure Nothing
getMult : Maybe Integer -> SourceEmptyRule RigCount getMult : Maybe Integer -> EmptyRule RigCount
getMult (Just 0) = pure erased getMult (Just 0) = pure erased
getMult (Just 1) = pure linear getMult (Just 1) = pure linear
getMult Nothing = pure top getMult Nothing = pure top
@ -522,7 +523,7 @@ mutual
let fc = MkFC fname start end let fc = MkFC fname start end
pure (!(getFn lhs), ImpossibleClause fc lhs) pure (!(getFn lhs), ImpossibleClause fc lhs)
where where
getFn : RawImp -> SourceEmptyRule Name getFn : RawImp -> EmptyRule Name
getFn (IVar _ n) = pure n getFn (IVar _ n) = pure n
getFn (IApp _ f a) = getFn f getFn (IApp _ f a) = getFn f
getFn (IAutoApp _ f a) = getFn f getFn (IAutoApp _ f a) = getFn f
@ -593,7 +594,7 @@ recordParam fname indents
<|> do symbol "{" <|> do symbol "{"
commit commit
start <- location start <- location
info <- the (SourceEmptyRule (PiInfo RawImp)) info <- the (EmptyRule (PiInfo RawImp))
(pure AutoImplicit <* keyword "auto" (pure AutoImplicit <* keyword "auto"
<|>(do <|>(do
keyword "default" keyword "default"
@ -662,10 +663,6 @@ logLevel
lvl <- intLit lvl <- intLit
pure (Just (topic, fromInteger lvl)) pure (Just (topic, fromInteger lvl))
builtinType : Rule BuiltinType
builtinType =
BuiltinNatural <$ exactIdent "Natural"
directive : FileName -> IndentInfo -> Rule ImpDecl directive : FileName -> IndentInfo -> Rule ImpDecl
directive fname indents directive fname indents
= do pragma "logging" = do pragma "logging"

View File

@ -3,9 +3,11 @@
module TTImp.ProcessBuiltin module TTImp.ProcessBuiltin
import Libraries.Data.Bool.Extra import Libraries.Data.Bool.Extra
import Data.Fin
import Libraries.Data.NameMap import Libraries.Data.NameMap
import Data.List import Data.List
import Core.CaseTree
import Core.Core import Core.Core
import Core.Context import Core.Context
import Core.Context.Log import Core.Context.Log
@ -16,24 +18,73 @@ import Core.UnifyState
import TTImp.TTImp import TTImp.TTImp
showDefType : Def -> String
showDefType None = "undefined"
showDefType (PMDef {}) = "function"
showDefType (ExternDef {}) = "external function"
showDefType (ForeignDef {}) = "foreign function"
showDefType (Builtin {}) = "builtin function"
showDefType (DCon {}) = "data constructor"
showDefType (TCon {}) = "type constructor"
showDefType (Hole {}) = "hole"
showDefType (BySearch {}) = "search"
showDefType (Guess {}) = "guess"
showDefType ImpBind = "bound name"
showDefType Delayed = "delayed"
||| Get the return type. ||| Get the return type.
getRetTy : {vars : _} -> Term vars -> Maybe (vars ** Term vars) getRetType : {vars : _} -> Term vars -> Maybe (vars ** Term vars)
getRetTy tm@(Bind _ x b scope) = case b of getRetType tm@(Bind _ x b scope) = case b of
Lam _ _ _ _ => Nothing Let _ _ val _ => getRetType $ subst {x} val scope
Let _ _ val _ => getRetTy $ subst {x} val scope Pi _ _ _ _ => getRetType scope
Pi _ _ _ _ => getRetTy scope
_ => Nothing _ => Nothing
getRetTy tm = Just (vars ** tm) getRetType tm = Just (vars ** tm)
||| Get the top level type constructor if there is one.
getTypeCons : {vars : _} -> Term vars -> Maybe Name
getTypeCons (Local _ _ _ p) = Just $ nameAt p
getTypeCons (Ref _ _ name) = Just name
getTypeCons (Meta {}) = Nothing
getTypeCons (Bind _ x b scope) = case b of
Let _ _ val _ => getTypeCons $ subst {x} val scope
_ => Nothing
getTypeCons (App _ fn _) = getTypeCons fn
getTypeCons _ = Nothing
getTypeArgs : {vars : _} -> Term vars -> List (vars ** Term vars)
getTypeArgs (Bind _ x b tm) = case b of
Let _ _ val _ => getTypeArgs $ subst {x} val tm
Pi _ _ _ arg => (_ ** arg) :: getTypeArgs tm
_ => []
getTypeArgs _ = []
getNEArgs : {vars : _} -> Term vars -> List (vars ** Term vars)
getNEArgs (Bind _ x b tm) = case b of
Let _ _ val _ => getNEArgs $ subst {x} val tm
Pi _ mul _ arg => if isErased mul
then getNEArgs tm
else (_ ** arg) :: getNEArgs tm
_ => []
getNEArgs _ = []
||| Get the first non-erased argument type. ||| Get the first non-erased argument type.
getFirstNETy : {vars : _} -> Term vars -> Maybe (vars ** Term vars) getFirstNEType : {vars : _} -> Term vars -> Maybe (vars ** Term vars)
getFirstNETy (Bind _ x b tm) = case b of getFirstNEType tm = case getNEArgs tm of
Let _ _ val _ => getFirstNETy $ subst {x} val tm [] => Nothing
arg :: _ => Just arg
||| Get the index of the first non-erased argument if it exists.
getNEIndex : (arity : Nat) -> Term vars -> Maybe (Fin arity)
getNEIndex ar (Bind _ x b tm) = case b of
Let _ _ val _ => getNEIndex ar $ subst {x} val tm
Pi _ mul _ arg => if isErased mul Pi _ mul _ arg => if isErased mul
then getFirstNETy tm then getNEIndex ar tm >>=
else Just (_ ** arg) \k => case strengthen (FS k) of
Left _ => Nothing
Right k' => Just k'
else natToFin 0 ar
_ => Nothing _ => Nothing
getFirstNETy tm = Nothing getNEIndex _ _ = Nothing
||| Do the terms match ignoring arguments to type constructors. ||| Do the terms match ignoring arguments to type constructors.
termConMatch : Term vs -> Term vs' -> Bool termConMatch : Term vs -> Term vs' -> Bool
@ -72,6 +123,11 @@ isStrict (PrimVal _ _) = True
isStrict (Erased _ _) = True isStrict (Erased _ _) = True
isStrict (TType _) = True isStrict (TType _) = True
getNatBuiltin : Ref Ctxt Defs => Name -> Core (Maybe NatBuiltin)
getNatBuiltin n = do
n' <- getFullName n
lookup n' . natTyNames . builtinTransforms <$> get Ctxt
||| Get the name and arity (of non-erased arguments only) of a list of names. ||| Get the name and arity (of non-erased arguments only) of a list of names.
||| `cons` should all be data constructors (`DCon`) otherwise it will throw an error. ||| `cons` should all be data constructors (`DCon`) otherwise it will throw an error.
getConsGDef : Context -> FC -> (cons : List Name) -> Core $ List (Name, GlobalDef) getConsGDef : Context -> FC -> (cons : List Name) -> Core $ List (Name, GlobalDef)
@ -84,8 +140,8 @@ getConsGDef c fc = traverse \n => do
||| Check a list of constructors has exactly ||| Check a list of constructors has exactly
||| 1 'Z'-like constructor ||| 1 'Z'-like constructor
||| and 1 `S`-like constructor, which has type `ty -> ty` or `ty arg -> `ty (f arg)`. ||| and 1 `S`-like constructor, which has type `ty -> ty` or `ty arg -> `ty (f arg)`.
checkCons : Context -> (cons : List (Name, GlobalDef)) -> (dataType : Name) -> FC -> Core NatBuiltin checkNatCons : Context -> (cons : List (Name, GlobalDef)) -> (dataType : Name) -> FC -> Core NatBuiltin
checkCons c cons ty fc = case !(foldr checkCon (pure (Nothing, Nothing)) cons) of checkNatCons c cons ty fc = case !(foldr checkCon (pure (Nothing, Nothing)) cons) of
(Just zero, Just succ) => pure $ MkNatBuiltin {zero, succ} (Just zero, Just succ) => pure $ MkNatBuiltin {zero, succ}
(Nothing, _) => throw $ GenericMsg fc $ "No 'Z'-like constructors for " ++ show ty ++ "." (Nothing, _) => throw $ GenericMsg fc $ "No 'Z'-like constructors for " ++ show ty ++ "."
(_, Nothing) => throw $ GenericMsg fc $ "No 'S'-like constructors for " ++ show ty ++ "." (_, Nothing) => throw $ GenericMsg fc $ "No 'S'-like constructors for " ++ show ty ++ "."
@ -102,12 +158,12 @@ checkCons c cons ty fc = case !(foldr checkCon (pure (Nothing, Nothing)) cons) o
checkTyS n gdef = do checkTyS n gdef = do
let type = gdef.type let type = gdef.type
erase = gdef.eraseArgs erase = gdef.eraseArgs
let Just (_ ** arg) = getFirstNETy type let Just (_ ** arg) = getFirstNEType type
| Nothing => throw $ InternalError "Expected a non-erased argument, found none." | Nothing => throw $ InternalError "Expected a non-erased argument, found none."
let Just (_ ** ret) = getRetTy type let Just (_ ** ret) = getRetType type
| Nothing => throw $ InternalError $ "Unexpected type " ++ show type | Nothing => throw $ InternalError $ "Unexpected type " ++ show type
unless (termConMatch arg ret) $ throw $ GenericMsg fc $ "Incorrect type for 'S'-like constructor for " ++ show ty ++ "." unless (termConMatch arg ret) $ throw $ GenericMsg fc $ "Incorrect type for 'S'-like constructor for " ++ show ty ++ "."
unless (isStrict arg) $ throw $ GenericMsg fc $ "Natural builtin does not support lazy types, as they can be potentially infinite." unless (isStrict arg) $ throw $ GenericMsg fc $ "Natural builtin does not support lazy types."
pure () pure ()
||| Check a constructor's arity and type. ||| Check a constructor's arity and type.
@ -116,7 +172,7 @@ checkCons c cons ty fc = case !(foldr checkCon (pure (Nothing, Nothing)) cons) o
checkCon (n, gdef) cons = do checkCon (n, gdef) cons = do
(zero, succ) <- cons (zero, succ) <- cons
let DCon _ arity _ = gdef.definition let DCon _ arity _ = gdef.definition
| def => throw $ GenericMsg fc $ "Expected data constructor, found:\n" ++ show def | def => throw $ GenericMsg fc $ "Expected data constructor, found:" ++ showDefType def
case arity `minus` length gdef.eraseArgs of case arity `minus` length gdef.eraseArgs of
0 => case zero of 0 => case zero of
Just _ => throw $ GenericMsg fc $ "Multiple 'Z'-like constructors for " ++ show ty ++ "." Just _ => throw $ GenericMsg fc $ "Multiple 'Z'-like constructors for " ++ show ty ++ "."
@ -129,45 +185,80 @@ checkCons c cons ty fc = case !(foldr checkCon (pure (Nothing, Nothing)) cons) o
_ => throw $ GenericMsg fc $ "Constructor " ++ show n ++ " doesn't match any pattern for Natural." _ => throw $ GenericMsg fc $ "Constructor " ++ show n ++ " doesn't match any pattern for Natural."
addBuiltinNat : addBuiltinNat :
{auto c : Ref Ctxt Defs} -> Ref Ctxt Defs =>
(ty : Name) -> NatBuiltin -> Core () (ty : Name) -> NatBuiltin -> Core ()
addBuiltinNat type cons = do addBuiltinNat type cons = do
log "builtin.Natural.addTransform" 10 $ "Add Builtin Natural transform for " ++ show type log "builtin.Natural.addTransform" 10
$ "Add %builtin Natural transform for " ++ show type ++ "."
update Ctxt $ record update Ctxt $ record
{ builtinTransforms.natTyNames $= insert type cons { builtinTransforms.natTyNames $= insert type cons
, builtinTransforms.natZNames $= insert cons.zero MkZERO , builtinTransforms.natZNames $= insert cons.zero MkZERO
, builtinTransforms.natSNames $= insert cons.succ MkSUCC , builtinTransforms.natSNames $= insert cons.succ MkSUCC
} }
addNatToInteger :
Ref Ctxt Defs =>
(fn : Name) ->
NatToInt ->
Core ()
addNatToInteger fn nToI = do
log "builtin.NaturalToInteger.addTransforms" 10
$ "Add %builtin NaturalToInteger transform for " ++ show fn ++ "."
update Ctxt $ record
{ builtinTransforms.natToIntegerFns $= insert fn nToI
}
||| Check a `%builtin Natural` pragma is correct. ||| Check a `%builtin Natural` pragma is correct.
processBuiltinNatural : processBuiltinNatural :
{auto c : Ref Ctxt Defs} -> Ref Ctxt Defs =>
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
Defs -> FC -> Name -> Core () Defs -> FC -> Name -> Core ()
processBuiltinNatural ds fc name = do processBuiltinNatural ds fc name = do
log "builtin.Natural" 5 $ "Processing Builtin Natural pragma for " ++ show name log "builtin.Natural" 5 $ "Processing %builtin Natural " ++ show name ++ "."
[(n, _, gdef)] <- lookupCtxtName name ds.gamma [(n, _, gdef)] <- lookupCtxtName name ds.gamma
| [] => throw $ UndefinedName fc name | [] => undefinedName fc name
| ns => throw $ AmbiguousName fc $ (\(n, _, _) => n) <$> ns | ns => throw $ AmbiguousName fc $ (\(n, _, _) => n) <$> ns
let TCon _ _ _ _ _ _ dcons _ = gdef.definition let TCon _ _ _ _ _ _ dcons _ = gdef.definition
| def => throw $ GenericMsg fc $ "Expected a type constructor, found:\n" ++ show def | def => throw $ GenericMsg fc
$ "Expected a type constructor, found " ++ showDefType def ++ "."
cons <- getConsGDef ds.gamma fc dcons cons <- getConsGDef ds.gamma fc dcons
cons <- checkCons ds.gamma cons n fc cons <- checkNatCons ds.gamma cons n fc
zero <- getFullName cons.zero zero <- getFullName cons.zero
succ <- getFullName cons.succ succ <- getFullName cons.succ
n <- getFullName name
addBuiltinNat n $ MkNatBuiltin {zero, succ} addBuiltinNat n $ MkNatBuiltin {zero, succ}
||| Check a `%builtin NaturalToInteger` pragma is correct.
processNatToInteger :
Ref Ctxt Defs =>
Defs -> FC -> Name -> Core ()
processNatToInteger ds fc fn = do
log "builtin.NaturalToInteger" 5 $ "Processing %builtin NaturalToInteger " ++ show fn ++ "."
[(n, _, gdef)] <- lookupCtxtName fn ds.gamma
| [] => undefinedName fc fn
| ns => throw $ AmbiguousName fc $ (\(n, _, _) => n) <$> ns
let PMDef _ args _ cases _ = gdef.definition
| def => throw $ GenericMsg fc
$ "Expected function definition, found " ++ showDefType def ++ "."
let [(_ ** arg)] = getNEArgs gdef.type
| [] => throw $ GenericMsg fc $ "No arguments found for " ++ show n ++ "."
| _ => throw $ GenericMsg fc $ "More than 1 non-erased arguments found for " ++ show n ++ "."
let Just tyCon = getTypeCons arg
| Nothing => throw $ GenericMsg fc
$ "No type constructor found for non-erased arguement of " ++ show n ++ "."
Just _ <- getNatBuiltin tyCon
| Nothing => throw $ GenericMsg fc $ "Non-erased argument is not a 'Nat'-like type."
let arity = length $ getTypeArgs gdef.type
let Just natIdx = getNEIndex arity gdef.type
| Nothing => throw $ InternalError "Couldn't find non-erased argument."
addNatToInteger n (MkNatToInt {arity, natIdx})
||| Check a `%builtin` pragma is correct. ||| Check a `%builtin` pragma is correct.
export export
processBuiltin : processBuiltin :
{auto c : Ref Ctxt Defs} -> Ref Ctxt Defs =>
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
NestedNames vars -> Env Term vars -> FC -> BuiltinType -> Name -> Core () NestedNames vars -> Env Term vars -> FC -> BuiltinType -> Name -> Core ()
processBuiltin nest env fc type name = do processBuiltin nest env fc type name = do
ds <- get Ctxt ds <- get Ctxt
case type of case type of
BuiltinNatural => processBuiltinNatural ds fc name BuiltinNatural => processBuiltinNatural ds fc name
_ => throw $ InternalError $ "%builtin " ++ show type ++ " not yet implemented." NaturalToInteger => processNatToInteger ds fc name
IntegerToNatural => throw $ InternalError "%builtin IntegerToNatural not yet implemented."

View File

@ -167,7 +167,7 @@ processTTImpFile : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} -> {auto u : Ref UST UState} ->
String -> Core Bool String -> Core Bool
processTTImpFile fname processTTImpFile fname
= do Right tti <- logTime "Parsing" $ coreLift $ parseFile fname = do Right (decor, tti) <- logTime "Parsing" $ coreLift $ parseFile fname
(do decls <- prog fname (do decls <- prog fname
eoi eoi
pure decls) pure decls)

View File

@ -464,13 +464,13 @@ checkClause {vars} mult vis totreq hashit n opts nest env (PatClause fc lhs_in r
pure (Right (MkClause env' lhstm' rhstm)) pure (Right (MkClause env' lhstm' rhstm))
-- TODO: (to decide) With is complicated. Move this into its own module? -- TODO: (to decide) With is complicated. Move this into its own module?
checkClause {vars} mult vis totreq hashit n opts nest env checkClause {vars} mult vis totreq hashit n opts nest env
(WithClause fc lhs_in wval_raw mprf flags cs) (WithClause ifc lhs_in wval_raw mprf flags cs)
= do (lhs, (vars' ** (sub', env', nest', lhspat, reqty))) <- = do (lhs, (vars' ** (sub', env', nest', lhspat, reqty))) <-
checkLHS False mult hashit n opts nest env fc lhs_in checkLHS False mult hashit n opts nest env ifc lhs_in
let wmode let wmode
= if isErased mult then InType else InExpr = if isErased mult then InType else InExpr
(wval, gwvalTy) <- wrapErrorC opts (InRHS fc !(getFullName (Resolved n))) $ (wval, gwvalTy) <- wrapErrorC opts (InRHS ifc !(getFullName (Resolved n))) $
elabTermSub n wmode opts nest' env' env sub' wval_raw Nothing elabTermSub n wmode opts nest' env' env sub' wval_raw Nothing
clearHoleLHS clearHoleLHS
@ -498,7 +498,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
-- to get the 'magic with' behaviour -- to get the 'magic with' behaviour
(wargs ** (scenv, var, binder)) <- bindWithArgs wvalTy ((,wval) <$> mprf) wvalEnv (wargs ** (scenv, var, binder)) <- bindWithArgs wvalTy ((,wval) <$> mprf) wvalEnv
let bnr = bindNotReq fc 0 env' withSub [] reqty let bnr = bindNotReq vfc 0 env' withSub [] reqty
let notreqns = fst bnr let notreqns = fst bnr
let notreqty = snd bnr let notreqty = snd bnr
@ -511,7 +511,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
(weakenNs (mkSizeOf wargs) notreqty)) (weakenNs (mkSizeOf wargs) notreqty))
let bNotReq = binder wtyScope let bNotReq = binder wtyScope
let Just (reqns, envns, wtype) = bindReq fc env' withSub [] bNotReq let Just (reqns, envns, wtype) = bindReq vfc env' withSub [] bNotReq
| Nothing => throw (InternalError "Impossible happened: With abstraction failure #4") | Nothing => throw (InternalError "Impossible happened: With abstraction failure #4")
-- list of argument names - 'Just' means we need to match the name -- list of argument names - 'Just' means we need to match the name
@ -526,11 +526,11 @@ checkClause {vars} mult vis totreq hashit n opts nest env
wname <- genWithName !(prettyName !(toFullNames (Resolved n))) wname <- genWithName !(prettyName !(toFullNames (Resolved n)))
widx <- addDef wname (record {flags $= (SetTotal totreq ::)} widx <- addDef wname (record {flags $= (SetTotal totreq ::)}
(newDef fc wname (if isErased mult then erased else top) (newDef vfc wname (if isErased mult then erased else top)
vars wtype vis None)) vars wtype vis None))
let toWarg : Maybe (PiInfo RawImp, Name) -> List (Maybe Name, RawImp) let toWarg : Maybe (PiInfo RawImp, Name) -> List (Maybe Name, RawImp)
:= flip maybe (\pn => [(Nothing, IVar fc (snd pn))]) $ := flip maybe (\pn => [(Nothing, IVar vfc (snd pn))]) $
(Nothing, wval_raw) :: (Nothing, wval_raw) ::
case mprf of case mprf of
Nothing => [] Nothing => []
@ -539,12 +539,12 @@ checkClause {vars} mult vis totreq hashit n opts nest env
let refl = IVar fc (NS builtinNS (UN "Refl")) in let refl = IVar fc (NS builtinNS (UN "Refl")) in
[(mprf, INamedApp fc refl (UN "x") wval_raw)] [(mprf, INamedApp fc refl (UN "x") wval_raw)]
let rhs_in = gapply (IVar fc wname) let rhs_in = gapply (IVar vfc wname)
$ map (\ nm => (Nothing, IVar fc nm)) envns $ map (\ nm => (Nothing, IVar vfc nm)) envns
++ concatMap toWarg wargNames ++ concatMap toWarg wargNames
log "declare.def.clause" 3 $ "Applying to with argument " ++ show rhs_in log "declare.def.clause" 3 $ "Applying to with argument " ++ show rhs_in
rhs <- wrapErrorC opts (InRHS fc !(getFullName (Resolved n))) $ rhs <- wrapErrorC opts (InRHS ifc !(getFullName (Resolved n))) $
checkTermSub n wmode opts nest' env' env sub' rhs_in checkTermSub n wmode opts nest' env' env sub' rhs_in
(gnf env' reqty) (gnf env' reqty)
@ -556,11 +556,14 @@ checkClause {vars} mult vis totreq hashit n opts nest env
nestname <- applyEnv env wname nestname <- applyEnv env wname
let nest'' = record { names $= (nestname ::) } nest let nest'' = record { names $= (nestname ::) } nest
let wdef = IDef fc wname cs' let wdef = IDef ifc wname cs'
processDecl [] nest'' env wdef processDecl [] nest'' env wdef
pure (Right (MkClause env' lhspat rhs)) pure (Right (MkClause env' lhspat rhs))
where where
vfc : FC
vfc = virtualiseFC ifc
bindWithArgs : bindWithArgs :
(wvalTy : Term xs) -> Maybe (Name, Term xs) -> (wvalTy : Term xs) -> Maybe (Name, Term xs) ->
(wvalEnv : Env Term xs) -> (wvalEnv : Env Term xs) ->
@ -576,13 +579,13 @@ checkClause {vars} mult vis totreq hashit n opts nest env
wargs = [wargn] wargs = [wargn]
scenv : Env Term (wargs ++ xs) scenv : Env Term (wargs ++ xs)
:= Pi fc top Explicit wvalTy :: wvalEnv := Pi vfc top Explicit wvalTy :: wvalEnv
var : Term (wargs ++ xs) var : Term (wargs ++ xs)
:= Local fc (Just False) Z First := Local vfc (Just False) Z First
binder : Term (wargs ++ xs) -> Term xs binder : Term (wargs ++ xs) -> Term xs
:= Bind fc wargn (Pi fc top Explicit wvalTy) := Bind vfc wargn (Pi vfc top Explicit wvalTy)
in pure (wargs ** (scenv, var, binder)) in pure (wargs ** (scenv, var, binder))
@ -592,7 +595,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env
let eqName = NS builtinNS (UN "Equal") let eqName = NS builtinNS (UN "Equal")
Just (TCon t ar _ _ _ _ _ _) <- lookupDefExact eqName (gamma defs) Just (TCon t ar _ _ _ _ _ _) <- lookupDefExact eqName (gamma defs)
| _ => throw (InternalError "Cannot find builtin Equal") | _ => throw (InternalError "Cannot find builtin Equal")
let eqTyCon = Ref fc (TyCon t ar) eqName let eqTyCon = Ref vfc (TyCon t ar) eqName
let wargn : Name let wargn : Name
wargn = MN "warg" 0 wargn = MN "warg" 0
@ -601,24 +604,24 @@ checkClause {vars} mult vis totreq hashit n opts nest env
wvalTy' := weaken wvalTy wvalTy' := weaken wvalTy
eqTy : Term (MN "warg" 0 :: xs) eqTy : Term (MN "warg" 0 :: xs)
:= apply fc eqTyCon := apply vfc eqTyCon
[ wvalTy' [ wvalTy'
, wvalTy' , wvalTy'
, weaken wval , weaken wval
, Local fc (Just False) Z First , Local vfc (Just False) Z First
] ]
scenv : Env Term (wargs ++ xs) scenv : Env Term (wargs ++ xs)
:= Pi fc top Implicit eqTy := Pi vfc top Implicit eqTy
:: Pi fc top Explicit wvalTy :: Pi vfc top Explicit wvalTy
:: wvalEnv :: wvalEnv
var : Term (wargs ++ xs) var : Term (wargs ++ xs)
:= Local fc (Just False) (S Z) (Later First) := Local vfc (Just False) (S Z) (Later First)
binder : Term (wargs ++ xs) -> Term xs binder : Term (wargs ++ xs) -> Term xs
:= \ t => Bind fc wargn (Pi fc top Explicit wvalTy) := \ t => Bind vfc wargn (Pi vfc top Explicit wvalTy)
$ Bind fc name (Pi fc top Implicit eqTy) t $ Bind vfc name (Pi vfc top Implicit eqTy) t
pure (wargs ** (scenv, var, binder)) pure (wargs ** (scenv, var, binder))

View File

@ -77,7 +77,7 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields
farg : IField -> farg : IField ->
(FC, Maybe Name, RigCount, PiInfo RawImp, RawImp) (FC, Maybe Name, RigCount, PiInfo RawImp, RawImp)
farg (MkIField fc c p n ty) = (fc, Just n, c, p, ty) farg (MkIField fc c p n ty) = (virtualiseFC fc, Just n, c, p, ty)
mkTy : List (FC, Maybe Name, RigCount, PiInfo RawImp, RawImp) -> mkTy : List (FC, Maybe Name, RigCount, PiInfo RawImp, RawImp) ->
RawImp -> RawImp RawImp -> RawImp
@ -86,7 +86,7 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields
= IPi fc c imp n argty (mkTy args ret) = IPi fc c imp n argty (mkTy args ret)
recTy : RawImp recTy : RawImp
recTy = apply (IVar fc tn) (map (\(n, c, p, tm) => (n, IVar EmptyFC n, p)) params) recTy = apply (IVar (virtualiseFC fc) tn) (map (\(n, c, p, tm) => (n, IVar EmptyFC n, p)) params)
where where
||| Apply argument to list of explicit or implicit named arguments ||| Apply argument to list of explicit or implicit named arguments
apply : RawImp -> List (Name, RawImp, PiInfo RawImp) -> RawImp apply : RawImp -> List (Name, RawImp, PiInfo RawImp) -> RawImp
@ -96,7 +96,8 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields
elabAsData : Name -> Core () elabAsData : Name -> Core ()
elabAsData cname elabAsData cname
= do let conty = mkTy paramTelescope $ = do let fc = virtualiseFC fc
let conty = mkTy paramTelescope $
mkTy (map farg fields) recTy mkTy (map farg fields) recTy
let con = MkImpTy EmptyFC EmptyFC cname !(bindTypeNames [] (map fst params ++ let con = MkImpTy EmptyFC EmptyFC cname !(bindTypeNames [] (map fst params ++
map fname fields ++ vars) conty) map fname fields ++ vars) conty)
@ -125,7 +126,9 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields
Env Term vs -> Term vs -> Env Term vs -> Term vs ->
Core () Core ()
elabGetters con done upds tyenv (Bind bfc n b@(Pi _ rc imp ty_chk) sc) elabGetters con done upds tyenv (Bind bfc n b@(Pi _ rc imp ty_chk) sc)
= if (n `elem` map fst params) || (n `elem` vars) = let rig = if isErased rc then erased else top
isVis = projVis vis
in if (n `elem` map fst params) || (n `elem` vars)
then elabGetters con then elabGetters con
(if imp == Explicit && not (n `elem` vars) (if imp == Explicit && not (n `elem` vars)
then S done else done) then S done else done)
@ -145,11 +148,14 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields
(map fst params ++ map fname fields ++ vars) $ (map fst params ++ map fname fields ++ vars) $
mkTy paramTelescope $ mkTy paramTelescope $
IPi bfc top Explicit (Just rname) recTy ty' IPi bfc top Explicit (Just rname) recTy ty'
log "declare.record.projection" 5 $ "Projection " ++ show rfNameNS ++ " : " ++ show projTy
processDecl [] nest env let mkProjClaim = \ nm =>
(IClaim bfc (if isErased rc let ty = MkImpTy EmptyFC EmptyFC nm projTy
then erased in IClaim bfc rig isVis [Inline] ty
else top) (projVis vis) [Inline] (MkImpTy EmptyFC EmptyFC rfNameNS projTy))
log "declare.record.projection" 5 $
"Projection " ++ show rfNameNS ++ " : " ++ show projTy
processDecl [] nest env (mkProjClaim rfNameNS)
-- Define the LHS and RHS -- Define the LHS and RHS
let lhs_exp let lhs_exp
@ -173,16 +179,15 @@ elabRecord {vars} eopts fc env nest newns vis tn params conName_in fields
when !isPrefixRecordProjections $ do -- beware: `!` is NOT boolean `not`! when !isPrefixRecordProjections $ do -- beware: `!` is NOT boolean `not`!
-- Claim the type. -- Claim the type.
-- we just reuse `projTy` defined above -- we just reuse `projTy` defined above
log "declare.record.projection.prefix" 5 $ "Prefix projection " ++ show unNameNS ++ " : " ++ show projTy log "declare.record.projection.prefix" 5 $
processDecl [] nest env "Prefix projection " ++ show unNameNS ++ " : " ++ show projTy
(IClaim bfc (if isErased rc processDecl [] nest env (mkProjClaim unNameNS)
then erased
else top) (projVis vis) [Inline] (MkImpTy EmptyFC EmptyFC unNameNS projTy))
-- Define the LHS and RHS -- Define the LHS and RHS
let lhs = IVar bfc unNameNS let lhs = IVar bfc unNameNS
let rhs = IVar bfc rfNameNS let rhs = IVar bfc rfNameNS
log "declare.record.projection.prefix" 5 $ "Prefix projection " ++ show lhs ++ " = " ++ show rhs log "declare.record.projection.prefix" 5 $
"Prefix projection " ++ show lhs ++ " = " ++ show rhs
processDecl [] nest env processDecl [] nest env
(IDef bfc unNameNS [PatClause bfc lhs rhs]) (IDef bfc unNameNS [PatClause bfc lhs rhs])

View File

@ -264,6 +264,8 @@ processType : {vars : _} ->
processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc nameFC n_in ty_raw) processType {vars} eopts nest env fc rig vis opts (MkImpTy tfc nameFC n_in ty_raw)
= do n <- inCurrentNS n_in = do n <- inCurrentNS n_in
addNameLoc nameFC n
log "declare.type" 1 $ "Processing " ++ show n log "declare.type" 1 $ "Processing " ++ show n
log "declare.type" 5 $ "Checking type decl " ++ show n ++ " : " ++ show ty_raw log "declare.type" 5 $ "Checking type decl " ++ show n ++ " : " ++ show ty_raw
idx <- resolveName n idx <- resolveName n

View File

@ -542,7 +542,7 @@ implicitsAs n defs ns tm
"\n In the type of " ++ show n ++ ": " ++ show ty ++ "\n In the type of " ++ show n ++ ": " ++ show ty ++
"\n Using locals: " ++ show ns ++ "\n Using locals: " ++ show ns ++
"\n Found implicits: " ++ show implicits "\n Found implicits: " ++ show implicits
pure $ impAs loc implicits (IVar loc nm) pure $ impAs (virtualiseFC loc) implicits (IVar loc nm)
where where
-- If there's an @{c} in the list of given implicits, that's the next -- If there's an @{c} in the list of given implicits, that's the next
-- autoimplicit, so don't rewrite the LHS and update the list of given -- autoimplicit, so don't rewrite the LHS and update the list of given
@ -715,10 +715,28 @@ getFC (IAs x _ _ _ _) = x
getFC (Implicit x _) = x getFC (Implicit x _) = x
getFC (IWithUnambigNames x _ _) = x getFC (IWithUnambigNames x _ _) = x
namespace ImpDecl
public export
getFC : ImpDecl -> FC
getFC (IClaim fc _ _ _ _) = fc
getFC (IData fc _ _) = fc
getFC (IDef fc _ _) = fc
getFC (IParameters fc _ _) = fc
getFC (IRecord fc _ _ _ ) = fc
getFC (INamespace fc _ _) = fc
getFC (ITransform fc _ _ _) = fc
getFC (IRunElabDecl fc _) = fc
getFC (IPragma _ _) = EmptyFC
getFC (ILog _) = EmptyFC
getFC (IBuiltin fc _ _) = fc
export export
apply : RawImp -> List RawImp -> RawImp apply : RawImp -> List RawImp -> RawImp
apply f [] = f apply f [] = f
apply f (x :: xs) = apply (IApp (getFC f) f x) xs apply f (x :: xs) =
let fFC = getFC f in
apply (IApp (fromMaybe fFC (mergeFC fFC (getFC x))) f x) xs
export export
gapply : RawImp -> List (Maybe Name, RawImp) -> RawImp gapply : RawImp -> List (Maybe Name, RawImp) -> RawImp
@ -745,18 +763,13 @@ getFn f = f
export export
TTC BuiltinType where TTC BuiltinType where
toBuf b BuiltinNatural = tag 0 toBuf b BuiltinNatural = tag 0
toBuf b NaturalPlus = tag 1 toBuf b NaturalToInteger = tag 1
toBuf b NaturalMult = tag 2 toBuf b IntegerToNatural = tag 2
toBuf b NaturalToInteger = tag 3
toBuf b IntegerToNatural = tag 4
fromBuf b = case !getTag of fromBuf b = case !getTag of
0 => pure BuiltinNatural 0 => pure BuiltinNatural
1 => pure NaturalPlus 1 => pure NaturalToInteger
2 => pure NaturalMult 2 => pure IntegerToNatural
3 => pure NaturalToInteger _ => corrupt "BuiltinType"
4 => pure IntegerToNatural
_ => corrupt "BuiltinType"
mutual mutual
export export

View File

@ -2,7 +2,9 @@ module TTImp.WithClause
import Core.Context import Core.Context
import Core.Context.Log import Core.Context.Log
import Core.Metadata
import Core.TT import Core.TT
import TTImp.BindImplicits import TTImp.BindImplicits
import TTImp.TTImp import TTImp.TTImp
import TTImp.Elab.Check import TTImp.Elab.Check
@ -15,30 +17,52 @@ import Data.Maybe
matchFail : FC -> Core a matchFail : FC -> Core a
matchFail loc = throw (GenericMsg loc "With clause does not match parent") matchFail loc = throw (GenericMsg loc "With clause does not match parent")
--- To be used on the lhs of a nested with clause to figure out a tight location
--- information to give to the generated LHS
getHeadLoc : RawImp -> Core FC
getHeadLoc (IVar fc _) = pure fc
getHeadLoc (IApp _ f _) = getHeadLoc f
getHeadLoc (IAutoApp _ f _) = getHeadLoc f
getHeadLoc (INamedApp _ f _ _) = getHeadLoc f
getHeadLoc t = throw (InternalError $ "Could not find head of LHS: " ++ show t)
addAlias : {auto m : Ref MD Metadata} ->
{auto c : Ref Ctxt Defs} ->
FC -> FC -> Core ()
addAlias from to =
whenJust (isConcreteFC from) $ \ from =>
whenJust (isConcreteFC to) $ \ to => do
log "ide-mode.highlighting.alias" 25 $
"Adding alias: " ++ show from ++ " -> " ++ show to
addSemanticAlias from to
mutual mutual
export export
getMatch : (lhs : Bool) -> RawImp -> RawImp -> getMatch : {auto m : Ref MD Metadata} ->
{auto c : Ref Ctxt Defs} ->
(lhs : Bool) -> RawImp -> RawImp ->
Core (List (String, RawImp)) Core (List (String, RawImp))
getMatch lhs (IBindVar to n) tm@(IBindVar from _)
= [(n, tm)] <$ addAlias from to
getMatch lhs (IBindVar _ n) tm = pure [(n, tm)] getMatch lhs (IBindVar _ n) tm = pure [(n, tm)]
getMatch lhs (Implicit _ _) tm = pure [] getMatch lhs (Implicit _ _) tm = pure []
getMatch lhs (IVar _ (NS ns n)) (IVar loc (NS ns' n')) getMatch lhs (IVar to (NS ns n)) (IVar from (NS ns' n'))
= if n == n' && isParentOf ns' ns then pure [] else matchFail loc = if n == n' && isParentOf ns' ns
getMatch lhs (IVar _ (NS ns n)) (IVar loc n') then [] <$ addAlias from to -- <$ decorateName loc nm
= if n == n' then pure [] else matchFail loc else matchFail from
getMatch lhs (IVar _ n) (IVar loc n') getMatch lhs (IVar to (NS ns n)) (IVar from n')
= if n == n' then pure [] else matchFail loc = if n == n'
then [] <$ addAlias from to -- <$ decorateName loc (NS ns n')
else matchFail from
getMatch lhs (IVar to n) (IVar from n')
= if n == n'
then [] <$ addAlias from to -- <$ decorateName loc n'
else matchFail from
getMatch lhs (IPi _ c p n arg ret) (IPi loc c' p' n' arg' ret') getMatch lhs (IPi _ c p n arg ret) (IPi loc c' p' n' arg' ret')
= if c == c' && samePiInfo p p' && n == n' = if c == c' && eqPiInfoBy (\_, _ => True) p p' && n == n'
then matchAll lhs [(arg, arg'), (ret, ret')] then matchAll lhs [(arg, arg'), (ret, ret')]
else matchFail loc else matchFail loc
where
samePiInfo : PiInfo RawImp -> PiInfo RawImp -> Bool
samePiInfo Explicit Explicit = True
samePiInfo Implicit Implicit = True
samePiInfo AutoImplicit AutoImplicit = True
samePiInfo (DefImplicit _) (DefImplicit _) = True
samePiInfo _ _ = False
-- TODO: Lam, Let, Case, Local, Update -- TODO: Lam, Let, Case, Local, Update
getMatch lhs (IApp _ f a) (IApp loc f' a') getMatch lhs (IApp _ f a) (IApp loc f' a')
= matchAll lhs [(f, f'), (a, a')] = matchAll lhs [(f, f'), (a, a')]
@ -72,7 +96,7 @@ mutual
-- one of them is okay -- one of them is okay
getMatch lhs (IAlternative fc _ as) (IAlternative _ _ as') getMatch lhs (IAlternative fc _ as) (IAlternative _ _ as')
= matchAny fc lhs (zip as as') = matchAny fc lhs (zip as as')
getMatch lhs (IAs _ _ _ (UN n) p) (IAs fc _ _ (UN n') p') getMatch lhs (IAs _ _ _ (UN n) p) (IAs _ fc _ (UN n') p')
= do ms <- getMatch lhs p p' = do ms <- getMatch lhs p p'
mergeMatches lhs ((n, IBindVar fc n') :: ms) mergeMatches lhs ((n, IBindVar fc n') :: ms)
getMatch lhs (IAs _ _ _ (UN n) p) p' getMatch lhs (IAs _ _ _ (UN n) p) p'
@ -87,14 +111,18 @@ mutual
else matchFail fc else matchFail fc
getMatch lhs pat spec = matchFail (getFC pat) getMatch lhs pat spec = matchFail (getFC pat)
matchAny : FC -> (lhs : Bool) -> List (RawImp, RawImp) -> matchAny : {auto m : Ref MD Metadata} ->
{auto c : Ref Ctxt Defs} ->
FC -> (lhs : Bool) -> List (RawImp, RawImp) ->
Core (List (String, RawImp)) Core (List (String, RawImp))
matchAny fc lhs [] = matchFail fc matchAny fc lhs [] = matchFail fc
matchAny fc lhs ((x, y) :: ms) matchAny fc lhs ((x, y) :: ms)
= catch (getMatch lhs x y) = catch (getMatch lhs x y)
(\err => matchAny fc lhs ms) (\err => matchAny fc lhs ms)
matchAll : (lhs : Bool) -> List (RawImp, RawImp) -> matchAll : {auto m : Ref MD Metadata} ->
{auto c : Ref Ctxt Defs} ->
(lhs : Bool) -> List (RawImp, RawImp) ->
Core (List (String, RawImp)) Core (List (String, RawImp))
matchAll lhs [] = pure [] matchAll lhs [] = pure []
matchAll lhs ((x, y) :: ms) matchAll lhs ((x, y) :: ms)
@ -102,7 +130,9 @@ mutual
mxy <- getMatch lhs x y mxy <- getMatch lhs x y
mergeMatches lhs (mxy ++ matches) mergeMatches lhs (mxy ++ matches)
mergeMatches : (lhs : Bool) -> List (String, RawImp) -> mergeMatches : {auto m : Ref MD Metadata} ->
{auto c : Ref Ctxt Defs} ->
(lhs : Bool) -> List (String, RawImp) ->
Core (List (String, RawImp)) Core (List (String, RawImp))
mergeMatches lhs [] = pure [] mergeMatches lhs [] = pure []
mergeMatches lhs ((n, tm) :: rest) mergeMatches lhs ((n, tm) :: rest)
@ -110,8 +140,9 @@ mutual
case lookup n rest' of case lookup n rest' of
Nothing => pure ((n, tm) :: rest') Nothing => pure ((n, tm) :: rest')
Just tm' => Just tm' =>
do ignore $ getMatch lhs tm tm' -- just need to know it succeeds do ignore $ getMatch lhs tm tm'
mergeMatches lhs rest -- ^ just need to know it succeeds
pure rest'
-- Get the arguments for the rewritten pattern clause of a with by looking -- Get the arguments for the rewritten pattern clause of a with by looking
-- up how the argument names matched -- up how the argument names matched
@ -138,11 +169,13 @@ getArgMatch ploc mode search warg ms (Just (_, nm))
export export
getNewLHS : {auto c : Ref Ctxt Defs} -> getNewLHS : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
FC -> (drop : Nat) -> NestedNames vars -> FC -> (drop : Nat) -> NestedNames vars ->
Name -> List (Maybe (PiInfo RawImp, Name)) -> Name -> List (Maybe (PiInfo RawImp, Name)) ->
RawImp -> RawImp -> Core RawImp RawImp -> RawImp -> Core RawImp
getNewLHS ploc drop nest wname wargnames lhs_raw patlhs getNewLHS iploc drop nest wname wargnames lhs_raw patlhs
= do (mlhs_raw, wrest) <- dropWithArgs drop patlhs = do let vploc = virtualiseFC iploc
(mlhs_raw, wrest) <- dropWithArgs drop patlhs
autoimp <- isUnboundImplicits autoimp <- isUnboundImplicits
setUnboundImplicits True setUnboundImplicits True
@ -154,15 +187,16 @@ getNewLHS ploc drop nest wname wargnames lhs_raw patlhs
log "declare.def.clause.with" 20 $ "Modified LHS (with implicits): " ++ show mlhs log "declare.def.clause.with" 20 $ "Modified LHS (with implicits): " ++ show mlhs
let (warg :: rest) = reverse wrest let (warg :: rest) = reverse wrest
| _ => throw (GenericMsg ploc "Badly formed 'with' clause") | _ => throw (GenericMsg iploc "Badly formed 'with' clause")
log "declare.def.clause.with" 5 $ show lhs ++ " against " ++ show mlhs ++ log "declare.def.clause.with" 5 $ show lhs ++ " against " ++ show mlhs ++
" dropping " ++ show (warg :: rest) " dropping " ++ show (warg :: rest)
ms <- getMatch True lhs mlhs ms <- getMatch True lhs mlhs
log "declare.def.clause.with" 5 $ "Matches: " ++ show ms log "declare.def.clause.with" 5 $ "Matches: " ++ show ms
let params = map (getArgMatch ploc (InLHS top) False warg ms) wargnames let params = map (getArgMatch vploc (InLHS top) False warg ms) wargnames
log "declare.def.clause.with" 5 $ "Parameters: " ++ show params log "declare.def.clause.with" 5 $ "Parameters: " ++ show params
let newlhs = apply (IVar ploc wname) (params ++ rest) hdloc <- getHeadLoc patlhs
let newlhs = apply (IVar hdloc wname) (params ++ rest)
log "declare.def.clause.with" 5 $ "New LHS: " ++ show newlhs log "declare.def.clause.with" 5 $ "New LHS: " ++ show newlhs
pure newlhs pure newlhs
where where
@ -174,11 +208,12 @@ getNewLHS ploc drop nest wname wargnames lhs_raw patlhs
pure (tm, arg :: rest) pure (tm, arg :: rest)
-- Shouldn't happen if parsed correctly, but there's no guarantee that -- Shouldn't happen if parsed correctly, but there's no guarantee that
-- inputs come from parsed source so throw an error. -- inputs come from parsed source so throw an error.
dropWithArgs _ _ = throw (GenericMsg ploc "Badly formed 'with' clause") dropWithArgs _ _ = throw (GenericMsg iploc "Badly formed 'with' clause")
-- Find a 'with' application on the RHS and update it -- Find a 'with' application on the RHS and update it
export export
withRHS : {auto c : Ref Ctxt Defs} -> withRHS : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
FC -> (drop : Nat) -> Name -> List (Maybe (PiInfo RawImp, Name)) -> FC -> (drop : Nat) -> Name -> List (Maybe (PiInfo RawImp, Name)) ->
RawImp -> RawImp -> RawImp -> RawImp ->
Core RawImp Core RawImp
@ -196,8 +231,9 @@ withRHS fc drop wname wargnames tm toplhs
updateWith fc tm (arg :: args) updateWith fc tm (arg :: args)
= do log "declare.def.clause.with" 10 $ "With-app: Matching " ++ show toplhs ++ " against " ++ show tm = do log "declare.def.clause.with" 10 $ "With-app: Matching " ++ show toplhs ++ " against " ++ show tm
ms <- getMatch False toplhs tm ms <- getMatch False toplhs tm
hdloc <- getHeadLoc tm
log "declare.def.clause.with" 10 $ "Result: " ++ show ms log "declare.def.clause.with" 10 $ "Result: " ++ show ms
let newrhs = apply (IVar fc wname) let newrhs = apply (IVar hdloc wname)
(map (getArgMatch fc InExpr True arg ms) wargnames) (map (getArgMatch fc InExpr True arg ms) wargnames)
log "declare.def.clause.with" 10 $ "With args for RHS: " ++ show wargnames log "declare.def.clause.with" 10 $ "With args for RHS: " ++ show wargnames
log "declare.def.clause.with" 10 $ "New RHS: " ++ show newrhs log "declare.def.clause.with" 10 $ "New RHS: " ++ show newrhs

View File

@ -47,7 +47,7 @@ yaffleMain : String -> List String -> Core ()
yaffleMain fname args yaffleMain fname args
= do defs <- initDefs = do defs <- initDefs
c <- newRef Ctxt defs c <- newRef Ctxt defs
m <- newRef MD initMetadata m <- newRef MD (initMetadata fname)
u <- newRef UST initUState u <- newRef UST initUState
d <- getDirs d <- getDirs
t <- processArgs args t <- processArgs args

View File

@ -153,4 +153,4 @@ repl
case runParser "(interactive)" Nothing inp command of case runParser "(interactive)" Nothing inp command of
Left err => do coreLift_ (printLn err) Left err => do coreLift_ (printLn err)
repl repl
Right cmd => when !(processCatch cmd) repl Right (decor, cmd) => when !(processCatch cmd) repl

1
support/chez/.gitignore vendored Normal file
View File

@ -0,0 +1 @@
support-sep.ss

30
support/chez/Makefile Normal file
View File

@ -0,0 +1,30 @@
include ../../config.mk
all: build
clean:
-$(RM) support-sep.ss
.PHONY: install build
build: support-sep.ss
install: build
mkdir -p ${PREFIX}/idris2-${IDRIS2_VERSION}/support/chez
install *.ss ${PREFIX}/idris2-${IDRIS2_VERSION}/support/chez
support-sep.ss: support.ss
# start library header
echo "(library (support) (export" > $@
# print the list of exports
cat support.ss \
| sed -n 's|(define (\?\([^ )]*\).*|\1|p' \
>> $@
echo ") (import (chezscheme))" >> $@
# copy the code
cat $< >> $@
# close the bracket
echo ") ; end of (library)" >> $@

View File

@ -19,7 +19,7 @@ import Test.Golden
-- Test cases -- Test cases
ttimpTests : TestPool ttimpTests : TestPool
ttimpTests = MkTestPool [] ttimpTests = MkTestPool "TTImp" []
[ "basic001", "basic002", "basic003", "basic004", "basic005" [ "basic001", "basic002", "basic003", "basic004", "basic005"
, "basic006" , "basic006"
, "coverage001", "coverage002" , "coverage001", "coverage002"
@ -34,7 +34,7 @@ ttimpTests = MkTestPool []
] ]
idrisTestsBasic : TestPool idrisTestsBasic : TestPool
idrisTestsBasic = MkTestPool [] idrisTestsBasic = MkTestPool "Fundamental language features" []
-- Fundamental language features -- Fundamental language features
["basic001", "basic002", "basic003", "basic004", "basic005", ["basic001", "basic002", "basic003", "basic004", "basic005",
"basic006", "basic007", "basic008", "basic009", "basic010", "basic006", "basic007", "basic008", "basic009", "basic010",
@ -50,7 +50,7 @@ idrisTestsBasic = MkTestPool []
"basic056", "basic057", "basic058", "basic059"] "basic056", "basic057", "basic058", "basic059"]
idrisTestsCoverage : TestPool idrisTestsCoverage : TestPool
idrisTestsCoverage = MkTestPool [] idrisTestsCoverage = MkTestPool "Coverage checking" []
-- Coverage checking -- Coverage checking
["coverage001", "coverage002", "coverage003", "coverage004", ["coverage001", "coverage002", "coverage003", "coverage004",
"coverage005", "coverage006", "coverage007", "coverage008", "coverage005", "coverage006", "coverage007", "coverage008",
@ -58,12 +58,12 @@ idrisTestsCoverage = MkTestPool []
"coverage013", "coverage014", "coverage015", "coverage016"] "coverage013", "coverage014", "coverage015", "coverage016"]
idrisTestsCasetree : TestPool idrisTestsCasetree : TestPool
idrisTestsCasetree = MkTestPool [] idrisTestsCasetree = MkTestPool "Case tree building" []
-- Case tree building -- Case tree building
["casetree001"] ["casetree001"]
idrisTestsError : TestPool idrisTestsError : TestPool
idrisTestsError = MkTestPool [] idrisTestsError = MkTestPool "Error messages" []
-- Error messages -- Error messages
["error001", "error002", "error003", "error004", "error005", ["error001", "error002", "error003", "error004", "error005",
"error006", "error007", "error008", "error009", "error010", "error006", "error007", "error008", "error009", "error010",
@ -74,7 +74,7 @@ idrisTestsError = MkTestPool []
"perror006", "perror007", "perror008"] "perror006", "perror007", "perror008"]
idrisTestsInteractive : TestPool idrisTestsInteractive : TestPool
idrisTestsInteractive = MkTestPool [] idrisTestsInteractive = MkTestPool "Interactive editing" []
-- Interactive editing support -- Interactive editing support
["interactive001", "interactive002", "interactive003", "interactive004", ["interactive001", "interactive002", "interactive003", "interactive004",
"interactive005", "interactive006", "interactive007", "interactive008", "interactive005", "interactive006", "interactive007", "interactive008",
@ -86,7 +86,7 @@ idrisTestsInteractive = MkTestPool []
"interactive029", "interactive030"] "interactive029", "interactive030"]
idrisTestsInterface : TestPool idrisTestsInterface : TestPool
idrisTestsInterface = MkTestPool [] idrisTestsInterface = MkTestPool "Interface" []
-- Interfaces -- Interfaces
["interface001", "interface002", "interface003", "interface004", ["interface001", "interface002", "interface003", "interface004",
"interface005", "interface006", "interface007", "interface008", "interface005", "interface006", "interface007", "interface008",
@ -97,7 +97,7 @@ idrisTestsInterface = MkTestPool []
"interface025"] "interface025"]
idrisTestsLinear : TestPool idrisTestsLinear : TestPool
idrisTestsLinear = MkTestPool [] idrisTestsLinear = MkTestPool "Quantities" []
-- QTT and linearity related -- QTT and linearity related
["linear001", "linear002", "linear003", -- "linear004" -- disabled due to requiring linearity subtyping ["linear001", "linear002", "linear003", -- "linear004" -- disabled due to requiring linearity subtyping
"linear005", "linear006", "linear007", "linear008", "linear005", "linear006", "linear007", "linear008",
@ -105,7 +105,7 @@ idrisTestsLinear = MkTestPool []
"linear013"] "linear013"]
idrisTestsLiterate : TestPool idrisTestsLiterate : TestPool
idrisTestsLiterate = MkTestPool [] idrisTestsLiterate = MkTestPool "Literate programming" []
-- Literate -- Literate
["literate001", "literate002", "literate003", "literate004", ["literate001", "literate002", "literate003", "literate004",
"literate005", "literate006", "literate007", "literate008", "literate005", "literate006", "literate007", "literate008",
@ -113,13 +113,13 @@ idrisTestsLiterate = MkTestPool []
"literate013", "literate014", "literate015", "literate016"] "literate013", "literate014", "literate015", "literate016"]
idrisTestsPerformance : TestPool idrisTestsPerformance : TestPool
idrisTestsPerformance = MkTestPool [] idrisTestsPerformance = MkTestPool "Performance" []
-- Performance: things which have been slow in the past, or which -- Performance: things which have been slow in the past, or which
-- pose interesting challenges for the elaborator -- pose interesting challenges for the elaborator
["perf001", "perf002", "perf003", "perf004", "perf005", "perf006"] ["perf001", "perf002", "perf003", "perf004", "perf005", "perf006"]
idrisTestsRegression : TestPool idrisTestsRegression : TestPool
idrisTestsRegression = MkTestPool [] idrisTestsRegression = MkTestPool "Various regressions" []
-- Miscellaneous regressions -- Miscellaneous regressions
["reg001", "reg002", "reg003", "reg004", "reg005", "reg006", "reg007", ["reg001", "reg002", "reg003", "reg004", "reg005", "reg006", "reg007",
"reg008", "reg009", "reg010", "reg011", "reg012", "reg013", "reg014", "reg008", "reg009", "reg010", "reg011", "reg012", "reg013", "reg014",
@ -129,7 +129,7 @@ idrisTestsRegression = MkTestPool []
"reg036", "reg037", "reg038", "reg039"] "reg036", "reg037", "reg038", "reg039"]
idrisTestsData : TestPool idrisTestsData : TestPool
idrisTestsData = MkTestPool [] idrisTestsData = MkTestPool "Data and record types" []
[-- Data types [-- Data types
"data001", "data001",
-- Records, access and dependent update -- Records, access and dependent update
@ -137,25 +137,28 @@ idrisTestsData = MkTestPool []
"record006", "record007"] "record006", "record007"]
idrisTestsBuiltin : TestPool idrisTestsBuiltin : TestPool
idrisTestsBuiltin = MkTestPool [] idrisTestsBuiltin = MkTestPool "Builtin types and functions" []
-- %builtin related tests for the frontend (type-checking) -- %builtin related tests for the frontend (type-checking)
["builtin001", "builtin002", "builtin003", "builtin004"] ["builtin001", "builtin002", "builtin003", "builtin004", "builtin005",
"builtin006", "builtin007", "builtin008", "builtin009"]
idrisTestsEvaluator : TestPool idrisTestsEvaluator : TestPool
idrisTestsEvaluator = MkTestPool [] idrisTestsEvaluator = MkTestPool "Evaluation" []
[ -- Evaluator [ -- Evaluator
"evaluator001", "evaluator002", "evaluator003", "evaluator004", "evaluator001", "evaluator002", "evaluator003", "evaluator004",
-- Unfortunately the behaviour of Double is platform dependent so the
-- following test is turned off.
-- "evaluator005",
-- Miscellaneous REPL -- Miscellaneous REPL
"interpreter001", "interpreter002", "interpreter003", "interpreter004", "interpreter001", "interpreter002", "interpreter003", "interpreter004",
"interpreter005", "interpreter006", "interpreter007"] "interpreter005", "interpreter006", "interpreter007"]
idrisTests : TestPool idrisTests : TestPool
idrisTests = MkTestPool [] idrisTests = MkTestPool "Misc" []
-- Documentation strings -- Documentation strings
["docs001", "docs002", ["docs001", "docs002",
-- Unfortunately the behaviour of Double is platform dependent so the -- Eta equality
-- following test is turned off. "eta001",
-- "evaluator005",
-- Modules and imports -- Modules and imports
"import001", "import002", "import003", "import004", "import005", "import001", "import002", "import003", "import004", "import005",
-- Implicit laziness, lazy evaluation -- Implicit laziness, lazy evaluation
@ -187,14 +190,14 @@ idrisTests = MkTestPool []
"pretty001"] "pretty001"]
typeddTests : TestPool typeddTests : TestPool
typeddTests = MkTestPool [] typeddTests = MkTestPool "Type Driven Development" []
[ "chapter01", "chapter02", "chapter03", "chapter04", "chapter05" [ "chapter01", "chapter02", "chapter03", "chapter04", "chapter05"
, "chapter06", "chapter07", "chapter08", "chapter09", "chapter10" , "chapter06", "chapter07", "chapter08", "chapter09", "chapter10"
, "chapter11", "chapter12", "chapter13", "chapter14" , "chapter11", "chapter12", "chapter13", "chapter14"
] ]
chezTests : TestPool chezTests : TestPool
chezTests = MkTestPool [Chez] chezTests = MkTestPool "Chez backend" [Chez]
[ "chez001", "chez002", "chez003", "chez004", "chez005", "chez006" [ "chez001", "chez002", "chez003", "chez004", "chez005", "chez006"
, "chez007", "chez008", "chez009", "chez010", "chez011", "chez012" , "chez007", "chez008", "chez009", "chez010", "chez011", "chez012"
, "chez013", "chez014", "chez015", "chez016", "chez017", "chez018" , "chez013", "chez014", "chez015", "chez016", "chez017", "chez018"
@ -202,6 +205,7 @@ chezTests = MkTestPool [Chez]
, "chez025", "chez026", "chez027", "chez028", "chez029", "chez030" , "chez025", "chez026", "chez027", "chez028", "chez029", "chez030"
, "chez031", "chez032" , "chez031", "chez032"
, "futures001" , "futures001"
, "bitops"
, "casts" , "casts"
, "newints" , "newints"
, "semaphores001" , "semaphores001"
@ -211,11 +215,11 @@ chezTests = MkTestPool [Chez]
] ]
refcTests : TestPool refcTests : TestPool
refcTests = MkTestPool [C] refcTests = MkTestPool "Reference counting C backend" [C]
[ "refc001" , "refc002" ] [ "refc001" , "refc002" ]
racketTests : TestPool racketTests : TestPool
racketTests = MkTestPool [Racket] racketTests = MkTestPool "Racket backend" [Racket]
[ "forkjoin001" [ "forkjoin001"
, "semaphores001", "semaphores002" , "semaphores001", "semaphores002"
, "futures001" , "futures001"
@ -230,12 +234,13 @@ racketTests = MkTestPool [Racket]
] ]
nodeTests : TestPool nodeTests : TestPool
nodeTests = MkTestPool [Node] nodeTests = MkTestPool "Node backend" [Node]
[ "node001", "node002", "node003", "node004", "node005", "node006" [ "node001", "node002", "node003", "node004", "node005", "node006"
, "node007", "node008", "node009", "node011", "node012", "node015" , "node007", "node008", "node009", "node011", "node012", "node015"
, "node017", "node018", "node019", "node021", "node022", "node023" , "node017", "node018", "node019", "node021", "node022", "node023"
, "node024", "node025" , "node024", "node025"
-- , "node14", "node020" -- , "node14", "node020"
, "bitops"
, "casts" , "casts"
, "newints" , "newints"
, "reg001" , "reg001"
@ -245,17 +250,17 @@ nodeTests = MkTestPool [Node]
] ]
ideModeTests : TestPool ideModeTests : TestPool
ideModeTests = MkTestPool [] ideModeTests = MkTestPool "IDE mode" []
[ "ideMode001", "ideMode002", "ideMode003", "ideMode004" [ "ideMode001", "ideMode002", "ideMode003", "ideMode004", "ideMode005"
] ]
preludeTests : TestPool preludeTests : TestPool
preludeTests = MkTestPool [] preludeTests = MkTestPool "Prelude library" []
[ "reg001" [ "reg001"
] ]
templateTests : TestPool templateTests : TestPool
templateTests = MkTestPool [] templateTests = MkTestPool "Test templates" []
[ "simple-test", "ttimp", "with-ipkg" [ "simple-test", "ttimp", "with-ipkg"
] ]
@ -265,7 +270,7 @@ templateTests = MkTestPool []
-- that only runs if all backends are -- that only runs if all backends are
-- available. -- available.
baseLibraryTests : TestPool baseLibraryTests : TestPool
baseLibraryTests = MkTestPool [Chez, Node] baseLibraryTests = MkTestPool "Base library" [Chez, Node]
[ "system_file001" [ "system_file001"
, "data_bits001" , "data_bits001"
, "system_info001" , "system_info001"
@ -273,12 +278,12 @@ baseLibraryTests = MkTestPool [Chez, Node]
-- same behavior as `baseLibraryTests` -- same behavior as `baseLibraryTests`
contribLibraryTests : TestPool contribLibraryTests : TestPool
contribLibraryTests = MkTestPool [Chez, Node] contribLibraryTests = MkTestPool "Contrib library" [Chez, Node]
[ "json_001" [ "json_001"
] ]
codegenTests : TestPool codegenTests : TestPool
codegenTests = MkTestPool [] codegenTests = MkTestPool "Code generation" []
[ "con001" [ "con001"
, "builtin001" , "builtin001"
] ]

View File

@ -4,12 +4,16 @@ threads ?= `nproc`
.PHONY: testbin test .PHONY: testbin test
test: test:
./build/exec/runtests $(IDRIS2) $(INTERACTIVE) --threads $(threads) --only $(only) ./build/exec/runtests $(IDRIS2) $(INTERACTIVE) --failure-file failures --threads $(threads) --only $(only)
retest:
./build/exec/runtests $(IDRIS2) $(INTERACTIVE) --failure-file failures --threads $(threads) --only-file failures --only $(only)
testbin: testbin:
${IDRIS2} --build tests.ipkg ${IDRIS2} --build tests.ipkg
clean: clean:
$(RM) failures
$(RM) -r build $(RM) -r build
$(RM) -r **/**/build $(RM) -r **/**/build
@find . -type f -name 'output' -exec rm -rf {} \; @find . -type f -name 'output' -exec rm -rf {} \;

View File

@ -0,0 +1,205 @@
import Data.List
--------------------------------------------------------------------------------
-- Int8
--------------------------------------------------------------------------------
Show Int8 where
show = prim__cast_Int8String
public export
Eq Int8 where
x == y = intToBool (prim__eq_Int8 x y)
Num Int8 where
(+) = prim__add_Int8
(*) = prim__mul_Int8
fromInteger = prim__cast_IntegerInt8
Neg Int8 where
(-) = prim__sub_Int8
negate = prim__sub_Int8 0
--------------------------------------------------------------------------------
-- Int16
--------------------------------------------------------------------------------
Show Int16 where
show = prim__cast_Int16String
public export
Eq Int16 where
x == y = intToBool (prim__eq_Int16 x y)
Num Int16 where
(+) = prim__add_Int16
(*) = prim__mul_Int16
fromInteger = prim__cast_IntegerInt16
Neg Int16 where
(-) = prim__sub_Int16
negate = prim__sub_Int16 0
--------------------------------------------------------------------------------
-- Int32
--------------------------------------------------------------------------------
Show Int32 where
show = prim__cast_Int32String
public export
Eq Int32 where
x == y = intToBool (prim__eq_Int32 x y)
Num Int32 where
(+) = prim__add_Int32
(*) = prim__mul_Int32
fromInteger = prim__cast_IntegerInt32
Neg Int32 where
(-) = prim__sub_Int32
negate = prim__sub_Int32 0
--------------------------------------------------------------------------------
-- Int64
--------------------------------------------------------------------------------
Show Int64 where
show = prim__cast_Int64String
public export
Eq Int64 where
x == y = intToBool (prim__eq_Int64 x y)
Num Int64 where
(+) = prim__add_Int64
(*) = prim__mul_Int64
fromInteger = prim__cast_IntegerInt64
Neg Int64 where
(-) = prim__sub_Int64
negate = prim__sub_Int64 0
--------------------------------------------------------------------------------
-- Tests
--------------------------------------------------------------------------------
showTpe : Type -> String
showTpe Bits16 = "Bits16"
showTpe Bits32 = "Bits32"
showTpe Bits64 = "Bits64"
showTpe Bits8 = "Bits8"
showTpe Int = "Int"
showTpe Int16 = "Int16"
showTpe Int32 = "Int32"
showTpe Int64 = "Int64"
showTpe Int8 = "Int8"
showTpe Integer = "Integer"
showTpe _ = "unknown type"
testOp : (a: Type) -> (Show a, Eq a)
=> (opName : String)
-> (op : a -> a -> a)
-> List (a,a,a)
-> List String
testOp a n op = mapMaybe doTest
where doTest : (a,a,a) -> Maybe String
doTest (x,y,res) =
let myRes = op x y
in if myRes == res then Nothing
else Just $ #"Invalid result for \#{n} on \#{showTpe a}. "#
++ #"Inputs: \#{show x}, \#{show y}. "#
++ #"Expected \#{show res} but got \#{show myRes}."#
results : List String
results =
testOp Int8 "shl" prim__shl_Int8
[(0,7,0),(1,1,2),(1,3,8),(1,7,-128),(-1,7,-128)]
++ testOp Int8 "shr" prim__shr_Int8
[(0,7,0),(1,1,0),(-128,1,-64),(127,3,15),(-1,3,-1)]
++ testOp Int8 "and" prim__and_Int8
[(127,0,0),(-128,0,0),(23,-1,23),(-128,-1,-128),(15,8,8)]
++ testOp Int8 "or" prim__or_Int8
[(127,0,127),(-128,-1,-1),(23,-1,-1),(15,64,79)]
++ testOp Int8 "xor" prim__xor_Int8
[(127,0,127),(-128,-1,127),(127,-1,-128),(15,64,79),(15,1,14)]
++ testOp Bits8 "shl" prim__shl_Bits8
[(0,7,0),(1,1,2),(1,3,8),(1,7,128),(255,7,128)]
++ testOp Bits8 "shr" prim__shr_Bits8
[(0,7,0),(1,1,0),(255,1,127),(127,3,15),(255,3,31)]
++ testOp Bits8 "and" prim__and_Bits8
[(127,0,0),(255,0,0),(23,255,23),(128,255,128),(15,8,8)]
++ testOp Bits8 "or" prim__or_Bits8
[(127,0,127),(128,255,255),(23,255,255),(15,64,79)]
++ testOp Bits8 "xor" prim__xor_Bits8
[(127,0,127),(128,255,127),(127,255,128),(15,64,79),(15,1,14)]
++ testOp Int16 "shl" prim__shl_Int16
[(0,15,0),(1,1,2),(1,4,16),(1,15,-0x8000),(-1,15,-0x8000)]
++ testOp Int16 "shr" prim__shr_Int16
[(0,15,0),(1,1,0),(-0x8000,1,-0x4000),(0x7fff,3,0x0fff),(-1,3,-1)]
++ testOp Int16 "and" prim__and_Int16
[(0x7fff,0,0),(-0x8000,0,0),(23,-1,23),(-0x8000,-1,-0x8000),(15,8,8)]
++ testOp Int16 "or" prim__or_Int16
[(127,0,127),(-0x8000,-1,-1),(23,-1,-1),(15,64,79)]
++ testOp Int16 "xor" prim__xor_Int16
[(127,0,127),(-0x8000,-1,0x7fff),(0x7fff,-1,-0x8000),(15,64,79),(15,1,14)]
++ testOp Bits16 "shl" prim__shl_Bits16
[(0,15,0),(1,1,2),(1,4,16),(1,15,0x8000),(0xffff,15,0x8000)]
++ testOp Bits16 "shr" prim__shr_Bits16
[(0,15,0),(1,1,0),(0xffff,1,0x7fff),(0x7fff,3,0x0fff),(0xffff,3,0x1fff)]
++ testOp Bits16 "and" prim__and_Bits16
[(0x7fff,0,0),(0xffff,0,0),(23,0xffff,23),(0x8000,0xffff,0x8000),(15,8,8)]
++ testOp Bits16 "or" prim__or_Bits16
[(0x7fff,0,0x7fff),(0x8000,0xffff,0xffff),(23,0xffff,0xffff),(15,64,79)]
++ testOp Bits16 "xor" prim__xor_Bits16
[(0x7fff,0,0x7fff),(0x8000,0xffff,0x7fff),(0x7fff,0xffff,0x8000),(15,64,79),(15,1,14)]
++ testOp Int32 "shl" prim__shl_Int32
[(0,31,0),(1,1,2),(1,4,16),(1,31,-0x80000000),(-1,31,-0x80000000)]
++ testOp Int32 "shr" prim__shr_Int32
[(0,31,0),(1,1,0),(-0x80000000,1,-0x40000000),(0x7fffffff,3,0x0fffffff),(-1,3,-1)]
++ testOp Int32 "and" prim__and_Int32
[(0x7fffffff,0,0),(-0x80000000,0,0),(23,-1,23),(-0x80000000,-1,-0x80000000),(31,8,8)]
++ testOp Int32 "or" prim__or_Int32
[(127,0,127),(-0x80000000,-1,-1),(23,-1,-1),(31,64,95)]
++ testOp Int32 "xor" prim__xor_Int32
[(127,0,127),(-0x80000000,-1,0x7fffffff),(0x7fffffff,-1,-0x80000000),(15,64,79),(15,1,14)]
++ testOp Bits32 "shl" prim__shl_Bits32
[(0,31,0),(1,1,2),(1,4,16),(1,31,0x80000000),(0xffffffff,31,0x80000000)]
++ testOp Bits32 "shr" prim__shr_Bits32
[(0,31,0),(1,1,0),(0xffffffff,1,0x7fffffff),(0x7fffffff,3,0x0fffffff),(0xffffffff,3,0x1fffffff)]
++ testOp Bits32 "and" prim__and_Bits32
[(0x7fffffff,0,0),(0xffffffff,0,0),(23,0xffffffff,23),(0x80000000,0xffffffff,0x80000000),(15,8,8)]
++ testOp Bits32 "or" prim__or_Bits32
[(0x7fffffff,0,0x7fffffff),(0x80000000,0xffffffff,0xffffffff),(23,0xffffffff,0xffffffff),(15,64,79)]
++ testOp Bits32 "xor" prim__xor_Bits32
[(0x7fffffff,0,0x7fffffff),(0x80000000,0xffffffff,0x7fffffff),(0x7fffffff,0xffffffff,0x80000000),(15,64,79),(15,1,14)]
++ testOp Int64 "shl" prim__shl_Int64
[(0,63,0),(1,1,2),(1,4,16),(1,63,-0x8000000000000000),(-1,63,-0x8000000000000000)]
++ testOp Int64 "shr" prim__shr_Int64
[(0,63,0),(1,1,0),(-0x8000000000000000,1,-0x4000000000000000),(0x7fffffffffffffff,3,0x0fffffffffffffff),(-1,3,-1)]
++ testOp Int64 "and" prim__and_Int64
[(0x7fffffffffffffff,0,0),(-0x8000000000000000,0,0),(23,-1,23),(-0x8000000000000000,-1,-0x8000000000000000),(63,8,8)]
++ testOp Int64 "or" prim__or_Int64
[(127,0,127),(-0x8000000000000000,-1,-1),(23,-1,-1),(63,64,127)]
++ testOp Int64 "xor" prim__xor_Int64
[(127,0,127),(-0x8000000000000000,-1,0x7fffffffffffffff),(0x7fffffffffffffff,-1,-0x8000000000000000),(15,64,79),(15,1,14)]
++ testOp Bits64 "shl" prim__shl_Bits64
[(0,63,0),(1,1,2),(1,4,16),(1,63,0x8000000000000000),(0xffffffffffffffff,63,0x8000000000000000)]
++ testOp Bits64 "shr" prim__shr_Bits64
[(0,63,0),(1,1,0),(0xffffffffffffffff,1,0x7fffffffffffffff),(0x7fffffffffffffff,3,0x0fffffffffffffff),(0xffffffffffffffff,3,0x1fffffffffffffff)]
++ testOp Bits64 "and" prim__and_Bits64
[(0x7fffffffffffffff,0,0),(0xffffffffffffffff,0,0),(23,0xffffffffffffffff,23),(0x8000000000000000,0xffffffffffffffff,0x8000000000000000),(15,8,8)]
++ testOp Bits64 "or" prim__or_Bits64
[(0x7fffffffffffffff,0,0x7fffffffffffffff),(0x8000000000000000,0xffffffffffffffff,0xffffffffffffffff),(23,0xffffffffffffffff,0xffffffffffffffff),(15,64,79)]
++ testOp Bits64 "xor" prim__xor_Bits64
[(0x7fffffffffffffff,0,0x7fffffffffffffff),(0x8000000000000000,0xffffffffffffffff,0x7fffffffffffffff),(0x7fffffffffffffff,0xffffffffffffffff,0x8000000000000000),(15,64,79),(15,1,14)]
main : IO ()
main = traverse_ putStrLn results

View File

@ -0,0 +1,2 @@
1/1: Building BitOps (BitOps.idr)
Main> Main> Bye for now!

2
tests/chez/bitops/input Normal file
View File

@ -0,0 +1,2 @@
:exec main
:q

3
tests/chez/bitops/run Normal file
View File

@ -0,0 +1,3 @@
$1 --no-banner --no-color --console-width 0 BitOps.idr < input
rm -rf build

View File

@ -1,5 +1,5 @@
Error: The given specifier was not accepted by any backend. Available backends: Error: The given specifier was not accepted by any backend. Available backends:
chez, racket, node, javascript, refc, gambit chez, chez-sep, racket, node, javascript, refc, gambit
Some backends have additional specifier rules, refer to their documentation. Some backends have additional specifier rules, refer to their documentation.
Specifiers.idr:29:1--30:35 Specifiers.idr:29:1--30:35
@ -7,7 +7,7 @@ Specifiers.idr:29:1--30:35
30 | plusRacketFail : Int -> Int -> Int 30 | plusRacketFail : Int -> Int -> Int
Error: The given specifier was not accepted by any backend. Available backends: Error: The given specifier was not accepted by any backend. Available backends:
chez, racket, node, javascript, refc, gambit chez, chez-sep, racket, node, javascript, refc, gambit
Some backends have additional specifier rules, refer to their documentation. Some backends have additional specifier rules, refer to their documentation.
Specifiers.idr:29:1--30:35 Specifiers.idr:29:1--30:35
@ -16,13 +16,13 @@ Specifiers.idr:29:1--30:35
Main> Loaded file Specifiers.idr Main> Loaded file Specifiers.idr
Specifiers> Error: The given specifier was not accepted by any backend. Available backends: Specifiers> Error: The given specifier was not accepted by any backend. Available backends:
chez, racket, node, javascript, refc, gambit chez, chez-sep, racket, node, javascript, refc, gambit
Some backends have additional specifier rules, refer to their documentation. Some backends have additional specifier rules, refer to their documentation.
Specifiers.idr:29:1--30:35 Specifiers.idr:29:1--30:35
Specifiers> [exec] Specifiers> Error: The given specifier was not accepted by any backend. Available backends: Specifiers> [exec] Specifiers> Error: The given specifier was not accepted by any backend. Available backends:
chez, racket, node, javascript, refc, gambit chez, chez-sep, racket, node, javascript, refc, gambit
Some backends have additional specifier rules, refer to their documentation. Some backends have additional specifier rules, refer to their documentation.
Specifiers.idr:29:1--30:35 Specifiers.idr:29:1--30:35

View File

@ -2,24 +2,32 @@ Dumping case trees to Main.cases
prim__add_Integer = [{arg:N}, {arg:N}]: (+Integer [!{arg:N}, !{arg:N}]) prim__add_Integer = [{arg:N}, {arg:N}]: (+Integer [!{arg:N}, !{arg:N}])
prim__sub_Integer = [{arg:N}, {arg:N}]: (-Integer [!{arg:N}, !{arg:N}]) prim__sub_Integer = [{arg:N}, {arg:N}]: (-Integer [!{arg:N}, !{arg:N}])
prim__mul_Integer = [{arg:N}, {arg:N}]: (*Integer [!{arg:N}, !{arg:N}]) prim__mul_Integer = [{arg:N}, {arg:N}]: (*Integer [!{arg:N}, !{arg:N}])
Main.main = [{ext:N}]: (Prelude.Interfaces.sum [(%con [cons] Builtin.MkPair Just 0 [(%con Prelude.Interfaces.MkFoldable Just 0 [(%lam acc (%lam elem (%lam func (%lam init (%lam input (Prelude.Types.foldr [!func, !init, !input])))))), (%lam elem (%lam acc (%lam func (%lam init (%lam input (Prelude.Types.foldl [!func, !init, !input])))))), (%lam elem (%lam {arg:N} (Prelude.Types.null [!{arg:N}]))), (%lam elem (%lam acc (%lam m (%lam {i_con:N} (%lam funcM (%lam init (%lam input (Prelude.Types.foldlM [!{i_con:N}, !funcM, !init, !input]))))))))]), (%con Prelude.Num.MkNum Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))])]), (Prelude.Types.rangeFromTo [(%con [cons] Builtin.MkPair Just 0 [(%con Prelude.Num.MkIntegral Just 0 [(%con Prelude.Num.MkNum Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.div [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.mod [!{arg:N}, !{arg:N}])))]), (%con [cons] Builtin.MkPair Just 0 [(%con Prelude.EqOrd.MkOrd Just 0 [(%con [cons] Prelude.EqOrd.MkEq Just 0 [(%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd./= [!{arg:N}, !{arg:N}])))]), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.compare [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.> [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.<= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.>= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.max [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.min [!{arg:N}, !{arg:N}])))]), (%con Prelude.Num.MkNeg Just 0 [(%con Prelude.Num.MkNum Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (Prelude.Num.negate [!{arg:N}])), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.- [!{arg:N}, !{arg:N}])))])])]), (cast-Integer-Int [1]), (cast-Integer-Int [100])])]) Main.main = [{ext:N}]: (Prelude.Interfaces.sum [(%con [cons] Builtin.MkPair Just 0 [(%con Prelude.Interfaces.MkFoldable Just 0 [(%lam acc (%lam elem (%lam func (%lam init (%lam input (Prelude.Types.foldr [!func, !init, !input])))))), (%lam elem (%lam acc (%lam func (%lam init (%lam input (Prelude.Types.foldl [!func, !init, !input])))))), (%lam elem (%lam {arg:N} (Prelude.Types.null [!{arg:N}]))), (%lam elem (%lam acc (%lam m (%lam {i_con:N} (%lam funcM (%lam init (%lam input (Prelude.Types.foldlM [!{i_con:N}, !funcM, !init, !input])))))))), (%lam elem (%lam {arg:N} (Prelude.Types.toList [!{arg:N}])))]), (%con Prelude.Num.MkNum Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))])]), (Prelude.Types.rangeFromTo [(%con [cons] Builtin.MkPair Just 0 [(%con Prelude.Num.MkIntegral Just 0 [(%con Prelude.Num.MkNum Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.div [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.mod [!{arg:N}, !{arg:N}])))]), (%con [cons] Builtin.MkPair Just 0 [(%con Prelude.EqOrd.MkOrd Just 0 [(%con [cons] Prelude.EqOrd.MkEq Just 0 [(%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd./= [!{arg:N}, !{arg:N}])))]), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.compare [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.> [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.<= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.>= [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.max [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (Prelude.EqOrd.min [!{arg:N}, !{arg:N}])))]), (%con Prelude.Num.MkNeg Just 0 [(%con Prelude.Num.MkNum Just 0 [(%lam {arg:N} (%lam {arg:N} (+Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (%lam {arg:N} (*Int [!{arg:N}, !{arg:N}]))), (%lam {arg:N} (cast-Integer-Int [!{arg:N}]))]), (%lam {arg:N} (Prelude.Num.negate [!{arg:N}])), (%lam {arg:N} (%lam {arg:N} (Prelude.Num.- [!{arg:N}, !{arg:N}])))])])]), (cast-Integer-Int [1]), (cast-Integer-Int [100])])])
Prelude.Basics.flip = [{arg:N}, {arg:N}, {arg:N}]: ((!{arg:N} [!{arg:N}]) [!{arg:N}]) Prelude.Basics.flip = [{arg:N}, {arg:N}, {arg:N}]: ((!{arg:N} [!{arg:N}]) [!{arg:N}])
Prelude.Basics.Nil = Constructor tag Just 0 arity 0
Prelude.Basics.:: = Constructor tag Just 1 arity 2
Builtin.snd = [{arg:N}]: (%case !{arg:N} [(%concase [cons] Builtin.MkPair Just 0 [{e:N}, {e:N}] !{e:N})] Nothing) Builtin.snd = [{arg:N}]: (%case !{arg:N} [(%concase [cons] Builtin.MkPair Just 0 [{e:N}, {e:N}] !{e:N})] Nothing)
Builtin.idris_crash = [{ext:N}]: (crash [___, !{ext:N}]) Builtin.idris_crash = [{ext:N}]: (crash [___, !{ext:N}])
Builtin.fst = [{arg:N}]: (%case !{arg:N} [(%concase [cons] Builtin.MkPair Just 0 [{e:N}, {e:N}] !{e:N})] Nothing) Builtin.fst = [{arg:N}]: (%case !{arg:N} [(%concase [cons] Builtin.MkPair Just 0 [{e:N}, {e:N}] !{e:N})] Nothing)
Builtin.believe_me = [{ext:N}]: (believe_me [___, ___, !{ext:N}]) Builtin.believe_me = [{ext:N}]: (believe_me [___, ___, !{ext:N}])
Prelude.Types.rangeFromTo = [{arg:N}, {arg:N}, {arg:N}]: (%case (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.MkOrd Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing) [(%constcase 1 (Prelude.Types.takeUntil [(%lam {arg:N} (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.MkOrd Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [!{arg:N}, (%lam {arg:N} (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.MkNeg Just 0 [{e:N}, {e:N}, {e:N}] (%case !{e:N} [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.MkNeg Just 0 [{e:N}, {e:N}, {e:N}] (%case !{e:N} [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing))] Nothing)]))] Nothing))] Nothing))])])), (%constcase 0 (%case (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.MkOrd Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing) [(%constcase 1 (Prelude.Types.takeUntil [(%lam {arg:N} (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.MkOrd Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [!{arg:N}, (%lam x (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.MkNeg Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!x]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.MkNeg Just 0 [{e:N}, {e:N}, {e:N}] (%case !{e:N} [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing))] Nothing)]))] Nothing))])])), (%constcase 0 (%con [cons] Prelude.Types.:: Just 1 [!{arg:N}, (%con [nil] Prelude.Types.Nil Just 0 [])]))] Nothing))] Nothing) Builtin.MkPair = Constructor tag Just 0 arity 2
Prelude.Types.null = [{arg:N}]: (%case !{arg:N} [(%concase [nil] Prelude.Types.Nil Just 0 [] (%delay Lazy 1)), (%concase [cons] Prelude.Types.:: Just 1 [{e:N}, {e:N}] (%delay Lazy 0))] Nothing) Prelude.Types.toList = [{ext:N}]: !{ext:N}
Prelude.Types.foldr = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase [nil] Prelude.Types.Nil Just 0 [] !{arg:N}), (%concase [cons] Prelude.Types.:: Just 1 [{e:N}, {e:N}] ((!{arg:N} [!{e:N}]) [(Prelude.Types.foldr [!{arg:N}, !{arg:N}, !{e:N}])]))] Nothing) Prelude.Types.rangeFromTo = [{arg:N}, {arg:N}, {arg:N}]: (%case (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.MkOrd Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing) [(%constcase 1 (Prelude.Types.takeUntil [(%lam {arg:N} (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.MkOrd Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [!{arg:N}, (%lam {arg:N} (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.MkNeg Just 0 [{e:N}, {e:N}, {e:N}] (%case !{e:N} [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.MkNeg Just 0 [{e:N}, {e:N}, {e:N}] (%case !{e:N} [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing))] Nothing)]))] Nothing))] Nothing))])])), (%constcase 0 (%case (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.MkOrd Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing) [(%constcase 1 (Prelude.Types.takeUntil [(%lam {arg:N} (%case (Builtin.fst [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.EqOrd.MkOrd Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}, {e:N}] ((!{e:N} [!{arg:N}]) [!{arg:N}]))] Nothing)), (Prelude.Types.countFrom [!{arg:N}, (%lam x (%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.MkNeg Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!x]) [(%case (Builtin.snd [(Builtin.snd [!{arg:N}])]) [(%concase Prelude.Num.MkNeg Just 0 [{e:N}, {e:N}, {e:N}] (%case !{e:N} [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [1]))] Nothing))] Nothing)]))] Nothing))])])), (%constcase 0 (%con [cons] Prelude.Basics.:: Just 1 [!{arg:N}, (%con [nil] Prelude.Basics.Nil Just 0 [])]))] Nothing))] Nothing)
Prelude.Types.foldl = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase [nil] Prelude.Types.Nil Just 0 [] !{arg:N}), (%concase [cons] Prelude.Types.:: Just 1 [{e:N}, {e:N}] (Prelude.Types.foldl [!{arg:N}, ((!{arg:N} [!{arg:N}]) [!{e:N}]), !{e:N}]))] Nothing) Prelude.Types.null = [{arg:N}]: (%case !{arg:N} [(%concase [nil] Prelude.Basics.Nil Just 0 [] (%delay Lazy 1)), (%concase [cons] Prelude.Basics.:: Just 1 [{e:N}, {e:N}] (%delay Lazy 0))] Nothing)
Prelude.Types.foldr = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase [nil] Prelude.Basics.Nil Just 0 [] !{arg:N}), (%concase [cons] Prelude.Basics.:: Just 1 [{e:N}, {e:N}] ((!{arg:N} [!{e:N}]) [(Prelude.Types.foldr [!{arg:N}, !{arg:N}, !{e:N}])]))] Nothing)
Prelude.Types.foldl = [{arg:N}, {arg:N}, {arg:N}]: (%case !{arg:N} [(%concase [nil] Prelude.Basics.Nil Just 0 [] !{arg:N}), (%concase [cons] Prelude.Basics.:: Just 1 [{e:N}, {e:N}] (Prelude.Types.foldl [!{arg:N}, ((!{arg:N} [!{arg:N}]) [!{e:N}]), !{e:N}]))] Nothing)
Prelude.Types.foldlM = [{arg:N}, {arg:N}, {arg:N}, {ext:N}]: (Prelude.Types.foldl [(%lam ma (%lam b (%case !{arg:N} [(%concase Prelude.Interfaces.MkMonad Just 0 [{e:N}, {e:N}, {e:N}] ((((!{e:N} [___]) [___]) [!ma]) [(%lam {eta:N} (Prelude.Basics.flip [!{arg:N}, !b, !{eta:N}]))]))] Nothing))), (%case !{arg:N} [(%concase Prelude.Interfaces.MkMonad Just 0 [{e:N}, {e:N}, {e:N}] (%case !{e:N} [(%concase Prelude.Interfaces.MkApplicative Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [___]) [!{arg:N}]))] Nothing))] Nothing), !{ext:N}]) Prelude.Types.foldlM = [{arg:N}, {arg:N}, {arg:N}, {ext:N}]: (Prelude.Types.foldl [(%lam ma (%lam b (%case !{arg:N} [(%concase Prelude.Interfaces.MkMonad Just 0 [{e:N}, {e:N}, {e:N}] ((((!{e:N} [___]) [___]) [!ma]) [(%lam {eta:N} (Prelude.Basics.flip [!{arg:N}, !b, !{eta:N}]))]))] Nothing))), (%case !{arg:N} [(%concase Prelude.Interfaces.MkMonad Just 0 [{e:N}, {e:N}, {e:N}] (%case !{e:N} [(%concase Prelude.Interfaces.MkApplicative Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [___]) [!{arg:N}]))] Nothing))] Nothing), !{ext:N}])
Prelude.Types.takeUntil = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%concase [cons] Prelude.Types.Stream.:: Just 0 [{e:N}, {e:N}] (%case (!{arg:N} [!{e:N}]) [(%constcase 1 (%con [cons] Prelude.Types.:: Just 1 [!{e:N}, (%con [nil] Prelude.Types.Nil Just 0 [])])), (%constcase 0 (%con [cons] Prelude.Types.:: Just 1 [!{e:N}, (Prelude.Types.takeUntil [!{arg:N}, (%force Inf !{e:N})])]))] Nothing))] Nothing) Prelude.Types.takeUntil = [{arg:N}, {arg:N}]: (%case !{arg:N} [(%concase [cons] Prelude.Types.Stream.:: Just 0 [{e:N}, {e:N}] (%case (!{arg:N} [!{e:N}]) [(%constcase 1 (%con [cons] Prelude.Basics.:: Just 1 [!{e:N}, (%con [nil] Prelude.Basics.Nil Just 0 [])])), (%constcase 0 (%con [cons] Prelude.Basics.:: Just 1 [!{e:N}, (Prelude.Types.takeUntil [!{arg:N}, (%force Inf !{e:N})])]))] Nothing))] Nothing)
Prelude.Types.prim__integerToNat = [{arg:N}]: (%case (%case (<=Integer [0, !{arg:N}]) [(%constcase 0 0)] Just 1) [(%constcase 1 (Builtin.believe_me [!{arg:N}])), (%constcase 0 0)] Nothing) Prelude.Types.prim__integerToNat = [{arg:N}]: (%case (%case (<=Integer [0, !{arg:N}]) [(%constcase 0 0)] Just 1) [(%constcase 1 (Builtin.believe_me [!{arg:N}])), (%constcase 0 0)] Nothing)
Prelude.Types.countFrom = [{arg:N}, {arg:N}]: (%con [cons] Prelude.Types.Stream.:: Just 0 [!{arg:N}, (%delay Inf (Prelude.Types.countFrom [(!{arg:N} [!{arg:N}]), !{arg:N}]))]) Prelude.Types.countFrom = [{arg:N}, {arg:N}]: (%con [cons] Prelude.Types.Stream.:: Just 0 [!{arg:N}, (%delay Inf (Prelude.Types.countFrom [(!{arg:N} [!{arg:N}]), !{arg:N}]))])
Prelude.Types.Stream.:: = Constructor tag Just 0 arity 2
Prelude.Num.negate = [{arg:N}]: (-Int [0, !{arg:N}]) Prelude.Num.negate = [{arg:N}]: (-Int [0, !{arg:N}])
Prelude.Num.mod = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.== [!{arg:N}, (cast-Integer-Int [0])]) [(%constcase 0 (%Int [!{arg:N}, !{arg:N}]))] Just (Builtin.idris_crash ["Unhandled input for Prelude.Num.case block in mod at Prelude/Num.idr:L:C--L:C"])) Prelude.Num.mod = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.== [!{arg:N}, (cast-Integer-Int [0])]) [(%constcase 0 (%Int [!{arg:N}, !{arg:N}]))] Just (Builtin.idris_crash ["Unhandled input for Prelude.Num.case block in mod at Prelude/Num.idr:L:C--L:C"]))
Prelude.Num.div = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.== [!{arg:N}, (cast-Integer-Int [0])]) [(%constcase 0 (/Int [!{arg:N}, !{arg:N}]))] Just (Builtin.idris_crash ["Unhandled input for Prelude.Num.case block in div at Prelude/Num.idr:L:C--L:C"])) Prelude.Num.div = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.== [!{arg:N}, (cast-Integer-Int [0])]) [(%constcase 0 (/Int [!{arg:N}, !{arg:N}]))] Just (Builtin.idris_crash ["Unhandled input for Prelude.Num.case block in div at Prelude/Num.idr:L:C--L:C"]))
Prelude.Num.- = [{ext:N}, {ext:N}]: (-Int [!{ext:N}, !{ext:N}]) Prelude.Num.- = [{ext:N}, {ext:N}]: (-Int [!{ext:N}, !{ext:N}])
Prelude.Num.MkNum = Constructor tag Just 0 arity 3
Prelude.Num.MkNeg = Constructor tag Just 0 arity 3
Prelude.Num.MkIntegral = Constructor tag Just 0 arity 3
Prelude.EqOrd.min = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]) [(%constcase 1 !{arg:N}), (%constcase 0 !{arg:N})] Nothing) Prelude.EqOrd.min = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]) [(%constcase 1 !{arg:N}), (%constcase 0 !{arg:N})] Nothing)
Prelude.EqOrd.max = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.> [!{arg:N}, !{arg:N}]) [(%constcase 1 !{arg:N}), (%constcase 0 !{arg:N})] Nothing) Prelude.EqOrd.max = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.> [!{arg:N}, !{arg:N}]) [(%constcase 1 !{arg:N}), (%constcase 0 !{arg:N})] Nothing)
Prelude.EqOrd.compare = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]) [(%constcase 1 0), (%constcase 0 (%case (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]) [(%constcase 1 1), (%constcase 0 2)] Nothing))] Nothing) Prelude.EqOrd.compare = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.< [!{arg:N}, !{arg:N}]) [(%constcase 1 0), (%constcase 0 (%case (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]) [(%constcase 1 1), (%constcase 0 2)] Nothing))] Nothing)
@ -29,7 +37,10 @@ Prelude.EqOrd.== = [{arg:N}, {arg:N}]: (%case (==Int [!{arg:N}, !{arg:N}]) [(%co
Prelude.EqOrd.< = [{arg:N}, {arg:N}]: (%case (<Int [!{arg:N}, !{arg:N}]) [(%constcase 0 0)] Just 1) Prelude.EqOrd.< = [{arg:N}, {arg:N}]: (%case (<Int [!{arg:N}, !{arg:N}]) [(%constcase 0 0)] Just 1)
Prelude.EqOrd.<= = [{arg:N}, {arg:N}]: (%case (<=Int [!{arg:N}, !{arg:N}]) [(%constcase 0 0)] Just 1) Prelude.EqOrd.<= = [{arg:N}, {arg:N}]: (%case (<=Int [!{arg:N}, !{arg:N}]) [(%constcase 0 0)] Just 1)
Prelude.EqOrd./= = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]) [(%constcase 1 0), (%constcase 0 1)] Nothing) Prelude.EqOrd./= = [{arg:N}, {arg:N}]: (%case (Prelude.EqOrd.== [!{arg:N}, !{arg:N}]) [(%constcase 1 0), (%constcase 0 1)] Nothing)
Prelude.Interfaces.sum = [{arg:N}, {ext:N}]: (%case (Builtin.fst [!{arg:N}]) [(%concase Prelude.Interfaces.MkFoldable Just 0 [{e:N}, {e:N}, {e:N}, {e:N}] (((((!{e:N} [___]) [___]) [(%lam {eta:N} (%lam {eta:N} (%case (Builtin.snd [!{arg:N}]) [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{eta:N}]) [!{eta:N}]))] Nothing)))]) [(%case (Builtin.snd [!{arg:N}]) [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [0]))] Nothing)]) [!{ext:N}]))] Nothing) Prelude.EqOrd.MkOrd = Constructor tag Just 0 arity 8
Prelude.EqOrd.MkEq = Constructor tag Just 0 arity 2
Prelude.Interfaces.sum = [{arg:N}, {ext:N}]: (%case (Builtin.fst [!{arg:N}]) [(%concase Prelude.Interfaces.MkFoldable Just 0 [{e:N}, {e:N}, {e:N}, {e:N}, {e:N}] (((((!{e:N} [___]) [___]) [(%lam {eta:N} (%lam {eta:N} (%case (Builtin.snd [!{arg:N}]) [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] ((!{e:N} [!{eta:N}]) [!{eta:N}]))] Nothing)))]) [(%case (Builtin.snd [!{arg:N}]) [(%concase Prelude.Num.MkNum Just 0 [{e:N}, {e:N}, {e:N}] (!{e:N} [0]))] Nothing)]) [!{ext:N}]))] Nothing)
Prelude.Interfaces.MkFoldable = Constructor tag Just 0 arity 5
PrimIO.unsafePerformIO = [{arg:N}]: (PrimIO.unsafeCreateWorld [(%lam w (PrimIO.unsafeDestroyWorld [___, (!{arg:N} [!w])]))]) PrimIO.unsafePerformIO = [{arg:N}]: (PrimIO.unsafeCreateWorld [(%lam w (PrimIO.unsafeDestroyWorld [___, (!{arg:N} [!w])]))])
PrimIO.unsafeDestroyWorld = [{arg:N}, {arg:N}]: !{arg:N} PrimIO.unsafeDestroyWorld = [{arg:N}, {arg:N}]: !{arg:N}
PrimIO.unsafeCreateWorld = [{arg:N}]: (!{arg:N} [%MkWorld]) PrimIO.unsafeCreateWorld = [{arg:N}]: (!{arg:N} [%MkWorld])

Some files were not shown because too many files have changed in this diff Show More