We intend to provide regular "minor version" releases of the Lean language at approximately monthly intervals. There is not yet a strong guarantee of backwards compatibility between versions, only an expectation that breaking changes will be documented in this file.
This file contains work-in-progress notes for the upcoming release, as well as previous stable releases. Please check the releases page for the current status of each version.
-
bv_decide
tactic. This release introduces a new tactic for proving goals involvingBitVec
andBool
. It reduces the goal to a SAT instance that is refuted by an external solver, and the resulting LRAT proof is checked in Lean. This is used to synthesize a proof of the goal by reflection. As this process uses verified algorithms, proofs generated by this tactic useLean.ofReduceBool
, so this tactic includes the Lean compiler as part of the trusted code base. The external solver CaDiCaL is included with Lean and does not need to be installed separately to make use ofbv_decide
.For example, we can use
bv_decide
to verify that a bit twiddling formula leaves at most one bit set:def popcount (x : BitVec 64) : BitVec 64 := let rec go (x pop : BitVec 64) : Nat → BitVec 64 | 0 => pop | n + 1 => go (x >>> 2) (pop + (x &&& 1)) n go x 0 64 example (x : BitVec 64) : popcount ((x &&& (x - 1)) ^^^ x) ≤ 1 := by simp only [popcount, popcount.go] bv_decide
When the external solver fails to refute the SAT instance generated by
bv_decide
, it can report a counterexample:/-- error: The prover found a counterexample, consider the following assignment: x = 0xffffffffffffffff#64 -/ #guard_msgs in example (x : BitVec 64) : x < x + 1 := by bv_decide
See
Lean.Elab.Tactic.BVDecide
for a more detailed overview, and look intests/lean/run/bv_*
for examples. -
simp
tactic- #4988 fixes a panic in the
reducePow
simproc. - #5071 exposes the
index
option to thedsimp
tactic, introduced tosimp
in #4202. - #5159 fixes a panic at
Fin.isValue
simproc. - #5167 and #5175 rename the
simpCtorEq
simproc toreduceCtorEq
and makes it optional. (See breaking changes.) - #5187 ensures
reduceCtorEq
is enabled in thenorm_cast
tactic. - #5073 modifies the simp debug trace messages to tag with "dpre" and "dpost" instead of "pre" and "post" when in definitional rewrite mode. #5054 explains the
reduce
steps fortrace.Debug.Meta.Tactic.simp
trace messages.
- #4988 fixes a panic in the
-
ext
tactic- #4996 reduces default maximum iteration depth from 1000000 to 100.
-
induction
tactic- #5117 fixes a bug where
let
bindings in minor premises wouldn't be counted correctly.
- #5117 fixes a bug where
-
omega
tactic- #5157 fixes a panic.
-
conv
tactic- #5149 improves
arg n
to handle subsingleton instance arguments.
- #5149 improves
-
#5044 upstreams the
#time
command. -
#5079 makes
#check
and#reduce
typecheck the elaborated terms. -
Incrementality
- #4974 fixes regression where we would not interrupt elaboration of previous document versions.
- #5004 fixes a performance regression.
- #5001 disables incremental body elaboration in presence of
where
clauses in declarations. - #5018 enables infotrees on the command line for ilean generation.
- #5040 and #5056 improve performance of info trees.
- #5090 disables incrementality in the
case .. | ..
tactic. - #5312 fixes a bug where changing whitespace after the module header could break subsequent commands.
-
Definitions
- #5016 and #5066 add
clean_wf
tactic to clean up tactic state indecreasing_by
. This can be disabled withset_option debug.rawDecreasingByGoal false
. - #5055 unifies equational theorems between structural and well-founded recursion.
- #5041 allows mutually recursive functions to use different parameter names among the “fixed parameter prefix”
- #4154 and #5109 add fine-grained equational lemmas for non-recursive functions. See breaking changes.
- #5129 unifies equation lemmas for recursive and non-recursive definitions. The
backward.eqns.deepRecursiveSplit
option can be set tofalse
to get the old behavior. See breaking changes. - #5141 adds
f.eq_unfold
lemmas. Now Lean produces the following zoo of rewrite rules:TheOption.map.eq_1 : Option.map f none = none Option.map.eq_2 : Option.map f (some x) = some (f x) Option.map.eq_def : Option.map f p = match o with | none => none | (some x) => some (f x) Option.map.eq_unfold : Option.map = fun f p => match o with | none => none | (some x) => some (f x)
f.eq_unfold
variant is especially useful to rewrite withrw
under binders. - #5136 fixes bugs in recursion over predicates.
- #5016 and #5066 add
-
Variable inclusion
- #5206 documents that
include
currently only applies to theorems.
- #5206 documents that
-
Elaboration
- #4926 fixes a bug where autoparam errors were associated to an incorrect source position.
- #4833 fixes an issue where cdot anonymous functions (e.g.
(· + ·)
) would not handle ambiguous notation correctly. Numbers the parameters, making this example expand asfun x1 x2 => x1 + x2
rather thanfun x x_1 => x + x_1
. - #5037 improves strength of the tactic that proves array indexing is in bounds.
- #5119 fixes a bug in the tactic that proves indexing is in bounds where it could loop in the presence of mvars.
- #5072 makes the structure type clickable in "not a field of structure" errors for structure instance notation.
- #4717 fixes a bug where mutual
inductive
commands could create terms that the kernel rejects. - #5142 fixes a bug where
variable
could fail when mixing binder updates and declarations.
-
Other fixes or improvements
- #5118 changes the definition of the
syntheticHole
parser so that hovering over_
in?_
gives the docstring for synthetic holes. - #5173 uses the emoji variant selector for ✅️,❌️,💥️ in messages, improving fonts selection.
- #5183 fixes a bug in
rename_i
where implementation detail hypotheses could be renamed.
- #5118 changes the definition of the
- #4821 resolves two language server bugs that especially affect Windows users. (1) Editing the header could result in the watchdog not correctly restarting the file worker, which would lead to the file seemingly being processed forever. (2) On an especially slow Windows machine, we found that starting the language server would sometimes not succeed at all. This PR also resolves an issue where we would not correctly emit messages that we received while the file worker is being restarted to the corresponding file worker after the restart.
- #5006 updates the user widget manual.
- #5193 updates the quickstart guide with the new display name for the Lean 4 extension ("Lean 4").
- #5185 fixes a bug where over time "import out of date" messages would accumulate.
- #4900 improves ilean loading performance by about a factor of two. Optimizes the JSON parser and the conversion from JSON to Lean data structures; see PR description for details.
- Other fixes or improvements
- #5031 localizes an instance in
Lsp.Diagnostics
.
- #5031 localizes an instance in
- #4976 introduces
@[app_delab]
, a macro for creating delaborators for particular constants. The@[app_delab ident]
syntax resolvesident
to its constant namename
and then expands to@[delab app.name]
. - #4982 fixes a bug where the pretty printer assumed structure projections were type correct (such terms can appear in type mismatch errors). Improves hoverability of
#print
output for structures. - #5218 and #5239 add
pp.exprSizes
debugging option. When true, each pretty printed expression is prefixed with[size a/b/c]
, wherea
is the size without sharing,b
is the actual size, andc
is the size with the maximum possible sharing.
-
#5020 swaps the parameters to
Membership.mem
. A purpose of this change is to make set-likeCoeSort
coercions to refer to the eta-expanded functionfun x => Membership.mem s x
, which can reduce in many computations. Another is that having thes
argument first leads to better discrimination tree keys. (See breaking changes.) -
Array
-
List
- #4995 upstreams
List.findIdx
lemmas. - #5029, #5048 and #5132 add
List.Sublist
lemmas, some upstreamed. #5077 fixes implicitness in refl/rfl lemma binders. addList.Sublist
theorems. - #5047 upstreams
List.Pairwise
lemmas. - #5053, #5124, and #5161 add
List.find?/findSome?/findIdx?
theorems. - #5039 adds
List.foldlRecOn
andList.foldrRecOn
recursion principles to prove things aboutList.foldl
andList.foldr
. - #5069 upstreams
List.Perm
. - #5092 and #5107 add
List.mergeSort
and a fast@[csimp]
implementation. - #5103 makes the simp lemmas for
List.subset
more aggressive. - #5106 changes the statement of
List.getLast?_cons
. - #5123 and #5158 add
List.range
andList.iota
lemmas. - #5130 adds
List.join
lemmas. - #5131 adds
List.append
lemmas. - #5152 adds
List.erase(|P|Idx)
lemmas. - #5127 makes miscellaneous lemma updates.
- #5153 and #5160 add lemmas about
List.attach
andList.pmap
. - #5164, #5177, and #5215 add
List.find?
andList.range'/range/iota
lemmas. - #5196 adds
List.Pairwise_erase
and related lemmas. - #5151 and #5163 improve confluence of
List
simp lemmas. #5105 and #5102 adjustList
simp lemmas. - #5178 removes
List.getLast_eq_iff_getLast_eq_some
as a simp lemma. - #5210 reverses the meaning of
List.getElem_drop
andList.getElem_drop'
. - #5214 moves
@[csimp]
lemmas earlier where possible.
- #4995 upstreams
-
Nat
andInt
- #5104 adds
Nat.add_left_eq_self
and relatives. - #5146 adds missing
Nat.and_xor_distrib_(left|right)
. - #5148 and #5190 improve
Nat
andInt
simp lemma confluence. - #5165 adjusts
Int
simp lemmas. - #5166 adds
Int
lemmas relatingneg
andemod
/mod
. - #5208 reverses the direction of the
Int.toNat_sub
simp lemma. - #5209 adds
Nat.bitwise
lemmas. - #5230 corrects the docstrings for integer division and modulus.
- #5104 adds
-
Option
-
BitVec
- #4889 adds
sshiftRight
bitblasting. - #4981 adds
Std.Associative
andStd.Commutative
instances forBitVec.[and|or|xor]
. - #4913 enables
missingDocs
error forBitVec
modules. - #4930 makes parameter names for
BitVec
more consistent. - #5098 adds
BitVec.intMin
. IntroducesboolToPropSimps
simp set for converting from boolean to propositional expressions. - #5200 and #5217 rename
BitVec.getLsb
toBitVec.getLsbD
, etc., to bring naming in line withList
/Array
/etc. - Theorems: #4977, #4951, #4667, #5007, #4997, #5083, #5081, #4392
- #4889 adds
-
UInt
- #4514 fixes naming convention for
UInt
lemmas.
- #4514 fixes naming convention for
-
Std.HashMap
andStd.HashSet
-
Std.Sat
(forbv_decide
)- #4933 adds definitions of SAT and CNF.
- #4953 defines "and-inverter graphs" (AIGs) as described in section 3 of Davis-Swords 2013.
-
Parsec
-
Thunk
- #4969 upstreams
Thunk.ext
.
- #4969 upstreams
-
IO
-
Other fixes or improvements
- #4945 adds
Array
,Bool
andProd
utilities from LeanSAT. - #4960 adds
Relation.TransGen.trans
. - #5012 states
WellFoundedRelation Nat
using<
, notNat.lt
. - #5011 uses
≠
instead ofNot (Eq ...)
inFin.ne_of_val_ne
. - #5197 upstreams
Fin.le_antisymm
. - #5042 reduces usage of
refine'
. - #5101 adds about
if-then-else
andOption
. - #5112 adds basic instances for
ULift
andPLift
. - #5133 and #5168 make fixes from running the simpNF linter over Lean.
- #5156 removes a bad simp lemma in
omega
theory. - #5155 improves confluence of
Bool
simp lemmas. - #5162 improves confluence of
Function.comp
simp lemmas. - #5191 improves confluence of
if-then-else
simp lemmas. - #5147 adds
@[elab_as_elim]
toQuot.rec
,Nat.strongInductionOn
andNat.casesStrongInductionOn
, and also renames the latter two toNat.strongRecOn
andNat.casesStrongRecOn
(deprecated in #5179). - #5180 disables some simp lemmas with bad discrimination tree keys.
- #5189 cleans up internal simp lemmas that had leaked.
- #5198 cleans up
allowUnsafeReducibility
. - #5229 removes unused lemmas from some
simp
tactics. - #5199 removes >6 month deprecations.
- #4945 adds
- Performance
- Some core algorithms have been rewritten in C++ for performance.
- #4934 has optimizations for the kernel's
Expr
equality test. - #4990 fixes bug in hashing for the kernel's
Expr
equality test. - #4935 and #4936 skip some
PreDefinition
transformations if they are not needed. - #5225 adds caching for visited exprs at
CheckAssignmentQuick
inExprDefEq
. - #5226 maximizes term sharing at
instantiateMVarDeclMVars
, used byrunTactic
.
- Diagnostics and profiling
- Other fixes or improvements
- #4921 cleans up
Expr.betaRev
. - #4940 fixes tests by not writing directly to stdout, which is unreliable now that elaboration and reporting are executed in separate threads.
- #4955 documents that
stderrAsMessages
is now the default on the command line as well. - #4647 adjusts documentation for building on macOS.
- #4987 makes regular mvar assignments take precedence over delayed ones in
instantiateMVars
. Normally delayed assignment metavariables are never directly assigned, but on errors Lean assignssorry
to unassigned metavariables. - #4967 adds linter name to errors when a linter crashes.
- #5043 cleans up command line snapshots logic.
- #5067 minimizes some imports.
- #5068 generalizes the monad for
addMatcherInfo
. - f71a1f adds missing test for #5126.
- #5201 restores a test.
- #3698 fixes a bug where label attributes did not pass on the attribute kind.
- Typos: #5080, #5150, #5202
- #4921 cleans up
- #3106 moves frontend to new snapshot architecture. Note that
Frontend.processCommand
andFrontendM
are no longer used by Lean core, but they will be preserved. - #4919 adds missing include in runtime for
AUTO_THREAD_FINALIZATION
feature on Windows. - #4941 adds more
LEAN_EXPORT
s for Windows. - #4911 improves formatting of CLI help text for the frontend.
- #4950 improves file reading and writing.
readBinFile
andreadFile
now only require two system calls (stat
+read
) instead of oneread
per 1024 byte chunk.Handle.getLine
andHandle.putStr
no longer get tripped up by NUL characters.
- #4971 handles the SIGBUS signal when detecting stack overflows.
- #5062 avoids overwriting existing signal handlers, like in rust-lang/rust#69685.
- #4860 improves workarounds for building on Windows. Splits
libleanshared
on Windows to avoid symbol limit, removes theLEAN_EXPORT
denylist workaround, adds missingLEAN_EXPORT
s. - #4952 output panics into Lean's redirected stderr, ensuring panics ARE visible as regular messages in the language server and properly ordered in relation to other messages on the command line.
- #4963 links LibUV.
- #5030 removes dead code.
- #4770 adds additional fields to the package configuration which will be used by Reservoir. See the PR description for details.
- #4914 and #4937 improve the release checklist.
- #4925 ignores stale leanpkg tests.
- #5003 upgrades
actions/cache
in CI. - #5010 sets
save-always
in cache actions in CI. - #5008 adds more libuv search patterns for the speedcenter.
- #5009 reduce number of runs in the speedcenter for "fast" benchmarks from 10 to 3.
- #5014 adjusts lakefile editing to use new
git
syntax inpr-release
workflow. - #5025 has
pr-release
workflow pass--retry
tocurl
. - #5022 builds MacOS Aarch64 release for PRs by default.
- #5045 adds libuv to the required packages heading in macos docs.
- #5034 fixes the install name of
libleanshared_1
on macOS. - #5051 fixes Windows stage 0.
- #5052 fixes 32bit stage 0 builds in CI.
- #5057 avoids rebuilding
leanmanifest
in each build. - #5099 makes
restart-on-label
workflow also filter by commit SHA. - #4325 adds CaDiCaL.
-
LibUV is now required to build Lean. This change only affects developers who compile Lean themselves instead of obtaining toolchains via
elan
. We have updated the official build instructions with information on how to obtain LibUV on our supported platforms. (#4963) -
Recursive definitions with a
decreasing_by
clause that begins withsimp_wf
may break. Try removingsimp_wf
or replacing it withsimp
. (#5016) -
The behavior of
rw [f]
wheref
is a non-recursive function defined by pattern matching changed.For example, preciously,
rw [Option.map]
would rewriteOption.map f o
tomatch o with …
. Now this rewrite fails because it will use the equational lemmas, and these require constructors – just like forList.map
.Remedies:
- Split on
o
before rewriting. - Use
rw [Option.map.eq_def]
, which rewrites any (saturated) application ofOption.map
. - Use
set_option backward.eqns.nonrecursive false
when defining the function in question. (#4154)
- Split on
-
The unified handling of equation lemmas for recursive and non-recursive functions can break existing code, as there now can be extra equational lemmas:
-
Explicit uses of
f.eq_2
might have to be adjusted if the numbering changed. -
Uses of
rw [f]
orsimp [f]
may no longer apply if they previously matched (and introduced amatch
statement), when the equational lemmas got more fine-grained.In this case either case analysis on the parameters before rewriting helps, or setting the option
backward.eqns.deepRecursiveSplit false
while defining the function.
-
-
The
reduceCtorEq
simproc is now optional, and it might need to be included in lists of simp lemmas, likesimp only [reduceCtorEq]
. This simproc is responsible for reducing equalities of constructors. (#5167) -
Nat.strongInductionOn
is nowNat.strongRecOn
andNat.caseStrongInductionOn
toNat.caseStrongRecOn
. (#5147) -
The parameters to
Membership.mem
have been swapped, which affects allMembership
instances. (#5020) -
The meanings of
List.getElem_drop
andList.getElem_drop'
have been reversed and the first is now a simp lemma. (#5210) -
The
Parsec
library has moved fromLean.Data.Parsec
toStd.Internal.Parsec
. TheParsec
type is now more general with a parameter for an iterable. Users parsing strings can migrate toParser
in theStd.Internal.Parsec.String
namespace, which also includes string-focused parsing combinators. (#4774) -
The
Lean
module has switched fromLean.HashMap
andLean.HashSet
toStd.HashMap
andStd.HashSet
(#4943).Lean.HashMap
andLean.HashSet
are now deprecated (#4954) and will be removed in a future release. Users ofLean
APIs that interact with hash maps, for exampleLean.Environment.const2ModIdx
, might encounter minor breakage due to the following changes fromLean.HashMap
toStd.HashMap
:- query functions use the term
get
instead offind
, (#4943) - the notation
map[key]
no longer returns an optional value but instead expects a proof that the key is present in the map. The previous behavior is available via themap[key]?
notation.
- query functions use the term
-
The variable inclusion mechanism has been changed. Like before, when a definition mentions a variable, Lean will add it as an argument of the definition, but now in theorem bodies, variables are not included based on usage in order to ensure that changes to the proof cannot change the statement of the overall theorem. Instead, variables are only available to the proof if they have been mentioned in the theorem header or in an
include
command or are instance implicit and depend only on such variables. Theomit
command can be used to omit included variables.See breaking changes below.
-
Recursive definitions
-
Structural recursion can now be explicitly requested using
termination_by structural x
in analogy to the existing
termination_by x
syntax that causes well-founded recursion to be used. #4542 -
#4672 fixes a bug that could lead to ill-typed terms.
-
The
termination_by?
syntax no longer forces the use of well-founded recursion, and when structural recursion is inferred, it will print the result using thetermination_by structural
syntax. -
Mutual structural recursion is now supported. This feature supports both mutual recursion over a non-mutual data type, as well as recursion over mutual or nested data types:
mutual def Even : Nat → Prop | 0 => True | n+1 => Odd n def Odd : Nat → Prop | 0 => False | n+1 => Even n end mutual inductive A | other : B → A | empty inductive B | other : A → B | empty end mutual def A.size : A → Nat | .other b => b.size + 1 | .empty => 0 def B.size : B → Nat | .other a => a.size + 1 | .empty => 0 end inductive Tree where | node : List Tree → Tree mutual def Tree.size : Tree → Nat | node ts => Tree.list_size ts def Tree.list_size : List Tree → Nat | [] => 0 | t::ts => Tree.size t + Tree.list_size ts end
Functional induction principles are generated for these functions as well (
A.size.induct
,A.size.mutual_induct
).Nested structural recursion is still not supported.
PRs: #4639, #4715, #4642, #4656, #4684, #4715, #4728, #4575, #4731, #4658, #4734, #4738, #4718, #4733, #4787, #4788, #4789, #4807, #4772
-
#4809 makes unnecessary
termination_by
clauses cause warnings, not errors. -
#4831 improves handling of nested structural recursion through non-recursive types.
-
#4839 improves support for structural recursive over inductive predicates when there are reflexive arguments.
-
-
simp
tactic- #4784 sets configuration
Simp.Config.implicitDefEqProofs
totrue
by default.
- #4784 sets configuration
-
omega
tactic -
decide
tactic- #4711 switches from using default transparency to at least default transparency when reducing the
Decidable
instance. - #4674 adds detailed feedback on
decide
tactic failure. It tells you whichDecidable
instances it unfolded, if it get stuck onEq.rec
it gives a hint about avoiding tactics when definingDecidable
instances, and if it gets stuck onClassical.choice
it gives hints about classical instances being in scope. During this process, it processesDecidable.rec
s and matches to pin blame on a non-reducing instance.
- #4711 switches from using default transparency to at least default transparency when reducing the
-
@[ext]
attribute- #4543 and #4762 make
@[ext]
realizeext_iff
theorems from userext
theorems. Fixes the attribute so that@[local ext]
and@[scoped ext]
are usable. The@[ext (iff := false)]
option can be used to turn offext_iff
realization. - #4694 makes "go to definition" work for the generated lemmas. Also adjusts the core library to make use of
ext_iff
generation. - #4710 makes
ext_iff
theorem preserve inst implicit binder types, rather than making all binder types implicit.
- #4543 and #4762 make
-
#eval
command- #4810 introduces a safer
#eval
command that prevents evaluation of terms that containsorry
. The motivation is that failing tactics, in conjunction with operations such as array accesses, can lead to the Lean process crashing. Users can use the new#eval!
command to use the previous unsafe behavior. (#4829 adjusts a test.)
- #4810 introduces a safer
-
#4447 adds
#discr_tree_key
and#discr_tree_simp_key
commands, for helping debug discrimination tree failures. The#discr_tree_key t
command prints the discrimination tree keys for a termt
(or, if it is a single identifier, the type of that constant). It uses the default configuration for generating keys. The#discr_tree_simp_key
command is similar to#discr_tree_key
, but treats the underlying type as one of a simp lemma, that is it transforms it into an equality and produces the key of the left-hand side.For example,
#discr_tree_key (∀ {a n : Nat}, bar a (OfNat.ofNat n)) -- bar _ (@OfNat.ofNat Nat _ _) #discr_tree_simp_key Nat.add_assoc -- @HAdd.hAdd Nat Nat Nat _ (@HAdd.hAdd Nat Nat Nat _ _ _) _
-
#4741 changes option parsing to allow user-defined options from the command line. Initial options are now re-parsed and validated after importing. Command line option assignments prefixed with
weak.
are silently discarded if the option name without the prefix does not exist. -
Deriving handlers
-
Metaprogramming
- #4593 adds
unresolveNameGlobalAvoidingLocals
. - #4618 deletes deprecated functions from 2022.
- #4642 adds
Meta.lambdaBoundedTelescope
. - #4731 adds
Meta.withErasedFVars
, to enter a context with some fvars erased from the local context. - #4777 adds assignment validation at
closeMainGoal
, preventing users from circumventing the occurs check for tactics such asexact
. - #4807 introduces
Lean.Meta.PProdN
module for packing and projecting nestedPProd
s. - #5170 fixes
Syntax.unsetTrailing
. A consequence of this is that "go to definition" now works on the last module name in animport
block (issue #4958).
- #4593 adds
- #4727 makes it so that responses to info view requests come as soon as the relevant tactic has finished execution.
- #4580 makes it so that whitespace changes do not invalidate imports, and so starting to type the first declaration after imports should no longer cause them to reload.
- #4780 fixes an issue where hovering over unimported builtin names could result in a panic.
- #4558 fixes the
pp.instantiateMVars
setting and changes the default value totrue
. - #4631 makes sure syntax nodes always run their formatters. Fixes an issue where if
ppSpace
appears in amacro
orelab
command then it does not format with a space. - #4665 fixes a bug where pretty printed signatures (for example in
#check
) were overly hoverable due topp.tagAppFns
being set. - #4724 makes
match
pretty printer be sensitive topp.explicit
, which makes hovering over amatch
in the Infoview show the underlying term. - #4764 documents why anonymous constructor notation isn't pretty printed with flattening.
- #4786 adjusts the parenthesizer so that only the parentheses are hoverable, implemented by having the parentheses "steal" the term info from the parenthesized expression.
- #4854 allows arbitrarily long sequences of optional arguments to be omitted from the end of applications, versus the previous conservative behavior of omitting up to one optional argument.
Nat
Int
- #4903 fixes performance of
HPow Int Nat Int
synthesis by rewriting it as aNatPow Int
instance.
- #4903 fixes performance of
UInt*
andFin
Option
GetElem
- #4603 adds
getElem_congr
to help with rewriting indices.
- #4603 adds
List
andArray
- Upstreamed from Batteries: #4586 upstreams
List.attach
andArray.attach
, #4697 upstreamsList.Subset
andList.Sublist
and API, #4706 upstreams basic material onList.Pairwise
andList.Nodup
, #4720 upstreams moreList.erase
API, #4836 and #4837 upstreamList.IsPrefix
/List.IsSuffix
/List.IsInfix
and addDecidable
instances, #4855 upstreamsList.tail
,List.findIdx
,List.indexOf
,List.countP
,List.count
, andList.range'
, #4856 upstreams more List lemmas, #4866 upstreamsList.pairwise_iff_getElem
, #4865 upstreamsList.eraseIdx
lemmas. - #4687 adjusts
List.replicate
simp lemmas and simprocs. - #4704 adds characterizations of
List.Sublist
. - #4707 adds simp normal form tests for
List.Pairwise
andList.Nodup
. - #4708 and #4815 reorganize lemmas on list getters.
- #4765 adds simprocs for literal array accesses such as
#[1,2,3,4,5][2]
. - #4790 removes typeclass assumptions for
List.Nodup.eraseP
. - #4801 adds efficient
usize
functions for array types. - #4820 changes
List.filterMapM
to run left-to-right. - #4835 fills in and cleans up gaps in List API.
- #4843, #4868, and #4877 correct
List.Subset
lemmas. - #4863 splits
Init.Data.List.Lemmas
into function-specific files. - #4875 fixes statement of
List.take_takeWhile
. - Lemmas: #4602, #4627, #4678 for
List.head
andlist.getLast
, #4723 forList.erase
, #4742
- Upstreamed from Batteries: #4586 upstreams
ByteArray
- #4582 eliminates
partial
fromByteArray.toList
andByteArray.findIdx?
.
- #4582 eliminates
BitVec
Std.HashMap
added:- #4583 adds
Std.HashMap
as a verified replacement forLean.HashMap
. See the PR for naming differences, but #4725 renamesHashMap.remove
toHashMap.erase
. - #4682 adds
Inhabited
instances. - #4732 improves
BEq
argument order in hash map lemmas. - #4759 makes lemmas resolve instances via unification.
- #4771 documents that hash maps should be used linearly to avoid expensive copies.
- #4791 removes
bif
from hash map lemmas, which is inconvenient to work with in practice. - #4803 adds more lemmas.
- #4583 adds
SMap
- #4690 upstreams
SMap.foldM
.
- #4690 upstreams
BEq
- #4607 adds
PartialEquivBEq
,ReflBEq
,EquivBEq
, andLawfulHashable
classes.
- #4607 adds
IO
- #4660 adds
IO.Process.Child.tryWait
.
- #4660 adds
- #4747, #4730, and #4756 add
×'
syntax forPProd
. Adds a delaborator forPProd
andMProd
values to pretty print as flattened angle bracket tuples. - Other fixes or improvements
- #4604 adds lemmas for cond.
- #4619 changes some definitions into theorems.
- #4616 fixes some names with duplicated namespaces.
- #4620 fixes simp lemmas flagged by the simpNF linter.
- #4666 makes the
Antisymm
class be aProp
. - #4621 cleans up unused arguments flagged by linter.
- #4680 adds imports for orphaned
Init
modules. - #4679 adds imports for orphaned
Std.Data
modules. - #4688 adds forward and backward directions of
not_exists
. - #4689 upstreams
eq_iff_true_of_subsingleton
. - #4709 fixes precedence handling for
Repr
instances for negative numbers forInt
andFloat
. - #4760 renames
TC
("transitive closure") toRelation.TransGen
. - #4842 fixes
List
deprecations. - #4852 upstreams some Mathlib attributes applied to lemmas.
- 93ac63 improves proof.
- #4862 and #4878 generalize the universe for
PSigma.exists
and rename it toExists.of_psigma_prop
. - Typos: #4737, 7d2155
- Docs: #4782, #4869, #4648
- Elaboration
- #4596 enforces
isDefEqStuckEx
atunstuckMVar
procedure, causing isDefEq to throw a stuck defeq exception if the metavariable was created in a previous level. This results in some better error messages, and it helpsrw
succeed in synthesizing instances (see issue #2736). - #4713 fixes deprecation warnings when there are overloaded symbols.
elab_as_elim
algorithm:- #4792 adds term elaborator for
Lean.Parser.Term.namedPattern
(e.g.n@(n' + 1)
) to report errors when used in non-pattern-matching contexts. - #4818 makes anonymous dot notation work when the expected type is a pi-type-valued type synonym.
- #4596 enforces
- Typeclass inference
- #4646 improves
synthAppInstances
, the function responsible for synthesizing instances for therw
andapply
tactics. Adds a synthesis loop to handle functions whose instances need to be synthesized in a complex order.
- #4646 improves
- Inductive types
- Definitions
- Diagnostics and profiling
- #4611 makes kernel diagnostics appear when
diagnostics
is enabled even if it is the only section. - #4753 adds missing
profileitM
functions. - #4754 adds
Lean.Expr.numObjs
to compute the number of allocated sub-expressions in a given expression, primarily for diagnosing performance issues. - #4769 adds missing
withTraceNode
s to improvetrace.profiler
output. - #4781 and #4882 make the "use
set_option diagnostics true
" message be conditional on current setting ofdiagnostics
.
- #4611 makes kernel diagnostics appear when
- Performance
- #4767, #4775, and #4887 add
ShareCommon.shareCommon'
for sharing common terms. In an example with 16 million subterms, it is 20 times faster than the oldshareCommon
procedure. - #4779 ensures
Expr.replaceExpr
preserves DAG structure inExpr
s. - #4783 documents performance issue in
Expr.replaceExpr
. - #4794, #4797, #4798 make
for_each
use precise cache. - #4795 makes
Expr.find?
andExpr.findExt?
use the kernel implementations. - #4799 makes
Expr.replace
use the kernel implementation. - #4871 makes
Expr.foldConsts
use a precise cache. - #4890 makes
expr_eq_fn
use a precise cache.
- #4767, #4775, and #4887 add
- Utilities
- #4453 upstreams
ToExpr FilePath
andcompile_time_search_path%
.
- #4453 upstreams
- Module system
- #4652 fixes handling of
const2ModIdx
infinalizeImport
, making it prefer the original module for a declaration when a declaration is re-declared.
- #4652 fixes handling of
- Kernel
- #4637 adds a check to prevent large
Nat
exponentiations from evaluating. Elaborator reduction is controlled by the optionexponentiation.threshold
. - #4683 updates comments in
kernel/declaration.h
, making sure they reflect the current Lean 4 types. - #4796 improves performance by using
replace
with a precise cache. - #4700 improves performance by fixing the implementation of move constructors and move assignment operators. Expression copying was taking 10% of total runtime in some workloads. See issue #4698.
- #4702 improves performance in
replace_rec_fn::apply
by avoiding expression copies. These copies represented about 13% of time spent insave_result
in some workloads. See the same issue.
- #4637 adds a check to prevent large
- Other fixes or improvements
- #4590 fixes a typo in some constants and
trace.profiler.useHeartbeats
. - #4617 add 'since' dates to
deprecated
attributes. - #4625 improves the robustness of the constructor-as-variable test.
- #4740 extends test with nice example reported on Zulip.
- #4766 moves
Syntax.hasIdent
to be available earlier and shakes dependencies. - #4881 splits out
Lean.Language.Lean.Types
. - #4893 adds
LEAN_EXPORT
forsharecommon
functions. - Typos: #4635, #4719, af40e6
- Docs: #4748 (
Command.Scope
)
- #4590 fixes a typo in some constants and
- #4661 moves
Std
fromlibleanshared
to much smallerlibInit_shared
. This fixes the Windows build. - #4668 fixes initialization, explicitly initializing
Std
inlean_initialize
. - #4746 adjusts
shouldExport
to exclude more symbols to get below Windows symbol limit. Some exceptions are added by #4884 and #4956 to support Verso. - #4778 adds
lean_is_exclusive_obj
(Lean.isExclusiveUnsafe
) andlean_set_external_data
. - #4515 fixes calling programs with spaces on Windows.
-
#4735 improves a number of elements related to Git checkouts, cloud releases, and related error handling.
- On error, Lake now prints all top-level logs. Top-level logs are those produced by Lake outside of the job monitor (e.g., when cloning dependencies).
- When fetching a remote for a dependency, Lake now forcibly fetches tags. This prevents potential errors caused by a repository recreating tags already fetched.
- Git error handling is now more informative.
- The builtin package facets
release
,optRelease
,extraDep
are now captions in the same manner as other facets. afterReleaseSync
andafterReleaseAsync
now fetchoptRelease
rather thanrelease
.- Added support for optional jobs, whose failure does not cause the whole build to failure. Now
optRelease
is such a job.
-
#4608 adds draft CI workflow when creating new projects.
-
#4847 adds CLI options to control log levels. The
--log-level=<lv>
controls the minimum log level Lake should output. For instance,--log-level=error
will only print errors (not warnings or info). Also, adds an analogous--fail-level
option to control the minimum log level for build failures. The existing--iofail
and--wfail
options are respectively equivalent to--fail-level=info
and--fail-level=warning
. -
Docs: #4853
- Workflows
- #4531 makes release trigger an update of
release.lean-lang.org
. - #4598 adjusts
pr-release
to the newlakefile.lean
syntax. - #4632 makes
pr-release
use the correct tag name. - #4638 adds ability to manually trigger nightly release.
- #4640 adds more debugging output for
restart-on-label
CI. - #4663 bumps up waiting for 10s to 30s for
restart-on-label
. - #4664 bumps versions for
actions/checkout
andactions/upload-artifacts
. - 582d6e bumps version for
actions/download-artifact
. - 6d9718 adds back dropped
check-stage3
. - 0768ad adds Jira sync (for FRO).
- #4830 adds support to report CI errors on FRO Zulip.
- #4838 adds trigger for
nightly_bump_toolchain
on mathlib4 upon nightly release. - abf420 fixes msys2.
- #4895 deprecates Nix-based builds and removes interactive components. Users who prefer the flake build should maintain it externally.
- #4531 makes release trigger an update of
- #4693, #4458, and #4876 update the release checklist.
- #4669 fixes the "max dynamic symbols" metric per static library.
- #4691 improves compatibility of
tests/list_simp
for retesting simp normal forms with Mathlib. - #4806 updates the quickstart guide.
- c02aa9 documents the triage team in the contribution guide.
-
For
@[ext]
-generatedext
andext_iff
lemmas, thex
andy
term arguments are now implicit. Furthermore these two lemmas are now protected (#4543). -
Now
trace.profiler.useHearbeats
istrace.profiler.useHeartbeats
(#4590). -
A bugfix in the structural recursion code may in some cases break existing code, when a parameter of the type of the recursive argument is bound behind indices of that type. This can usually be fixed by reordering the parameters of the function (#4672).
-
Now
List.filterMapM
sequences monadic actions left-to-right (#4820). -
The effect of the
variable
command on proofs oftheorem
s has been changed. Whether such section variables are accessible in the proof now depends only on the theorem signature and other top-level commands, not on the proof itself. This change ensures that- the statement of a theorem is independent of its proof. In other words, changes in the proof cannot change the theorem statement.
- tactics such as
induction
cannot accidentally include a section variable. - the proof can be elaborated in parallel to subsequent declarations in a future version of Lean.
The effect of
variable
s on the theorem header as well as on other kinds of declarations is unchanged.Specifically, section variables are included if they
- are directly referenced by the theorem header,
- are included via the new
include
command in the current section and not subsequently mentioned in anomit
statement, - are directly referenced by any variable included by these rules, OR
- are instance-implicit variables that reference only variables included by these rules.
For porting, a new option
deprecated.oldSectionVars
is included to locally switch back to the old behavior.
-
split
tactic:- #4401 improves the strategy
split
uses to generalize discriminants of matches and addstrace.split.failure
trace class for diagnosing issues.
- #4401 improves the strategy
-
rw
tactic: -
simp
tactic:- #4430 adds
dsimproc
s forif
expressions (ite
anddite
). - #4434 improves heuristics for unfolding. Equational lemmas now have priorities where more-specific equationals lemmas are tried first before a possible catch-all.
- #4481 fixes an issue where function-valued
OfNat
numeric literals would become denormalized. - #4467 fixes an issue where dsimp theorems might not apply to literals.
- #4484 fixes the source position for the warning for deprecated simp arguments.
- #4258 adds docstrings for
dsimp
configuration. - #4567 improves the accuracy of used simp lemmas reported by
simp?
. - fb9727 adds (but does not implement) the simp configuration option
implicitDefEqProofs
, which will enable includingrfl
-theorems in proof terms.
- #4430 adds
-
omega
tactic:- #4360 makes the tactic generate error messages lazily, improving its performance when used in tactic combinators.
-
bv_omega
tactic:- #4579 works around changes to the definition of
Fin.sub
in this release.
- #4579 works around changes to the definition of
-
#4490 sets up groundwork for a tactic index in generated documentation, as there was in Lean 3. See PR description for details.
-
Commands
- #4370 makes the
variable
command fully elaborate binders during validation, fixing an issue where some errors would be reported only at the next declaration. - #4408 fixes a discrepancy in universe parameter order between
theorem
anddef
declarations. - #4493 and
#4482 fix a discrepancy in the elaborators for
theorem
,def
, andexample
, makingProp
-valuedexample
s and other definition commands elaborate liketheorem
s. - 8f023b, 3c4d6b and 0783d0 change the
#reduce
command to be able to control what gets reduced. For example,#reduce (proofs := true) (types := false) e
reduces both proofs and types in the expressione
. By default, neither proofs or types are reduced. - #4489 fixes an elaboration bug in
#check_tactic
. - #4505 adds support for
open _root_.<namespace>
.
- #4370 makes the
-
Options
- #4576 adds the
debug.byAsSorry
option. Settingset_option debug.byAsSorry true
causes allby ...
terms to elaborate assorry
. - 7b56eb and d8e719 add the
debug.skipKernelTC
option. Settingset_option debug.skipKernelTC true
turns off kernel typechecking. This is meant for temporarily working around kernel performance issues, and it compromises soundness since buggy tactics may produce invalid proofs, which will not be caught if this option is set to true.
- #4576 adds the
-
#4301 adds a linter to flag situations where a local variable's name is one of the argumentless constructors of its type. This can arise when a user either doesn't open a namespace or doesn't add a dot or leading qualifier, as in the following:
inductive Tree (α : Type) where | leaf | branch (left : Tree α) (val : α) (right : Tree α) def depth : Tree α → Nat | leaf => 0
With this linter, the
leaf
pattern is highlighted as a local variable whose name overlaps with the constructorTree.leaf
.The linter can be disabled with
set_option linter.constructorNameAsVariable false
.Additionally, the error message that occurs when a name in a pattern that takes arguments isn't valid now suggests similar names that would be valid. This means that the following definition:
def length (list : List α) : Nat := match list with | nil => 0 | cons x xs => length xs + 1
now results in the following warning:
warning: Local variable 'nil' resembles constructor 'List.nil' - write '.nil' (with a dot) or 'List.nil' to use the constructor. note: this linter can be disabled with `set_option linter.constructorNameAsVariable false`
and error:
invalid pattern, constructor or constant marked with '[match_pattern]' expected Suggestion: 'List.cons' is similar
-
Metaprogramming
- #4454 adds public
Name.isInternalDetail
function for filtering declarations using naming conventions for internal names.
- #4454 adds public
-
Other fixes or improvements
- #4443 makes the watchdog be more resilient against badly behaving clients.
- #4433 restores fallback pretty printers when context is not available, and documents
addMessageContext
. - #4556 introduces
pp.maxSteps
option and sets the default value ofpp.deepTerms
tofalse
. Together, these keep excessively large or deep terms from overwhelming the Infoview.
- #4560 splits
GetElem
class intoGetElem
andGetElem?
. This enables removingDecidable
instance arguments fromGetElem.getElem?
andGetElem.getElem!
, improving their rewritability. See the docstrings for these classes for more information. Array
List
- #4469 and #4475 improve the organization of the
List
API. - #4470 improves the
List.set
andList.concat
API. - #4472 upstreams lemmas about
List.filter
from Batteries. - #4473 adjusts
@[simp]
attributes. - #4488 makes
List.getElem?_eq_getElem
be a simp lemma. - #4487 adds missing
List.replicate
API. - #4521 adds lemmas about
List.map
. - #4500 changes
List.length_cons
to useas.length + 1
instead ofas.length.succ
. - #4524 fixes the statement of
List.filter_congr
. - #4525 changes binder explicitness in
List.bind_map
. - #4550 adds
maximum?_eq_some_iff'
andminimum?_eq_some_iff?
.
- #4469 and #4475 improve the organization of the
- #4400 switches the normal forms for indexing
List
andArray
toxs[n]
andxs[n]?
. HashMap
- #4372 fixes linearity in
HashMap.insert
andHashMap.erase
, leading to a 40% speedup in a replace-heavy workload.
- #4372 fixes linearity in
Option
Nat
Int
- #3850 adds complete div/mod simprocs for
Int
.
- #3850 adds complete div/mod simprocs for
String
/Char
Fin
- #4421 adjusts
Fin.sub
to be more performant in definitional equality checks.
- #4421 adjusts
Prod
BitVec
Std
library- #4499 introduces
Std
, a library situated betweenInit
andLean
, providing functionality not in the prelude both to Lean's implementation and to external users.
- #4499 introduces
- Other fixes or improvements
- #4391 makes
getBitVecValue?
recognizeBitVec.ofNatLt
. - #4410 adjusts
instantiateMVars
algorithm to zeta reducelet
expressions while beta reducing instantiated metavariables. - #4420 fixes occurs check for metavariable assignments to also take metavariable types into account.
- #4425 fixes
forEachModuleInDir
to iterate over each Lean file exactly once. - #3886 adds support to build Lean core oleans using Lake.
- Defeq and WHNF algorithms
- Typeclass inference
- #4530 fixes handling of metavariables when caching results at
synthInstance?
.
- #4530 fixes handling of metavariables when caching results at
- Elaboration
- #4426 makes feature where the "don't know how to synthesize implicit argument" error reports the name of the argument more reliable.
- #4497 fixes a name resolution bug for generalized field notation (dot notation).
- #4536 blocks the implicit lambda feature for
(e :)
notation. - #4562 makes it be an error for there to be two functions with the same name in a
where
/let rec
block.
- Recursion principles
- #4549 refactors
findRecArg
, extractingwithRecArgInfo
. Errors are now reported in parameter order rather than the order they are tried (non-indices are tried first). For every argument, it will say why it wasn't tried, even if the reason is obvious (e.g. a fixed prefix or isProp
-typed, etc.).
- #4549 refactors
- Porting core C++ to Lean
- Documentation
- #4501 adds a more-detailed docstring for
PersistentEnvExtension
.
- #4501 adds a more-detailed docstring for
- Other fixes or improvements
- #4382 removes
@[inline]
attribute fromNameMap.find?
, which caused respecialization at each call site. - 5f9ded improves output of
trace.Elab.snapshotTree
. - #4424 removes "you might need to open '{dir}' in your editor" message that is now handled by Lake and the VS Code extension.
- #4451 improves the performance of
CollectMVars
andFindMVar
. - #4479 adds missing
DecidableEq
andRepr
instances for intermediate structures used by theBitVec
andFin
simprocs. - #4492 adds tests for a previous
isDefEq
issue. - 9096d6 removes
PersistentHashMap.size
. - #4508 fixes
@[implemented_by]
for functions defined by well-founded recursion. - #4509 adds additional tests for
apply?
tactic. - d6eab3 fixes a benchmark.
- #4563 adds a workaround for a bug in
IndPredBelow.mkBelowMatcher
.
- #4382 removes
- Cleanup: #4380, #4431, #4494, e8f768, de2690, d3a756, #4404, #4537.
- d85d3d fixes criterion for tail-calls in ownership calculation.
- #3963 adds validation of UTF-8 at the C++-to-Lean boundary in the runtime.
- #4512 fixes missing unboxing in interpreter when loading initialized value.
- #4477 exposes the compiler flags for the bundled C compiler (clang).
-
#4384 deprecates
inputFile
and replaces it withinputBinFile
andinputTextFile
. UnlikeinputBinFile
(andinputFile
),inputTextFile
normalizes line endings, which helps ensure text file traces are platform-independent. -
#4371 simplifies dependency resolution code.
-
#4439 touches up the Lake configuration DSL and makes other improvements: string literals can now be used instead of identifiers for names, avoids using French quotes in
lake new
andlake init
templates, changes theexe
template to useMain
for the main module, improves themath
template error iflean-toolchain
fails to download, and downgrades unknown configuration fields from an error to a warning to improve cross-version compatibility. -
#4496 tweaks
require
syntax and updates docs. Nowrequire
in TOML for a package name such asdoc-gen4
does not need French quotes. -
#4485 fixes a bug where package versions in indirect dependencies would take precedence over direct dependencies.
-
#4478 fixes a bug where Lake incorrectly included the module dynamic library in a platform-independent trace.
-
#4529 fixes some issues with bad import errors. A bad import in an executable no longer prevents the executable's root module from being built. This also fixes a problem where the location of a transitive bad import would not been shown. The root module of the executable now respects
nativeFacets
. -
#4564 fixes a bug where non-identifier script names could not be entered on the CLI without French quotes.
-
#4566 addresses a few issues with precompiled libraries.
- Fixes a bug where Lake would always precompile the package of a module.
- If a module is precompiled, it now precompiles its imports. Previously, it would only do this if imported.
-
#4495, #4692, #4849 add a new type of
require
that fetches package metadata from a registry API endpoint (e.g. Reservoir) and then clones a Git package using the information provided. To require such a dependency, the new syntax is:require <scope> / <pkg-name> [@ git <rev>] -- Examples: require "leanprover" / "doc-gen4" require "leanprover-community" / "proofwidgets" @ git "v0.0.39"
Or in TOML:
[[require]] name = "<pkg-name>" scope = "<scope>" rev = "<rev>"
Unlike with Git dependencies, Lake can make use of the richer information provided by the registry to determine the default branch of the package. This means for repositories of packages like
doc-gen4
which have a default branch that is notmaster
, Lake will now use said default branch (e.g., indoc-gen4
's case,main
).Lake also supports configuring the registry endpoint via an environment variable:
RESERVIOR_API_URL
. Thus, any server providing a similar interface to Reservoir can be used as the registry. Further configuration options paralleling those of Cargo's Alternative Registries and Source Replacement will come in the future.
- #4427 uses Namespace runners for CI for
leanprover/lean4
. - #4440 fixes speedcenter tests in CI.
- #4441 fixes that workflow change would break CI for unrebased PRs.
- #4442 fixes Wasm release-ci.
- 6d265b fixes for
github.event.pull_request.merge_commit_sha
sometimes not being available. - 16cad2 adds optimization for CI to not fetch complete history.
- #4544 causes releases to be marked as prerelease on GitHub.
- #4446 switches Lake to using
src/lake/lakefile.toml
to avoid needing to load a version of Lake to build Lake. - Nix
Char.csize
is replaced byChar.utf8Size
(#4357).- Library lemmas now are in terms of
(· == a)
over(a == ·)
(#3056). - Now the normal forms for indexing into
List
andArray
isxs[n]
andxs[n]?
rather than using functions likeList.get
(#4400). - Sometimes terms created via a sequence of unifications will be more eta reduced than before and proofs will require adaptation (#4387).
- The
GetElem
class has been split into two; see the docstrings forGetElem
andGetElem?
for more information (#4560).
- Definition transparency
- #4053 adds the
seal
andunseal
commands, which make definitions locally be irreducible or semireducible. - #4061 marks functions defined by well-founded recursion with
@[irreducible]
by default, which should prevent the expensive and often unfruitful unfolding of such definitions (see breaking changes below).
- #4053 adds the
- Incrementality
- #3940 extends incremental elaboration into various steps inside of declarations: definition headers, bodies, and tactics. .
- 250994
and 67338b
add
@[incremental]
attribute to mark an elaborator as supporting incremental elaboration. - #4259 improves resilience by ensuring incremental commands and tactics are reached only in supported ways.
- #4268 adds special handling for
:= by
so that stray tokens in tactic blocks do not inhibit incrementality. - #4308 adds incremental
have
tactic. - #4340 fixes incorrect info tree reuse.
- #4364 adds incrementality for careful command macros such as
set_option in theorem
,theorem foo.bar
, andlemma
. - #4395 adds conservative fix for whitespace handling to avoid incremental reuse leading to goals in front of the text cursor being shown.
- #4407 fixes non-incremental commands in macros blocking further incremental reporting.
- #4436 fixes incremental reporting when there are nested tactics in terms.
- #4459 adds incrementality support for
next
andif
tactics. - #4554 disables incrementality for tactics in terms in tactics.
- Functional induction
- #4135 ensures that the names used for functional induction are reserved.
- #4327 adds support for structural recursion on reflexive types.
For example,
inductive Many (α : Type u) where | none : Many α | more : α → (Unit → Many α) → Many α def Many.map {α β : Type u} (f : α → β) : Many α → Many β | .none => .none | .more x xs => .more (f x) (fun _ => (xs ()).map f) #check Many.map.induct /- Many.map.induct {α β : Type u} (f : α → β) (motive : Many α → Prop) (case1 : motive Many.none) (case2 : ∀ (x : α) (xs : Unit → Many α), motive (xs ()) → motive (Many.more x xs)) : ∀ (a : Many α), motive a -/
- #3903 makes the Lean frontend normalize all line endings to LF before processing. This lets Lean be insensitive to CRLF vs LF line endings, improving the cross-platform experience and making Lake hashes be faithful to what Lean processes.
- #4130 makes the tactic framework be able to recover from runtime errors (for example, deterministic timeouts or maximum recursion depth errors).
split
tacticapply
tactic- #3929 makes error message for
apply
show implicit arguments in unification errors as needed. ModifiesMessageData
type (see breaking changes below).
- #3929 makes error message for
cases
tactic- #4224 adds support for unification of offsets such as
x + 20000 = 20001
incases
tactic.
- #4224 adds support for unification of offsets such as
omega
tacticsimp
tactic-
#4176 makes names of erased lemmas clickable.
-
#4208 adds a pretty printer for discrimination tree keys.
-
#4202 adds
Simp.Config.index
configuration option, which controls whether to use the full discrimination tree when selecting candidate simp lemmas. Whenindex := false
, only the head function is taken into account, like in Lean 3. This feature can help users diagnose tricky simp failures or issues in code from libraries developed using Lean 3 and then ported to Lean 4.In the following example, it will report that
foo
is a problematic theorem.opaque f : Nat → Nat → Nat @[simp] theorem foo : f x (x, y).2 = y := by sorry example : f a b ≤ b := by set_option diagnostics true in simp (config := { index := false }) /- [simp] theorems with bad keys foo, key: f _ (@Prod.mk ℕ ℕ _ _).2 -/
With the information above, users can annotate theorems such as
foo
usingno_index
for problematic subterms. Example:opaque f : Nat → Nat → Nat @[simp] theorem foo : f x (no_index (x, y).2) = y := by sorry example : f a b ≤ b := by simp -- `foo` is still applied with `index := true`
-
#4274 prevents internal
match
equational theorems from appearing in simp trace. -
#4177 and #4359 make
simp
continue even if a simp lemma does not elaborate, if the tactic state is in recovery mode. -
#4341 fixes panic when applying
@[simp]
to malformed theorem syntax. -
#4345 fixes
simp
so that it does not use the forward version of a user-specified backward theorem. -
#4352 adds missing
dsimp
simplifications for fixed parameters of generated congruence theorems. -
#4362 improves trace messages for
simp
so that constants are hoverable.
-
- Elaboration
- #4046 makes subst notation (
he ▸ h
) try rewriting in both directions even when there is no expected type available. - #3328 adds support for identifiers in autoparams (for example,
rfl
in(h : x = y := by exact rfl)
). - #4096 changes how the type in
let
andhave
is elaborated, requiring that any tactics in the type be evaluated before proceeding, improving performance. - #4215 ensures the expression tree elaborator commits to the computed "max type" for the entire arithmetic expression.
- #4267 cases signature elaboration errors to show even if there are parse errors in the body.
- #4368 improves error messages when numeric literals fail to synthesize an
OfNat
instance, including special messages warning when the expected type of the numeral can be a proposition. - #4643 fixes issue leading to nested error messages and info trees vanishing, where snapshot subtrees were not restored on reuse.
- #4657 calculates error suppression per snapshot, letting elaboration errors appear even when there are later parse errors (RFC #3556).
- #4046 makes subst notation (
- Metaprogramming
- Work toward implementing
grind
tactic- 0a515e
and #4164
add
grind_norm
andgrind_norm_proc
attributes and@[grind_norm]
theorems. - #4170, #4221,
and #4249 create
grind
preprocessor and core module. - #4235 and d6709e
add special
cases
tactic togrind
along with@[grind_cases]
attribute to mark types that thiscases
tactic should automatically apply to. - #4243 adds special
injection?
tactic togrind
.
- 0a515e
and #4164
add
- Other fixes or improvements
- #4065 fixes a bug in the
Nat.reduceLeDiff
simproc. - #3969 makes deprecation warnings activate even for generalized field notation ("dot notation").
- #4132 fixes the
sorry
term so that it does not activate the implicit lambda feature - 9803c5
and 47c8e3
move
cdot
andcalc
parsers toLean
namespace. - #4252 fixes the
case
tactic so that it is usable in macros by having it erase macro scopes from the tag. - 26b671
and cc33c3
extract
haveId
syntax. - #4335 fixes bugs in partial
calc
tactic when there is mdata or metavariables. - #4329 makes
termination_by?
report unused each unused parameter as_
.
- #4065 fixes a bug in the
- Docs: #4238, #4294, #4338.
- #4066 fixes features like "Find References" when browsing core Lean sources.
- #4254 allows embedding user widgets in structured messages. Companion PR is vscode-lean4#449.
- #4445 makes watchdog more resilient against badly behaving clients.
- #4059 upstreams many
List
andArray
operations and theorems from Batteries. - #4055 removes the unused
Inhabited
instance forSubtype
. - #3967 adds dates in existing
@[deprecated]
attributes. - #4231 adds boilerplate
Char
,UInt
, andFin
theorems. - #4205 fixes the
MonadStore
type classes to usesemiOutParam
. - #4350 renames
IsLawfulSingleton
toLawfulSingleton
. Nat
Array
- #4074 improves the functional induction principle
Array.feraseIdx.induct
.
- #4074 improves the functional induction principle
List
- #4172 removes
@[simp]
fromList.length_pos
.
- #4172 removes
Option
BitVec
- Theorems: #3920, #4095, #4075, #4148, #4165, #4178, #4200, #4201, #4298, #4299, #4257, #4179, #4321, #4187.
- #4193 adds simprocs for reducing
x >>> i
andx <<< i
wherei
is a bitvector literal. - #4194 adds simprocs for reducing
(x <<< i) <<< j
and(x >>> i) >>> j
wherei
andj
are natural number literals. - #4229 redefines
rotateLeft
/rotateRight
to use modulo reduction of shift offset. - 0d3051 makes
<num>#<term>
bitvector literal notation global.
Char
/String
HashMap
- #4248 fixes implicitness of typeclass arguments in
HashMap.ofList
.
- #4248 fixes implicitness of typeclass arguments in
IO
- #4036 adds
IO.Process.getCurrentDir
andIO.Process.setCurrentDir
for adjusting the current process's working directory.
- #4036 adds
- Cleanup: #4077, #4189, #4304.
- Docs: #4001, #4166, #4332.
- Defeq and WHNF algorithms
- Definition transparency
- #4052 adds validation to application of
@[reducible]
/@[semireducible]
/@[irreducible]
attributes (withlocal
/scoped
modifiers as well). Settingset_option allowUnsafeReductibility true
turns this validation off.
- #4052 adds validation to application of
- Inductive types
- Diagnostics and profiling
- Typeclass resolution
- #4119 fixes multiple issues with TC caching interacting with
synthPendingDepth
, addsmaxSynthPendingDepth
option with default value1
. - #4210 ensures local instance cache does not contain multiple copies of the same instance.
- #4216 fix handling of metavariables, to avoid needing to set the option
backward.synthInstance.canonInstances
tofalse
.
- #4119 fixes multiple issues with TC caching interacting with
- Other fixes or improvements
- #4080 fixes propagation of state for
Lean.Elab.Command.liftCoreM
andLean.Elab.Command.liftTermElabM
. - #3944 makes the
Repr
deriving handler be consistent betweenstructure
andinductive
for how types and proofs are erased. - #4113 propagates
maxHeartbeats
to kernel to control "(kernel) deterministic timeout" error. - #4125 reverts #3970 (monadic generalization of
FindExpr
). - #4128 catches stack overflow in auto-bound implicits feature.
- #4129 adds
tryCatchRuntimeEx
combinator to replacecatchRuntimeEx
reader state. - #4155 simplifies the expression canonicalizer.
- #4151 and #4369 add many missing trace classes.
- #4185 makes congruence theorem generators clean up type annotations of argument types.
- #4192 fixes restoration of infotrees when auto-bound implicit feature is activated, fixing a pretty printing error in hovers and strengthening the unused variable linter.
- dfb496 fixes
declareBuiltin
to allow it to be called multiple times per declaration. - #4569 fixes an issue introduced in a merge conflict, where the interrupt exception was swallowed by some
tryCatchRuntimeEx
uses. - #4584 (backported as b056a0) adapts kernel interruption to the new cancellation system.
- Cleanup: #4112, #4126, #4091, #4139, #4153.
- Tests: 030406, #4133.
- #4080 fixes propagation of state for
- #4100 improves reset/reuse algorithm; it now runs a second pass relaxing the constraint that reused memory cells must only be for the exact same constructor.
- #2903 fixes segfault in old compiler from mishandling
noConfusion
applications. - #4311 fixes bug in constant folding.
- #3915 documents the runtime memory layout for inductive types.
- #4518 makes trace reading more robust. Lake now rebuilds if trace files are invalid or unreadable and is backwards compatible with previous pure numeric traces.
- #4057 adds support for docstrings on
require
commands. - #4088 improves hovers for
family_def
andlibrary_data
commands. - #4147 adds default
README.md
to package templates - #4261 extends
lake test
help page, adds help page forlake check-test
, addslake lint
and tag@[lint_driver]
, adds support for specifying test and lint drivers from dependencies, addstestDriverArgs
andlintDriverArgs
options, adds support for library test drivers, makeslake check-test
andlake check-lint
only load the package without dependencies. - #4270 adds
lake pack
andlake unpack
for packing and unpacking Lake build artifacts from an archive. - #4083
Switches the manifest format to use
major.minor.patch
semantic versions. Major version increments indicate breaking changes (e.g., new required fields and semantic changes to existing fields). Minor version increments (after0.x
) indicate backwards-compatible extensions (e.g., adding optional fields, removing fields). This change is backwards compatible. Lake will still successfully read old manifests with numeric versions. It will treat the numeric versionN
as semantic version0.N.0
. Lake will also accept manifest versions with-
suffixes (e.g.,x.y.z-foo
) and then ignore the suffix. - #4273 adds a lift from
JobM
toFetchM
for backwards compatibility reasons. - #4351 fixes
LogIO
-to-CliM
-lifting performance issues. - #4343 make Lake store the dependency trace for a build in the cached build long and then verifies that it matches the trace of the current build before replaying the log.
- #4402 moves the cached log into the trace file (no more
.log.json
). This means logs are no longer cached on fatal errors and this ensures that an out-of-date log is not associated with an up-to-date trace. Separately,.hash
file generation was changed to be more reliable as well. The.hash
files are deleted as part of the build and always regenerate with--rehash
. - Other fixes or improvements
- #3984 adds a script (
script/rebase-stage0.sh
) forgit rebase -i
that automatically updates each stage0. - #4108 finishes renamings from transition to Std to Batteries.
- #4109 adjusts the Github bug template to mention testing using live.lean-lang.org.
- #4136 makes CI rerun only when
full-ci
label is added or removed. - #4175 and 72b345
switch to using
#guard_msgs
to run tests as much as possible. - #3125 explains the Lean4
pygments
lexer. - #4247 sets up a procedure for preparing release notes.
- #4032 modernizes build instructions and workflows.
- #4255 moves some expensive checks from merge queue to releases.
- #4265 adds aarch64 macOS as native compilation target for CI.
- f05a82 restores macOS aarch64 install suffix in CI
- #4317 updates build instructions for macOS.
- #4333 adjusts workflow to update Batteries in manifest when creating
lean-pr-testing-NNNN
Mathlib branches. - #4355 simplifies
lean4checker
step of release checklist. - #4361 adds installing elan to
pr-release
CI step. - #4628 fixes the Windows build, which was missing an exported symbol.
While most changes could be considered to be a breaking change, this section makes special note of API changes.
Nat.zero_or
andNat.or_zero
have been swapped (#4094).IsLawfulSingleton
is nowLawfulSingleton
(#4350).- The
BitVec
literal notation is now<num>#<term>
rather than<term>#<term>
, and it is global rather than scoped. UseBitVec.ofNat w x
rather thanx#w
whenx
is a not a numeric literal (0d3051). BitVec.rotateLeft
andBitVec.rotateRight
now take the shift modulo the bitwidth (#4229).- These are no longer simp lemmas:
List.length_pos
(#4172),Option.bind_eq_some
(#4314). - Types in
let
andhave
(both the expressions and tactics) may fail to elaborate due to new restrictions on what sorts of elaboration problems may be postponed (#4096). In particular, tactics embedded in the type will no longer make use of the type ofvalue
in expressions such aslet x : type := value; body
. - Now functions defined by well-founded recursion are marked with
@[irreducible]
by default (#4061). Existing proofs that hold by definitional equality (e.g.rfl
) can be rewritten to explicitly unfold the function definition (usingsimp
,unfold
,rw
), or the recursive function can be temporarily made semireducible (usingunseal f in
before the command), or the function definition itself can be marked as@[semireducible]
to get the previous behavior. - Due to #3929:
-
The
MessageData.ofPPFormat
constructor has been removed. Its functionality has been split into two:- for lazy structured messages, please use
MessageData.lazy
; - for embedding
Format
orFormatWithInfos
, useMessageData.ofFormatWithInfos
.
An example migration can be found in #3929.
- for lazy structured messages, please use
-
The
MessageData.ofFormat
constructor has been turned into a function. If you need to inspectMessageData
, you can pattern-match onMessageData.ofFormatWithInfos
.
-
-
Functional induction principles. #3432, #3620, #3754, #3762, #3738, #3776, #3898.
Derived from the definition of a (possibly mutually) recursive function, a functional induction principle is created that is tailored to proofs about that function.
For example from:
def ackermann : Nat → Nat → Nat | 0, m => m + 1 | n+1, 0 => ackermann n 1 | n+1, m+1 => ackermann n (ackermann (n + 1) m)
we get
ackermann.induct (motive : Nat → Nat → Prop) (case1 : ∀ (m : Nat), motive 0 m) (case2 : ∀ (n : Nat), motive n 1 → motive (Nat.succ n) 0) (case3 : ∀ (n m : Nat), motive (n + 1) m → motive n (ackermann (n + 1) m) → motive (Nat.succ n) (Nat.succ m)) (x x : Nat) : motive x x
It can be used in the
induction
tactic using theusing
syntax:induction n, m using ackermann.induct
-
The termination checker now recognizes more recursion patterns without an explicit
termination_by
. In particular the idiom of counting up to an upper bound, as indef Array.sum (arr : Array Nat) (i acc : Nat) : Nat := if _ : i < arr.size then Array.sum arr (i+1) (acc + arr[i]) else acc
is recognized without having to say
termination_by arr.size - i
. -
#3629, #3655, #3747: Adds
@[induction_eliminator]
and@[cases_eliminator]
attributes to be able to define custom eliminators for theinduction
andcases
tactics, replacing the@[eliminator]
attribute. Gives custom eliminators forNat
so thatinduction
andcases
put goal states into terms of0
andn + 1
rather thanNat.zero
andNat.succ n
. Added optiontactic.customEliminators
to control whether to use custom eliminators. Added a hack forrcases
/rintro
/obtain
to use the custom eliminator forNat
. -
Shorter instances names. There is a new algorithm for generating names for anonymous instances. Across Std and Mathlib, the median ratio between lengths of new names and of old names is about 72%. With the old algorithm, the longest name was 1660 characters, and now the longest name is 202 characters. The new algorithm's 95th percentile name length is 67 characters, versus 278 for the old algorithm. While the new algorithm produces names that are 1.2% less unique, it avoids cross-project collisions by adding a module-based suffix when it does not refer to declarations from the same "project" (modules that share the same root). #3089 and #3934.
-
8d2adf Importing two different files containing proofs of the same theorem is no longer considered an error. This feature is particularly useful for theorems that are automatically generated on demand (e.g., equational theorems).
-
84b091 Lean now generates an error if the type of a theorem is not a proposition.
-
Definition transparency. 47a343.
@[reducible]
,@[semireducible]
, and@[irreducible]
are now scoped and able to be set for imported declarations. -
simp
/dsimp
- #3607 enables kernel projection reduction in
dsimp
- b24fbf
and acdb00:
dsimproc
command to define defeq-preserving simplification procedures. - #3624 makes
dsimp
normalize raw nat literals asOfNat.ofNat
applications. - #3628 makes
simp
correctly handleOfScientific.ofScientific
literals. - #3654 makes
dsimp?
report used simprocs. - dee074 fixes equation theorem
handling in
simp
for non-recursive definitions. - #3819 improved performance when simp encounters a loop.
- #3821 fixes discharger/cache interaction.
- #3824 keeps
simp
from breakingChar
literals. - #3838 allows
Nat
instances matching to be more lenient. - #3870 documentation for
simp
configuration options. - #3972 fixes simp caching.
- #4044 improves cache behavior for "well-behaved" dischargers.
- #3607 enables kernel projection reduction in
-
omega
-
rfl
-
#3719 upstreams the
rw?
tactic, with fixes and improvements in #3783, #3794, #3911. -
conv
-
#guard_msgs
- #3617 introduces whitespace protection using the
⏎
character. - #3883:
The
#guard_msgs
command now has options to change whitespace normalization and sensitivity to message ordering. For example,#guard_msgs (whitespace := lax) in cmd
collapses whitespace before checking messages, and#guard_msgs (ordering := sorted) in cmd
sorts the messages in lexicographic order before checking. - #3931 adds an unused variables ignore function for
#guard_msgs
. - #3912 adds a diff between the expected and actual outputs. This feature is currently
disabled by default, but can be enabled with
set_option guard_msgs.diff true
. Depending on user feedback, this option may default totrue
in a future version of Lean.
- #3617 introduces whitespace protection using the
-
do
notation- #3820 makes it an error to lift
(<- ...)
out of a pureif ... then ... else ...
- #3820 makes it an error to lift
-
Lazy discrimination trees
- #3610 fixes a name collision for
LazyDiscrTree
that could lead to cache poisoning. - #3677 simplifies and fixes
LazyDiscrTree
handling forexact?
/apply?
. - #3685 moves general
exact?
/apply?
functionality intoLazyDiscrTree
. - #3769 has lemma selection improvements for
rw?
andLazyDiscrTree
. - #3818 improves ordering of matches.
- #3610 fixes a name collision for
-
#3590 adds
inductive.autoPromoteIndices
option to be able to disable auto promotion of indices in theinductive
command. -
Miscellaneous bug fixes and improvements
- #3606 preserves
cache
anddischargeDepth
fields inLean.Meta.Simp.Result.mkEqSymm
. - #3633 makes
elabTermEnsuringType
respecterrToSorry
, improving error recovery of thehave
tactic. - #3647 enables
noncomputable unsafe
definitions, for deferring implementations until later. - #3672 adjust namespaces of tactics.
- #3725 fixes
Ord
derive handler for indexed inductive types with unused alternatives. - #3893 improves performance of derived
Ord
instances. - #3771 changes error reporting for failing tactic macros. Improves
rfl
error message. - #3745 fixes elaboration of generalized field notation if the object of the notation is an optional parameter.
- #3799 makes commands such as
universe
,variable
,namespace
, etc. require that their argument appear in a later column. Commands that can optionally parse anident
or parse any number ofident
s generally should require that theident
usecolGt
. This keeps typos in commands from being interpreted as identifiers. - #3815 lets the
split
tactic be used for writing code. - #3822 adds missing info in
induction
tactic forwith
clauses of the form| cstr a b c => ?_
. - #3806 fixes
withSetOptionIn
combinator. - #3844 removes unused
trace.Elab.syntax
option. - #3896 improves hover and go-to-def for
attribute
command. - #3989 makes linter options more discoverable.
- #3916 fixes go-to-def for syntax defined with
@[builtin_term_parser]
. - #3962 fixes how
solveByElim
handlessymm
lemmas, makingexact?
/apply?
usable again. - #3968 improves the
@[deprecated]
attribute, adding(since := "<date>")
field. - #3768 makes
#print
command show structure fields. - #3974 makes
exact?%
behave likeby exact?
rather thanby apply?
. - #3994 makes elaboration of
he ▸ h
notation more predictable. - #3991 adjusts transparency for
decreasing_trivial
macros. - #4092 improves performance of
binop%
andbinrel%
expression tree elaborators.
- #3606 preserves
-
Docs: #3748, #3796, #3800, #3874, #3863, #3862, #3891, #3873, #3908, #3872.
- #3602 enables
import
auto-completions. - #3608 fixes issue leanprover/vscode-lean4#392. Diagnostic ranges had an off-by-one error that would misplace goal states for example.
- #3014 introduces snapshot trees, foundational work for incremental tactics and parallelism. #3849 adds basic incrementality API.
- #3271 adds support for server-to-client requests.
- #3656 fixes jump to definition when there are conflicting names from different files. Fixes issue #1170.
- #3691, #3925, #3932 keep semantic tokens synchronized (used for semantic highlighting), with performance improvements.
- #3247 and #3730 add diagnostics to run "Restart File" when a file dependency is saved.
- #3722 uses the correct module names when displaying references.
- #3728 makes errors in header reliably appear and makes the "Import out of date" warning be at "hint" severity. #3739 simplifies the text of this warning.
- #3778 fixes #3462, where info nodes from before the cursor would be used for computing completions.
- #3985 makes trace timings appear in Infoview.
- #3797 fixes the hovers over binders so that they show their types.
- #3640 and #3735: Adds attribute
@[pp_using_anonymous_constructor]
to make structures pretty print as⟨x, y, z⟩
rather than as{a := x, b := y, c := z}
. This attribute is applied toSigma
,PSigma
,PProd
,Subtype
,And
, andFin
. - #3749
Now structure instances pretty print with parent structures' fields inlined.
That is, if
B
extendsA
, then{ toA := { x := 1 }, y := 2 }
now pretty prints as{ x := 1, y := 2 }
. Setting optionpp.structureInstances.flatten
to false turns this off. - #3737, #3744
and #3750:
Option
pp.structureProjections
is renamed topp.fieldNotation
, and there is now a suboptionpp.fieldNotation.generalized
to enable pretty printing function applications using generalized field notation (defaults to true). Field notation can be disabled on a function-by-function basis using the@[pp_nodot]
attribute. The notation is not used for theorems. - #4071 fixes interaction between app unexpanders and
pp.fieldNotation.generalized
- #3625 makes
delabConstWithSignature
(used by#check
) have the ability to put arguments "after the colon" to avoid printing inaccessible names. - #3798,
#3978,
#3798:
Adds options
pp.mvars
(default: true) andpp.mvars.withType
(default: false). Whenpp.mvars
is false, expression metavariables pretty print as?_
and universe metavariables pretty print as_
. Whenpp.mvars.withType
is true, expression metavariables pretty print with a type ascription. These can be set when using#guard_msgs
to make tests not depend on the particular names of metavariables. - #3917 makes binders hoverable and gives them docstrings.
- #4034 makes hovers for RHS terms in
match
expressions in the Infoview reliably show the correct term.
-
Bool
/Prop
-
Nat
-
Int
- Theorems: #3890
-
UInt
s- #3960 improves performance of upcasting.
-
Array
andSubarray
-
List
- #3785 upstreams tail-recursive List operations and
@[csimp]
lemmas.
- #3785 upstreams tail-recursive List operations and
-
BitVec
-
String
-
IO
- #4097 adds
IO.getTaskState
which returns whether a task is finished, actively running, or waiting on other Tasks to finish.
- #4097 adds
-
Refactors
- #3605 reduces imports for
Init.Data.Nat
andInit.Data.Int
. - #3613 reduces imports for
Init.Omega.Int
. - #3634 upstreams
Std.Data.Nat
and #3635 upstreamsStd.Data.Int
. - #3790 reduces more imports for
omega
. - #3694 extends
GetElem
interface withgetElem!
andgetElem?
to simplify containers likeRBMap
. - #3865 renames
Option.toMonad
(see breaking changes below). - #3882 unifies
lexOrd
withcompareLex
.
- #3605 reduces imports for
-
Other fixes or improvements
- #3765 makes
Quotient.sound
be atheorem
. - #3645 fixes
System.FilePath.parent
in the case of absolute paths. - #3660
ByteArray.toUInt64LE!
andByteArray.toUInt64BE!
were swapped. - #3881, #3887 fix linearity issues in
HashMap.insertIfNew
,HashSet.erase
, andHashMap.erase
. TheHashMap.insertIfNew
fix improvesimport
performance. - #3830 ensures linearity in
Parsec.many*Core
. - #3930 adds
FS.Stream.isTty
field. - #3866 deprecates
Option.toBool
in favor ofOption.isSome
. - #3975 upstreams
Data.List.Init
andData.Array.Init
material from Std. - #3942 adds instances that make
ac_rfl
work without Mathlib. - #4010 changes
Fin.induction
to use structural induction. - 02753f
fixes bug in
reduceLeDiff
simproc. - #4097
adds
IO.TaskState
andIO.getTaskState
to get the task from the Lean runtime's task manager.
- #3765 makes
-
Docs: #3615, #3664, #3707, #3734, #3868, #3861, #3869, #3858, #3856, #3857, #3867, #3864, #3860, #3859, #3871, #3919.
- Defeq and WHNF algorithms
- #3616 gives better support for reducing
Nat.rec
expressions. - #3774 add tracing for "non-easy" WHNF cases.
- #3807 fixes an
isDefEq
performance issue, now trying structure eta after lazy delta reduction. - #3816 fixes
.yesWithDeltaI
behavior to prevent increasing transparency level when reducing projections. - #3837 improves heuristic at
isDefEq
. - #3965 improves
isDefEq
for constraints of the formt.i =?= s.i
. - #3977 improves
isDefEqProj
. - #3981 adds universe constraint approximations to be able to solve
u =?= max u ?v
using?v = u
. These approximations are only applied when universe constraints cannot be postponed anymore. - #4004 improves
isDefEqProj
during typeclass resolution. - #4012 adds
backward.isDefEq.lazyProjDelta
andbackward.isDefEq.lazyWhnfCore
backwards compatibility flags.
- #3616 gives better support for reducing
- Kernel
- Discrimination trees
- Typeclass instance synthesis
- Definition processing
- #3661, #3767 changes automatically generated equational theorems to be named
using suffix
.eq_<idx>
instead of._eq_<idx>
, and.eq_def
instead of._unfold
. (See breaking changes below.) #3675 adds a mechanism to reserve names. #3803 fixes reserved name resolution inside namespaces and fixes handling ofmatch
er declarations and equation lemmas. - #3662 causes auxiliary definitions nested inside theorems to become
def
s if they are not proofs. - #4006 makes proposition fields of
structure
s be theorems. - #4018 makes it an error for a theorem to be
extern
. - #4047 improves performance making equations for well-founded recursive definitions.
- #3661, #3767 changes automatically generated equational theorems to be named
using suffix
- Refactors
- #3614 avoids unfolding in
Lean.Meta.evalNat
. - #3621 centralizes functionality for
Fix
/GuessLex
/FunInd
in theArgsPacker
module. - #3186 rewrites the UnusedVariable linter to be more performant.
- #3589 removes coercion from
String
toName
(see breaking changes below). - #3237 removes the
lines
field fromFileMap
. - #3951 makes msg parameter to
throwTacticEx
optional.
- #3614 avoids unfolding in
- Diagnostics
- #4016, #4019,
#4020, #4030,
#4031,
c3714b,
#4049 adds
set_option diagnostics true
for diagnostic counters. Tracks number of unfolded declarations, instances, reducible declarations, used instances, recursor reductions,isDefEq
heuristic applications, among others. This option is suggested in exceptional situations, such as at deterministic timeout and maximum recursion depth. - 283587
adds diagnostic information for
simp
. - #4043 adds diagnostic information for congruence theorems.
- #4048 display diagnostic information
for
set_option diagnostics true in <tactic>
andset_option diagnostics true in <term>
.
- #4016, #4019,
#4020, #4030,
#4031,
c3714b,
#4049 adds
- Other features
- #3800 adds environment extension to record which definitions use structural or well-founded recursion.
- #3801
trace.profiler
can now export to Firefox Profiler. - #3918, #3953 adds
@[builtin_doc]
attribute to make docs and location of a declaration available as a builtin. - #3939 adds the
lean --json
CLI option to print messages as JSON. - #3075 improves
test_extern
command. - #3970 gives monadic generalization of
FindExpr
.
- Docs: #3743, #3921, #3954.
- Other fixes: #3622, #3726, #3823, #3897, #3964, #3946, #4007, #4026.
- #3632 makes it possible to allocate and free thread-local runtime resources for threads not started by Lean itself.
- #3627 improves error message about compacting closures.
- #3692 fixes deadlock in
IO.Promise.resolve
. - #3753 catches error code from
MoveFileEx
on Windows. - #4028 fixes a double
reset
bug inResetReuse
transformation. - 6e731b
removes
interpreter
copy constructor to avoid potential memory safety issues.
-
TOML Lake configurations. #3298, #4104.
Lake packages can now use TOML as a alternative configuration file format instead of Lean. If the default
lakefile.lean
is missing, Lake will also look for alakefile.toml
. The TOML version of the configuration supports a restricted set of the Lake configuration options, only including those which can easily mapped to a TOML data structure. The TOML syntax itself fully compiles with the TOML v1.0.0 specification.As part of the introduction of this new feature, we have been helping maintainers of some major packages within the ecosystem switch to this format. For example, the following is Aesop's new
lakefile.toml
:leanprover-community/aesop/lakefile.toml
name = "aesop" defaultTargets = ["Aesop"] testRunner = "test" precompileModules = false [[require]] name = "batteries" git = "https://github.com/leanprover-community/batteries" rev = "main" [[lean_lib]] name = "Aesop" [[lean_lib]] name = "AesopTest" globs = ["AesopTest.+"] leanOptions = {linter.unusedVariables = false} [[lean_exe]] name = "test" srcDir = "scripts"
To assist users who wish to transition their packages between configuration file formats, there is also a new
lake translate-config
command for migrating to/from TOML.Running
lake translate-config toml
will produce alakefile.toml
version of a package'slakefile.lean
. Any configuration options unsupported by the TOML format will be discarded during translation, but the originallakefile.lean
will remain so that you can verify the translation looks good before deleting it. -
Build progress overhaul. #3835, #4115, #4127, #4220, #4232, #4236.
Builds are now managed by a top-level Lake build monitor, this makes the output of Lake builds more standardized and enables producing prettier and more configurable progress reports.
As part of this change, job isolation has improved. Stray I/O and other build related errors in custom targets are now properly isolated and caught as part of their job. Import errors no longer cause Lake to abort the entire build and are instead localized to the build jobs of the modules in question.
Lake also now uses ANSI escape sequences to add color and produce progress lines that update in-place; this can be toggled on and off using
--ansi
/--no-ansi
.--wfail
and--iofail
options have been added that causes a build to fail if any of the jobs log a warning (--wfail
) or produce any output or log information messages (--iofail
). Unlike some other build systems, these options do NOT convert these logs into errors, and Lake does not abort jobs on such a log (i.e., dependent jobs will still continue unimpeded). -
lake test
. #3779.Lake now has a built-in
test
command which will run a script or executable labelled@[test_runner]
(in Lean) or defined as thetestRunner
(in TOML) in the root package.Lake also provides a
lake check-test
command which will exit with code0
if the package has a properly configured test runner or error with1
otherwise. -
lake lean
. #3793.The new command
lake lean <file> [-- <args...>]
functions likelake env lean <file> <args...>
, except that it builds the imports offile
before runninglean
. This makes it very useful for running test or example code that imports modules that are not guaranteed to have been built beforehand. -
Miscellaneous bug fixes and improvements
- #3609
LEAN_GITHASH
environment variable to override the detected Git hash for Lean when computing traces, useful for testing custom builds of Lean. - #3795 improves relative package directory path normalization in the pre-rename check.
- #3957 fixes handling of packages that appear multiple times in a dependency tree.
- #3999 makes it an error for there to be a mismatch between a package name and what it is required as. Also adds a special message for the
std
-to-batteries
rename. - #4033 fixes quiet mode.
- #3609
-
Docs: #3704.
-
#3600 runs nix-ci more uniformly.
-
#3612 avoids argument limits when building on Windows.
-
#3682 builds Lean's
.o
files in parallel to rest of core. -
#3601 changes the way Lean is built on Windows (see breaking changes below). As a result, Lake now dynamically links executables with
supportInterpreter := true
on Windows tolibleanshared.dll
andlibInit_shared.dll
. Therefore, such executables will not run unless those shared libraries are co-located with the executables or part ofPATH
. Running the executable vialake exe
will ensure these libraries are part ofPATH
.In a related change, the signature of the
nativeFacets
Lake configuration options has changed from a staticArray
to a function(shouldExport : Bool) → Array
. See its docstring or Lake's README for further details on the changed option. -
#3690 marks "Build matrix complete" as canceled if the build is canceled.
-
#3700, #3702, #3701, #3834, #3923: fixes and improvements for std and mathlib CI.
-
#3712 fixes
nix build .
on macOS. -
#3717 replaces
shell.nix
in devShell withflake.nix
. -
#3971 prevents stage0 changes via the merge queue.
-
#3979 adds handling for
changes-stage0
label. -
#3952 adds a script to summarize GitHub issues.
-
18a699 fixes asan linking
-
Due to the major Lake build refactor, code using the affected parts of the Lake API or relying on the previous output format of Lake builds is likely to have been broken. We have tried to minimize the breakages and, where possible, old definitions have been marked
@[deprecated]
with a reference to the new alternative. -
Executables configured with
supportInterpreter := true
on Windows should now be run vialake exe
to function properly. -
Automatically generated equational theorems are now named using suffix
.eq_<idx>
instead of._eq_<idx>
, and.eq_def
instead of._unfold
. Example:
def fact : Nat → Nat
| 0 => 1
| n+1 => (n+1) * fact n
theorem ex : fact 0 = 1 := by unfold fact; decide
#check fact.eq_1
-- fact.eq_1 : fact 0 = 1
#check fact.eq_2
-- fact.eq_2 (n : Nat) : fact (Nat.succ n) = (n + 1) * fact n
#check fact.eq_def
/-
fact.eq_def :
∀ (x : Nat),
fact x =
match x with
| 0 => 1
| Nat.succ n => (n + 1) * fact n
-/
-
The coercion from
String
toName
was removed. Previously, it wasName.mkSimple
, which does not separate strings at dots, but experience showed that this is not always the desired coercion. For the previous behavior, manually insert a call toName.mkSimple
. -
The
Subarray
fieldsas
,h₁
andh₂
have been renamed toarray
,start_le_stop
, andstop_le_array_size
, respectively. This more closely follows standard Lean conventions. Deprecated aliases for the field projections were added; these will be removed in a future release. -
The change to the instance name algorithm (described above) can break projects that made use of the auto-generated names.
-
Option.toMonad
has been renamed toOption.getM
and the unneeded[Monad m]
instance argument has been removed.
-
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 formlet 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 language (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 available in the core language.- 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, while
HAdd.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 bynomatch ...
andnofun
. #3279 and #3286
Other improvements:
- several bug fixes for
simp
: - fixes for
match
expressions: - improve
termination_by
error messages #3255 - fix
rename_i
in macros, fixes #3553 #3581 - fix excessive resource usage in
generalize
, fixes #3524 #3575 - an equation lemma with autoParam arguments fails to rewrite, fixing #2243 #3316
add_decl_doc
should check that declarations are local #3311- instantiate the types of inductives with the right parameters, closing #3242 #3246
- New simprocs for many basic types. #3407
Lake fixes:
- Backport of #3552 fixing a performance regression in server startup.
-
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 simplified 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` constructor 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 whole
mutual
block. - Therefore, the function name no longer has to be mentioned in the hint.
- If the function has a
where
clause, thetermination_by
anddecreasing_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
after the whole
-
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, use
all_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 owndecreasing_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- allow generalization in
let
#3060, fixing #3065 - reducing out-of-bounds
swap!
should returna
, not `default`` #3197, fixing #3196 - derive
BEq
on structure withProp
-fields #3191, fixing #3140 - refine through more
casesOnApp
/matcherApp
#3176, fixing #3175 - do not strip dotted components from lean module names #2994, fixing #2999
- fix
deriving
only deriving the first declaration for some handlers #3058, fixing #3057 - do not instantiate metavariables in kabstract/rw for disallowed occurrences #2539, fixing #2538
- hover info for
cases h : ...
#3084
-
Modify the lexical syntax of string literals to have string gaps, which are escape sequences of the form
"\" newline whitespace*
. These have the interpretation 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 library 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. With
set_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.
-
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.
-
simp [f]
does not unfold partial applications off
anymore. See issue #2042. To fix proofs affected by this change, useunfold f
orsimp (config := { unfoldPartialApp := true }) [f]
. -
By default,
simp
will no longer try to use Decidable instances to rewrite terms. In particular, not all decidable goals will be closed bysimp
, and thedecide
tactic may be useful in such cases. Thedecide
simp configuration option can be used to locally restore the oldsimp
behavior, as insimp (config := {decide := true})
; this includes using Decidable instances to verify side goals such as numeric inequalities. -
Many bug fixes:
- Add left/right actions to term tree coercion elaborator and make `^`` a right action
- Fix for #2775, don't catch max recursion depth errors
- Reduction of
Decidable
instances very slow when usingcases
tactic simp
not rewriting in bindersimp
unfoldinglet
even withzeta := false
optionsimp
(with beta/zeta disabled) and discrimination trees- unknown free variable introduced by
rw ... at h
dsimp
doesn't userfl
theorems which consist of an unapplied constantdsimp
does not close reflexive equality goals if they are wrapped in metadatarw [h]
usesh
from the environment in preference toh
from the local context- missing
withAssignableSyntheticOpaque
forassumption
tactic - ignoring default value for field warning
-
Cancel outstanding tasks on document edit in the language server.
Lake:
- Sensible defaults for
lake new MyProject math
- Changed
postUpdate?
configuration option to apost_update
declaration. See thepost_update
syntax docstring for more information on the new syntax. - A manifest is automatically created on workspace load if one does not exists..
- The
:=
syntax for configuration declarations (i.e.,package
,lean_lib
, andlean_exe
) has been deprecated. For example,package foo := {...}
is deprecated. - support for overriding package URLs via
LAKE_PKG_URL_MAP
- Moved the default build directory (e.g.,
build
), default packages directory (e.g.,lake-packages
), and the compiled configuration (e.g.,lakefile.olean
) into a new dedicated directory for Lake outputs,.lake
. The cloud release build archives are also stored here, fixing #2713. - Update manifest format to version 7 (see lean4#2801 for details on the changes).
- Deprecate the
manifestFile
field of a package configuration. - There is now a more rigorous check on
lakefile.olean
compatibility (see #2842 for more details).
- isDefEq cache for terms not containing metavariables..
- Make
Environment.mk
andEnvironment.add
private, and addreplay
as a safer alternative. IO.Process.output
no longer inherits the standard input of the caller.- Do not inhibit caching of default-level
match
reduction. - List the valid case tags when the user writes an invalid one.
- The derive handler for
DecidableEq
now handles mutual inductive types. - Show path of failed import in Lake.
- Fix linker warnings on macOS.
- Lake: Add
postUpdate?
package configuration option. Used by a package to specify some code which should be run after a successfullake update
of the package or one of its downstream dependencies. (lake#185) - Improvements to Lake startup time (#2572, #2573)
refine e
now replaces the main goal with metavariables which were created during elaboration ofe
and no longer captures pre-existing metavariables that occur ine
(#2502).- This is accomplished via changes to
withCollectingNewGoalsFrom
, which also affectselabTermWithHoles
,refine'
,calc
(tactic), andspecialize
. Likewise, all of these now only include newly-created metavariables in their output. - Previously, both newly-created and pre-existing metavariables occurring in
e
were returned inconsistently in different edge cases, causing duplicated goals in the infoview (issue #2495), erroneously closed goals (issue #2434), and unintuitive behavior due torefine e
capturing previously-created goals appearing unexpectedly ine
(no issue; see PR).
- This is accomplished via changes to
-
The error positioning on missing tokens has been improved. In particular, this should make it easier to spot errors in incomplete tactic proofs.
-
After elaborating a configuration file, Lake will now cache the configuration to a
lakefile.olean
. Subsequent runs of Lake will import this OLean instead of elaborating the configuration file. This provides a significant performance improvement (benchmarks indicate that using the OLean cuts Lake's startup time in half), but there are some important details to keep in mind:- Lake will regenerate this OLean after each modification to the
lakefile.lean
orlean-toolchain
. You can also force a reconfigure by passing the new--reconfigure
/-R
option tolake
. - Lake configuration options (i.e.,
-K
) will be fixed at the moment of elaboration. Setting these options whenlake
is using the cached configuration will have no effect. To change options, runlake
with-R
/--reconfigure
. - The
lakefile.olean
is a local configuration and should not be committed to Git. Therefore, existing Lake packages need to add it to their.gitignore
.
- Lake will regenerate this OLean after each modification to the
-
The signature of
Lake.buildO
has changed,args
has been split intoweakArgs
andtraceArgs
.traceArgs
are included in the input trace andweakArgs
are not. See Lake's FFI example for a demonstration of how to adapt to this change. -
The signatures of
Lean.importModules
,Lean.Elab.headerToImports
, andLean.Elab.parseImports
have changed from takingList Import
toArray Import
. -
There is now an
occs
field in the configuration object for therewrite
tactic, allowing control of which occurrences of a pattern should be rewritten. This was previously a separate argument forLean.MVarId.rewrite
, and this has been removed in favour of an additional field ofRewrite.Config
. It was not previously accessible from user tactics.
-
Lean.Meta.getConst?
has been renamed. We have renamedgetConst?
togetUnfoldableConst?
(andgetConstNoEx?
togetUnfoldableConstNoEx?
). These were not intended to be part of the public API, but downstream projects had been using them (sometimes expecting different behaviour) incorrectly instead ofLean.getConstInfo
. -
dsimp
/simp
/simp_all
now fail by default if they make no progress.This can be overridden with the
(config := { failIfUnchanged := false })
option. This change was made to ease manual use ofsimp
(with complicated goals it can be hard to tell if it was effective) and to allow easier flow control in tactics internally usingsimp
. See the summary discussion on zulip for more details. -
simp_all
now preserves order of hypotheses.In order to support the
failIfUnchanged
configuration option fordsimp
/simp
/simp_all
the waysimp_all
replaces hypotheses has changed. In particular it is now more likely to preserve the order of hypotheses. Seesimp_all
reorders hypotheses unnecessarily. (Previously all non-dependent propositional hypotheses were reverted and reintroduced. Now only such hypotheses which were changed, or which come after a changed hypothesis, are reverted and reintroduced. This has the effect of preserving the ordering amongst the non-dependent propositional hypotheses, but now any dependent or non-propositional hypotheses retain their position amongst the unchanged non-dependent propositional hypotheses.) This may affect proofs that userename_i
,case ... =>
, ornext ... =>
. -
this
is now a regular identifier again that is implicitly introduced by anonymoushave :=
for the remainder of the tactic block. It used to be a keyword that was visible in all scopes and led to unexpected behavior when explicitly used as a binder name. -
Make
calc
require the sequence of relation/proof-s to have the same indentation, and addcalc
alternative syntax allowing underscores_
in the first relation.The flexible indentation in
calc
was often used to align the relation symbols:example (x y : Nat) : (x + y) * (x + y) = x * x + y * x + x * y + y * y := calc (x + y) * (x + y) = (x + y) * x + (x + y) * y := by rw [Nat.mul_add] -- improper indentation _ = x * x + y * x + (x + y) * y := by rw [Nat.add_mul] _ = x * x + y * x + (x * y + y * y) := by rw [Nat.add_mul] _ = x * x + y * x + x * y + y * y := by rw [←Nat.add_assoc]
This is no longer legal. The new syntax puts the first term right after the
calc
and each step has the same indentation:example (x y : Nat) : (x + y) * (x + y) = x * x + y * x + x * y + y * y := calc (x + y) * (x + y) _ = (x + y) * x + (x + y) * y := by rw [Nat.mul_add] _ = x * x + y * x + (x + y) * y := by rw [Nat.add_mul] _ = x * x + y * x + (x * y + y * y) := by rw [Nat.add_mul] _ = x * x + y * x + x * y + y * y := by rw [←Nat.add_assoc]
-
Update Lake to latest prerelease.
-
Make go-to-definition on a typeclass projection application go to the instance(s).
-
Add
linter.deprecated
option to silence deprecation warnings. -
simp
can track information and can print an equivalentsimp only
. PR #1626. -
Enforce uniform indentation in tactic blocks / do blocks. See issue #1606.
-
Moved
AssocList
,HashMap
,HashSet
,RBMap
,RBSet
,PersistentArray
,PersistentHashMap
,PersistentHashSet
to the Lean package. The standard library contains versions that will evolve independently to simplify bootstrapping process. -
Standard library moved to the std4 GitHub repository.
-
InteractiveGoals
now has information that a client infoview can use to show what parts of the goal have changed after applying a tactic. PR #1610. -
Add
[inheritDoc]
attribute. PR #1480. -
Expose that
panic = default
. PR #1614. -
New code generator project has started.
-
Remove description argument from
register_simp_attr
. PR #1566. -
Many new doc strings have been added to declarations at
Init
.
-
Update Lake to v4.0.0. See the v4.0.0 release notes for detailed changes.
-
Mutual declarations in different namespaces are now supported. Example:
mutual def Foo.boo (x : Nat) := match x with | 0 => 1 | x + 1 => 2*Boo.bla x def Boo.bla (x : Nat) := match x with | 0 => 2 | x+1 => 3*Foo.boo x end
A
namespace
is automatically created for the common prefix. Example:mutual def Tst.Foo.boo (x : Nat) := ... def Tst.Boo.bla (x : Nat) := ... end
expands to
namespace Tst mutual def Foo.boo (x : Nat) := ... def Boo.bla (x : Nat) := ... end end Tst
-
Allow users to install their own
deriving
handlers for existing type classes. See example at Simple.lean. -
Add tactic
congr (num)?
. See doc string for additional details. -
match
-syntax notation now checks for unused alternatives. See issue #1371. -
Auto-completion for structure instance fields. Example:
example : Nat × Nat := { f -- HERE }
fst
now appears in the list of auto-completion suggestions. -
Auto-completion for dotted identifier notation. Example:
example : Nat := .su -- HERE
succ
now appears in the list of auto-completion suggestions. -
nat_lit
is not needed anymore when declaringOfNat
instances. See issues #1389 and #875. Example:inductive Bit where | zero | one instance inst0 : OfNat Bit 0 where ofNat := Bit.zero instance : OfNat Bit 1 where ofNat := Bit.one example : Bit := 0 example : Bit := 1
-
Add
[elabAsElim]
attribute (it is calledelab_as_eliminator
in Lean 3). Motivation: simplify the Mathlib port to Lean 4. -
Trans
type class now accepts relations inType u
. See this Zulip issue. -
Accept unescaped keywords as inductive constructor names. Escaping can often be avoided at use sites via dot notation.
inductive MyExpr | let : ... def f : MyExpr → MyExpr | .let ... => .let ...
-
Throw an error message at parametric local instances such as
[Nat -> Decidable p]
. The type class resolution procedure cannot use this kind of local instance because the parameter does not have a forward dependency. This check can be disabled usingset_option checkBinderAnnotations false
. -
Add option
pp.showLetValues
. When set tofalse
, the info view hides the value oflet
-variables in a goal. By default, it istrue
when visualizing tactic goals, andfalse
otherwise. See issue #1345 for additional details. -
Add option
warningAsError
. When set to true, warning messages are treated as errors. -
Support dotted notation and named arguments in patterns. Example:
def getForallBinderType (e : Expr) : Expr := match e with | .forallE (binderType := type) .. => type | _ => panic! "forall expected"
-
"jump-to-definition" now works for function names embedded in the following attributes
@[implementedBy funName]
,@[tactic parserName]
,@[termElab parserName]
,@[commandElab parserName]
,@[builtinTactic parserName]
,@[builtinTermElab parserName]
, and@[builtinCommandElab parserName]
. See issue #1350. -
Improve
MVarId
methods discoverability. See issue #1346. We still have to add similar methods forFVarId
,LVarId
,Expr
, and other objects. Many existing methods have been marked as deprecated. -
Add attribute
[deprecated]
for marking deprecated declarations. Examples:def g (x : Nat) := x + 1 -- Whenever `f` is used, a warning message is generated suggesting to use `g` instead. @[deprecated g] def f (x : Nat) := x + 1 #check f 0 -- warning: `f` has been deprecated, use `g` instead -- Whenever `h` is used, a warning message is generated. @[deprecated] def h (x : Nat) := x + 1 #check h 0 -- warning: `h` has been deprecated
-
Add type
LevelMVarId
(and abbreviationLMVarId
) for universe level metavariable ids. Motivation: prevent meta-programmers from mixing up universe and expression metavariable ids. -
Improve
calc
term and tactic. See issue #1342. -
Relaxed antiquotation parsing further reduces the need for explicit
$x:p
antiquotation kind annotations. -
Add support for computed fields in inductives. Example:
inductive Exp | var (i : Nat) | app (a b : Exp) with @[computedField] hash : Exp → Nat | .var i => i | .app a b => a.hash * b.hash + 1
The result of the
Exp.hash
function is then stored as an extra "computed" field in the.var
and.app
constructors;Exp.hash
accesses this field and thus runs in constant time (even on dag-like values). -
Update
a[i]
notation. It is now based on the typeclassclass GetElem (cont : Type u) (idx : Type v) (elem : outParam (Type w)) (dom : outParam (cont → idx → Prop)) where getElem (xs : cont) (i : idx) (h : dom xs i) : Elem
The notation
a[i]
is now defined as followsmacro:max x:term noWs "[" i:term "]" : term => `(getElem $x $i (by get_elem_tactic))
The proof that
i
is a valid index is synthesized using the tacticget_elem_tactic
. For example, the typeArray α
has the following instancesinstance : GetElem (Array α) Nat α fun xs i => LT.lt i xs.size where ... instance : GetElem (Array α) USize α fun xs i => LT.lt i.toNat xs.size where ...
You can use the notation
a[i]'h
to provide the proof manually. Two other notations were introduced:a[i]!
anda[i]?
, Fora[i]!
, a panic error message is produced at runtime ifi
is not a valid index.a[i]?
has typeOption α
, anda[i]?
evaluates tonone
if the indexi
is not valid. The three new notations are defined as follows:@[inline] def getElem' [GetElem cont idx elem dom] (xs : cont) (i : idx) (h : dom xs i) : elem := getElem xs i h @[inline] def getElem! [GetElem cont idx elem dom] [Inhabited elem] (xs : cont) (i : idx) [Decidable (dom xs i)] : elem := if h : _ then getElem xs i h else panic! "index out of bounds" @[inline] def getElem? [GetElem cont idx elem dom] (xs : cont) (i : idx) [Decidable (dom xs i)] : Option elem := if h : _ then some (getElem xs i h) else none macro:max x:term noWs "[" i:term "]" noWs "?" : term => `(getElem? $x $i) macro:max x:term noWs "[" i:term "]" noWs "!" : term => `(getElem! $x $i) macro x:term noWs "[" i:term "]'" h:term:max : term => `(getElem' $x $i $h)
See discussion on Zulip. Examples:
example (a : Array Int) (i : Nat) : Int := a[i] -- Error: failed to prove index is valid ... example (a : Array Int) (i : Nat) (h : i < a.size) : Int := a[i] -- Ok example (a : Array Int) (i : Nat) : Int := a[i]! -- Ok example (a : Array Int) (i : Nat) : Option Int := a[i]? -- Ok example (a : Array Int) (h : a.size = 2) : Int := a[0]'(by rw [h]; decide) -- Ok example (a : Array Int) (h : a.size = 2) : Int := have : 0 < a.size := by rw [h]; decide have : 1 < a.size := by rw [h]; decide a[0] + a[1] -- Ok example (a : Array Int) (i : USize) (h : i.toNat < a.size) : Int := a[i] -- Ok
The
get_elem_tactic
is defined asmacro "get_elem_tactic" : tactic => `(first | get_elem_tactic_trivial | fail "failed to prove index is valid, ..." )
The
get_elem_tactic_trivial
auxiliary tactic can be extended usingmacro_rules
. By default, it triestrivial
,simp_arith
, and a special case forFin
. In the future, it will also trylinarith
. You can extendget_elem_tactic_trivial
usingmy_tactic
as followsmacro_rules | `(tactic| get_elem_tactic_trivial) => `(tactic| my_tactic)
Note that
Idx
's type inGetElem
does not depend onCont
. So, you cannot write the instanceinstance : GetElem (Array α) (Fin ??) α fun xs i => ...
, but the Lean library comes equipped with the following auxiliary instance:instance [GetElem cont Nat elem dom] : GetElem cont (Fin n) elem fun xs i => dom xs i where getElem xs i h := getElem xs i.1 h
and helper tactic
macro_rules | `(tactic| get_elem_tactic_trivial) => `(tactic| apply Fin.val_lt_of_le; get_elem_tactic_trivial; done)
Example:
example (a : Array Nat) (i : Fin a.size) := a[i] -- Ok example (a : Array Nat) (h : n ≤ a.size) (i : Fin n) := a[i] -- Ok
-
Better support for qualified names in recursive declarations. The following is now supported:
namespace Nat def fact : Nat → Nat | 0 => 1 | n+1 => (n+1) * Nat.fact n end Nat
-
Add support for
CommandElabM
monad at#eval
. Example:import Lean open Lean Elab Command #eval do let id := mkIdent `foo elabCommand (← `(def $id := 10)) #eval foo -- 10
-
Try to elaborate
do
notation even if the expected type is not available. We still delay elaboration when the expected type is not available. This change is particularly useful when writing examples such as#eval do IO.println "hello" IO.println "world"
That is, we don't have to use the idiom
#eval show IO _ from do ...
anymore. Note that auto monadic lifting is less effective when the expected type is not available. Monadic polymorphic functions (e.g.,ST.Ref.get
) also require the expected type. -
On Linux, panics now print a backtrace by default, which can be disabled by setting the environment variable
LEAN_BACKTRACE
to0
. Other platforms are TBD. -
The
group(·)
syntax
combinator is now introduced automatically where necessary, such as when using multiple parsers inside(...)+
. -
Add "Typed Macros": syntax trees produced and accepted by syntax antiquotations now remember their syntax kinds, preventing accidental production of ill-formed syntax trees and reducing the need for explicit
:kind
antiquotation annotations. See PR for details. -
Aliases of protected definitions are protected too. Example:
protected def Nat.double (x : Nat) := 2*x namespace Ex export Nat (double) -- Add alias Ex.double for Nat.double end Ex open Ex #check Ex.double -- Ok #check double -- Error, `Ex.double` is alias for `Nat.double` which is protected
-
Use
IO.getRandomBytes
to initialize random seed forIO.rand
. See discussion at this PR. -
Improve dot notation and aliases interaction. See discussion on Zulip for additional details. Example:
def Set (α : Type) := α → Prop def Set.union (s₁ s₂ : Set α) : Set α := fun a => s₁ a ∨ s₂ a def FinSet (n : Nat) := Fin n → Prop namespace FinSet export Set (union) -- FinSet.union is now an alias for `Set.union` end FinSet example (x y : FinSet 10) : FinSet 10 := x.union y -- Works
-
ext
andenter
conv tactics can now go inside let-declarations. Example:example (g : Nat → Nat) (y : Nat) (h : let x := y + 1; g (0+x) = x) : g (y + 1) = y + 1 := by conv at h => enter [x, 1, 1]; rw [Nat.zero_add] /- g : Nat → Nat y : Nat h : let x := y + 1; g x = x ⊢ g (y + 1) = y + 1 -/ exact h
-
Add
zeta
conv tactic to expand let-declarations. Example:example (h : let x := y + 1; 0 + x = y) : False := by conv at h => zeta; rw [Nat.zero_add] /- y : Nat h : y + 1 = y ⊢ False -/ simp_arith at h
-
Improve namespace resolution. See issue #1224. Example:
import Lean open Lean Parser Elab open Tactic -- now opens both `Lean.Parser.Tactic` and `Lean.Elab.Tactic`
-
Rename
constant
command toopaque
. See discussion at Zulip. -
Extend
induction
andcases
syntax: multiple left-hand-sides in a single alternative. This extension is very similar to the one implemented formatch
expressions. Examples:inductive Foo where | mk1 (x : Nat) | mk2 (x : Nat) | mk3 def f (v : Foo) := match v with | .mk1 x => x + 1 | .mk2 x => 2*x + 1 | .mk3 => 1 theorem f_gt_zero : f v > 0 := by cases v with | mk1 x | mk2 x => simp_arith! -- New feature used here! | mk3 => decide
-
Add unnamed antiquotation
$_
for use in syntax quotation patterns. -
Add unused variables linter. Feedback welcome!
-
Lean now generates an error if the body of a declaration body contains a universe parameter that does not occur in the declaration type, nor is an explicit parameter. Examples:
/- The following declaration now produces an error because `PUnit` is universe polymorphic, but the universe parameter does not occur in the function type `Nat → Nat` -/ def f (n : Nat) : Nat := let aux (_ : PUnit) : Nat := n + 1 aux ⟨⟩ /- The following declaration is accepted because the universe parameter was explicitly provided in the function signature. -/ def g.{u} (n : Nat) : Nat := let aux (_ : PUnit.{u}) : Nat := n + 1 aux ⟨⟩
-
Add
subst_vars
tactic. -
Fix
autoParam
in structure fields lost in multiple inheritance.. -
Add
[eliminator]
attribute. It allows users to specify default recursor/eliminators for theinduction
andcases
tactics. It is an alternative for theusing
notation. Example:@[eliminator] protected def recDiag {motive : Nat → Nat → Sort u} (zero_zero : motive 0 0) (succ_zero : (x : Nat) → motive x 0 → motive (x + 1) 0) (zero_succ : (y : Nat) → motive 0 y → motive 0 (y + 1)) (succ_succ : (x y : Nat) → motive x y → motive (x + 1) (y + 1)) (x y : Nat) : motive x y := let rec go : (x y : Nat) → motive x y | 0, 0 => zero_zero | x+1, 0 => succ_zero x (go x 0) | 0, y+1 => zero_succ y (go 0 y) | x+1, y+1 => succ_succ x y (go x y) go x y termination_by go x y => (x, y) def f (x y : Nat) := match x, y with | 0, 0 => 1 | x+1, 0 => f x 0 | 0, y+1 => f 0 y | x+1, y+1 => f x y termination_by f x y => (x, y) example (x y : Nat) : f x y > 0 := by induction x, y <;> simp [f, *]
-
Add support for
casesOn
applications to structural and well-founded recursion modules. This feature is useful when writing definitions using tactics. Example:inductive Foo where | a | b | c | pair: Foo × Foo → Foo def Foo.deq (a b : Foo) : Decidable (a = b) := by cases a <;> cases b any_goals apply isFalse Foo.noConfusion any_goals apply isTrue rfl case pair a b => let (a₁, a₂) := a let (b₁, b₂) := b exact match deq a₁ b₁, deq a₂ b₂ with | isTrue h₁, isTrue h₂ => isTrue (by rw [h₁,h₂]) | isFalse h₁, _ => isFalse (fun h => by cases h; cases (h₁ rfl)) | _, isFalse h₂ => isFalse (fun h => by cases h; cases (h₂ rfl))
-
Option
is again a monad. The auxiliary typeOptionM
has been removed. See Zulip thread. -
Improve
split
tactic. It used to fail onmatch
expressions of the formmatch h : e with ...
wheree
is not a free variable. The failure used to occur during generalization. -
New encoding for
match
-expressions that use theh :
notation for discriminants. The information is not lost during delaboration, and it is the foundation for a bettersplit
tactic. at delaboration time. Example:#print Nat.decEq /- protected def Nat.decEq : (n m : Nat) → Decidable (n = m) := fun n m => match h : Nat.beq n m with | true => isTrue (_ : n = m) | false => isFalse (_ : ¬n = m) -/
-
exists
tactic is now takes a comma separated list of terms. -
Add
dsimp
anddsimp!
tactics. They guarantee the result term is definitionally equal, and only applyrfl
-theorems. -
Fix binder information for
match
patterns that use definitions tagged with[matchPattern]
(e.g.,Nat.add
). We now have proper binder information for the variabley
in the following example.def f (x : Nat) : Nat := match x with | 0 => 1 | y + 1 => y
-
(Fix) the default value for structure fields may now depend on the structure parameters. Example:
structure Something (i: Nat) where n1: Nat := 1 n2: Nat := 1 + i def s : Something 10 := {} example : s.n2 = 11 := rfl
-
Apply
rfl
theorems at thedsimp
auxiliary method used bysimp
.dsimp
can be used anywhere in an expression because it preserves definitional equality. -
Refine auto bound implicit feature. It does not consider anymore unbound variables that have the same name of a declaration being defined. Example:
def f : f → Bool := -- Error at second `f` fun _ => true inductive Foo : List Foo → Type -- Error at second `Foo` | x : Foo []
Before this refinement, the declarations above would be accepted and the second
f
andFoo
would be treated as auto implicit variables. That is,f : {f : Sort u} → f → Bool
, andFoo : {Foo : Type u} → List Foo → Type
. -
Fix syntax highlighting for recursive declarations. Example
inductive List (α : Type u) where | nil : List α -- `List` is not highlighted as a variable anymore | cons (head : α) (tail : List α) : List α def List.map (f : α → β) : List α → List β | [] => [] | a::as => f a :: map f as -- `map` is not highlighted as a variable anymore
-
Add
autoUnfold
option toLean.Meta.Simp.Config
, and the following macrossimp!
forsimp (config := { autoUnfold := true })
simp_arith!
forsimp (config := { autoUnfold := true, arith := true })
simp_all!
forsimp_all (config := { autoUnfold := true })
simp_all_arith!
forsimp_all (config := { autoUnfold := true, arith := true })
When the
autoUnfold
is set to true,simp
tries to unfold the following kinds of definition- Recursive definitions defined by structural recursion.
- Non-recursive definitions where the body is a
match
-expression. This kind of definition is only unfolded if thematch
can be reduced. Example:
def append (as bs : List α) : List α := match as with | [] => bs | a :: as => a :: append as bs theorem append_nil (as : List α) : append as [] = as := by induction as <;> simp_all! theorem append_assoc (as bs cs : List α) : append (append as bs) cs = append as (append bs cs) := by induction as <;> simp_all!
-
Add
save
tactic for creating checkpoints more conveniently. Example:example : <some-proposition> := by tac_1 tac_2 save tac_3 ...
is equivalent to
example : <some-proposition> := by checkpoint tac_1 tac_2 tac_3 ...
-
Remove support for
{}
annotation from inductive datatype constructors. This annotation was barely used, and we can control the binder information for parameter bindings using the new inductive family indices to parameter promotion. Example: the following declaration using{}
inductive LE' (n : Nat) : Nat → Prop where | refl {} : LE' n n -- Want `n` to be explicit | succ : LE' n m → LE' n (m+1)
can now be written as
inductive LE' : Nat → Nat → Prop where | refl (n : Nat) : LE' n n | succ : LE' n m → LE' n (m+1)
In both cases, the inductive family has one parameter and one index. Recall that the actual number of parameters can be retrieved using the command
#print
. -
Remove support for
{}
annotation in thestructure
command. -
Several improvements to LSP server. Examples: "jump to definition" in mutually recursive sections, fixed incorrect hover information in "match"-expression patterns, "jump to definition" for pattern variables, fixed auto-completion in function headers, etc.
-
In
macro ... xs:p* ...
and similar macro bindings of combinators,xs
now has the correct typeArray Syntax
-
Identifiers in syntax patterns now ignore macro scopes during matching.
-
Improve binder names for constructor auto implicit parameters. Example, given the inductive datatype
inductive Member : α → List α → Type u | head : Member a (a::as) | tail : Member a bs → Member a (b::bs)
before:
#check @Member.head -- @Member.head : {x : Type u_1} → {a : x} → {as : List x} → Member a (a :: as)
now:
#check @Member.head -- @Member.head : {α : Type u_1} → {a : α} → {as : List α} → Member a (a :: as)
-
Improve error message when constructor parameter universe level is too big.
-
Add support for
for h : i in [start:stop] do ..
whereh : i ∈ [start:stop]
. This feature is useful for proving termination of functions such as:inductive Expr where | app (f : String) (args : Array Expr) def Expr.size (e : Expr) : Nat := Id.run do match e with | app f args => let mut sz := 1 for h : i in [: args.size] do -- h.upper : i < args.size sz := sz + size (args.get ⟨i, h.upper⟩) return sz
-
Add tactic
case'
. It is similar tocase
, but does not admit the goal on failure. For example, the new tactic is useful when writing tactic scripts where we need to usecase'
atfirst | ... | ...
, and we want to take the next alternative whencase'
fails. -
Add tactic macro
macro "stop" s:tacticSeq : tactic => `(repeat sorry)
See discussion on Zulip.
-
When displaying goals, we do not display inaccessible proposition names if they do not have forward dependencies. We still display their types. For example, the goal
case node.inl.node β : Type u_1 b : BinTree β k : Nat v : β left : Tree β key : Nat value : β right : Tree β ihl : BST left → Tree.find? (Tree.insert left k v) k = some v ihr : BST right → Tree.find? (Tree.insert right k v) k = some v h✝ : k < key a✝³ : BST left a✝² : ForallTree (fun k v => k < key) left a✝¹ : BST right a✝ : ForallTree (fun k v => key < k) right ⊢ BST left
is now displayed as
case node.inl.node β : Type u_1 b : BinTree β k : Nat v : β left : Tree β key : Nat value : β right : Tree β ihl : BST left → Tree.find? (Tree.insert left k v) k = some v ihr : BST right → Tree.find? (Tree.insert right k v) k = some v : k < key : BST left : ForallTree (fun k v => k < key) left : BST right : ForallTree (fun k v => key < k) right ⊢ BST left
-
The hypothesis name is now optional in the
by_cases
tactic. -
Fix inconsistency between
syntax
and kind names. The node kindsnumLit
,charLit
,nameLit
,strLit
, andscientificLit
are now callednum
,char
,name
,str
, andscientific
respectively. Example: we now writemacro_rules | `($n:num) => `("hello")
instead of
macro_rules | `($n:numLit) => `("hello")
-
(Experimental) New
checkpoint <tactic-seq>
tactic for big interactive proofs. -
Rename tactic
nativeDecide
=>native_decide
. -
Antiquotations are now accepted in any syntax. The
incQuotDepth
syntax
parser is therefore obsolete and has been removed. -
Renamed tactic
nativeDecide
=>native_decide
. -
"Cleanup" local context before elaborating a
match
alternative right-hand-side. Examples:example (x : Nat) : Nat := match g x with | (a, b) => _ -- Local context does not contain the auxiliary `_discr := g x` anymore example (x : Nat × Nat) (h : x.1 > 0) : f x > 0 := by match x with | (a, b) => _ -- Local context does not contain the `h✝ : x.fst > 0` anymore
-
Improve
let
-pattern (andhave
-pattern) macro expansion. In the following example,example (x : Nat × Nat) : f x > 0 := by let (a, b) := x done
The resulting goal is now
... |- f (a, b) > 0
instead of... |- f x > 0
. -
Add cross-compiled aarch64 Linux and aarch64 macOS releases.
-
Add tutorial-like examples to our documentation, rendered using LeanInk+Alectryon.
-
simp
now takes user-defined simp-attributes. You can define a newsimp
attribute by creating a file (e.g.,MySimp.lean
) containingimport Lean open Lean.Meta initialize my_ext : SimpExtension ← registerSimpAttr `my_simp "my own simp attribute"
If you don't need to access
my_ext
, you can also use the macroimport Lean register_simp_attr my_simp "my own simp attribute"
Recall that the new
simp
attribute is not active in the Lean file where it was defined. Here is a small example using the new feature.import MySimp def f (x : Nat) := x + 2 def g (x : Nat) := x + 1 @[my_simp] theorem f_eq : f x = x + 2 := rfl @[my_simp] theorem g_eq : g x = x + 1 := rfl example : f x + g x = 2*x + 3 := by simp_arith [my_simp]
-
Extend
match
syntax: multiple left-hand-sides in a single alternative. Example:def fib : Nat → Nat | 0 | 1 => 1 | n+2 => fib n + fib (n+1)
This feature was discussed at issue 371. It was implemented as a macro expansion. Thus, the following is accepted.
inductive StrOrNum where | S (s : String) | I (i : Int) def StrOrNum.asString (x : StrOrNum) := match x with | I a | S a => toString a
-
Improve
#eval
command. Now, when it fails to synthesize aLean.MetaEval
instance for the result type, it reduces the type and tries again. The following example now works without additional annotationsdef Foo := List Nat def test (x : Nat) : Foo := [x, x+1, x+2] #eval test 4
-
rw
tactic can now apply auto-generated equation theorems for a given definition. Example:example (a : Nat) (h : n = 1) : [a].length = n := by rw [List.length] trace_state -- .. |- [].length + 1 = n rw [List.length] trace_state -- .. |- 0 + 1 = n rw [h]
-
Extend dot-notation
x.field
for arrow types. If type ofx
is an arrow, we look up forFunction.field
. For example, givenf : Nat → Nat
andg : Nat → Nat
,f.comp g
is now notation forFunction.comp f g
. -
The new
.<identifier>
notation is now also accepted where a function type is expected.example (xs : List Nat) : List Nat := .map .succ xs example (xs : List α) : Std.RBTree α ord := xs.foldl .insert ∅
-
Support notation
let <pattern> := <expr> | <else-case>
indo
blocks. -
Remove support for "auto"
pure
. In the Zulip thread, the consensus seemed to be that "auto"pure
is more confusing than it's worth. -
Remove restriction in
congr
theorems that all function arguments on the left-hand-side must be free variables. For example, the following theorem is now a validcongr
theorem.@[congr] theorem dep_congr [DecidableEq ι] {p : ι → Set α} [∀ i, Inhabited (p i)] : ∀ {i j} (h : i = j) (x : p i) (y : α) (hx : x = y), Pi.single (f := (p ·)) i x = Pi.single (f := (p ·)) j ⟨y, hx ▸ h ▸ x.2⟩ :=
-
Improve elaboration postponement heuristic when expected type is a metavariable. Lean now reduces the expected type before performing the test.
-
Remove deprecated leanpkg in favor of Lake now bundled with Lean.
-
Various improvements to go-to-definition & find-all-references accuracy.
-
Auto generated congruence lemmas with support for casts on proofs and
Decidable
instances (see wishlist). -
Rename option
autoBoundImplicitLocal
=>autoImplicit
. -
Relax auto-implicit restrictions. The command
set_option relaxedAutoImplicit false
disables the relaxations. -
contradiction
tactic now closes the goal if there is aFalse.elim
application in the target. -
Renamed tatic
byCases
=>by_cases
(motivation: enforcing naming convention). -
Local instances occurring in patterns are now considered by the type class resolution procedure. Example:
def concat : List ((α : Type) × ToString α × α) → String | [] => "" | ⟨_, _, a⟩ :: as => toString a ++ concat as
-
Notation for providing the motive for
match
expressions has changed. before:match x, rfl : (y : Nat) → x = y → Nat with | 0, h => ... | x+1, h => ...
now:
match (motive := (y : Nat) → x = y → Nat) x, rfl with | 0, h => ... | x+1, h => ...
With this change, the notation for giving names to equality proofs in
match
-expressions is not whitespace sensitive anymore. That is, we can now writematch h : sort.swap a b with | (r₁, r₂) => ... -- `h : sort.swap a b = (r₁, r₂)`
-
(generalizing := true)
is the default behavior formatch
expressions even if the expected type is not a proposition. In the following example, we used to have to include(generalizing := true)
manually.inductive Fam : Type → Type 1 where | any : Fam α | nat : Nat → Fam Nat example (a : α) (x : Fam α) : α := match x with | Fam.any => a | Fam.nat n => n
-
We now use
PSum
(instead ofSum
) when compiling mutually recursive definitions using well-founded recursion. -
Better support for parametric well-founded relations. See issue #1017. This change affects the low-level
termination_by'
hint because the fixed prefix of the function parameters in not "packed" anymore when constructing the well-founded relation type. For example, in the following definition,as
is part of the fixed prefix, and is not packed anymore. In previous versions, thetermination_by'
term would be written asmeasure fun ⟨as, i, _⟩ => as.size - i
def sum (as : Array Nat) (i : Nat) (s : Nat) : Nat := if h : i < as.size then sum as (i+1) (s + as.get ⟨i, h⟩) else s termination_by' measure fun ⟨i, _⟩ => as.size - i
-
Add
while <cond> do <do-block>
,repeat <do-block>
, andrepeat <do-block> until <cond>
macros fordo
-block. These macros are based onpartial
definitions, and consequently are useful only for writing programs we don't want to prove anything about. -
Add
arith
option toSimp.Config
, the macrosimp_arith
expands tosimp (config := { arith := true })
. OnlyNat
and linear arithmetic is currently supported. Example:example : 0 < 1 + x ∧ x + y + 2 ≥ y + 1 := by simp_arith
-
Add
fail <string>?
tactic that always fail. -
Add support for acyclicity at dependent elimination. See issue #1022.
-
Add
trace <string>
tactic for debugging purposes. -
Add nontrivial
SizeOf
instance for typesUnit → α
, and add support for them in the auto-generatedSizeOf
instances for user-defined inductive types. For example, given the inductive datatypeinductive LazyList (α : Type u) where | nil : LazyList α | cons (hd : α) (tl : LazyList α) : LazyList α | delayed (t : Thunk (LazyList α)) : LazyList α
we now have
sizeOf (LazyList.delayed t) = 1 + sizeOf t
instead ofsizeOf (LazyList.delayed t) = 2
. -
Add support for guessing (very) simple well-founded relations when proving termination. For example, the following function does not require a
termination_by
annotation anymore.def Array.insertAtAux (i : Nat) (as : Array α) (j : Nat) : Array α := if h : i < j then let as := as.swap! (j-1) j; insertAtAux i as (j-1) else as
-
Add support for
for h : x in xs do ...
notation whereh : x ∈ xs
. This is mainly useful for showing termination. -
Auto implicit behavior changed for inductive families. An auto implicit argument occurring in inductive family index is also treated as an index (IF it is not fixed, see next item). For example
inductive HasType : Index n → Vector Ty n → Ty → Type where
is now interpreted as
inductive HasType : {n : Nat} → Index n → Vector Ty n → Ty → Type where
-
To make the previous feature more convenient to use, we promote a fixed prefix of inductive family indices to parameters. For example, the following declaration is now accepted by Lean
inductive Lst : Type u → Type u | nil : Lst α | cons : α → Lst α → Lst α
and
α
inLst α
is a parameter. The actual number of parameters can be inspected using the command#print Lst
. This feature also makes sure we still accept the declarationinductive Sublist : List α → List α → Prop | slnil : Sublist [] [] | cons l₁ l₂ a : Sublist l₁ l₂ → Sublist l₁ (a :: l₂) | cons2 l₁ l₂ a : Sublist l₁ l₂ → Sublist (a :: l₁) (a :: l₂)
-
Added auto implicit "chaining". Unassigned metavariables occurring in the auto implicit types now become new auto implicit locals. Consider the following example:
inductive HasType : Fin n → Vector Ty n → Ty → Type where | stop : HasType 0 (ty :: ctx) ty | pop : HasType k ctx ty → HasType k.succ (u :: ctx) ty
ctx
is an auto implicit local in the two constructors, and it has typectx : Vector Ty ?m
. Without auto implicit "chaining", the metavariable?m
will remain unassigned. The new feature creates yet another implicit localn : Nat
and assignsn
to?m
. So, the declaration above is shorthand forinductive HasType : {n : Nat} → Fin n → Vector Ty n → Ty → Type where | stop : {ty : Ty} → {n : Nat} → {ctx : Vector Ty n} → HasType 0 (ty :: ctx) ty | pop : {n : Nat} → {k : Fin n} → {ctx : Vector Ty n} → {ty : Ty} → HasType k ctx ty → HasType k.succ (u :: ctx) ty
-
Eliminate auxiliary type annotations (e.g,
autoParam
andoptParam
) from recursor minor premises and projection declarations. Consider the following examplestructure A := x : Nat h : x = 1 := by trivial example (a : A) : a.x = 1 := by have aux := a.h -- `aux` has now type `a.x = 1` instead of `autoParam (a.x = 1) auto✝` exact aux example (a : A) : a.x = 1 := by cases a with | mk x h => -- `h` has now type `x = 1` instead of `autoParam (x = 1) auto✝` assumption
-
We now accept overloaded notation in patterns, but we require the set of pattern variables in each alternative to be the same. Example:
inductive Vector (α : Type u) : Nat → Type u | nil : Vector α 0 | cons : α → Vector α n → Vector α (n+1) infix:67 " :: " => Vector.cons -- Overloading the `::` notation def head1 (x : List α) (h : x ≠ []) : α := match x with | a :: as => a -- `::` is `List.cons` here def head2 (x : Vector α (n+1)) : α := match x with | a :: as => a -- `::` is `Vector.cons` here
-
New notation
.<identifier>
based on Swift. The namespace is inferred from the expected type. See issue #944. Examples:def f (x : Nat) : Except String Nat := if x > 0 then .ok x else .error "x is zero" namespace Lean.Elab open Lsp def identOf : Info → Option (RefIdent × Bool) | .ofTermInfo ti => match ti.expr with | .const n .. => some (.const n, ti.isBinder) | .fvar id .. => some (.fvar id, ti.isBinder) | _ => none | .ofFieldInfo fi => some (.const fi.projName, false) | _ => none def isImplicit (bi : BinderInfo) : Bool := bi matches .implicit end Lean.Elab