Added support for GHC 8.8.2 [Issue #4285].
Removed support for GHC 7.10.3.
Interface files are now written in directory _build/VERSION/agda/
at
the project root (the closest enclosing directory where an .agda-lib
file is present). If there is no project root then the interface file
is written alongside the module it corresponds to.
The flag --local-interfaces
forces Agda to revert back to storing
interface files alongside module files no matter what.
Agda now uses the default RTS options -M3.5G -I0
. If
you run Agda on a 32-bit system or a system with less than 8GB of
RAM, it is recommended to set the RTS options explicitly to a lower
value by running agda
with option +RTS -M1.2G -RTS
(for example) or by setting the GHCRTS enviroment variable. See the
GHC User's Guide
for more information.
If Agda is compiled using GHC 8.4 or later, then one can expect to see substantially lower memory consumption [Issues #4457 and #4316].
This is due to the use of "compact regions".
CHANGELOG.md
was split. Changes to previous versions of Agda
are in the directory doc/release-notes
.New pragma WARNING_ON_IMPORT
to let module authors raise a warning
when a module is imported. This can be use to tell users deprecations.
New option --confluence-check
(off by default) enables confluence
checking of user-defined rewrite rules (this only has an effect when
--rewriting
is also enabled).
New option --no-projection-like
to turn off the analysis whether a
type signature likens that of a projection.
Projection-likeness is an optimization that reduces the size of
terms by dropping parameter-like reconstructible function arguments.
Thus, it is advisable to leave this optimization on, the flag is
meant for debugging Agda.
Option --no-forcing
is now a pragma option, i.e., the forcing analysis
can be switched off on a per-file basis via
{-# OPTIONS --no-forcing #-}
at the beginning of the file [Issue #3872].
New pragma option --no-flat-split
disables pattern matching on @♭
arguments.
New pragma option --allow-incomplete-matches
. It is similar to
--allow-unsolved-metas
: modules containing partial function definitions
can be imported. Its local equivalent is the NON_COVERING
pragma to
be placed before the function (or the block of mutually defined functions)
which the user knows to be partial.
Option --interaction-json
now brings more information about goals,
unsolved metas, warnings, errors.
It also displays pretty-printed terms.
New pragma option --keep-pattern-variables
to prevent case
splitting from replacing variables with dot patterns.
Pragma {-# ETA <record name> #-}
is no longer considered --safe
.
See [Issue #4450].
New pragma options --subtyping
and --no-subtyping
(default) to
turn on/off subtyping rules globally [see
Issue_#4474]. Currently,
this includes subtyping for irrelevance, erasure, and flat
modalities. Additionally, --subtyping
is implied by
--cumulativity
(see below). --subtyping
is currently NOT implied
by --sized-types
, and subtyping for sized types is used even when
--subtyping
is not enabled.
New profiling options to measure time spent per module or top-level definition.
-v profile.modules:10
prints a breakdown per top-level module-v profile.definitions:10
prints a breakdown per top-level
definition infix 3.14 _<_
Note that this includes a respective change in the reflected Agda syntax.
renaming
directive,
see
Issue #1346. Example: open M using (_∙_)
open M renaming (_∙_ to infixl 10 _*_)
After this, _∙_
is in scope with its original fixity, and as _*_
as left
associative operator of precedence 10.
Implicit non-dependent function spaces {A} → B
and {{A}} → B
are now supported.
Idiom brackets
Idiom brackets can accommodate none or multiple applications separated by a vertical bar |
if there are two additional operations
empty : ∀ {A} → F A
_<|>_ : ∀ {A} → F A → F A → F A
i.e. an Alternative type class in Haskell. As usual, the new idiom brackets desugar before scope checking.
Idiom brackets with multiple applications
(| e₁ a₁ .. aₙ | e₂ a₁ .. aₘ | .. | eₖ a₁ .. aₗ |)
expand to (assuming right associative _<|>_
)
(pure e₁ <*> a₁ <*> .. <*> aₙ) <|> ((pure e₂ <*> a₁ <*> .. <*> aₘ) <|> (pure eₖ <*> a₁ <*> .. <*> aₗ))
Idiom brackets with no application (|)
or ⦇⦈
are equivalent to empty
.
Users can now match on irrefutable patterns on the LHS using a
pattern-matching with
. An expression of the form:
f xs with p1 <- e1 | ... | pn <- en
with q1 <- f1 | ... | qm <- fm = rhs
is translated to nested with
clauses, essentially equivalent to:
f xs with e1 | ... | en
... | p1 | ... | pn
with f1 | ... | fm
... | q1 | ... | qm = rhs
Users can now use record patterns in telescope and lambda abstractions. The type of the second projection from a dependent pair is the prototypical example It can be defined as follows:
snd : ((a , _) : Σ A B) → B a
And this second projection can be implemented with a lamba-abstraction using one of these irrefutable patterns:
snd = λ (a , b) → b
Using an as-pattern, users can get a name for the value as well as for its subparts. We can for instance prove that any pair is equal to the pairing of its first and second projections:
eta : (p@(a , b) : Σ A B) → p ≡ (a , b)
eta p = refl
Absurd match in a do block
The last expression in a do block can now also be an absurd match () <- f
.
Named where
modules are now in scope in the rhs of the clause (see
Issue #4050). Example:
record Wrap : Set₂ where
field wrapped : Set₁
test : Wrap
test = record { M }
module M where
wrapped : Set₁
wrapped = Set
{{-
is now lexed as { {-
rather than {{ -
,
see Issue #3962.
Syntax for large numbers: you can now separate groups of 3 digits using _
.
e.g. write 1_000_000
instead of 1000000
.
quoteGoal
and quoteContext
are no longer keywords.
Record constructors can no longer be qualified by the record module. (See Issue #4189.)
record Foo : Set where
constructor foo
works = foo
fails = Foo.foo
codata
definitions have been removed from the concrete syntax
Previously they got accepted syntactically, but resulted in errors.
Imports can now be anonymous.
(See Issue_#3727.)
For example, the following will not bring Agda.Builtin.Unit
into scope:
open import Agda.Builtin.Unit as _
blah :: ⊤
blah = tt
For instance,
-- A.agda
module A where
record R : Set₁ where
field f : Set
-- B.agda
module B where
import A
-- C.agda
module C where
import B
fails : Set → _
fails X = record {f = X} -- import A required to infer record type R
New modality @♭/@flat
(previously only available in the branch "flat").
An idempotent comonadic modality modeled after spatial/crisp type theory.
See Flat Modality
in the documentation for more.
@0
/ @erased
).
Terms marked as erased cannot influence computations and are erased
at run time
[Issue #3855]. See
Run-time
Irrelevance
in the documentation for more information.Note that this feature can cause previously solved metavariables to become unsolved even in code that doesn't use run-time erasure (see issue #4174).
f : .A → A
, Agda no longer accepts f
at type A →
A
. Instead, Agda accepts λ x → f x : A → A
. The same holds for
erasure (@0
) and flat (@♭
) modalities. Consequently, it may be
required to eta-expand certain functions in order to make old code
work with Agda 2.6.1. Alternatively, enabling the new --subtyping
flag will restore the old behaviour but might negatively impact
typechecking performance.--cumulativity
When the --cumulativity
flag is enabled, Agda uses the subtyping
rule Set i =< Set j
whenever i =< j
. For example, in
addition to its usual type Set
, Nat
also has the type
Set₁
and even Set i
for any i : Level
. More information
about this new option can be found in section
Cumulativity
of the user manual.
with
are
no longer accepted as terminating. See
Issue #59 for why this
feature was originally introduced and
#3604 for why it had to
be removed.The easiest way to fix termination problems caused by with
is to abstract
over the offending recursive call before any other with
s. For example
data D : Set where
[_] : Nat → D
fails : D → Nat
fails [ zero ] = zero
fails [ suc n ] with some-stuff
... | _ = fails [ n ]
This fails termination because the relation between [ suc n ]
and [ n ]
is lost since the generated with-function only gets passed n
. To fix it we
can abstract over the recursive call:
fixed : D → Nat
fixed [ zero ] = zero
fixed [ suc n ] with fixed [ n ] | some-stuff
... | rec | _ = rec
If the function takes more arguments you might need to abstract over a partial application to just the structurally recursive argument. For instance,
fails : Nat → D → Nat
fails _ [ zero ] = zero
fails _ [ suc n ] with some-stuff
... | m = fails m [ n ]
fixed : Nat → D → Nat
fixed _ [ zero ] = zero
fixed _ [ suc n ] with (λ m → fixed m [ n ]) | some-stuff
... | rec | m = rec m
A possible complication is that later with
-abstractions might change the
type of the abstracted recursive call:
T : D → Set
suc-T : ∀ {n} → T [ n ] → T [ suc n ]
zero-T : T [ zero ]
fails : (d : D) → T d
fails [ zero ] = zero-T
fails [ suc n ] with some-stuff
... | _ with [ n ]
... | z = suc-T (fails [ n ])
still-fails : (d : D) → T d
still-fails [ zero ] = zero-T
still-fails [ suc n ] with still-fails [ n ] | some-stuff
... | rec | _ with [ n ]
... | z = suc-T rec -- Type error because rec : T z
To solve this problem you can add rec
to the with-abstraction messing up
its type. This will prevent it from having its type changed:
fixed : (d : D) → T d
fixed [ zero ] = zero-T
fixed [ suc n ] with fixed [ n ] | some-stuff
... | rec | _ with rec | [ n ]
... | _ | z = suc-T rec
For example, consider the following example using a dependent
coinductive record Tree
:
data Fin : Nat → Set where
fzero : ∀ n → Fin (suc n)
fsuc : ∀ n (i : Fin n) → Fin (suc n)
toNat : ∀ n → Fin n → Nat
toNat .(suc n) (fzero n) = zero
toNat .(suc n) (fsuc n i) = suc (toNat n i)
record Tree : Set where
coinductive
field label : Nat
child : Fin label → Tree
open Tree
tree : Nat → Tree
tree n .label = n
tree n .child i = tree (n + toNat _ i)
Agda solves the underscore by tree n .label
, which is a corecursive
call in a non-guarded position, violating the guardedness criterion.
This lead to a complaint of the termination checker.
Now this call is reduced to n
first using the non-recursive clause
tree n .label = n
, which leaves us only with the guarded call
tree (n + toNat n i)
, and the termination checker is happy.
Note: Similar false positives arose already for non-recursive dependent records, e.g., when trying to define an inhabitant of the Σ-type by copattern matching on the projects. See Issue_#2068 for a non-recursive example.
Agda will no longer reduce irrelevant definitions and definitions
with a type in Prop
. This does not have an effect on the
semantics, but should lead to improved performance (see Issues
#4115,
#4118,
#4120,
#4122).
Terms of a type in Prop
are now printed as _
. To show the actual
term, you can use the --show-irrelevant
flag (see
Issue #3337.
--rewriting
) with data or record types as
the head symbol are no longer allowed (see
Issue #3846).You can declare tactics to be used to solve a particular implicit argument using the following syntax:
example : {@(tactic f) x : A} → B
where f : Term → TC ⊤
. At calls to example
, f
is called on the
metavariable inserted for x
. f
can be an arbitrary term and may depend on
previous arguments to the function. For instance,
example₂ : (depth : Nat) {@(tactic search depth) x : A} → B
Record fields can also be annotated with a tactic, allowing them to be omitted in constructor applications, record constructions and co-pattern matches:
record Example : Set where
constructor mkExample
field x : A
@(tactic solveP x) {y} : P x
where solveP : (x : A) → Term → TC ⊤
is a tactic that tries to
prove P x
[Issue #4124].
quoteGoal
and quoteContext
has been
removed. primWord64ToNatInjective : ∀ a b → primWord64ToNat a ≡ primWord64ToNat b → a ≡ b
primFloatToWord64 : Float → Word64
primFloatToWord64Injective : ∀ a b → primFloatToWord64 a ≡ primFloatToWord64 b → a ≡ b
primMetaToNat : Meta → Nat
primMetaToNatInjective : ∀ a b → primMetaToNat a ≡ primMetaToNat b → a ≡ b
primQNameToWord64s : Name → Word64 × Word64
primQNameToWord64sInjective : ∀ a b → primQNameToWord64s a ≡ primQNameToWord64s b → a ≡ b
These can be used to define safe decidable propositional equality, see Issue agda-stdlib#698.
primShowNat : Nat → String
placed in Agda.Builtin.String.
IO
has been declared strictly positive in both its
level and type argument. f : {a : Level} {A : Set a} (a : A) → A -- warning raised: repeated a
g : {a : Level} {A : Set a} → (a : A) → A -- warning not raised: two distinct telescopes
Note that this warning is turned off by default (you can use
-WShadowingInTelescope
or --warning ShadowingInTelescope
to turn
it on, -Wall
would also naturally work).
Agda input method: new key bindings \ G h
and \ G H
for η
and
H
(capital η)
[Issue #3856].
Syntax highlighting: in literate modes, the pure texts (other than Agda code and the code-text separators) are no longer highlighted (it was highlighted as comments before). This somehow provides more information about how Agda lexes literate files.
Agda now also displays the values of let-bound variables in the context instead of just their types [Issue #4199].
Agda will now try to preserve the ellipsis (...
) during case
splitting when possible. To manually expand the ellipsis, you may
ask Agda to case split on the special identifier .
.
[Issue #2589]
Agda will now also show variables named _
in the context if they
are instance arguments (see
#4307). Instance
arguments are now also marked as (instance)
in the context. Example:
f : {{_ : A}} → A
f = ?
Agda will now display the goal as follows:
Goal: A
————————————————————————————————————————————————————————————
_ : A (instance)
It is now possible to ask Agda to terminate itself after any
previously invoked commands have completed, by giving a prefix
argument to agda2-term
.
The command agda2-measure-load-time
has been removed.
data I : Set where
bar : I
{-# FOREIGN GHC data I = Bar #-}
{-# COMPILE GHC I = data I (Bar) #-}
data S : Set where
foo : I → S
{-# FOREIGN GHC data S = Foo I #-}
{-# COMPILE GHC S = data S (Foo) #-}
Previously [Issue #2921],
the last binding was incorrect, since the argument of
singleton type I
was erased from the constructor foo
during
compilation. The required shape of S
was previously
{-# FOREIGN GHC data S = Foo #-}
i.e., constructor Foo
had to have no arguments.
For the sake of transparency, Haskell constructors bound to Agda constructors now take the same arguments. This is especially important if Haskell bindings are to be produced automatically by third party tool.
It is also possible to write, say, hide=true
instead of hide
,
and hide=false
means that the hide
option should not be used.
Furthermore the same option can be given multiple times, in which
case later choices take precedence over earlier ones.
number
.When the option number
is used an equation number is generated for
the code listing. The number is set to the right, centered
vertically. By default the number is set in parentheses, but this
can be changed by redefining \AgdaFormatCodeNumber
.
The option can optionally be given an argument: when number=l
is
used a label l
, referring to the code listing, is generated. It is
possible to use this option several times with different labels.
The option has no effect if used together with hide
, inline
or
inline*
.
Removed module Agda.Utils.HashMap
. It only re-exported Data.HashMap.Strict
from the package unordered-containers
. Use Data.HashMap.Strict
instead.
Removed module Agda.Utils.Char
. It used to provide functions converting a
Char
in base 8, 10, and 16 to the corresponding Int
. Use digitToInt
in
Data.Char
instead. The rest of module was about Unicode test which was not
used.
Agda.Utils.List
no longer provides headMaybe
.
Use listToMaybe
in Data.Maybe
instead.
Agda.Utils.Either
no longer provides mapEither
. Use bimap
in
Data.Bifunctor
instead.
Agda.Utils.Map
no longer provides unionWithM
, insertWithKeyM
,
allWithKey
, unzip
, and unzip3
.
For 2.6.1, the following issues were also closed (see bug tracker):
stack exec
for GHC backendrenaming
clause--no-unicode
option producing unicode variable namesdata .. where
definitionsimport … hiding …
should be documentedHiding … has no effect
should be improved~/.agda
?with
abstraction fails with HIT constructors in the goalThe following previously closed issues were reopened: