Releases: leanprover/lean4
v4.7.0
Changes since v4.6.0 (from RELEASES.md)
-
simp
andrw
now use instance arguments found by unification,
rather than always resynthesizing. For backwards compatibility, the original behaviour is
available viaset_option tactic.skipAssignedInstances false
.
#3507 and
#3509. -
When the
pp.proofs
is false, now omitted proofs use⋯
rather than_
,
which gives a more helpful error message when copied from the Infoview.
Thepp.proofs.threshold
option lets small proofs always be pretty printed.
#3241. -
pp.proofs.withType
is now set to false by default to reduce noise in the info view. -
The pretty printer for applications now handles the case of over-application itself when applying app unexpanders.
In particular, the| `($_ $a $b $xs*) => `(($a + $b) $xs*)
case of anapp_unexpander
is no longer necessary.
#3495. -
New
simp
(anddsimp
) configuration option:zetaDelta
. It isfalse
by default.
Thezeta
option is stilltrue
by default, but their meaning has changed.- When
zeta := true
,simp
anddsimp
reduce terms of the form
let x := val; e[x]
intoe[val]
. - When
zetaDelta := true
,simp
anddsimp
will expand let-variables in
the context. For example, suppose the context containsx := val
. Then,
any occurrence ofx
is replaced withval
.
See issue #2682 for additional details. Here are some examples:
example (h : z = 9) : let x := 5; let y := 4; x + y = z := by intro x simp /- New goal: h : z = 9; x := 5 |- x + 4 = z -/ rw [h] example (h : z = 9) : let x := 5; let y := 4; x + y = z := by intro x -- Using both `zeta` and `zetaDelta`. simp (config := { zetaDelta := true }) /- New goal: h : z = 9; x := 5 |- 9 = z -/ rw [h] example (h : z = 9) : let x := 5; let y := 4; x + y = z := by intro x simp [x] -- asks `simp` to unfold `x` /- New goal: h : z = 9; x := 5 |- 9 = z -/ rw [h] example (h : z = 9) : let x := 5; let y := 4; x + y = z := by intro x simp (config := { zetaDelta := true, zeta := false }) /- New goal: h : z = 9; x := 5 |- let y := 4; 5 + y = z -/ rw [h]
- When
-
When adding new local theorems to
simp
, the system assumes that the function application arguments
have been annotated withno_index
. This modification, which addresses issue #2670,
restores the Lean 3 behavior that users expect. With this modification, the following examples are now operational:example {α β : Type} {f : α × β → β → β} (h : ∀ p : α × β, f p p.2 = p.2) (a : α) (b : β) : f (a, b) b = b := by simp [h] example {α β : Type} {f : α × β → β → β} (a : α) (b : β) (h : f (a,b) (a,b).2 = (a,b).2) : f (a, b) b = b := by simp [h]
In both cases,
h
is applicable becausesimp
does not index f-arguments anymore when addingh
to thesimp
-set.
It's important to note, however, that global theorems continue to be indexed in the usual manner. -
Improved the error messages produced by the
decide
tactic. #3422 -
Improved auto-completion performance. #3460
-
Improved initial language server startup performance. #3552
-
Changed call hierarchy to sort entries and strip private header from names displayed in the call hierarchy. #3482
-
There is now a low-level error recovery combinator in the parsing framework, primarily intended for DSLs. #3413
-
You can now write
termination_by?
after a declaration to see the automatically inferred
termination argument, and turn it into atermination_by …
clause using the “Try this” widget or a code action. #3514 -
A large fraction of
Std
has been moved into the Lean repository.
This was motivated by:- Making universally useful tactics such as
ext
,by_cases
,change at
,
norm_cast
,rcases
,simpa
,simp?
,omega
, andexact?
available to all users of Lean, without imports. - Minimizing the syntactic changes between plain Lean and Lean with
import Std
. - Simplifying the development process for the basic data types
Nat
,Int
,Fin
(and variants such asUInt64
),List
,Array
,
andBitVec
as we begin making the APIs and simp normal forms for these types
more complete and consistent. - Laying the groundwork for the Std roadmap, as a library focused on
essential datatypes not provided by the core langauge (e.g.RBMap
)
and utilities such as basic IO.
While we have achieved most of our initial aims inv4.7.0-rc1
,
some upstreaming will continue over the coming months.
- Making universally useful tactics such as
-
The
/
and%
notations inInt
now useInt.ediv
andInt.emod
(i.e. the rounding conventions have changed).
PreviouslyStd
overrode these notations, so this is no change for users ofStd
.
There is now kernel support for these functions.
#3376. -
omega
, our integer linear arithmetic tactic, is now availabe in the core langauge.- It is supplemented by a preprocessing tactic
bv_omega
which can solve goals aboutBitVec
which naturally translate into linear arithmetic problems.
#3435. omega
now has support forFin
#3427,
the<<<
operator #3433.- During the port
omega
was modified to no longer identify atoms up to definitional equality
(so in particular it can no longer proveid x ≤ x
). #3525.
This may cause some regressions.
We plan to provide a general purpose preprocessing tactic later, or anomega!
mode. omega
is now invoked in Lean's automation for termination proofs
#3503 as well as in
array indexing proofs #3515.
This automation will be substantially revised in the medium term,
and whileomega
does help automate some proofs, we plan to make this much more robust.
- It is supplemented by a preprocessing tactic
-
The library search tactics
exact?
andapply?
that were originally in
Mathlib are now available in Lean itself. These use the implementation using
lazy discrimination trees fromStd
, and thus do not require a disk cache but
have a slightly longer startup time. The order used for selection lemmas has
changed as well to favor goals purely based on how many terms in the head
pattern match the current goal. -
The
solve_by_elim
tactic has been ported fromStd
to Lean so that library
search can use it. -
New
#check_tactic
and#check_simp
commands have been added. These are
useful for checking tactics (particularlysimp
) behave as expected in test
suites. -
Previously, app unexpanders would only be applied to entire applications. However, some notations produce
functions, and these functions can be given additional arguments. The solution so far has been to write app unexpanders so that they can take an arbitrary number of additional arguments. However this leads to misleading hover information in the Infoview. For example, whileHAdd.hAdd f g 1
pretty prints as(f + g) 1
, hovering overf + g
showsf
. There is no way to fix the situation from within an app unexpander; the expression position forHAdd.hAdd f g
is absent, and app unexpanders cannot register TermInfo.This commit changes the app delaborator to try running app unexpanders on every prefix of an application, from longest to shortest prefix. For efficiency, it is careful to only try this when app delaborators do in fact exist for the head constant, and it also ensures arguments are only delaborated once. Then, in
(f + g) 1
, thef + g
gets TermInfo registered for that subexpression, making it properly hoverable.
Breaking changes:
Lean.withTraceNode
and variants got a strongerMonadAlwaysExcept
assumption to
fix trace trees not being built on elaboration runtime exceptions. Instances for most elaboration
monads built onEIO Exception
should be synthesized automatically.- The
match ... with.
andfun.
notations previously in Std have been replaced by
nomatch ...
andnofun
. #3279 and #3286
Other improvements:
- several bug fixes for
simp
: - fixes for
match
expressions:- fix regression with builtin literals #3521...
v4.7.0-rc2
Identical to v4.7.0-rc1
, but with the backport of #3610, fixing a bug when exact?
is used in Mathlib.
v4.7.0-rc1
Changes since v4.6.0 (from RELEASES.md)
-
simp
andrw
now use instance arguments found by unification,
rather than always resynthesizing. For backwards compatibility, the original behaviour is
available viaset_option tactic.skipAssignedInstances false
.
#3507 and
#3509. -
When the
pp.proofs
is false, now omitted proofs use⋯
rather than_
,
which gives a more helpful error message when copied from the Infoview.
Thepp.proofs.threshold
option lets small proofs always be pretty printed.
#3241. -
pp.proofs.withType
is now set to false by default to reduce noise in the info view. -
The pretty printer for applications now handles the case of over-application itself when applying app unexpanders.
In particular, the| `($_ $a $b $xs*) => `(($a + $b) $xs*)
case of anapp_unexpander
is no longer necessary.
#3495. -
New
simp
(anddsimp
) configuration option:zetaDelta
. It isfalse
by default.
Thezeta
option is stilltrue
by default, but their meaning has changed.- When
zeta := true
,simp
anddsimp
reduce terms of the form
let x := val; e[x]
intoe[val]
. - When
zetaDelta := true
,simp
anddsimp
will expand let-variables in
the context. For example, suppose the context containsx := val
. Then,
any occurrence ofx
is replaced withval
.
See issue #2682 for additional details. Here are some examples:
example (h : z = 9) : let x := 5; let y := 4; x + y = z := by intro x simp /- New goal: h : z = 9; x := 5 |- x + 4 = z -/ rw [h] example (h : z = 9) : let x := 5; let y := 4; x + y = z := by intro x -- Using both `zeta` and `zetaDelta`. simp (config := { zetaDelta := true }) /- New goal: h : z = 9; x := 5 |- 9 = z -/ rw [h] example (h : z = 9) : let x := 5; let y := 4; x + y = z := by intro x simp [x] -- asks `simp` to unfold `x` /- New goal: h : z = 9; x := 5 |- 9 = z -/ rw [h] example (h : z = 9) : let x := 5; let y := 4; x + y = z := by intro x simp (config := { zetaDelta := true, zeta := false }) /- New goal: h : z = 9; x := 5 |- let y := 4; 5 + y = z -/ rw [h]
- When
-
When adding new local theorems to
simp
, the system assumes that the function application arguments
have been annotated withno_index
. This modification, which addresses issue #2670,
restores the Lean 3 behavior that users expect. With this modification, the following examples are now operational:example {α β : Type} {f : α × β → β → β} (h : ∀ p : α × β, f p p.2 = p.2) (a : α) (b : β) : f (a, b) b = b := by simp [h] example {α β : Type} {f : α × β → β → β} (a : α) (b : β) (h : f (a,b) (a,b).2 = (a,b).2) : f (a, b) b = b := by simp [h]
In both cases,
h
is applicable becausesimp
does not index f-arguments anymore when addingh
to thesimp
-set.
It's important to note, however, that global theorems continue to be indexed in the usual manner. -
Improved the error messages produced by the
decide
tactic. #3422 -
Improved auto-completion performance. #3460
-
Improved initial language server startup performance. #3552
-
Changed call hierarchy to sort entries and strip private header from names displayed in the call hierarchy. #3482
-
There is now a low-level error recovery combinator in the parsing framework, primarily intended for DSLs. #3413
-
You can now write
termination_by?
after a declaration to see the automatically inferred
termination argument, and turn it into atermination_by …
clause using the “Try this” widget or a code action. #3514 -
A large fraction of
Std
has been moved into the Lean repository.
This was motivated by:- Making universally useful tactics such as
ext
,by_cases
,change at
,
norm_cast
,rcases
,simpa
,simp?
,omega
, andexact?
available to all users of Lean, without imports. - Minimizing the syntactic changes between plain Lean and Lean with
import Std
. - Simplifying the development process for the basic data types
Nat
,Int
,Fin
(and variants such asUInt64
),List
,Array
,
andBitVec
as we begin making the APIs and simp normal forms for these types
more complete and consistent. - Laying the groundwork for the Std roadmap, as a library focused on
essential datatypes not provided by the core langauge (e.g.RBMap
)
and utilities such as basic IO.
While we have achieved most of our initial aims inv4.7.0-rc1
,
some upstreaming will continue over the coming months.
- Making universally useful tactics such as
-
The
/
and%
notations inInt
now useInt.ediv
andInt.emod
(i.e. the rounding conventions have changed).
PreviouslyStd
overrode these notations, so this is no change for users ofStd
.
There is now kernel support for these functions.
#3376. -
omega
, our integer linear arithmetic tactic, is now availabe in the core langauge.- It is supplemented by a preprocessing tactic
bv_omega
which can solve goals aboutBitVec
which naturally translate into linear arithmetic problems.
#3435. omega
now has support forFin
#3427,
the<<<
operator #3433.- During the port
omega
was modified to no longer identify atoms up to definitional equality
(so in particular it can no longer proveid x ≤ x
). #3525.
This may cause some regressions.
We plan to provide a general purpose preprocessing tactic later, or anomega!
mode. omega
is now invoked in Lean's automation for termination proofs
#3503 as well as in
array indexing proofs #3515.
This automation will be substantially revised in the medium term,
and whileomega
does help automate some proofs, we plan to make this much more robust.
- It is supplemented by a preprocessing tactic
-
The library search tactics
exact?
andapply?
that were originally in
Mathlib are now available in Lean itself. These use the implementation using
lazy discrimination trees fromStd
, and thus do not require a disk cache but
have a slightly longer startup time. The order used for selection lemmas has
changed as well to favor goals purely based on how many terms in the head
pattern match the current goal. -
The
solve_by_elim
tactic has been ported fromStd
to Lean so that library
search can use it. -
New
#check_tactic
and#check_simp
commands have been added. These are
useful for checking tactics (particularlysimp
) behave as expected in test
suites. -
Previously, app unexpanders would only be applied to entire applications. However, some notations produce
functions, and these functions can be given additional arguments. The solution so far has been to write app unexpanders so that they can take an arbitrary number of additional arguments. However this leads to misleading hover information in the Infoview. For example, whileHAdd.hAdd f g 1
pretty prints as(f + g) 1
, hovering overf + g
showsf
. There is no way to fix the situation from within an app unexpander; the expression position forHAdd.hAdd f g
is absent, and app unexpanders cannot register TermInfo.This commit changes the app delaborator to try running app unexpanders on every prefix of an application, from longest to shortest prefix. For efficiency, it is careful to only try this when app delaborators do in fact exist for the head constant, and it also ensures arguments are only delaborated once. Then, in
(f + g) 1
, thef + g
gets TermInfo registered for that subexpression, making it properly hoverable.
Breaking changes:
Lean.withTraceNode
and variants got a strongerMonadAlwaysExcept
assumption to
fix trace trees not being built on elaboration runtime exceptions. Instances for most elaboration
monads built onEIO Exception
should be synthesized automatically.- The
match ... with.
andfun.
notations previously in Std have been replaced by
nomatch ...
andnofun
. #3279 and #3286
Other improvements:
- several bug fixes for
simp
: - fixes for
match
expressions:- fix regression with builtin literals #3521...
v4.6.1
v4.6.0
Changes since v4.5.0 (from RELEASES.md)
-
Add custom simplification procedures (aka
simproc
s) tosimp
. Simprocs can be triggered by the simplifier on a specified term-pattern. Here is an small example:import Lean.Meta.Tactic.Simp.BuiltinSimprocs.Nat def foo (x : Nat) : Nat := x + 10 /-- The `simproc` `reduceFoo` is invoked on terms that match the pattern `foo _`. -/ simproc reduceFoo (foo _) := /- A term of type `Expr → SimpM Step -/ fun e => do /- The `Step` type has three constructors: `.done`, `.visit`, `.continue`. * The constructor `.done` instructs `simp` that the result does not need to be simplied further. * The constructor `.visit` instructs `simp` to visit the resulting expression. * The constructor `.continue` instructs `simp` to try other simplification procedures. All three constructors take a `Result`. The `.continue` contructor may also take `none`. `Result` has two fields `expr` (the new expression), and `proof?` (an optional proof). If the new expression is definitionally equal to the input one, then `proof?` can be omitted or set to `none`. -/ /- `simp` uses matching modulo reducibility. So, we ensure the term is a `foo`-application. -/ unless e.isAppOfArity ``foo 1 do return .continue /- `Nat.fromExpr?` tries to convert an expression into a `Nat` value -/ let some n ← Nat.fromExpr? e.appArg! | return .continue return .done { expr := Lean.mkNatLit (n+10) }
We disable simprocs support by using the command
set_option simprocs false
. This command is particularly useful when porting files to v4.6.0.
Simprocs can be scoped, manually added tosimp
commands, and suppressed using-
. They are also supported bysimp?
.simp only
does not execute anysimproc
. Here are some examples for thesimproc
defined above.example : x + foo 2 = 12 + x := by set_option simprocs false in /- This `simp` command does not make progress since `simproc`s are disabled. -/ fail_if_success simp simp_arith example : x + foo 2 = 12 + x := by /- `simp only` must not use the default simproc set. -/ fail_if_success simp only simp_arith example : x + foo 2 = 12 + x := by /- `simp only` does not use the default simproc set, but we can provide simprocs as arguments. -/ simp only [reduceFoo] simp_arith example : x + foo 2 = 12 + x := by /- We can use `-` to disable `simproc`s. -/ fail_if_success simp [-reduceFoo] simp_arith
The command
register_simp_attr <id>
now creates asimp
and asimproc
set with the name<id>
. The following command instructs Lean to insert thereduceFoo
simplification procedure into the setmy_simp
. If no set is specified, Lean uses the defaultsimp
set.simproc [my_simp] reduceFoo (foo _) := ...
-
The syntax of the
termination_by
anddecreasing_by
termination hints is overhauled:- They are now placed directly after the function they apply to, instead of
after the wholemutual
block. - Therefore, the function name no longer has to be mentioned in the hint.
- If the function has a
where
clause, thetermination_by
and
decreasing_by
for that function come before thewhere
. The
functions in thewhere
clause can have their own termination hints, each
following the corresponding definition. - The
termination_by
clause can only bind “extra parameters”, that are not
already bound by the function header, but are bound in a lambda (:= fun x y z =>
) or in patterns (| x, n + 1 => …
). These extra parameters used to
be understood as a suffix of the function parameters; now it is a prefix.
Migration guide: In simple cases just remove the function name, and any
variables already bound at the header.def foo : Nat → Nat → Nat := … -termination_by foo a b => a - b +termination_by a b => a - b
or
def foo : Nat → Nat → Nat := … -termination_by _ a b => a - b +termination_by a b => a - b
If the parameters are bound in the function header (before the
:
), remove them as well:def foo (a b : Nat) : Nat := … -termination_by foo a b => a - b +termination_by a - b
Else, if there are multiple extra parameters, make sure to refer to the right
ones; the bound variables are interpreted from left to right, no longer from
right to left:def foo : Nat → Nat → Nat → Nat | a, b, c => … -termination_by foo b c => b +termination_by a b => b
In the case of a
mutual
block, place the termination arguments (without the
function name) next to the function definition:-mutual -def foo : Nat → Nat → Nat := … -def bar : Nat → Nat := … -end -termination_by - foo a b => a - b - bar a => a +mutual +def foo : Nat → Nat → Nat := … +termination_by a b => a - b +def bar : Nat → Nat := … +termination_by a => a +end
Similarly, if you have (mutual) recursion through
where
orlet rec
, the
termination hints are now placed directly after the function they apply to:-def foo (a b : Nat) : Nat := … - where bar (x : Nat) : Nat := … -termination_by - foo a b => a - b - bar x => x +def foo (a b : Nat) : Nat := … +termination_by a - b + where + bar (x : Nat) : Nat := … + termination_by x -def foo (a b : Nat) : Nat := - let rec bar (x : Nat) : Nat := … - … -termination_by - foo a b => a - b - bar x => x +def foo (a b : Nat) : Nat := + let rec bar (x : Nat) : Nat := … + termination_by x + … +termination_by a - b
In cases where a single
decreasing_by
clause applied to multiple mutually
recursive functions before, the tactic now has to be duplicated. - They are now placed directly after the function they apply to, instead of
-
The semantics of
decreasing_by
changed; the tactic is applied to all
termination proof goals together, not individually.This helps when writing termination proofs interactively, as one can focus
each subgoal individually, for example using·
. Previously, the given
tactic script had to work for all goals, and one had to resort to tactic
combinators likefirst
:def foo (n : Nat) := … foo e1 … foo e2 … -decreasing_by -simp_wf -first | apply something_about_e1; … - | apply something_about_e2; … +decreasing_by +all_goals simp_wf +· apply something_about_e1; … +· apply something_about_e2; …
To obtain the old behaviour of applying a tactic to each goal individually,
useall_goals
:def foo (n : Nat) := … -decreasing_by some_tactic +decreasing_by all_goals some_tactic
In the case of mutual recursion each
decreasing_by
now applies to just its
function. If some functions in a recursive group do not have their own
decreasing_by
, the defaultdecreasing_tactic
is used. If the same tactic
ought to be applied to multiple functions, thedecreasing_by
clause has to
be repeated at each of these functions. -
Modify
InfoTree.context
to facilitate augmenting it with partial contexts while elaborating a command. This breaks backwards compatibility with all downstream projects that traverse theInfoTree
manually instead of going through the functions inInfoUtils.lean
, as well as those manually creating and savingInfoTree
s. See PR #3159 for how to migrate your code. -
Add language server support for call hierarchy requests (PR #3082). The change to the .ilean format in this PR means that projects must be fully rebuilt once in order to generate .ilean files with the new format before features like "find references" work correctly again.
-
Structure instances with multiple sources (for example
{a, b, c with x := 0}
) now have their fields filled from these sources
in strict left-to-right order. Furthermore, the structure instance elaborator now aggressively use sources to fill in subobject
fields, which prevents unnecessary eta expansion of the sources,
and hence greatly reduces the reliance on costly structure eta reduction. This has a large impact on mathlib,
reducing total CPU instructions by 3% and enabling impactful refactors like leanprover-community/mathlib4#8386
which reduces the build time by almost 20%.
See PR #2478 and RFC #2451. -
Add pretty printer settings to omit deeply nested terms (
pp.deepTerms false
andpp.deepTerms.threshold
) (PR #3201) -
Add pretty printer options
pp.numeralTypes
andpp.natLit
.
Whenpp.numeralTypes
is true, then natural number literals, integer literals, and rational number literals
are pretty printed with type ascriptions, such as(2 : Rat)
,(-2 : Rat)
, and(-2 / 3 : Rat)
.
Whenpp.natLit
is true, then raw natural number literals are pretty printed asnat_lit 2
.
PR #2933 and RFC #3021.
Lake updates:
Other improvements:
v4.6.0-rc1
- Add custom simplification procedures (aka
simproc
s) tosimp
. Simprocs can be triggered by the simplifier on a specified term-pattern. Here is an small example:
import Lean.Meta.Tactic.Simp.BuiltinSimprocs.Nat
def foo (x : Nat) : Nat :=
x + 10
/--
The `simproc` `reduceFoo` is invoked on terms that match the pattern `foo _`.
-/
simproc reduceFoo (foo _) :=
/- A term of type `Expr → SimpM Step -/
fun e => do
/-
The `Step` type has three constructors: `.done`, `.visit`, `.continue`.
* The constructor `.done` instructs `simp` that the result does
not need to be simplied further.
* The constructor `.visit` instructs `simp` to visit the resulting expression.
* The constructor `.continue` instructs `simp` to try other simplification procedures.
All three constructors take a `Result`. The `.continue` contructor may also take `none`.
`Result` has two fields `expr` (the new expression), and `proof?` (an optional proof).
If the new expression is definitionally equal to the input one, then `proof?` can be omitted or set to `none`.
-/
/- `simp` uses matching modulo reducibility. So, we ensure the term is a `foo`-application. -/
unless e.isAppOfArity ``foo 1 do
return .continue
/- `Nat.fromExpr?` tries to convert an expression into a `Nat` value -/
let some n ← Nat.fromExpr? e.appArg!
| return .continue
return .done { expr := Lean.mkNatLit (n+10) }
We disable simprocs support by using the command set_option simprocs false
. This command is particularly useful when porting files to v4.6.0.
Simprocs can be scoped, manually added to simp
commands, and suppressed using -
. They are also supported by simp?
. simp only
does not execute any simproc
. Here are some examples for the simproc
defined above.
example : x + foo 2 = 12 + x := by
set_option simprocs false in
/- This `simp` command does not make progress since `simproc`s are disabled. -/
fail_if_success simp
simp_arith
example : x + foo 2 = 12 + x := by
/- `simp only` must not use the default simproc set. -/
fail_if_success simp only
simp_arith
example : x + foo 2 = 12 + x := by
/-
`simp only` does not use the default simproc set,
but we can provide simprocs as arguments. -/
simp only [reduceFoo]
simp_arith
example : x + foo 2 = 12 + x := by
/- We can use `-` to disable `simproc`s. -/
fail_if_success simp [-reduceFoo]
simp_arith
The command register_simp_attr <id>
now creates a simp
and a simproc
set with the name <id>
. The following command instructs Lean to insert the reduceFoo
simplification procedure into the set my_simp
. If no set is specified, Lean uses the default simp
set.
simproc [my_simp] reduceFoo (foo _) := ...
-
The syntax of the
termination_by
anddecreasing_by
termination hints is overhauled:- They are now placed directly after the function they apply to, instead of
after the wholemutual
block. - Therefore, the function name no longer has to be mentioned in the hint.
- If the function has a
where
clause, thetermination_by
and
decreasing_by
for that function come before thewhere
. The
functions in thewhere
clause can have their own termination hints, each
following the corresponding definition. - The
termination_by
clause can only bind “extra parameters”, that are not
already bound by the function header, but are bound in a lambda (:= fun x y z =>
) or in patterns (| x, n + 1 => …
). These extra parameters used to
be understood as a suffix of the function parameters; now it is a prefix.
Migration guide: In simple cases just remove the function name, and any
variables already bound at the header.def foo : Nat → Nat → Nat := … -termination_by foo a b => a - b +termination_by a b => a - b
or
def foo : Nat → Nat → Nat := … -termination_by _ a b => a - b +termination_by a b => a - b
If the parameters are bound in the function header (before the
:
), remove them as well:def foo (a b : Nat) : Nat := … -termination_by foo a b => a - b +termination_by a - b
Else, if there are multiple extra parameters, make sure to refer to the right
ones; the bound variables are interpreted from left to right, no longer from
right to left:def foo : Nat → Nat → Nat → Nat | a, b, c => … -termination_by foo b c => b +termination_by a b => b
In the case of a
mutual
block, place the termination arguments (without the
function name) next to the function definition:-mutual -def foo : Nat → Nat → Nat := … -def bar : Nat → Nat := … -end -termination_by - foo a b => a - b - bar a => a +mutual +def foo : Nat → Nat → Nat := … +termination_by a b => a - b +def bar : Nat → Nat := … +termination_by a => a +end
Similarly, if you have (mutual) recursion through
where
orlet rec
, the
termination hints are now placed directly after the function they apply to:-def foo (a b : Nat) : Nat := … - where bar (x : Nat) : Nat := … -termination_by - foo a b => a - b - bar x => x +def foo (a b : Nat) : Nat := … +termination_by a - b + where + bar (x : Nat) : Nat := … + termination_by x -def foo (a b : Nat) : Nat := - let rec bar (x : Nat) : Nat := … - … -termination_by - foo a b => a - b - bar x => x +def foo (a b : Nat) : Nat := + let rec bar (x : Nat) : Nat := … + termination_by x + … +termination_by a - b
In cases where a single
decreasing_by
clause applied to multiple mutually
recursive functions before, the tactic now has to be duplicated. - They are now placed directly after the function they apply to, instead of
-
The semantics of
decreasing_by
changed; the tactic is applied to all
termination proof goals together, not individually.This helps when writing termination proofs interactively, as one can focus
each subgoal individually, for example using·
. Previously, the given
tactic script had to work for all goals, and one had to resort to tactic
combinators likefirst
:def foo (n : Nat) := … foo e1 … foo e2 … -decreasing_by -simp_wf -first | apply something_about_e1; … - | apply something_about_e2; … +decreasing_by +all_goals simp_wf +· apply something_about_e1; … +· apply something_about_e2; …
To obtain the old behaviour of applying a tactic to each goal individually,
useall_goals
:def foo (n : Nat) := … -decreasing_by some_tactic +decreasing_by all_goals some_tactic
In the case of mutual recursion each
decreasing_by
now applies to just its
function. If some functions in a recursive group do not have their own
decreasing_by
, the defaultdecreasing_tactic
is used. If the same tactic
ought to be applied to multiple functions, thedecreasing_by
clause has to
be repeated at each of these functions. -
Modify
InfoTree.context
to facilitate augmenting it with partial contexts while elaborating a command. This breaks backwards compatibility with all downstream projects that traverse theInfoTree
manually instead of going through the functions inInfoUtils.lean
, as well as those manually creating and savingInfoTree
s. See PR #3159 for how to migrate your code. -
Add language server support for call hierarchy requests (PR #3082). The change to the .ilean format in this PR means that projects must be fully rebuilt once in order to generate .ilean files with the new format before features like "find references" work correctly again.
-
Structure instances with multiple sources (for example
{a, b, c with x := 0}
) now have their fields filled from these sources
in strict left-to-right order. Furthermore, the structure instance elaborator now aggressively use sources to fill in subobject
fields, which prevents unnecessary eta expansion of the sources,
and hence greatly reduces the reliance on costly structure eta reduction. This has a large impact on mathlib,
reducing total CPU instructions by 3% and enabling impactful refactors like leanprover-community/mathlib4#8386
which reduces the build time by almost 20%.
See PR #2478 and RFC #2451. -
Add pretty printer settings to omit deeply nested terms (
pp.deepTerms false
andpp.deepTerms.threshold
) (PR #3201) -
Add pretty printer options
pp.numeralTypes
andpp.natLit
.
Whenpp.numeralTypes
is true, then natural number literals, integer literals, and rational number literals
are pretty printed with type ascriptions, such as(2 : Rat)
,(-2 : Rat)
, and(-2 / 3 : Rat)
.
Whenpp.natLit
is true, then raw natural number literals are pretty printed asnat_lit 2
.
PR #2933 and RFC #3021.
Lake updates:
Other improvements:
- make
intro
be aware oflet_fun
#3115 - produce simpler proof terms in
rw
#3121 - fuse nested
mkCongrArg
calls in proofs generated bysimp
#3203 induction using
followed by a general term [#3188](https://github....
v4.5.0
Changes since v4.4.0 (from RELEASES.md)
-
Modify the lexical syntax of string literals to have string gaps, which are escape sequences of the form
"\" newline whitespace*
.
These have the interpetation of an empty string and allow a string to flow across multiple lines without introducing additional whitespace.
The following is equivalent to"this is a string"
."this is \ a string"
-
Add raw string literal syntax. For example,
r"\n"
is equivalent to"\\n"
, with no escape processing.
To include double quote characters in a raw string one can add sufficiently many#
characters before and after
the bounding"
s, as inr#"the "the" is in quotes"#
for"the \"the\" is in quotes"
.
PR #2929 and issue #1422. -
The low-level
termination_by'
clause is no longer supported.Migration guide: Use
termination_by
instead, e.g.:-termination_by' measure (fun ⟨i, _⟩ => as.size - i) +termination_by i _ => as.size - i
If the well-founded relation you want to use is not the one that the
WellFoundedRelation
type class would infer for your termination argument,
you can useWellFounded.wrap
from the std libarary to explicitly give one:-termination_by' ⟨r, hwf⟩ +termination_by x => hwf.wrap x
-
Support snippet edits in LSP
TextEdit
s. SeeLean.Lsp.SnippetString
for more details. -
Deprecations and changes in the widget API.
Widget.UserWidgetDefinition
is deprecated in favour ofWidget.Module
. The annotation@[widget]
is deprecated in favour of@[widget_module]
. To migrate a definition of typeUserWidgetDefinition
, remove thename
field and replace the type withWidget.Module
. Removing thename
results in a title bar no longer being drawn above your panel widget. To add it back, draw it as part of the component using<details open=true><summary class='mv2 pointer'>{name}</summary>{rest_of_widget}</details>
. See an example migration here.- The new command
show_panel_widgets
allows displaying always-on and locally-on panel widgets. RpcEncodable
widget props can now be stored in the infotree.- See RFC 2963 for more details and motivation.
-
If no usable lexicographic order can be found automatically for a termination proof, explain why.
See feat: GuessLex: if no measure is found, explain why. -
Option to print inferred termination argument.
Withset_option showInferredTerminationBy true
you will get messages likeInferred termination argument: termination_by ackermann n m => (sizeOf n, sizeOf m)
for automatically generated
termination_by
clauses. -
More detailed error messages for invalid mutual blocks.
-
Multiple improvements to the output of
simp?
andsimp_all?
. -
Tactics with
withLocation *
no longer fail if they close the main goal. -
Implementation of a
test_extern
command for writing tests for@[extern]
and@[implemented_by]
functions.
Usage isimport Lean.Util.TestExtern test_extern Nat.add 17 37
The head symbol must be the constant with the
@[extern]
or@[implemented_by]
attribute. The return type must have aDecidableEq
instance.
Bug fixes for
#2853, #2953, #2966,
#2971, #2990, #3094.
Bug fix for eager evaluation of default value in Option.getD
.
Avoid panic in leanPosToLspPos
when file source is unavailable.
Improve short-circuiting behavior for List.all
and List.any
.
Several Lake bug fixes: #3036, #3064, #3069.
What's Changed
- doc: Improve docstrings around Array.mk,.data,.toList by @nomeata in #2771
- chore: make PR title check work as a merge_group check by @nomeata in #2987
- chore: trim CI set by default by @nomeata in #2986
- chore: update stage0 by @mhuisi in #2992
- test: termination checking and duplicated terms by @nomeata in #2993
- refactor: CasesOnApp.refineThrough can return a lambda, not an open term by @nomeata in #2974
- chore: begin development cycle for v4.5.0 by @semorrison in #2995
- doc: VS Code dev setup by @Kha in #2961
- doc: fix typos by @marcusrossel in #2996
- fix: find macOS system libraries in leanc by @Kha in #2997
- chore: CI: in quick mode, only Nix build runs the tests by @nomeata in #2998
- chore: CI: update github-script by @Kha in #3002
- chore: run CI on new labels by @nomeata in #3003
- chore: fix CPP warnings about
static_assert
by @nomeata in #3005 - chore: improve tests/lean/copy-produced by @nomeata in #3006
- fix: missing
whnf
inmkBelowBinder
andmkMotiveBinder
by @arthur-adjedj in #2991 - doc: fix recent issue links in RELEASES.md by @tydeu in #3000
- test: expand tests/lean/issue2981.lean a bit by @nomeata in #3007
- chore: run tests with full-ci by @nomeata in #3009
- refactor: rewrite TerminationHint elaborators by @nomeata in #2958
- chore: Nix CI: update setup by @Kha in #3015
- doc: In testing doc, suggest
make
to pick up new tests by @nomeata in #2815 - chore: add vscode cmake configuration by @eric-wieser in #3008
- doc: typo Runnign by @nomeata in #3018
- refactor: WF.Fix: gather subgoals by @nomeata in #3017
- feat: GuessLex: if no measure is found, explain why by @nomeata in #2960
- feat: GuessLex: print inferred termination argument by @nomeata in #3012
- doc: add migration guide for per-package server options by @mhuisi in #3025
- feat: detail error message about invalid mutual blocks by @odanoburu in #2949
- doc: fix MetavarContext markdown by @nomeata in #3026
- chore: set
warningAsError
in CI only by @Kha in #3030 - fix: WF.Fix: deduplicate subsumed goals before running tactic by @nomeata in #3024
- feat: string gaps for continuing string literals across multiple lines by @kmill in #2821
- fix: GuessLex: deduplicate recursive calls by @nomeata in #3004
- doc: document constructors of
TransparencyMode
by @eric-wieser in #3037 - chore: update stage0 by @nomeata in #3041
- chore: remove obsolete comment in test by @nomeata in #3044
- doc: explain how to use custom lexers in the latest minted by @eric-wieser in #3047
- chore: stage0 autoupdater action by @nomeata in #3042
- fix: Option.getD eagerly evaluates dflt by @digama0 in #3043
- feat: drop support for termination_by' by @nomeata in #3033
- chore: make List.all and List.any short-circuit by @semorrison in #2972
- chore: fix superfluous lemmas in simp.trace by @semorrison in #2923
- fix: omit fvars from simp_all? theorem list by @JLimperg in #2969
- chore: refactor pr release workflow by @nomeata in #3020
- chore: withLocation * should not fail if it closes the main goal by @semorrison in #2917
- fix: move Lean.List.toSMap to List.toSMap by @digama0 in #3035
- fix:
run_task
/deactivate_task
race condition on `m...
v4.5.0-rc1
Changes since v4.4.0
(from RELEASE.md):
-
Modify the lexical syntax of string literals to have string gaps, which are escape sequences of the form
"\" newline whitespace*
.
These have the interpetation of an empty string and allow a string to flow across multiple lines without introducing additional whitespace.
The following is equivalent to"this is a string"
."this is \ a string"
-
Add raw string literal syntax. For example,
r"\n"
is equivalent to"\\n"
, with no escape processing.
To include double quote characters in a raw string one can add sufficiently many#
characters before and after
the bounding"
s, as inr#"the "the" is in quotes"#
for"the \"the\" is in quotes"
.
PR #2929 and issue #1422. -
The low-level
termination_by'
clause is no longer supported.Migration guide: Use
termination_by
instead, e.g.:-termination_by' measure (fun ⟨i, _⟩ => as.size - i) +termination_by go i _ => as.size - i
If the well-founded relation you want to use is not the one that the
WellFoundedRelation
type class would infer for your termination argument,
you can useWellFounded.wrap
from the std libarary to explicitly give one:-termination_by' ⟨r, hwf⟩ +termination_by _ x => hwf.wrap x
-
Support snippet edits in LSP
TextEdit
s. SeeLean.Lsp.SnippetString
for more details. -
Deprecations and changes in the widget API.
Widget.UserWidgetDefinition
is deprecated in favour ofWidget.Module
. The annotation@[widget]
is deprecated in favour of@[widget_module]
. To migrate a definition of typeUserWidgetDefinition
, remove thename
field and replace the type withWidget.Module
. Removing thename
results in a title bar no longer being drawn above your panel widget. To add it back, draw it as part of the component using<details open=true><summary class='mv2 pointer'>{name}</summary>{rest_of_widget}</details>
. See an example migration here.- The new command
show_panel_widgets
allows displaying always-on and locally-on panel widgets. RpcEncodable
widget props can now be stored in the infotree.- See RFC 2963 for more details and motivation.
-
If no usable lexicographic order can be found automatically for a termination proof, explain why.
See feat: GuessLex: if no measure is found, explain why. -
Option to print inferred termination argument.
Withset_option showInferredTerminationBy true
you will get messages likeInferred termination argument: termination_by ackermann n m => (sizeOf n, sizeOf m)
for automatically generated
termination_by
clauses. -
More detailed error messages for invalid mutual blocks.
-
Multiple improvements to the output of
simp?
andsimp_all?
. -
Tactics with
withLocation *
no longer fail if they close the main goal. -
Implementation of a
test_extern
command for writing tests for@[extern]
and@[implemented_by]
functions.
Usage isimport Lean.Util.TestExtern test_extern Nat.add 17 37
The head symbol must be the constant with the
@[extern]
or@[implemented_by]
attribute. The return type must have aDecidableEq
instance.
Bug fixes for
#2853, #2953, #2966,
#2971, #2990, #3094.
Bug fix for eager evaluation of default value in Option.getD
.
Avoid panic in leanPosToLspPos
when file source is unavailable.
Improve short-circuiting behavior for List.all
and List.any
.
Several Lake bug fixes: #3036, #3064, #3069.
Full Changelog: v4.4.0...v4.5.0-rc1
v4.4.0
Changes since v4.3.0
(from RELEASES.md)
-
Lake and the language server now support per-package server options using the
moreServerOptions
config field, as well as options that apply to both the language server andlean
using theleanOptions
config field. Setting either of these fields instead ofmoreServerArgs
ensures that viewing files from a dependency uses the options for that dependency. Additionally,moreServerArgs
is being deprecated in favor of themoreGlobalServerArgs
field. See PR #2858.A Lakefile with the following deprecated package declaration:
def moreServerArgs := #[ "-Dpp.unicode.fun=true" ] def moreLeanArgs := moreServerArgs package SomePackage where moreServerArgs := moreServerArgs moreLeanArgs := moreLeanArgs
... can be updated to the following package declaration to use per-package options:
package SomePackage where leanOptions := #[⟨`pp.unicode.fun, true⟩]
Bug fixes for #2628, #2883 #2810, #2925, and #2914.
Lake:
lake init .
and a barelake init
and will now use the current directory as the package name. #2890lake new
andlake init
will now produce errors on invalid package names such as..
,foo/bar
,Init
,Lean
,Lake
, andMain
. See issue #2637 and PR #2890.lean_lib
no longer converts its name to upper camel case (e.g.,lean_lib bar
will include modules namedbar.*
rather thanBar.*
). See issue #2567 and PR #2889.- Lean and Lake now properly support non-identifier library names (e.g.,
lake new 123-hello
andimport «123Hello»
now work correctly). See issue #2865 and PR #2889. - Lake now filters the environment extensions loaded from a compiled configuration (
lakefile.olean
) to include only those relevant to Lake's workspace loading process. This resolves segmentation faults caused by environment extension type mismatches (e.g., when defining custom elaborators viaelab
in configurations). See issue #2632 and PR #2896. - Cloud releases will now properly be re-unpacked if the build directory is removed. See PR #2928.
- Lake's
math
template has been simplified. See PR #2930. lake exe <target>
now parsestarget
like a build target (as the help text states it should) rather than as a basic name. For example,lake exe @mathlib/runLinter
should now work. See PR #2932.lake new foo.bar [std]
now generates executables namedfoo-bar
andlake new foo.bar exe
properly createsfoo/bar.lean
. See PR #2932.- Later packages and libraries in the dependency tree are now preferred over earlier ones. That is, the later ones "shadow" the earlier ones. Such an ordering is more consistent with how declarations generally work in programming languages. This will break any package that relied on the previous ordering. See issue #2548 and PR #2937.
- Executable roots are no longer mistakenly treated as importable. They will no longer be picked up by
findModule?
. See PR #2937.
Full Changelog: v4.3.0...v4.4.0
v4.4.0-rc1
Changes since v4.3.0 (from RELEASES.md):
-
Lake and the language server now support per-package server options using the
moreServerOptions
config field, as well as options that apply to both the language server andlean
using theleanOptions
config field. Setting either of these fields instead ofmoreServerArgs
ensures that viewing files from a dependency uses the options for that dependency. Additionally,moreServerArgs
is being deprecated in favor of themoreGlobalServerArgs
field. See PR #2858.A Lakefile with the following deprecated package declaration:
def moreServerArgs := #[ "-Dpp.unicode.fun=true" ] def moreLeanArgs := moreServerArgs package SomePackage where moreServerArgs := moreServerArgs moreLeanArgs := moreLeanArgs
... can be updated to the following package declaration to use per-package options:
package SomePackage where leanOptions := #[⟨`pp.unicode.fun, true⟩]
Bug fixes for #2628, #2883,
#2810, #2925, and #2914.
Lake:
lake init .
and a barelake init
and will now use the current directory as the package name. #2890lake new
andlake init
will now produce errors on invalid package names such as..
,foo/bar
,Init
,Lean
,Lake
, andMain
. See issue #2637 and PR #2890.lean_lib
no longer converts its name to upper camel case (e.g.,lean_lib bar
will include modules namedbar.*
rather thanBar.*
). See issue #2567 and PR #2889.- Lean and Lake now properly support non-identifier library names (e.g.,
lake new 123-hello
andimport «123Hello»
now work correctly). See issue #2865 and PR #2889. - Lake now filters the environment extensions loaded from a compiled configuration (
lakefile.olean
) to include only those relevant to Lake's workspace loading process. This resolves segmentation faults caused by environment extension type mismatches (e.g., when defining custom elaborators viaelab
in configurations). See issue #2632 and PR #2896. - Cloud releases will now properly be re-unpacked if the build directory is removed. See PR #2928.
- Lake's
math
template has been simplified. See PR #2930. lake exe <target>
now parsestarget
like a build target (as the help text states it should) rather than as a basic name. For example,lake exe @mathlib/runLinter
should now work. See PR #2932.lake new foo.bar [std]
now generates executables namedfoo-bar
andlake new foo.bar exe
properly createsfoo/bar.lean
. See PR #2932.- Later packages and libraries in the dependency tree are now preferred over earlier ones. That is, the later ones "shadow" the earlier ones. Such an ordering is more consistent with how declarations generally work in programming languages. This will break any package that relied on the previous ordering. See issue #2548 and PR #2937.
- Executable roots are no longer mistakenly treated as importable. They will no longer be picked up by
findModule?
. See PR #2937.
Pull requests merged:
- chore: update release notes after v4.3.0-rc2 by @semorrison in #2891
- chore: Issue template: Suggest
#eval Lean.versionString
by @nomeata in #2884 - fix: support non-identifier library names by @tydeu in #2888
- refactor: lake: use plain lib name for root and native name by @tydeu in #2889
- feat: bare
lake init
& validated pkg names by @tydeu in #2890 - Allow trailing comma in tuples, lists, and tactics by @vleni in #2643
- chore: syntax documentation for
universe
,open
,export
,variable
by @AdrienChampion in #2645 - fix: lake: whitelist loaded config olean env exts by @tydeu in #2896
- chore: ignore forgotten Lake test artifacts by @tydeu in #2886
- doc: fix typos by @marcusrossel in #2915
- fix: DecidableEq deriving handler could not handle fields whose types start with an implicit argument by @kmill in #2918
- doc: release notes for recent lake fixes by @tydeu in #2906
- fix: PackMutual: Deal with extra arguments by @nomeata in #2892
- feat:
Lean.MVarId.cleanup
configuration by @kmill in #2919 - feat: helpful error message about supportInterpreter by @semorrison in #2912
- doc: Adjust contributor's docs to squash merging by @nomeata in #2927
- feat: rename request handler by @digama0 in #2462
- doc: mention
dite
inite
docstring by @AdrienChampion in #2924 - fix:
eq_refl
tactic’s name iseqRefl
by @nomeata in #2911 - refactor: WF.Fix: Pass all remaining goals to
Term.reportUnsolvedGoals
by @nomeata in #2922 - chore: run CI on merge_group by @semorrison in #2948
- fix: most-recently-nightly-tag does not assume a 'nightly' remote by @semorrison in #2947
- fix: Use whnf for mutual recursion with types hiding
→
by @nomeata in #2926 - chore: CI: pin macos-11 to work around 12.7.1 breakage by @Kha in #2946
- fix: Float RecApp out of applications by @nomeata in #2818
- fix:
PackMutual
: Eta-Expand as needed by @nomeata in #2902 - chore: script/most-recent-nightly-tag uses https rather than ssh repo URL by @semorrison in #2951
- chore: Run CI on all PRs, even base ≠ master by @nomeata in #2955
- doc: Markdown fixes in Lean.Expr by @nomeata in #2956
- feat: import auto-completion by @mhuisi in #2904
- chore: remove unused MonadBacktrack instance for SimpM by @semorrison in #2943
- feat:
pp.beta
to apply beta reduction when pretty printing by @kmill in #2864 - feat: per-package server options by @mhuisi in #2858
- fix: ignore errors on
IO.FS.Handle
finalization by @Kha in #2935 - feat: embed and check githash in .olean by @Kha in #2766
- fix: missing withContext in simp trace by @digama0 in #2053
- feat: Add MatcherApp. and CasesOnApp.refineThrough by @nomeata in #2882
- fix: untar cloud release if no build dir by @tydeu in #2928
- fix: lake: proper exe targets & pkg generation by @tydeu in #2932
- refactor: reverse pkg/lib search & no exe roots in import by @tydeu in #2937
- feat: Guess lexicographic order for well-founded recursion by @nomeata in #2874
- refactor: lake: simplify math template & test it by @tydeu in #2930
- doc: release notes for recent lake changes by @tydeu in #2938
- feat: WF.GuessLex: If there is only one plausible measure, use it by @nomeata in #2954
- fix: remove unnecessary step in pr-release.yml by @semorrison in #2976
- chore: Check PR title, not commit, for commit convention by @nomeata in #2978
- chore: CI: Create an all-builds-ok job by @nomeata in https://github.com/leanprover/lean4/pul...