Reflection

Builtin types

Names

The built-in QNAME type represents quoted names and comes equipped with equality, ordering, and a show function.

  1. postulate Name : Set
  2. {-# BUILTIN QNAME Name #-}
  3. primitive
  4. primQNameEquality : Name Name Bool
  5. primQNameLess : Name Name Bool
  6. primShowQName : Name String

The fixity of a name can also be retrived.

  1. primitive
  2. primQNameFixity : Name Fixity

To define a decidable propositional equality with the option --safe, one can use the conversion to a pair of built-in 64-bit machine words

  1. primitive
  2. primQNameToWord64s : Name Σ Word64 _ Word64)

with the injectivity proof in the Properties module.:

  1. primitive
  2. primQNameToWord64sInjective : a b primQNameToWord64s a primQNameToWord64s b a b

Name literals are created using the quote keyword and can appear both in terms and in patterns

  1. nameOfNat : Name
  2. nameOfNat = quote Nat
  3. isNat : Name Bool
  4. isNat (quote Nat) = true
  5. isNat _ = false

Note that the name being quoted must be in scope.

Metavariables

Metavariables are represented by the built-in AGDAMETA type. They have primitive equality, ordering, show, and conversion to Nat:

  1. postulate Meta : Set
  2. {-# BUILTIN AGDAMETA Meta #-}
  3. primitive
  4. primMetaEquality : Meta Meta Bool
  5. primMetaLess : Meta Meta Bool
  6. primShowMeta : Meta String
  7. primMetaToNat : Meta Nat

Builtin metavariables show up in reflected terms. In Properties, there is a proof of injectivity of primMetaToNat

  1. primitive
  2. primMetaToNatInjective : a b primMetaToNat a primMetaToNat b a b

which can be used to define a decidable propositional equality with the option --safe.

Literals

Literals are mapped to the built-in AGDALITERAL datatype. Given the appropriate built-in binding for the types Nat, Float, etc, the AGDALITERAL datatype has the following shape:

  1. data Literal : Set where
  2. nat : (n : Nat) Literal
  3. word64 : (n : Word64) Literal
  4. float : (x : Float) Literal
  5. char : (c : Char) Literal
  6. string : (s : String) Literal
  7. name : (x : Name) Literal
  8. meta : (x : Meta) Literal
  9. {-# BUILTIN AGDALITERAL Literal #-}
  10. {-# BUILTIN AGDALITNAT nat #-}
  11. {-# BUILTIN AGDALITWORD64 word64 #-}
  12. {-# BUILTIN AGDALITFLOAT float #-}
  13. {-# BUILTIN AGDALITCHAR char #-}
  14. {-# BUILTIN AGDALITSTRING string #-}
  15. {-# BUILTIN AGDALITQNAME name #-}
  16. {-# BUILTIN AGDALITMETA meta #-}

Arguments

Arguments can be (visible), {hidden}, or {{instance}}:

  1. data Visibility : Set where
  2. visible hidden instance : Visibility
  3. {-# BUILTIN HIDING Visibility #-}
  4. {-# BUILTIN VISIBLE visible #-}
  5. {-# BUILTIN HIDDEN hidden #-}
  6. {-# BUILTIN INSTANCE instance #-}

Arguments can be relevant or irrelevant:

  1. data Relevance : Set where
  2. relevant irrelevant : Relevance
  3. {-# BUILTIN RELEVANCE Relevance #-}
  4. {-# BUILTIN RELEVANT relevant #-}
  5. {-# BUILTIN IRRELEVANT irrelevant #-}

Arguments also have a quantity:

  1. data Quantity : Set where
  2. quantity-0 quantity : Quantity
  3. {-# BUILTIN QUANTITY Quantity #-}
  4. {-# BUILTIN QUANTITY-0 quantity-0 #-}
  5. {-# BUILTIN QUANTITY quantity #-}

Relevance and quantity are combined into a modality:

  1. data Modality : Set where
  2. modality : (r : Relevance) (q : Quantity) Modality
  3. {-# BUILTIN MODALITY Modality #-}
  4. {-# BUILTIN MODALITY-CONSTRUCTOR modality #-}

The visibility and the modality characterise the behaviour of an argument:

  1. data ArgInfo : Set where
  2. arg-info : (v : Visibility) (m : Modality) ArgInfo
  3. data Arg (A : Set) : Set where
  4. arg : (i : ArgInfo) (x : A) Arg A
  5. {-# BUILTIN ARGINFO ArgInfo #-}
  6. {-# BUILTIN ARGARGINFO arg-info #-}
  7. {-# BUILTIN ARG Arg #-}
  8. {-# BUILTIN ARGARG arg #-}

Name abstraction

  1. data Abs (A : Set) : Set where
  2. abs : (s : String) (x : A) Abs A
  3. {-# BUILTIN ABS Abs #-}
  4. {-# BUILTIN ABSABS abs #-}

Terms

Terms, sorts, patterns, and clauses are mutually recursive and mapped to the AGDATERM, AGDASORT, AGDAPATTERN and AGDACLAUSE built-ins respectively. Types are simply terms. Terms and patterns use de Bruijn indices to represent variables.

  1. data Term : Set
  2. data Sort : Set
  3. data Pattern : Set
  4. data Clause : Set
  5. Type = Term
  6. Telescope = List String λ _ Arg Type)
  7. data Term where
  8. var : (x : Nat) (args : List (Arg Term)) Term
  9. con : (c : Name) (args : List (Arg Term)) Term
  10. def : (f : Name) (args : List (Arg Term)) Term
  11. lam : (v : Visibility) (t : Abs Term) Term
  12. pat-lam : (cs : List Clause) (args : List (Arg Term)) Term
  13. pi : (a : Arg Type) (b : Abs Type) Term
  14. agda-sort : (s : Sort) Term
  15. lit : (l : Literal) Term
  16. meta : (x : Meta) List (Arg Term) Term
  17. unknown : Term -- Treated as '_' when unquoting.
  18. data Sort where
  19. set : (t : Term) Sort -- A Set of a given (possibly neutral) level.
  20. lit : (n : Nat) Sort -- A Set of a given concrete level.
  21. prop : (t : Term) Sort -- A Prop of a given (possibly neutral) level.
  22. propLit : (n : Nat) Sort -- A Prop of a given concrete level.
  23. inf : (n : Nat) Sort -- Setωi of a given concrete level i.
  24. unknown : Sort
  25. data Pattern where
  26. con : (c : Name) (ps : List (Arg Pattern)) Pattern
  27. dot : (t : Term) Pattern
  28. var : (x : Nat ) Pattern
  29. lit : (l : Literal) Pattern
  30. proj : (f : Name) Pattern
  31. absurd : (x : Nat) Pattern -- Absurd patterns have de Bruijn indices
  32. data Clause where
  33. clause : (tel : Telescope) (ps : List (Arg Pattern)) (t : Term) Clause
  34. absurd-clause : (tel : Telescope) (ps : List (Arg Pattern)) Clause
  35. {-# BUILTIN AGDATERM Term #-}
  36. {-# BUILTIN AGDASORT Sort #-}
  37. {-# BUILTIN AGDAPATTERN Pattern #-}
  38. {-# BUILTIN AGDACLAUSE Clause #-}
  39. {-# BUILTIN AGDATERMVAR var #-}
  40. {-# BUILTIN AGDATERMCON con #-}
  41. {-# BUILTIN AGDATERMDEF def #-}
  42. {-# BUILTIN AGDATERMMETA meta #-}
  43. {-# BUILTIN AGDATERMLAM lam #-}
  44. {-# BUILTIN AGDATERMEXTLAM pat-lam #-}
  45. {-# BUILTIN AGDATERMPI pi #-}
  46. {-# BUILTIN AGDATERMSORT agda-sort #-}
  47. {-# BUILTIN AGDATERMLIT lit #-}
  48. {-# BUILTIN AGDATERMUNSUPPORTED unknown #-}
  49. {-# BUILTIN AGDASORTSET set #-}
  50. {-# BUILTIN AGDASORTLIT lit #-}
  51. {-# BUILTIN AGDASORTPROP prop #-}
  52. {-# BUILTIN AGDASORTPROPLIT propLit #-}
  53. {-# BUILTIN AGDASORTINF inf #-}
  54. {-# BUILTIN AGDASORTUNSUPPORTED unknown #-}
  55. {-# BUILTIN AGDAPATCON con #-}
  56. {-# BUILTIN AGDAPATDOT dot #-}
  57. {-# BUILTIN AGDAPATVAR var #-}
  58. {-# BUILTIN AGDAPATLIT lit #-}
  59. {-# BUILTIN AGDAPATPROJ proj #-}
  60. {-# BUILTIN AGDAPATABSURD absurd #-}
  61. {-# BUILTIN AGDACLAUSECLAUSE clause #-}
  62. {-# BUILTIN AGDACLAUSEABSURD absurd-clause #-}

Absurd lambdas λ () are quoted to extended lambdas with an absurd clause.

The built-in constructors AGDATERMUNSUPPORTED and AGDASORTUNSUPPORTED are translated to meta variables when unquoting.

Declarations

There is a built-in type AGDADEFINITION representing definitions. Values of this type is returned by the AGDATCMGETDEFINITION built-in described below.

  1. data Definition : Set where
  2. function : (cs : List Clause) Definition
  3. data-type : (pars : Nat) (cs : List Name) Definition -- parameters and constructors
  4. record-type : (c : Name) (fs : List (Arg Name)) -- c: name of record constructor
  5. Definition -- fs: fields
  6. data-cons : (d : Name) Definition -- d: name of data type
  7. axiom : Definition
  8. prim-fun : Definition
  9. {-# BUILTIN AGDADEFINITION Definition #-}
  10. {-# BUILTIN AGDADEFINITIONFUNDEF function #-}
  11. {-# BUILTIN AGDADEFINITIONDATADEF data-type #-}
  12. {-# BUILTIN AGDADEFINITIONRECORDDEF record-type #-}
  13. {-# BUILTIN AGDADEFINITIONDATACONSTRUCTOR data-cons #-}
  14. {-# BUILTIN AGDADEFINITIONPOSTULATE axiom #-}
  15. {-# BUILTIN AGDADEFINITIONPRIMITIVE prim-fun #-}

Type errors

Type checking computations (see below) can fail with an error, which is a list of ErrorParts. This allows metaprograms to generate nice errors without having to implement pretty printing for reflected terms.

  1. -- Error messages can contain embedded names and terms.
  2. data ErrorPart : Set where
  3. strErr : String ErrorPart
  4. termErr : Term ErrorPart
  5. pattErr : Pattern ErrorPart
  6. nameErr : Name ErrorPart
  7. {-# BUILTIN AGDAERRORPART ErrorPart #-}
  8. {-# BUILTIN AGDAERRORPARTSTRING strErr #-}
  9. {-# BUILTIN AGDAERRORPARTTERM termErr #-}
  10. {-# BUILTIN AGDAERRORPARTNAME nameErr #-}

Blockers

A blocker represents a set of metavariables that impedes the progress of a reflective computation. Using a blocker containing all the metas in (for example) a term traversed by a macro is a lot more efficient than blocking on individual metas as they are encountered.

  1. data Blocker : Set where
  2. blockerAny : List Blocker Blocker
  3. blockerAll : List Blocker Blocker
  4. blockerMeta : Meta Blocker
  5. {-# BUILTIN AGDABLOCKER Blocker #-}
  6. {-# BUILTIN AGDABLOCKERANY blockerAny #-}
  7. {-# BUILTIN AGDABLOCKERALL blockerAll #-}
  8. {-# BUILTIN AGDABLOCKERMETA blockerMeta #-}

Type checking computations

Metaprograms, i.e. programs that create other programs, run in a built-in type checking monad TC:

  1. postulate
  2. TC : {a} Set a Set a
  3. returnTC : {a} {A : Set a} A TC A
  4. bindTC : {a b} {A : Set a} {B : Set b} TC A (A TC B) TC B
  5. {-# BUILTIN AGDATCM TC #-}
  6. {-# BUILTIN AGDATCMRETURN returnTC #-}
  7. {-# BUILTIN AGDATCMBIND bindTC #-}

The TC monad provides an interface to the Agda type checker using the following primitive operations:

  1. postulate
  2. -- Unify two terms, potentially solving metavariables in the process.
  3. unify : Term Term TC
  4. -- Throw a type error. Can be caught by catchTC.
  5. typeError : {a} {A : Set a} List ErrorPart TC A
  6. -- Block a type checking computation on a blocker. This will abort
  7. -- the computation and restart it (from the beginning) when the
  8. -- blocker has been solved.
  9. blockTC : {a} {A : Set a} Blocker TC A
  10. -- Prevent current solutions of metavariables from being rolled back in
  11. -- case 'blockOnMeta' is called.
  12. commitTC : TC
  13. -- Backtrack and try the second argument if the first argument throws a
  14. -- type error.
  15. catchTC : {a} {A : Set a} TC A TC A TC A
  16. -- Infer the type of a given term
  17. inferType : Term TC Type
  18. -- Check a term against a given type. This may resolve implicit arguments
  19. -- in the term, so a new refined term is returned. Can be used to create
  20. -- new metavariables: newMeta t = checkType unknown t
  21. checkType : Term Type TC Term
  22. -- Compute the normal form of a term.
  23. normalise : Term TC Term
  24. -- Compute the weak head normal form of a term.
  25. reduce : Term TC Term
  26. -- Get the current context. Returns the context in reverse order, so that
  27. -- it is indexable by deBruijn index. Note that the types in the context are
  28. -- valid in the rest of the context. To use in the current context they need
  29. -- to be weakened by 1 + their position in the list.
  30. getContext : TC Telescope
  31. -- Extend the current context with a variable of the given type and its name.
  32. extendContext : {a} {A : Set a} String Arg Type TC A TC A
  33. -- Set the current context relative to the context the TC computation
  34. -- is invoked from. Takes a context telescope entries in reverse
  35. -- order, as given by `getContext`. Each type should be valid in the
  36. -- context formed by the remaining elements in the list.
  37. inContext : {a} {A : Set a} Telescope TC A TC A
  38. -- Quote a value, returning the corresponding Term.
  39. quoteTC : {a} {A : Set a} A TC Term
  40. -- Unquote a Term, returning the corresponding value.
  41. unquoteTC : {a} {A : Set a} Term TC A
  42. -- Quote a value in Setω, returning the corresponding Term
  43. quoteωTC : {A : Setω} A TC Term
  44. -- Create a fresh name.
  45. freshName : String TC Name
  46. -- Declare a new function of the given type. The function must be defined
  47. -- later using 'defineFun'. Takes an Arg Name to allow declaring instances
  48. -- and irrelevant functions. The Visibility of the Arg must not be hidden.
  49. declareDef : Arg Name Type TC
  50. -- Declare a new postulate of the given type. The Visibility of the Arg
  51. -- must not be hidden. It fails when executed from command-line with --safe
  52. -- option.
  53. declarePostulate : Arg Name Type TC
  54. -- Declare a new datatype. The second argument is the number of parameters.
  55. -- The third argument is the type of the datatype, i.e. its parameters and
  56. -- indices. The datatype must be defined later using 'defineData'.
  57. declareData : Name Nat Type TC
  58. -- Define a declared datatype. The datatype must have been declared using
  59. -- 'declareData`. The second argument is a list of pairs in which each pair
  60. -- is the name of a constructor and its type.
  61. defineData : Name → List (Σ Name (λ _ → Type)) → TC ⊤
  62. -- Define a declared function. The function may have been declared using
  63. -- 'declareDef' or with an explicit type signature in the program.
  64. defineFun : Name → List Clause → TC ⊤
  65. -- Get the type of a defined name relative to the current
  66. -- module. Replaces 'primNameType'.
  67. getType : Name → TC Type
  68. -- Get the definition of a defined name relative to the current
  69. -- module. Replaces 'primNameDefinition'.
  70. getDefinition : Name → TC Definition
  71. -- Check if a name refers to a macro
  72. isMacro : Name → TC Bool
  73. -- Generate FOREIGN pragma with specified backend and top-level backend-dependent text.
  74. pragmaForeign : String → String → TC ⊤
  75. -- Generate COMPILE pragma with specified backend, associated name and backend-dependent text.
  76. pragmaCompile : String → Name → String → TC ⊤
  77. -- Change the behaviour of inferType, checkType, quoteTC, getContext
  78. -- to normalise (or not) their results. The default behaviour is no
  79. -- normalisation.
  80. withNormalisation : ∀ {a} {A : Set a} → Bool → TC A → TC A
  81. askNormalisation : TC Bool
  82. -- If 'true', makes the following primitives to reconstruct hidden arguments:
  83. -- getDefinition, normalise, reduce, inferType, checkType and getContext
  84. withReconstructed : ∀ {a} {A : Set a} → Bool → TC A → TC A
  85. askReconstructed : TC Bool
  86. -- Whether implicit arguments at the end should be turned into metavariables
  87. withExpandLast : ∀ {a} {A : Set a} → Bool → TC A → TC A
  88. askExpandLast : TC Bool
  89. -- White/blacklist specific definitions for reduction while executing the TC computation
  90. -- 'true' for whitelist, 'false' for blacklist
  91. withReduceDefs : ∀ {a} {A : Set a} → (Σ Bool λ _ → List Name) → TC A → TC A
  92. askReduceDefs : TC (Σ Bool λ _ → List Name)
  93. -- Prints the third argument to the debug buffer in Emacs
  94. -- if the verbosity level (set by the -v flag to Agda)
  95. -- is higher than the second argument. Note that Level 0 and 1 are printed
  96. -- to the info buffer instead. For instance, giving -v a.b.c:10 enables
  97. -- printing from debugPrint "a.b.c.d" 10 msg.
  98. debugPrint : String → Nat → List ErrorPart → TC ⊤
  99. -- Return the formatted string of the argument using the internal pretty printer.
  100. formatErrorParts : List ErrorPart → TC String
  101. -- Fail if the given computation gives rise to new, unsolved
  102. -- "blocking" constraints.
  103. noConstraints : ∀ {a} {A : Set a} → TC A → TC A
  104. -- Run the given TC action and return the first component. Resets to
  105. -- the old TC state if the second component is 'false', or keep the
  106. -- new TC state if it is 'true'.
  107. runSpeculative : ∀ {a} {A : Set a} → TC (Σ A λ _ → Bool) → TC A
  108. -- Get a list of all possible instance candidates for the given meta
  109. -- variable (it does not have to be an instance meta).
  110. getInstances : Meta → TC (List Term)
  111. {-# BUILTIN AGDATCMUNIFY unify #-}
  112. {-# BUILTIN AGDATCMTYPEERROR typeError #-}
  113. {-# BUILTIN AGDATCMBLOCK blockTC #-}
  114. {-# BUILTIN AGDATCMCATCHERROR catchTC #-}
  115. {-# BUILTIN AGDATCMINFERTYPE inferType #-}
  116. {-# BUILTIN AGDATCMCHECKTYPE checkType #-}
  117. {-# BUILTIN AGDATCMNORMALISE normalise #-}
  118. {-# BUILTIN AGDATCMREDUCE reduce #-}
  119. {-# BUILTIN AGDATCMGETCONTEXT getContext #-}
  120. {-# BUILTIN AGDATCMEXTENDCONTEXT extendContext #-}
  121. {-# BUILTIN AGDATCMINCONTEXT inContext #-}
  122. {-# BUILTIN AGDATCMQUOTETERM quoteTC #-}
  123. {-# BUILTIN AGDATCMUNQUOTETERM unquoteTC #-}
  124. {-# BUILTIN AGDATCMQUOTEOMEGATERM quoteωTC #-}
  125. {-# BUILTIN AGDATCMFRESHNAME freshName #-}
  126. {-# BUILTIN AGDATCMDECLAREDEF declareDef #-}
  127. {-# BUILTIN AGDATCMDECLAREPOSTULATE declarePostulate #-}
  128. {-# BUILTIN AGDATCMDECLAREDATA declareData #-}
  129. {-# BUILTIN AGDATCMDEFINEDATA defineData #-}
  130. {-# BUILTIN AGDATCMDEFINEFUN defineFun #-}
  131. {-# BUILTIN AGDATCMGETTYPE getType #-}
  132. {-# BUILTIN AGDATCMGETDEFINITION getDefinition #-}
  133. {-# BUILTIN AGDATCMCOMMIT commitTC #-}
  134. {-# BUILTIN AGDATCMISMACRO isMacro #-}
  135. {-# BUILTIN AGDATCMWITHNORMALISATION withNormalisation #-}
  136. {-# BUILTIN AGDATCMWITHRECONSTRUCTED withReconstructed #-}
  137. {-# BUILTIN AGDATCMWITHEXPANDLAST withExpandLast #-}
  138. {-# BUILTIN AGDATCMWITHREDUCEDEFS withReduceDefs #-}
  139. {-# BUILTIN AGDATCMASKNORMALISATION askNormalisation #-}
  140. {-# BUILTIN AGDATCMASKRECONSTRUCTED askReconstructed #-}
  141. {-# BUILTIN AGDATCMASKEXPANDLAST askExpandLast #-}
  142. {-# BUILTIN AGDATCMASKREDUCEDEFS askReduceDefs #-}
  143. {-# BUILTIN AGDATCMDEBUGPRINT debugPrint #-}
  144. {-# BUILTIN AGDATCMNOCONSTRAINTS noConstraints #-}
  145. {-# BUILTIN AGDATCMRUNSPECULATIVE runSpeculative #-}
  146. {-# BUILTIN AGDATCMGETINSTANCES getInstances #-}

Metaprogramming

There are three ways to run a metaprogram (TC computation). To run a metaprogram in a term position you use a macro. To run metaprograms to create top-level definitions you can use the unquoteDecl and unquoteDef primitives (see Unquoting Declarations).

Macros

Macros are functions of type t₁ → t₂ → .. → Term → TC ⊤ that are defined in a macro block. The last argument is supplied by the type checker and will be the representation of a metavariable that should be instantiated with the result of the macro.

Macro application is guided by the type of the macro, where Term and Name arguments are quoted before passed to the macro. Arguments of any other type are preserved as-is.

For example, the macro application f u v w where f : Term → Name → Bool → Term → TC ⊤ desugars into:

  1. unquote (f (quoteTerm u) (quote v) w)

where quoteTerm u takes a u of arbitrary type and returns its representation in the Term data type, and unquote m runs a computation in the TC monad. Specifically, when checking unquote m : A for some type A the type checker proceeds as follows:

  • Check m : Term → TC ⊤.

  • Create a fresh metavariable hole : A.

  • Let qhole : Term be the quoted representation of hole.

  • Execute m qhole.

  • Return (the now hopefully instantiated) hole.

Reflected macro calls are constructed using the def constructor, so given a macro g : Term → TC ⊤ the term def (quote g) [] unquotes to a macro call to g.

Note

The quoteTerm and unquote primitives are available in the language, but it is recommended to avoid using them in favour of macros.

Limitations:

  • Macros cannot be recursive. This can be worked around by defining the recursive function outside the macro block and have the macro call the recursive function.

Silly example:

  1. macro
  2. plus-to-times : Term Term TC
  3. plus-to-times (def (quote _+_) (a b [])) hole =
  4. unify hole (def (quote _*_) (a b []))
  5. plus-to-times v hole = unify hole v
  6. thm : (a b : Nat) plus-to-times (a + b) a * b
  7. thm a b = refl

Macros lets you write tactics that can be applied without any syntactic overhead. For instance, suppose you have a solver:

  1. magic : Type Term

that takes a reflected goal and outputs a proof (when successful). You can then define the following macro:

  1. macro
  2. by-magic : Term TC
  3. by-magic hole =
  4. bindTC (inferType hole) λ goal
  5. unify hole (magic goal)

This lets you apply the magic tactic as a normal function:

  1. thm : ¬ P NP
  2. thm = by-magic

Tactic Arguments

You can declare tactics to be used to solve a particular implicit argument using a @(tactic t) annotation. The provided tactic should be a term t : Term → TC ⊤. For instance,

  1. defaultTo : {A : Set} (x : A) Term TC
  2. defaultTo x hole = bindTC (quoteTC x) (unify hole)
  3. f : {@(tactic defaultTo true) x : Bool} Bool
  4. f {x} = x
  5. test-f : f true
  6. test-f = refl

At calls to f, defaultTo true is called on the metavariable inserted for x if it is not given explicitly. The tactic can depend on previous arguments to the function. For instance,

  1. g : (x : Nat) {@(tactic defaultTo x) y : Nat} Nat
  2. g x {y} = x + y
  3. test-g : g 4 8
  4. test-g = refl

Record fields can also be annotated with a tactic, allowing them to be omitted in constructor applications, record constructions and co-pattern matches:

  1. record Bools : Set where
  2. constructor mkBools
  3. field fst : Bool
  4. @(tactic defaultTo fst) {snd} : Bool
  5. open Bools
  6. tt tt tt tt : Bools
  7. tt = mkBools true {true}
  8. tt = mkBools true
  9. tt = record{ fst = true }
  10. tt .fst = true
  11. test-tt : tt tt tt [] tt tt tt []
  12. test-tt = refl

Unquoting Declarations

While macros let you write metaprograms to create terms, it is also useful to be able to create top-level definitions. You can do this from a macro using the declareDef, declareData, defineFun and defineData primitives, but there is no way to bring such definitions into scope. For this purpose there are two top-level primitives unquoteDecl and unquoteDef that runs a TC computation in a declaration position. They both have the same form for declaring function definitions:

  1. unquoteDecl x .. x = m
  2. unquoteDef x .. x = m

except that the list of names can be empty for unquoteDecl, but not for unquoteDef. In both cases m should have type TC ⊤. The main difference between the two is that unquoteDecl requires m to both declare (with declareDef) and define (with defineFun) the xᵢ whereas unquoteDef expects the xᵢ to be already declared. In other words, unquoteDecl brings the xᵢ into scope, but unquoteDef requires them to already be in scope.

In m the xᵢ stand for the names of the functions being defined (i.e. xᵢ : Name) rather than the actual functions.

One advantage of unquoteDef over unquoteDecl is that unquoteDef is allowed in mutual blocks, allowing mutually recursion between generated definitions and hand-written definitions.

Example usage:

  1. arg : {A : Set} Visibility A Arg A
  2. arg v = arg (arg-info v (modality relevant quantity-ω))
  3. -- Defining: id-name {A} x = x
  4. defId : (id-name : Name) TC
  5. defId id-name = do
  6. defineFun id-name
  7. [ clause
  8. ( ("A" , arg visible (agda-sort (lit 0)))
  9. ("x" , arg visible (var 0 []))
  10. [])
  11. ( arg hidden (var 1)
  12. arg visible (var 0)
  13. [] )
  14. (var 0 [])
  15. ]
  16. id : {A : Set} (x : A) A
  17. unquoteDef id = defId id
  18. mkId : (id-name : Name) TC
  19. mkId id-name = do
  20. ty quoteTC ({A : Set} (x : A) A)
  21. declareDef (arg visible id-name) ty
  22. defId id-name
  23. unquoteDecl id = mkId id

Another form of unquoteDecl is used to declare data types:

  1. unquoteDecl data x constructor c .. c = m

m is a metaprogram required to declare and define a data type x and its constructors c₁ to cₙ using declareData and defineData.

System Calls

It is possible to run system calls as part of a metaprogram, using the execTC builtin. You can use this feature to implement type providers, or to call external solvers. For instance, the following example calls /bin/echo from Agda:

  1. postulate
  2. execTC : (exe : String) (args : List String) (stdIn : String)
  3. TC Nat _ Σ String _ String)))
  4. {-# BUILTIN AGDATCMEXEC execTC #-}
  5. macro
  6. echo : List String Term TC
  7. echo args hole = do
  8. (exitCode , (stdOut , stdErr)) execTC "echo" args ""
  9. unify hole (lit (string stdOut))
  10. _ : echo ("hello" "world" []) "hello world\n"
  11. _ = refl

The execTC builtin takes three arguments: the basename of the executable (e.g., "echo"), a list of arguments, and the contents of the standard input. It returns a triple, consisting of the exit code (as a natural number), the contents of the standard output, and the contents of the standard error.

It would be ill-advised to allow Agda to make arbitrary system calls. Hence, the feature must be activated by passing the --allow-exec option, either on the command-line or using a pragma. (Note that --allow-exec is incompatible with --safe.) Furthermore, Agda can only call executables which are listed in the list of trusted executables, ~/.agda/executables. For instance, to run the example above, you must add /bin/echo to this file:

  1. # contents of ~/.agda/executables
  2. /bin/echo

The executable can then be called by passing its basename to execTC, subtracting the .exe on Windows.