Skip to content

Commit 1968db8

Browse files
authoredOct 29, 2024··
Merge pull request #39 from leanprover-community/fast_unusedTransitiveImports
feat: fast unused transitive imports
2 parents 0ea83a6 + 793740f commit 1968db8

File tree

6 files changed

+99
-5
lines changed

6 files changed

+99
-5
lines changed
 

‎ImportGraph/Imports.lean

+3-5
Original file line numberDiff line numberDiff line change
@@ -176,15 +176,13 @@ end Lean.NameMap
176176
Returns a `List (Name × List Name)` with a key for each module `n` in `amongst`,
177177
whose corresponding value is the list of modules `m` in `amongst` which are transitively imported by `n`,
178178
but no declaration in `n` makes use of a declaration in `m`.
179-
180-
The current implementation is too slow to run on the entirety of Mathlib,
181-
although it should be fine for any sequential chain of imports in Mathlib.
182179
-/
183-
def unusedTransitiveImports (amongst : List Name) : CoreM (List (Name × List Name)) := do
180+
def unusedTransitiveImports (amongst : List Name) (verbose : Bool := false) : CoreM (List (Name × List Name)) := do
184181
let env ← getEnv
185182
let transitiveImports := env.importGraph.transitiveClosure
183+
let transitivelyRequired ← env.transitivelyRequiredModules' amongst verbose
186184
amongst.mapM fun n => do return (n,
187-
let unused := (transitiveImports.find? n).getD {} \ (← env.transitivelyRequiredModules n)
185+
let unused := (transitiveImports.find? n).getD {} \ (transitivelyRequired.find? n |>.getD {})
188186
amongst.filter (fun m => unused.contains m))
189187

190188
/--

‎ImportGraph/RequiredModules.lean

+51
Original file line numberDiff line numberDiff line change
@@ -115,6 +115,57 @@ def Environment.transitivelyRequiredModules (env : Environment) (module : Name)
115115
|>.filter (env.getModuleFor? · = some module)
116116
(NameSet.ofList constants).transitivelyRequiredModules env
117117

118+
/--
119+
Computes all the modules transitively required by the specified modules.
120+
Should be equivalent to calling `transitivelyRequiredModules` on each module, but shares more of the work.
121+
-/
122+
partial def Environment.transitivelyRequiredModules' (env : Environment) (modules : List Name) (verbose : Bool := false) :
123+
CoreM (NameMap NameSet) := do
124+
let N := env.header.moduleNames.size
125+
let mut c2m : NameMap (BitVec N) := {}
126+
let mut pushed : NameSet := {}
127+
let mut result : NameMap NameSet := {}
128+
for m in modules do
129+
if verbose then
130+
IO.println s!"Processing module {m}"
131+
let mut r : BitVec N := 0
132+
for n in env.header.moduleData[(env.header.moduleNames.getIdx? m).getD 0]!.constNames do
133+
if ! n.isInternal then
134+
-- This is messy: Mathlib is big enough that writing a recursive function causes a stack overflow.
135+
-- So we use an explicit stack instead. We visit each constant twice:
136+
-- once to record the constants transitively used by it,
137+
-- and again to record the modules which defined those constants.
138+
let mut stack : List (Name × Option NameSet) := [⟨n, none⟩]
139+
pushed := pushed.insert n
140+
while !stack.isEmpty do
141+
match stack with
142+
| [] => panic! "Stack is empty"
143+
| (c, used?) :: tail =>
144+
stack := tail
145+
match used? with
146+
| none =>
147+
if !c2m.contains c then
148+
let used := (← getConstInfo c).getUsedConstantsAsSet
149+
stack := ⟨c, some used⟩ :: stack
150+
for u in used do
151+
if !pushed.contains u then
152+
stack := ⟨u, none⟩ :: stack
153+
pushed := pushed.insert u
154+
| some used =>
155+
let usedModules : NameSet :=
156+
used.fold (init := {}) (fun s u => if let some m := env.getModuleFor? u then s.insert m else s)
157+
let transitivelyUsed : BitVec N :=
158+
used.fold (init := toBitVec usedModules) (fun s u => s ||| ((c2m.find? u).getD 0))
159+
c2m := c2m.insert c transitivelyUsed
160+
r := r ||| ((c2m.find? n).getD 0)
161+
result := result.insert m (toNameSet r)
162+
return result
163+
where
164+
toBitVec {N : Nat} (s : NameSet) : BitVec N :=
165+
s.fold (init := 0) (fun b n => b ||| BitVec.twoPow _ ((env.header.moduleNames.getIdx? n).getD 0))
166+
toNameSet {N : Nat} (b : BitVec N) : NameSet :=
167+
env.header.moduleNames.zipWithIndex.foldl (init := {}) (fun s (n, i) => if b.getLsbD i then s.insert n else s)
168+
118169
/--
119170
Return the names of the modules in which constants used in the current file were defined.
120171
+23
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,23 @@
1+
import ImportGraph.Imports
2+
3+
open Lean
4+
5+
def Core.withImportModules (modules : Array Name) {α} (f : CoreM α) : IO α := do
6+
searchPathRef.set compile_time_search_path%
7+
unsafe Lean.withImportModules (modules.map (fun m => {module := m})) {} (trustLevel := 1024)
8+
fun env => Prod.fst <$> Core.CoreM.toIO
9+
(ctx := { fileName := "<CoreM>", fileMap := default }) (s := { env := env }) do f
10+
11+
/--
12+
`lake exe unused_transitive_imports m1 m2 ...`
13+
14+
For each specified module `m`, prints those `n` from the argument list which are imported, but transitively unused by `m`.
15+
-/
16+
def main (args : List String) : IO UInt32 := do
17+
let (flags, args) := args.partition (fun s => s.startsWith "-")
18+
let mut modules := args.map (fun s => s.toName)
19+
Core.withImportModules modules.toArray do
20+
let r ← unusedTransitiveImports modules (verbose := flags.contains "-v" || flags.contains "--verbose")
21+
for (n, u) in r do
22+
IO.println s!"{n}: {u}"
23+
return 0

‎ImportGraphTest/Imports.lean

+9
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,6 @@
11
import ImportGraph.Imports
22
import ImportGraph.RequiredModules
3+
import ImportGraphTest.Used
34

45
open Lean
56

@@ -36,6 +37,7 @@ elab "#unused_transitive_imports" names:ident* : command => do
3637
logInfo <| s!"Transitively unused imports of {n}:\n{"\n".intercalate (u.map (fun i => s!" {i}"))}"
3738

3839
-- This test case can be removed after nightly-2024-10-24, because these imports have been cleaned up.
40+
-- It should be replaced with another test case!
3941
/--
4042
info: Transitively unused imports of Init.Control.StateRef:
4143
Init.System.IO
@@ -46,6 +48,13 @@ info: Transitively unused imports of Init.System.IO:
4648
#guard_msgs in
4749
#unused_transitive_imports Init.Control.StateRef Init.System.IO Init.Control.Reader Init.Control.Basic
4850

51+
/--
52+
info: Transitively unused imports of ImportGraphTest.Used:
53+
ImportGraphTest.Unused
54+
-/
55+
#guard_msgs in
56+
#unused_transitive_imports ImportGraphTest.Used ImportGraphTest.Unused Init.Control.Reader
57+
4958
-- This is a spurious unused transitive import, because it relies on notation from `Init.Core`.
5059
/--
5160
info: Transitively unused imports of Init.Control.Basic:

‎README.md

+6
Original file line numberDiff line numberDiff line change
@@ -54,6 +54,12 @@ There are a few commands implemented, which help you analysing the imports of a
5454
(Must be run at the end of the file. Tactics and macros may result in incorrect output.)
5555
* `#find_home decl`: suggests files higher up the import hierarchy to which `decl` could be moved.
5656

57+
## Other executables
58+
59+
`lake exe unused_transitive_imports m1 m2 ...`
60+
61+
For each specified module `m`, prints those `n` from the argument list which are imported, but transitively unused by `m`.
62+
5763
## Installation
5864

5965
The installation works exactly like for any [Lake package](https://reservoir.lean-lang.org/),

‎lakefile.toml

+7
Original file line numberDiff line numberDiff line change
@@ -28,5 +28,12 @@ root = "Main"
2828
# Remove this line if you do not need such functionality.
2929
supportInterpreter = true
3030

31+
# `lake exe unused_transitive_imports` prints unused transitive imports from amongst a given list of modules.
32+
[[lean_exe]]
33+
name = "unused_transitive_imports"
34+
root = "ImportGraph.UnusedTransitiveImports"
35+
supportInterpreter = true
36+
3137
[[lean_lib]]
3238
name = "ImportGraphTest"
39+

0 commit comments

Comments
 (0)
Please sign in to comment.