Skip to content

Commit

Permalink
perf: rewrite UnusedVariables lint (#3186)
Browse files Browse the repository at this point in the history
This is a rewrite of the `UnusedVariables` lint to inline and simplify
many of the dependent functions to try to improve the performance of
this lint, which quite often shows up in perf reports.

* The mvar assignment scanning is one of the most expensive parts of the
process, so we do two things to improve this:
  * Lazily perform the scan only if we need it
* Use an object-pointer hashmap to ensure that we don't have quadratic
behavior when there are many mvar assignments with slight differences.
* The dependency on `Lean.Server` is removed, meaning we don't need to
do the LSP conversion stuff anymore. The main logic of reference finding
is inlined.
* We take `fvarAliases` into account, and union together fvars which are
aliases of a base fvar. (It would be great if we had `UnionFind` here.)

More docs will be added once we confirm an actual perf improvement.

---------

Co-authored-by: Sebastian Ullrich <[email protected]>
  • Loading branch information
digama0 and Kha authored Mar 21, 2024
1 parent 164689f commit 49f66dc
Show file tree
Hide file tree
Showing 12 changed files with 396 additions and 167 deletions.
4 changes: 2 additions & 2 deletions src/Init/ByCases.lean
Original file line number Diff line number Diff line change
Expand Up @@ -21,9 +21,9 @@ macro_rules

/-! ## if-then-else -/

@[simp] theorem if_true {h : Decidable True} (t e : α) : ite True t e = t := if_pos trivial
@[simp] theorem if_true {_ : Decidable True} (t e : α) : ite True t e = t := if_pos trivial

@[simp] theorem if_false {h : Decidable False} (t e : α) : ite False t e = e := if_neg id
@[simp] theorem if_false {_ : Decidable False} (t e : α) : ite False t e = e := if_neg id

theorem ite_id [Decidable c] {α} (t : α) : (if c then t else t) = t := by split <;> rfl

Expand Down
1 change: 1 addition & 0 deletions src/Init/Control/ExceptCps.lean
Original file line number Diff line number Diff line change
Expand Up @@ -18,6 +18,7 @@ namespace ExceptCpsT
def run {ε α : Type u} [Monad m] (x : ExceptCpsT ε m α) : m (Except ε α) :=
x _ (fun a => pure (Except.ok a)) (fun e => pure (Except.error e))

set_option linter.unusedVariables false in -- `s` unused
@[always_inline, inline]
def runK {ε α : Type u} (x : ExceptCpsT ε m α) (s : ε) (ok : α → m β) (error : ε → m β) : m β :=
x _ ok error
Expand Down
3 changes: 2 additions & 1 deletion src/Init/Data/Ord.lean
Original file line number Diff line number Diff line change
Expand Up @@ -119,6 +119,7 @@ class Ord (α : Type u) where

export Ord (compare)

set_option linter.unusedVariables false in -- allow specifying `ord` explicitly
/--
Compare `x` and `y` by comparing `f x` and `f y`.
-/
Expand Down Expand Up @@ -215,7 +216,7 @@ protected def opposite (ord : Ord α) : Ord α where
/--
`ord.on f` compares `x` and `y` by comparing `f x` and `f y` according to `ord`.
-/
protected def on (ord : Ord β) (f : α → β) : Ord α where
protected def on (_ : Ord β) (f : α → β) : Ord α where
compare := compareOn f

/--
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Compiler/LCNF/JoinPoints.lean
Original file line number Diff line number Diff line change
Expand Up @@ -346,7 +346,7 @@ We call this whenever we enter a new local function. It clears both the
current join point and the list of candidates since we can't lift join
points outside of functions as explained in `mergeJpContextIfNecessary`.
-/
def withNewFunScope (decl : FunDecl) (x : ExtendM α): ExtendM α := do
def withNewFunScope (x : ExtendM α): ExtendM α := do
withReader (fun ctx => { ctx with currentJp? := none, candidates := {} }) do
withNewScope do
x
Expand Down Expand Up @@ -412,7 +412,7 @@ where
withNewCandidate decl.fvarId do
return Code.updateFun! code decl (← go k)
| .fun decl k =>
let decl ← withNewFunScope decl do
let decl ← withNewFunScope do
decl.updateValue (← go decl.value)
withNewCandidate decl.fvarId do
return Code.updateFun! code decl (← go k)
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Elab/Term.lean
Original file line number Diff line number Diff line change
Expand Up @@ -665,7 +665,7 @@ def mkTypeMismatchError (header? : Option String) (e : Expr) (eType : Expr) (exp
return m!"{header}{← mkHasTypeButIsExpectedMsg eType expectedType}"

def throwTypeMismatchError (header? : Option String) (expectedType : Expr) (eType : Expr) (e : Expr)
(f? : Option Expr := none) (extraMsg? : Option MessageData := none) : TermElabM α := do
(f? : Option Expr := none) (_extraMsg? : Option MessageData := none) : TermElabM α := do
/-
We ignore `extraMsg?` for now. In all our tests, it contained no useful information. It was
always of the form:
Expand Down
Loading

0 comments on commit 49f66dc

Please sign in to comment.