Added support for GHC 8.0.2 and 8.2.1.
Removed support for GHC 7.6.3.
Markdown support for literate Agda [PR #2357].
Files ending in .lagda.md
will be parsed as literate Markdown files.
```
or ```agda
in its own line, and end with
```
, also in its own line.<!--
and -->
).The dot in front of an inaccessible pattern can now be skipped if the pattern consists entirely of constructors or literals. For example:
open import Agda.Builtin.Bool
data D : Bool → Set where
c : D true
f : (x : Bool) → D x → Bool
f true c = true
Before this change, you had to write f .true c = true
.
agda
test : Nat → Set
test zero with zero
test _ | _ = Nat
test (suc x) with zero
test _ | _ = Nat
We do not have to spell out the pattern of the parent clause
(zero
/ suc x
) in the with-clause if we do not need the
pattern variables. Note that x
is not in scope in the
with-clause!A more elaborate example, which cannot be reduced to
an ellipsis ...
:
record R : Set where
coinductive -- disallow matching
field f : Bool
n : Nat
data P (r : R) : Nat → Set where
fTrue : R.f r ≡ true → P r zero
nSuc : P r (suc (R.n r))
data Q : (b : Bool) (n : Nat) → Set where
true! : Q true zero
suc! : ∀{b n} → Q b (suc n)
test : (r : R) {n : Nat} (p : P r n) → Q (R.f r) n
test r nSuc = suc!
test r (fTrue p) with R.f r
test _ (fTrue ()) | false
test _ _ | true = true! -- underscore instead of (isTrue _)
This is useful for case splitting on the result inside an expression: given
record _×_ (A B : Set) : Set where
field
π₁ : A
π₂ : B
open _×_
one may case split on the result (C-c C-c RET) in a hole
λ { → {!!}}
of type A × B to produce
λ { .π₁ → {!!} ; .π₂ → {!!}}
Records with a field of an empty type are now recognized as empty by Agda. In particular, they can be matched against with an absurd pattern (). For example:
data ⊥ : Set where
record Empty : Set where
field absurdity : ⊥
magic : Empty → ⊥
magic ()
Injective pragmas.
Injective pragmas can be used to mark a definition as injective for the
pattern matching unifier. This can be used as a version of
--injective-type-constructors
that only applies to specific datatypes.
For example:
open import Agda.Builtin.Equality
data Fin : Nat → Set where
zero : {n : Nat} → Fin (suc n)
suc : {n : Nat} → Fin n → Fin (suc n)
{-# INJECTIVE Fin #-}
Fin-injective : {m n : Nat} → Fin m ≡ Fin n → m ≡ n
Fin-injective refl = refl
Aside from datatypes, this pragma can also be used to mark other definitions as being injective (for example postulates).
Metavariables can no longer be instantiated during case splitting. This means Agda will refuse to split instead of taking the first constructor it finds. For example:
open import Agda.Builtin.Nat
data Vec (A : Set) : Nat → Set where
nil : Vec A 0
cons : {n : Nat} → A → Vec A n → Vec A (suc n)
foo : Vec Nat _ → Nat
foo x = {!x!}
In Agda 2.5.2, case splitting on x
produced the single clause
foo nil = {!!}
, but now Agda refuses to split.
debugPrint
. debugPrint : String → Nat → List ErrorPart → TC ⊤
This maps to the internal function reportSDoc
. Debug output is enabled with
the -v
flag at the command line, or in an OPTIONS
pragma. For instance,
giving -v a.b.c:10
enables printing from debugPrint "a.b.c.d" 10 msg
. In the
Emacs mode, debug output ends up in the *Agda debug*
buffer.
BUILTIN REFL is now superfluous, subsumed by BUILTIN EQUALITY [Issue #2389].
BUILTIN EQUALITY is now more liberal [Issue #2386]. It accepts, among others, the following new definitions of equality: ```agda -- Non-universe polymorphic: data ≡ {A : Set} (x : A) : A → Set where refl : x ≡ x
-- ... with explicit argument to refl; data ≡ {A : Set} : (x y : A) → Set where refl : {x : A} → x ≡ x
-- ... even visible data ≡ {A : Set} : (x y : A) → Set where refl : (x : A) → x ≡ x
-- Equality in a different universe than domain: -- (also with explicit argument to refl) data ≡ {a} {A : Set a} (x : A) : A → Set where refl : x ≡ x
The standard definition is still:
```agda
-- Equality in same universe as domain:
data _≡_ {a} {A : Set a} (x : A) : A → Set a where
refl : x ≡ x
If your file is named Bla.agda
, then the following content
is rejected.
foo = Set
module Bla where
bar = Set
Before the fix of this issue, Agda would add the missing module
header module Bla where
at the top of the file.
However, in this particular case it is more likely the user
put the declaration foo = Set
before the module start in error.
Now you get the error
Illegal declaration(s) before top-level module
if the following conditions are met:
1. There is at least one non-import declaration or non-toplevel pragma
before the start of the first module.
2. The module has the same name as the file.
3. The module is the only module at this level
(may have submodules, of course).
If you should see this error, insert a top-level module before the illegal declarations, or move them inside the existing module.
New warnings:
Unreachable clauses give rise to a simple warning. They are highlighted in gray.
Incomplete patterns are non-fatal warnings: it is possible to keep interacting with the file (the reduction will simply be stuck on arguments not matching any pattern). The definition with incomplete patterns are highlighted in wheat.
Clauses which do not hold definitionally are now highlighted in white smoke.
Fewer commands have the side effect that the buffer is saved.
Aborting commands.
Now one can (try to) abort an Agda command by using C-c C-x C-a
or
a menu entry. The effect is similar to that of restarting Agda (C-c
C-x C-r
), but some state is preserved, which could mean that it
takes less time to reload the module.
Warning: If a command is aborted while it is writing data to disk
(for instance .agdai
files or Haskell files generated by the GHC
backend), then the resulting files may be corrupted. Note also that
external commands (like GHC) are not aborted, and their output may
continue to be sent to the Emacs mode.
New bindings for the Agda input method:
All the bold digits are now available. The naming scheme is \Bx
for digit x
.
Typing \:
you can now get a whole slew of colons.
(The Agda input method originally only bound the standard unicode colon, which looks deceptively like the normal colon.)
Case splitting now preserves underscores. [Issue #819]
data ⊥ : Set where
test : {A B : Set} → A → ⊥ → B
test _ x = {! x !}
Splitting on x
yields
test _ ()
Interactively expanding ellipsis. [Issue #2589] An ellipsis in a with-clause can be expanded by splitting on "variable" "." (dot).
test0 : Nat → Nat
test0 x with zero
... | q = {! . !} -- C-c C-c
Splitting on dot here yields:
test0 x | q = ?
New command to check an expression against the type of the hole
it is in and see what it elaborates to.
[Issue #2700]
This is useful to determine e.g. what solution typeclass resolution yields.
The command is bound to C-c C-;
and respects the C-u
modifier.
record Pointed (A : Set) : Set where
field point : A
it : ∀ {A : Set} {{x : A}} → A
it {{x}} = x
instance _ = record { point = 3 - 4 }
_ : Pointed Nat
_ = {! it !} -- C-u C-u C-c C-;
yields
Goal: Pointed Nat
Elaborates to: record { point = 0 }
agda2-give
is called with a prefix, then giving is forced,
i.e., the safety checks are skipped,
including positivity, termination, and double type-checking.
[Issue #2730]Invoke forced giving with key sequence C-u C-c C-SPC
.
name
field in an .agda-lib
file is now optional.
[Issue #2708]This feature is convenient if you just want to specify the dependencies
and include pathes for your local project in an .agda-lib
file.
Naturally, libraries without names cannot be depended on.
The compiler pragmas (COMPILED
, COMPILED_DATA
, etc.) have been unified across
backends into two new pragmas:
{-# COMPILE <Backend> <Name> <Text> #-}
{-# FOREIGN <Backend> <Text> #-}
The old pragmas still work, but will emit a warning if used. They will be removed completely in Agda 2.6.
The translation of old pragmas into new ones is as follows:
Old | New |
---|---|
{-# COMPILED f e #-} |
{-# COMPILE GHC f = e #-} |
{-# COMPILED_TYPE A T #-} |
{-# COMPILE GHC A = type T #-} |
{-# COMPILED_DATA A D C1 .. CN #-} |
{-# COMPILE GHC A = data D (C1 \| .. \| CN) #-} |
{-# COMPILED_DECLARE_DATA #-} |
obsolete, removed |
{-# COMPILED_EXPORT f g #-} |
{-# COMPILE GHC f as g #-} |
{-# IMPORT M #-} |
{-# FOREIGN GHC import qualified M #-} |
{-# HASKELL code #-} |
{-# FOREIGN GHC code #-} |
{-# COMPILED_UHC f e #-} |
{-# COMPILE UHC f = e #-} |
{-# COMPILED_DATA_UHC A D C1 .. CN #-} |
{-# COMPILE UHC A = data D (C1 \| .. \| CN) #-} |
{-# IMPORT_UHC M #-} |
{-# FOREIGN UHC __IMPORT__ M #-} |
{-# COMPILED_JS f e #-} |
{-# COMPILE JS f = e #-} |
The COMPILED pragma (and the corresponding COMPILE GHC pragma) is now also allowed for functions. This makes it possible to have both an Agda implementation and a native Haskell runtime implementation.
The GHC file header pragmas LANGUAGE
, OPTIONS_GHC
, and INCLUDE
inside a FOREIGN GHC
pragma are recognized and printed correctly
at the top of the generated Haskell file.
[Issue #2712]
The UHC backend has been moved to its own repository [https://github.com/agda/agda-uhc] and is no longer part of the Agda distribution.
The (now deprecated) IMPORT and IMPORT_UHC pragmas no longer cause import statements in modules importing the module containing the pragma.
The same is true for the corresponding FOREIGN pragmas.
There is a new API in Agda.Compiler.Backend
for creating stand-alone
backends using Agda as a library. This allows prospective backend writers to
experiment with new backends without having to change the Agda code base.
Symbolic anchors look like
<a id="test1">
<a id="M.bla">
while other anchors just give the character position in the file:
<a id="42">
Top-level module names do not get a symbolic anchor, since the position of a top-level module is defined to be the beginning of the file.
Example:
module Issue2604 where -- Character position anchor
test1 : Set₁ -- Issue2604.html#test1
test1 = bla
where
bla = Set -- Character position anchor
test2 : Set₁ -- Issue2604.html#test2
test2 = bla
where
bla = Set -- Character position anchor
test3 : Set₁ -- Issue2604.html#test3
test3 = bla
module M where -- Issue2604.html#M
bla = Set -- Issue2604.html#M.bla
module NamedModule where -- Issue2604.html#NamedModule
test4 : Set₁ -- Issue2604.html#NamedModule.test4
test4 = M.bla
module _ where -- Character position anchor
test5 : Set₁ -- Character position anchor
test5 = M.bla
Agda now uses an encoding that amounts to first converting the
module names to UTF-8, and then percent-encoding the resulting
bytes. For instance, HTML for the module Σ
is placed in
%CE%A3.html
.
A constraint on the indentation of the first token t on a line is determined as follows:
Note that if any token in L or E belongs to a previous code
block, then the constraint may not be satisfied unless (say) the
AgdaAlign
environment is used in an appropriate way.
If custom settings are used, for instance if \AgdaIndent
is
redefined, then the constraint discussed above may not be satisfied.
(Note that the meaning of the \AgdaIndent
command's argument has
changed, and that the command is now used in a different way in the
generated LaTeX files.)
Examples:
Here C
is indented further than B
:
postulate
A B
C : Set
Here C
is not (necessarily) indented further than B
, because
X
shadows B
:
postulate
A B : Set
X
C : Set
The new rule is inspired by, but not identical to, the one used by lhs2TeX's poly mode (see Section 8.4 of the manual for lhs2TeX version 1.17).
Some spacing issues [#2353, #2441, #2733, #2740] have been fixed.
The user can now control the typesetting of (certain) individual tokens
by redefining the \AgdaFormat
command. Example:
```latex
\usepackage{ifthen}
% Insert extra space before some tokens. \DeclareRobustCommand{\AgdaFormat}[2]{%
\ifthenelse{
\equal{#1}{≡⟨} \OR
\equal{#1}{≡⟨⟩} \OR
\equal{#1}{∎}
}{\ }{}#2}
Note the use of `\DeclareRobustCommand`. The first argument to
`\AgdaFormat` is the token, and the second argument the thing to
be typeset.
* One can now instruct the agda package not to select any fonts.
If the `nofontsetup` option is used, then some font packages are
loaded, but specific fonts are not selected:
```latex
\usepackage[nofontsetup]{agda}
The height is controlled by the length \AgdaEmptySkip
, which by
default is \baselineskip
.
+̲
, containing +
and a
combining character, as having length two. However, it seems more
reasonable to treat it as having length one, as it occupies a single
column, if displayed "properly" using a monospace font. The new flag
--count-clusters
is an attempt at fixing this. When this flag is
enabled the backend counts "extended grapheme
clusters"
rather than code points.Note that this fix is not perfect: a single extended grapheme cluster might be displayed in different ways by different programs, and might, in some cases, occupy more than one column. Here are some examples of extended grapheme clusters, all of which are treated as a single character by the alignment algorithm:
│ │
│+̲│
│Ö̂│
│நி│
│ᄀힰᇹ│
│ᄀᄀᄀᄀᄀᄀힰᇹᇹᇹᇹᇹᇹ│
│ │
Note also that the layout machinery does not count extended grapheme
clusters, but code points. The following code is syntactically
correct, but if --count-clusters
is used, then the LaTeX backend
does not align the two field
keywords:
record +̲ : Set₁ where field A : Set
field B : Set
The --count-clusters
flag is not enabled in all builds of Agda,
because the implementation depends on the
ICU library, the installation of
which could cause extra trouble for some users. The presence of this
flag is controlled by the Cabal flag enable-cluster-counting
.
When this variant of the backend is used the top-level module is not type-checked, only scope-checked. This implies that some highlighting information is not available. For instance, overloaded constructors are not resolved.
QuickLaTeX can be invoked from the Emacs mode, or using agda
--latex --only-scope-checking
. If the module has already been
type-checked successfully, then this information is reused; in this
case QuickLaTeX behaves like the regular LaTeX backend.
The --only-scope-checking
flag can also be used independently, but
it is perhaps unclear what purpose that would serve. (The flag can
currently not be combined with --html
, --dependency-graph
or
--vim
.) The flag is not allowed in safe mode.
--safe
option is now a valid pragma.This makes it possible to declare a module as being part of the safe
subset of the language by stating {-# OPTIONS --safe #-}
at the top
of the corresponding file. Incompatibilities between the --safe
option
and other options or language constructs are non-fatal errors.
--no-main
option is now a valid pragma.One can now suppress the compiler warning about a missing main function by putting
{-# OPTIONS --no-main #-}
on top of the file.
--warning=MODE
(or -W MODE
) for
setting the warning mode. Current options are
warn
for displaying warnings (default)error
for turning warnings into errorsignore
for not displaying warningsFor 2.5.3, the following issues have been fixed (see bug tracker):
Size< i
to gain the necessary informationC-c C-z
) panics on pattern synonyms--without-K
is outdatedagda --latex
produces invalid LaTeX when there are block commentsexact-split
documentation is outdated and incompletename
field in .agda-lib files mandatory?