Skip to content

Releases: leanprover/lean4

v4.8.0-rc2

22 May 00:07
Compare
Choose a tag to compare
v4.8.0-rc2 Pre-release
Pre-release

Update to v4.8.0-rc1 with improvements to lake's build monitor.

v4.8.0-rc1

02 May 11:26
Compare
Choose a tag to compare
v4.8.0-rc1 Pre-release
Pre-release
chore: set release flag

v4.7.0

03 Apr 07:28
Compare
Choose a tag to compare

Changes since v4.6.0 (from RELEASES.md)

  • simp and rw now use instance arguments found by unification,
    rather than always resynthesizing. For backwards compatibility, the original behaviour is
    available via set_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.
    The pp.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 an app_unexpander is no longer necessary.
    #3495.

  • New simp (and dsimp) configuration option: zetaDelta. It is false by default.
    The zeta option is still true by default, but their meaning has changed.

    • When zeta := true, simp and dsimp reduce terms of the form
      let x := val; e[x] into e[val].
    • When zetaDelta := true, simp and dsimp will expand let-variables in
      the context. For example, suppose the context contains x := val. Then,
      any occurrence of x is replaced with val.

    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 adding new local theorems to simp, the system assumes that the function application arguments
    have been annotated with no_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 because simp does not index f-arguments anymore when adding h to the simp-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 a termination_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:

    1. Making universally useful tactics such as ext, by_cases, change at,
      norm_cast, rcases, simpa, simp?, omega, and exact?
      available to all users of Lean, without imports.
    2. Minimizing the syntactic changes between plain Lean and Lean with import Std.
    3. Simplifying the development process for the basic data types
      Nat, Int, Fin (and variants such as UInt64), List, Array,
      and BitVec as we begin making the APIs and simp normal forms for these types
      more complete and consistent.
    4. Laying the groundwork for the Std roadmap, as a library focused on
      essential datatypes not provided by the core langauge (e.g. RBMap)
      and utilities such as basic IO.
      While we have achieved most of our initial aims in v4.7.0-rc1,
      some upstreaming will continue over the coming months.
  • The / and % notations in Int now use Int.ediv and Int.emod
    (i.e. the rounding conventions have changed).
    Previously Std overrode these notations, so this is no change for users of Std.
    There is now kernel support for these functions.
    #3376.

  • omega, our integer linear arithmetic tactic, is now availabe in the core langauge.

    • It is supplemented by a preprocessing tactic bv_omega which can solve goals about BitVec
      which naturally translate into linear arithmetic problems.
      #3435.
    • omega now has support for Fin #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 prove id x ≤ x). #3525.
      This may cause some regressions.
      We plan to provide a general purpose preprocessing tactic later, or an omega! 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 while omega does help automate some proofs, we plan to make this much more robust.
  • The library search tactics exact? and apply? that were originally in
    Mathlib are now available in Lean itself. These use the implementation using
    lazy discrimination trees from Std, 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 from Std 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 (particularly simp) 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 over f + g shows f. There is no way to fix the situation from within an app unexpander; the expression position for HAdd.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, the f + g gets TermInfo registered for that subexpression, making it properly hoverable.

    #3375

Breaking changes:

  • Lean.withTraceNode and variants got a stronger MonadAlwaysExcept assumption to
    fix trace trees not being built on elaboration runtime exceptions. Instances for most elaboration
    monads built on EIO Exception should be synthesized automatically.
  • The match ... with. and fun. notations previously in Std have been replaced by
    nomatch ... and nofun. #3279 and #3286

Other improvements:

  • several bug fixes for simp:
    • we should not crash when simp loops #3269
    • simp gets stuck on autoParam #3315
    • simp fails when custom discharger makes no progress #3317
    • simp fails to discharge autoParam premises even when it can reduce them to True #3314
    • simp? suggests generated equations lemma names, fixes #3547 #3573
  • fixes for match expressions:
    • fix regression with builtin literals #3521...
Read more

v4.7.0-rc2

06 Mar 03:58
Compare
Choose a tag to compare
v4.7.0-rc2 Pre-release
Pre-release

Identical to v4.7.0-rc1, but with the backport of #3610, fixing a bug when exact? is used in Mathlib.

v4.7.0-rc1

04 Mar 14:03
Compare
Choose a tag to compare
v4.7.0-rc1 Pre-release
Pre-release

Changes since v4.6.0 (from RELEASES.md)

  • simp and rw now use instance arguments found by unification,
    rather than always resynthesizing. For backwards compatibility, the original behaviour is
    available via set_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.
    The pp.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 an app_unexpander is no longer necessary.
    #3495.

  • New simp (and dsimp) configuration option: zetaDelta. It is false by default.
    The zeta option is still true by default, but their meaning has changed.

    • When zeta := true, simp and dsimp reduce terms of the form
      let x := val; e[x] into e[val].
    • When zetaDelta := true, simp and dsimp will expand let-variables in
      the context. For example, suppose the context contains x := val. Then,
      any occurrence of x is replaced with val.

    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 adding new local theorems to simp, the system assumes that the function application arguments
    have been annotated with no_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 because simp does not index f-arguments anymore when adding h to the simp-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 a termination_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:

    1. Making universally useful tactics such as ext, by_cases, change at,
      norm_cast, rcases, simpa, simp?, omega, and exact?
      available to all users of Lean, without imports.
    2. Minimizing the syntactic changes between plain Lean and Lean with import Std.
    3. Simplifying the development process for the basic data types
      Nat, Int, Fin (and variants such as UInt64), List, Array,
      and BitVec as we begin making the APIs and simp normal forms for these types
      more complete and consistent.
    4. Laying the groundwork for the Std roadmap, as a library focused on
      essential datatypes not provided by the core langauge (e.g. RBMap)
      and utilities such as basic IO.
      While we have achieved most of our initial aims in v4.7.0-rc1,
      some upstreaming will continue over the coming months.
  • The / and % notations in Int now use Int.ediv and Int.emod
    (i.e. the rounding conventions have changed).
    Previously Std overrode these notations, so this is no change for users of Std.
    There is now kernel support for these functions.
    #3376.

  • omega, our integer linear arithmetic tactic, is now availabe in the core langauge.

    • It is supplemented by a preprocessing tactic bv_omega which can solve goals about BitVec
      which naturally translate into linear arithmetic problems.
      #3435.
    • omega now has support for Fin #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 prove id x ≤ x). #3525.
      This may cause some regressions.
      We plan to provide a general purpose preprocessing tactic later, or an omega! 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 while omega does help automate some proofs, we plan to make this much more robust.
  • The library search tactics exact? and apply? that were originally in
    Mathlib are now available in Lean itself. These use the implementation using
    lazy discrimination trees from Std, 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 from Std 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 (particularly simp) 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 over f + g shows f. There is no way to fix the situation from within an app unexpander; the expression position for HAdd.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, the f + g gets TermInfo registered for that subexpression, making it properly hoverable.

    #3375

Breaking changes:

  • Lean.withTraceNode and variants got a stronger MonadAlwaysExcept assumption to
    fix trace trees not being built on elaboration runtime exceptions. Instances for most elaboration
    monads built on EIO Exception should be synthesized automatically.
  • The match ... with. and fun. notations previously in Std have been replaced by
    nomatch ... and nofun. #3279 and #3286

Other improvements:

  • several bug fixes for simp:
    • we should not crash when simp loops #3269
    • simp gets stuck on autoParam #3315
    • simp fails when custom discharger makes no progress #3317
    • simp fails to discharge autoParam premises even when it can reduce them to True #3314
    • simp? suggests generated equations lemma names, fixes #3547 #3573
  • fixes for match expressions:
    • fix regression with builtin literals #3521...
Read more

v4.6.1

04 Mar 13:46
Compare
Choose a tag to compare

Changes since v4.6.0 (from RELEASES.md)

  • Backport of #3552 fixing a performance regression
    in server startup.

v4.6.0

28 Feb 23:47
Compare
Choose a tag to compare

Changes since v4.5.0 (from RELEASES.md)

  • Add custom simplification procedures (aka simprocs) to simp. Simprocs can be triggered by the simplifier on a specified term-pattern. Here is an small example:

    import Lean.Meta.Tactic.Simp.BuiltinSimprocs.Nat
    
    def foo (x : Nat) : Nat :=
      x + 10
    
    /--
    The `simproc` `reduceFoo` is invoked on terms that match the pattern `foo _`.
    -/
    simproc reduceFoo (foo _) :=
      /- A term of type `Expr → SimpM Step -/
      fun e => do
        /-
        The `Step` type has three constructors: `.done`, `.visit`, `.continue`.
        * The constructor `.done` instructs `simp` that the result does
          not need to be simplied further.
        * The constructor `.visit` instructs `simp` to visit the resulting expression.
        * The constructor `.continue` instructs `simp` to try other simplification procedures.
    
        All three constructors take a `Result`. The `.continue` contructor may also take `none`.
        `Result` has two fields `expr` (the new expression), and `proof?` (an optional proof).
         If the new expression is definitionally equal to the input one, then `proof?` can be omitted or set to `none`.
        -/
        /- `simp` uses matching modulo reducibility. So, we ensure the term is a `foo`-application. -/
        unless e.isAppOfArity ``foo 1 do
          return .continue
        /- `Nat.fromExpr?` tries to convert an expression into a `Nat` value -/
        let some n ← Nat.fromExpr? e.appArg!
          | return .continue
        return .done { expr := Lean.mkNatLit (n+10) }

    We disable simprocs support by using the command set_option simprocs false. This command is particularly useful when porting files to v4.6.0.
    Simprocs can be scoped, manually added to simp commands, and suppressed using -. They are also supported by simp?. simp only does not execute any simproc. Here are some examples for the simproc defined above.

    example : x + foo 2 = 12 + x := by
      set_option simprocs false in
        /- This `simp` command does not make progress since `simproc`s are disabled. -/
        fail_if_success simp
      simp_arith
    
    example : x + foo 2 = 12 + x := by
      /- `simp only` must not use the default simproc set. -/
      fail_if_success simp only
      simp_arith
    
    example : x + foo 2 = 12 + x := by
      /-
      `simp only` does not use the default simproc set,
      but we can provide simprocs as arguments. -/
      simp only [reduceFoo]
      simp_arith
    
    example : x + foo 2 = 12 + x := by
      /- We can use `-` to disable `simproc`s. -/
      fail_if_success simp [-reduceFoo]
      simp_arith

    The command register_simp_attr <id> now creates a simp and a simproc set with the name <id>. The following command instructs Lean to insert the reduceFoo simplification procedure into the set my_simp. If no set is specified, Lean uses the default simp set.

    simproc [my_simp] reduceFoo (foo _) := ...
  • The syntax of the termination_by and decreasing_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, the termination_by and
      decreasing_by for that function come before the where. The
      functions in the where 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 or let 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.

  • 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 like first:

     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 own
    decreasing_by, the default decreasing_tactic is used. If the same tactic
    ought to be applied to multiple functions, the decreasing_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 the InfoTree manually instead of going through the functions in InfoUtils.lean, as well as those manually creating and saving InfoTrees. 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 and pp.deepTerms.threshold) (PR #3201)

  • Add pretty printer options pp.numeralTypes and pp.natLit.
    When pp.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).
    When pp.natLit is true, then raw natural number literals are pretty printed as nat_lit 2.
    PR #2933 and RFC #3021.

Lake updates:

  • improved platform information & control #3226
  • lake update from unsupported manifest versions #3149

Other improvements:

  • make intro be aware of let_fun #3115
  • produce simpler proof terms in rw #3121
  • fuse nested `mkCongrArg...
Read more

v4.6.0-rc1

01 Feb 23:36
Compare
Choose a tag to compare
v4.6.0-rc1 Pre-release
Pre-release
  • Add custom simplification procedures (aka simprocs) to simp. Simprocs can be triggered by the simplifier on a specified term-pattern. Here is an small example:
import Lean.Meta.Tactic.Simp.BuiltinSimprocs.Nat

def foo (x : Nat) : Nat :=
  x + 10

/--
The `simproc` `reduceFoo` is invoked on terms that match the pattern `foo _`.
-/
simproc reduceFoo (foo _) :=
  /- A term of type `Expr → SimpM Step -/
  fun e => do
    /-
    The `Step` type has three constructors: `.done`, `.visit`, `.continue`.
    * The constructor `.done` instructs `simp` that the result does
      not need to be simplied further.
    * The constructor `.visit` instructs `simp` to visit the resulting expression.
    * The constructor `.continue` instructs `simp` to try other simplification procedures.

    All three constructors take a `Result`. The `.continue` contructor may also take `none`.
    `Result` has two fields `expr` (the new expression), and `proof?` (an optional proof).
     If the new expression is definitionally equal to the input one, then `proof?` can be omitted or set to `none`.
    -/
    /- `simp` uses matching modulo reducibility. So, we ensure the term is a `foo`-application. -/
    unless e.isAppOfArity ``foo 1 do
      return .continue
    /- `Nat.fromExpr?` tries to convert an expression into a `Nat` value -/
    let some n ← Nat.fromExpr? e.appArg!
      | return .continue
    return .done { expr := Lean.mkNatLit (n+10) }

We disable simprocs support by using the command set_option simprocs false. This command is particularly useful when porting files to v4.6.0.
Simprocs can be scoped, manually added to simp commands, and suppressed using -. They are also supported by simp?. simp only does not execute any simproc. Here are some examples for the simproc defined above.

example : x + foo 2 = 12 + x := by
  set_option simprocs false in
    /- This `simp` command does not make progress since `simproc`s are disabled. -/
    fail_if_success simp
  simp_arith

example : x + foo 2 = 12 + x := by
  /- `simp only` must not use the default simproc set. -/
  fail_if_success simp only
  simp_arith

example : x + foo 2 = 12 + x := by
  /-
  `simp only` does not use the default simproc set,
  but we can provide simprocs as arguments. -/
  simp only [reduceFoo]
  simp_arith

example : x + foo 2 = 12 + x := by
  /- We can use `-` to disable `simproc`s. -/
  fail_if_success simp [-reduceFoo]
  simp_arith

The command register_simp_attr <id> now creates a simp and a simproc set with the name <id>. The following command instructs Lean to insert the reduceFoo simplification procedure into the set my_simp. If no set is specified, Lean uses the default simp set.

simproc [my_simp] reduceFoo (foo _) := ...
  • The syntax of the termination_by and decreasing_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, the termination_by and
      decreasing_by for that function come before the where. The
      functions in the where 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 or let 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.

  • 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 like first:

     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 own
    decreasing_by, the default decreasing_tactic is used. If the same tactic
    ought to be applied to multiple functions, the decreasing_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 the InfoTree manually instead of going through the functions in InfoUtils.lean, as well as those manually creating and saving InfoTrees. 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 and pp.deepTerms.threshold) (PR #3201)

  • Add pretty printer options pp.numeralTypes and pp.natLit.
    When pp.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).
    When pp.natLit is true, then raw natural number literals are pretty printed as nat_lit 2.
    PR #2933 and RFC #3021.

Lake updates:

  • improved platform information & control #3226
  • lake update from unsupported manifest versions #3149

Other improvements:

  • make intro be aware of let_fun #3115
  • produce simpler proof terms in rw #3121
  • fuse nested mkCongrArg calls in proofs generated by simp #3203
  • induction using followed by a general term [#3188](https://github....
Read more

v4.5.0

01 Feb 03:02
Compare
Choose a tag to compare

Changes since v4.4.0 (from RELEASES.md)

  • Modify the lexical syntax of string literals to have string gaps, which are escape sequences of the form "\" newline whitespace*.
    These have the interpetation of an empty string and allow a string to flow across multiple lines without introducing additional whitespace.
    The following is equivalent to "this is a string".

    "this is \
       a string"

    PR #2821 and RFC #2838.

  • 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 in r#"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 use WellFounded.wrap from the std libarary to explicitly give one:

    -termination_by' ⟨r, hwf⟩
    +termination_by x => hwf.wrap x
  • Support snippet edits in LSP TextEdits. See Lean.Lsp.SnippetString for more details.

  • Deprecations and changes in the widget API.

    • Widget.UserWidgetDefinition is deprecated in favour of Widget.Module. The annotation @[widget] is deprecated in favour of @[widget_module]. To migrate a definition of type UserWidgetDefinition, remove the name field and replace the type with Widget.Module. Removing the name 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 like

    Inferred 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? and simp_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 is

    import 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 a DecidableEq instance.

Bug fixes for
#2853, #2953, #2966,
#2971, #2990, #3094.

Bug fix for eager evaluation of default value in Option.getD.
Avoid panic in leanPosToLspPos when file source is unavailable.
Improve short-circuiting behavior for List.all and List.any.

Several Lake bug fixes: #3036, #3064, #3069.

What's Changed

Read more

v4.5.0-rc1

21 Dec 23:06
Compare
Choose a tag to compare
v4.5.0-rc1 Pre-release
Pre-release

Changes since v4.4.0 (from RELEASE.md):

  • Modify the lexical syntax of string literals to have string gaps, which are escape sequences of the form "\" newline whitespace*.
    These have the interpetation of an empty string and allow a string to flow across multiple lines without introducing additional whitespace.
    The following is equivalent to "this is a string".

    "this is \
       a string"

    PR #2821 and RFC #2838.

  • 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 in r#"the "the" is in quotes"# for "the \"the\" is in quotes".
    PR #2929 and issue #1422.

  • The low-level termination_by' clause is no longer supported.

    Migration guide: Use termination_by instead, e.g.:

    -termination_by' measure (fun ⟨i, _⟩ => as.size - i)
    +termination_by go i _ => as.size - i

    If the well-founded relation you want to use is not the one that the
    WellFoundedRelation type class would infer for your termination argument,
    you can use WellFounded.wrap from the std libarary to explicitly give one:

    -termination_by' ⟨r, hwf⟩
    +termination_by _ x => hwf.wrap x
  • Support snippet edits in LSP TextEdits. See Lean.Lsp.SnippetString for more details.

  • Deprecations and changes in the widget API.

    • Widget.UserWidgetDefinition is deprecated in favour of Widget.Module. The annotation @[widget] is deprecated in favour of @[widget_module]. To migrate a definition of type UserWidgetDefinition, remove the name field and replace the type with Widget.Module. Removing the name 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 like

    Inferred 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? and simp_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 is

    import 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 a DecidableEq instance.

Bug fixes for
#2853, #2953, #2966,
#2971, #2990, #3094.

Bug fix for eager evaluation of default value in Option.getD.
Avoid panic in leanPosToLspPos when file source is unavailable.
Improve short-circuiting behavior for List.all and List.any.

Several Lake bug fixes: #3036, #3064, #3069.

Full Changelog: v4.4.0...v4.5.0-rc1