Skip to content

Commit

Permalink
fix: unused variable false positive under match + assumption
Browse files Browse the repository at this point in the history
  • Loading branch information
Kha committed Sep 13, 2024
1 parent adfd6c0 commit 79943be
Show file tree
Hide file tree
Showing 4 changed files with 23 additions and 7 deletions.
2 changes: 1 addition & 1 deletion src/Lean/Elab/InfoTree/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -183,7 +183,7 @@ def UserWidgetInfo.format (info : UserWidgetInfo) : Format :=
f!"UserWidget {info.id}\n{Std.ToFormat.format <| info.props.run' {}}"

def FVarAliasInfo.format (info : FVarAliasInfo) : Format :=
f!"FVarAlias {info.userName.eraseMacroScopes}"
f!"FVarAlias {info.userName.eraseMacroScopes}: {info.id.name} -> {info.baseId.name}"

def FieldRedeclInfo.format (ctx : ContextInfo) (info : FieldRedeclInfo) : Format :=
f!"FieldRedecl @ {formatStxRange ctx info.stx}"
Expand Down
9 changes: 7 additions & 2 deletions src/Lean/Linter/UnusedVariables.lean
Original file line number Diff line number Diff line change
Expand Up @@ -447,7 +447,10 @@ def unusedVariables : Linter where
let fvarAliases : Std.HashMap FVarId FVarId := s.fvarAliases.fold (init := {}) fun m id baseId =>
m.insert id (followAliases s.fvarAliases baseId)

let getCanonVar (id : FVarId) : FVarId := fvarAliases.getD id id

-- Collect all non-alias fvars corresponding to `fvarUses` by resolving aliases in the list.
-- Unlike `s.fvarUses`, `fvarUsesRef` is guaranteed to contain no aliases.
let fvarUsesRef ← IO.mkRef <| fvarAliases.fold (init := s.fvarUses) fun fvarUses id baseId =>
if fvarUses.contains id then fvarUses.insert baseId else fvarUses

Expand All @@ -461,7 +464,7 @@ def unusedVariables : Linter where
let fvarUses ← fvarUsesRef.get
-- If any of the `fvar`s corresponding to this declaration is (an alias of) a variable in
-- `fvarUses`, then it is used
if aliases.any fun id => fvarUses.contains (fvarAliases.getD id id) then continue
if aliases.any fun id => fvarUses.contains (getCanonVar id) then continue
-- If this is a global declaration then it is (potentially) used after the command
if s.constDecls.contains range then continue

Expand Down Expand Up @@ -493,10 +496,12 @@ def unusedVariables : Linter where
if !initializedMVars then
-- collect additional `fvarUses` from tactic assignments
visitAssignments (← IO.mkRef {}) fvarUsesRef s.assignments
-- Resolve potential aliases again to preserve `fvarUsesRef` invariant
fvarUsesRef.modify fun fvarUses => fvarUses.toList.map getCanonVar |> .ofList
initializedMVars := true
let fvarUses ← fvarUsesRef.get
-- Redo the initial check because `fvarUses` could be bigger now
if aliases.any fun id => fvarUses.contains (fvarAliases.getD id id) then continue
if aliases.any fun id => fvarUses.contains (getCanonVar id) then continue

-- If we made it this far then the variable is unused and not ignored
unused := unused.push (declStx, userName)
Expand Down
8 changes: 4 additions & 4 deletions tests/lean/1018unknowMVarIssue.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -55,8 +55,8 @@ a : α
• Fam2.any : Fam2 α α @ ⟨9, 4⟩†-⟨9, 12⟩†
• α : Type @ ⟨9, 4⟩†-⟨9, 12⟩†
• a (isBinder := true) : α @ ⟨8, 2⟩†-⟨10, 19⟩†
• FVarAlias a
• FVarAlias α
• FVarAlias a: _uniq.636 -> _uniq.312
• FVarAlias α: _uniq.635 -> _uniq.310
• ?m x α a : α @ ⟨9, 18⟩-⟨9, 19⟩ @ Lean.Elab.Term.elabHole
• [.] Fam2.nat : none @ ⟨10, 4⟩-⟨10, 12⟩
• Fam2.nat : Nat → Fam2 Nat Nat @ ⟨10, 4⟩-⟨10, 12⟩
Expand All @@ -70,8 +70,8 @@ a : α
• Fam2.nat n : Fam2 Nat Nat @ ⟨10, 4⟩†-⟨10, 14⟩
• n (isBinder := true) : Nat @ ⟨10, 13⟩-⟨10, 14⟩
• a (isBinder := true) : Nat @ ⟨8, 2⟩†-⟨10, 19⟩†
• FVarAlias a
• FVarAlias n
• FVarAlias a: _uniq.667 -> _uniq.312
• FVarAlias n: _uniq.666 -> _uniq.310
• n : Nat @ ⟨10, 18⟩-⟨10, 19⟩ @ Lean.Elab.Term.elabIdent
• [.] n : some Nat @ ⟨10, 18⟩-⟨10, 19⟩
• n : Nat @ ⟨10, 18⟩-⟨10, 19⟩
Expand Down
11 changes: 11 additions & 0 deletions tests/lean/linterUnusedVariables.lean
Original file line number Diff line number Diff line change
Expand Up @@ -255,3 +255,14 @@ def ignoreEverything : Lean.Linter.IgnoreFunction :=
fun _ _ _ => true

def ignored (x : Nat) := 0

inductive A where
| intro : Nat → A

def A.out : A → Nat
| .intro n => n

/-! `h` is used indirectly via an alias introduced by `match` that is used only via the mvar ctx -/
theorem problematicAlias (n : A) (i : Nat) (h : i ≤ n.out) : i ≤ n.out :=
match n with
| .intro _ => by assumption

0 comments on commit 79943be

Please sign in to comment.