Implement %macro flag

A %macro must always be fully applied. Whenever the elaborator
encounters a %macro application (except in a function LHS) it evaluates
the application and sends the result to %runElab. So:

%macro
foo : args -> Elab TT
...
def = foo a b c

is equivalent to

foo : args -> Elab TT
...
def = %runElab foo a b c
This commit is contained in:
Edwin Brady 2020-06-01 17:55:54 +01:00
parent adc449acb3
commit d5b5af91d1
3 changed files with 15 additions and 3 deletions

View File

@ -14,6 +14,9 @@ Language changes:
+ API defined in `Language.Reflection`, including functions for getting types
of global names, constructors of data types, and adding new top level
declarations
+ Implemented `%macro` function flag, to remove the syntactic noise of
invoking elaborator scripts. This means the function must always
be fully applied, and is run under `%runElab`
Changes since Idris 2 v0.1.0
----------------------------

View File

@ -126,8 +126,13 @@ expandAmbigName mode nest env orig args (IVar fc x) exp
mkTerm : Bool -> EState vars -> Name -> GlobalDef -> RawImp
mkTerm prim est n def
= wrapDot prim est mode n (map (snd . snd) args)
(definition def) (buildAlt (IVar fc n) args)
= let tm = wrapDot prim est mode n (map (snd . snd) args)
(definition def) (buildAlt (IVar fc n) args) in
if Macro `elem` flags def
then case mode of
InLHS _ => tm
_ => IRunElab fc (ICoerced fc tm)
else tm
mkAlt : Bool -> EState vars -> (Name, Int, GlobalDef) -> RawImp
mkAlt prim est (fullname, i, gdef)

View File

@ -6,5 +6,9 @@ powerFn : Nat -> TTImp
powerFn Z = `(const 1)
powerFn (S k) = `(\x => mult x (~(powerFn k) x))
%macro
power : Nat -> Elab TT
power n = check (powerFn n)
cube : Nat -> Nat
cube = %runElab check (powerFn 3)
cube = power 3