├── .github └── workflows │ └── build.yml ├── .gitignore ├── .gitpod.yml ├── ImportGraph.lean ├── ImportGraph ├── Cli.lean ├── CurrentModule.lean ├── Gexf.lean ├── Imports.lean ├── Lean │ └── Name.lean ├── RequiredModules.lean └── UnusedTransitiveImports.lean ├── ImportGraphTest.lean ├── ImportGraphTest ├── .gitignore ├── AnotherFileWithTransitiveImports.lean ├── Dot.lean ├── FileWithTransitiveImports.lean ├── Imports.lean ├── Unused.lean ├── Used.lean └── expected.dot ├── LICENSE ├── Main.lean ├── README.md ├── html-template ├── .gitignore ├── LICENSE_source ├── README.md ├── index.html └── vendor │ ├── graphology-LICENSE │ ├── graphology-library.min.js │ ├── graphology.min.js │ ├── sigma-LICENSE │ └── sigma.min.js ├── lake-manifest.json ├── lakefile.toml └── lean-toolchain /.github/workflows/build.yml: -------------------------------------------------------------------------------- 1 | name: Build 2 | run-name: Build the project 3 | on: 4 | push: 5 | branches: 6 | - main 7 | pull_request: 8 | branches: 9 | - main 10 | workflow_dispatch: 11 | jobs: 12 | build: 13 | runs-on: ubuntu-latest 14 | steps: 15 | - uses: actions/checkout@v4 16 | - uses: leanprover/lean-action@v1 17 | with: 18 | check-reservoir-eligibility: true 19 | # use setup from lean-action to perform the following steps 20 | - name: verify `lake exe graph` works 21 | run: | 22 | lake exe graph 23 | rm import_graph.dot 24 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | /build 2 | /lakefile.olean 3 | /lake-packages/* 4 | .lake/ 5 | graph.pdf 6 | -------------------------------------------------------------------------------- /.gitpod.yml: -------------------------------------------------------------------------------- 1 | # This is run when starting a Gitpod workspace on this repository 2 | 3 | image: leanprovercommunity/gitpod4 4 | 5 | vscode: 6 | extensions: 7 | - leanprover.lean4 # install the Lean 4 VS Code extension 8 | 9 | tasks: 10 | - init: | 11 | elan self update -------------------------------------------------------------------------------- /ImportGraph.lean: -------------------------------------------------------------------------------- 1 | import ImportGraph.Cli 2 | import ImportGraph.Imports 3 | import ImportGraph.CurrentModule 4 | import ImportGraph.Lean.Name 5 | import ImportGraph.RequiredModules 6 | -------------------------------------------------------------------------------- /ImportGraph/Cli.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Kim Morrison. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Kim Morrison 5 | -/ 6 | import Cli.Basic 7 | import ImportGraph.CurrentModule 8 | import ImportGraph.Imports 9 | import ImportGraph.Lean.Name 10 | import ImportGraph.Gexf 11 | 12 | open Cli 13 | 14 | open Lean 15 | open ImportGraph 16 | 17 | /-- 18 | Write an import graph, represented as a `NameMap (Array Name)` to the ".dot" graph format. 19 | * Nodes in the `unused` set will be shaded light gray. 20 | * If `markedPackage` is provided: 21 | * Nodes which start with the `markedPackage` will be highlighted in green and drawn closer together. 22 | * Edges from `directDeps` into the module are highlighted in green 23 | * Nodes in `directDeps` are marked with a green border and green text. 24 | -/ 25 | def asDotGraph 26 | (graph : NameMap (Array Name)) 27 | (unused : NameSet := {}) 28 | (header := "import_graph") 29 | (markedPackage : Option Name := none) 30 | (directDeps : NameSet := {}) 31 | (from_ to : NameSet := {}): 32 | String := Id.run do 33 | 34 | let mut lines := #[s!"digraph \"{header}\" " ++ "{"] 35 | for (n, is) in graph do 36 | let shape := if from_.contains n then "invhouse" else if to.contains n then "house" else "ellipse" 37 | if markedPackage.isSome ∧ directDeps.contains n then 38 | -- note: `fillcolor` defaults to `color` if not specified 39 | let fill := if unused.contains n then "#e0e0e0" else "white" 40 | lines := lines.push s!" \"{n}\" [style=filled, fontcolor=\"#4b762d\", color=\"#71b144\", fillcolor=\"{fill}\", penwidth=2, shape={shape}];" 41 | else if unused.contains n then 42 | lines := lines.push s!" \"{n}\" [style=filled, fillcolor=\"#e0e0e0\", shape={shape}];" 43 | else if isInModule markedPackage n then 44 | -- mark node 45 | lines := lines.push s!" \"{n}\" [style=filled, fillcolor=\"#96ec5b\", shape={shape}];" 46 | else 47 | lines := lines.push s!" \"{n}\" [shape={shape}];" 48 | -- Then add edges 49 | for i in is do 50 | if isInModule markedPackage n then 51 | if isInModule markedPackage i then 52 | -- draw the main project close together 53 | lines := lines.push s!" \"{i}\" -> \"{n}\" [weight=100];" 54 | else 55 | -- mark edges into the main project 56 | lines := lines.push s!" \"{i}\" -> \"{n}\" [penwidth=2, color=\"#71b144\"];" 57 | else 58 | lines := lines.push s!" \"{i}\" -> \"{n}\";" 59 | lines := lines.push "}" 60 | return "\n".intercalate lines.toList 61 | 62 | open Lean Core System 63 | 64 | open IO.FS IO.Process Name in 65 | /-- Implementation of the import graph command line program. -/ 66 | def importGraphCLI (args : Cli.Parsed) : IO UInt32 := do 67 | -- file extensions that should be created 68 | let extensions : Std.HashSet String := match args.variableArgsAs! String with 69 | | #[] => {"dot"} 70 | | outputs => outputs.foldl (fun acc (o : String) => 71 | match FilePath.extension o with 72 | | none => acc.insert "dot" 73 | | some "gexf" => acc.insert "gexf" 74 | | some "html" => acc.insert "gexf" 75 | -- currently all other formats are handled by passing the `.dot` file to 76 | -- graphviz 77 | | some _ => acc.insert "dot" ) {} 78 | 79 | let to ← match args.flag? "to" with 80 | | some to => pure <| to.as! (Array ModuleName) 81 | | none => pure #[← getCurrentModule] 82 | let from? : Option (Array Name) := match args.flag? "from" with 83 | | some fr => some <| fr.as! (Array ModuleName) 84 | | none => none 85 | initSearchPath (← findSysroot) 86 | 87 | unsafe Lean.enableInitializersExecution 88 | let outFiles ← try unsafe withImportModules (to.map ({module := ·})) {} (trustLevel := 1024) fun env => do 89 | let toModule := ImportGraph.getModule to[0]! 90 | let mut graph := env.importGraph 91 | let unused ← 92 | match args.flag? "to" with 93 | | some _ => 94 | let ctx := { options := {}, fileName := "", fileMap := default } 95 | let state := { env } 96 | let used ← Prod.fst <$> (CoreM.toIO (env.transitivelyRequiredModules' to.toList) ctx state) 97 | let used := used.foldl (init := NameSet.empty) (fun s _ t => s ∪ t) 98 | pure <| graph.foldl (fun acc n _ => if used.contains n then acc else acc.insert n) NameSet.empty 99 | | none => pure NameSet.empty 100 | if let Option.some f := from? then 101 | graph := graph.downstreamOf (NameSet.ofArray f) 102 | let includeLean := args.hasFlag "include-lean" 103 | let includeStd := args.hasFlag "include-std" || includeLean 104 | let includeDeps := args.hasFlag "include-deps" || includeStd 105 | -- Note: `includeDirect` does not imply `includeDeps`! 106 | -- e.g. if the package contains `import Lean`, the node `Lean` will be included with 107 | -- `--include-direct`, but not included with `--include-deps`. 108 | let includeDirect := args.hasFlag "include-direct" 109 | 110 | -- `directDeps` contains files which are not in the package 111 | -- but directly imported by a file in the package 112 | let directDeps : NameSet := graph.foldl (init := .empty) (fun acc n deps => 113 | if toModule.isPrefixOf n then 114 | deps.filter (!toModule.isPrefixOf ·) |>.foldl (init := acc) NameSet.insert 115 | else 116 | acc) 117 | 118 | let filter (n : Name) : Bool := 119 | toModule.isPrefixOf n || 120 | bif isPrefixOf `Std n then includeStd else 121 | bif isPrefixOf `Lean n || isPrefixOf `Init n then includeLean else 122 | includeDeps 123 | let filterDirect (n : Name) : Bool := 124 | includeDirect ∧ directDeps.contains n 125 | 126 | graph := graph.filterMap (fun n i => 127 | if filter n then 128 | -- include node regularly 129 | (i.filter (fun m => filterDirect m || filter m)) 130 | else if filterDirect n then 131 | -- include node as direct dependency; drop any further deps. 132 | some #[] 133 | else 134 | -- not included 135 | none) 136 | if args.hasFlag "exclude-meta" then 137 | -- Mathlib-specific exclusion of tactics 138 | let filterMathlibMeta : Name → Bool := fun n => ( 139 | isPrefixOf `Mathlib.Tactic n ∨ 140 | isPrefixOf `Mathlib.Lean n ∨ 141 | isPrefixOf `Mathlib.Mathport n ∨ 142 | isPrefixOf `Mathlib.Util n) 143 | graph := graph.filterGraph filterMathlibMeta (replacement := `«Mathlib.Tactics») 144 | if !args.hasFlag "show-transitive" then 145 | graph := graph.transitiveReduction 146 | 147 | let markedPackage : Option Name := if args.hasFlag "mark-package" then toModule else none 148 | 149 | -- Create all output files that are requested 150 | let mut outFiles : Std.HashMap String String := {} 151 | if extensions.contains "dot" then 152 | let dotFile := asDotGraph graph (unused := unused) (markedPackage := markedPackage) (directDeps := directDeps) 153 | (to := NameSet.ofArray to) (from_ := NameSet.ofArray (from?.getD #[])) 154 | outFiles := outFiles.insert "dot" dotFile 155 | if extensions.contains "gexf" then 156 | -- filter out the top node as it makes the graph less pretty 157 | let graph₂ := match args.flag? "to" with 158 | | none => graph.filter (fun n _ => ! if to.contains `Mathlib then #[`Mathlib, `Mathlib.Tactic].contains n else to.contains n) 159 | | some _ => graph 160 | let gexfFile := Graph.toGexf graph₂ toModule env 161 | outFiles := outFiles.insert "gexf" gexfFile 162 | return outFiles 163 | 164 | catch err => 165 | -- TODO: try to build `to` first, so this doesn't happen 166 | throw <| IO.userError <| s!"{err}\nIf the error above says `object file ... does not exist`, " ++ 167 | s!"try if `lake build {" ".intercalate (to.toList.map Name.toString)}` fixes the issue" 168 | throw err 169 | 170 | match args.variableArgsAs! String with 171 | | #[] => writeFile "import_graph.dot" (outFiles["dot"]!) 172 | | outputs => for o in outputs do 173 | let fp : FilePath := o 174 | match fp.extension with 175 | | none 176 | | "dot" => writeFile fp (outFiles["dot"]!) 177 | | "gexf" => IO.FS.writeFile fp (outFiles["gexf"]!) 178 | | "html" => 179 | let gexfFile := (outFiles["gexf"]!) 180 | -- use `html-template/index.html` and insert any dependencies to make it 181 | -- a stand-alone HTML file. 182 | -- note: changes in `index.html` might need to be reflected here! 183 | let exeDir := (FilePath.parent (← IO.appPath) |>.get!) / ".." / ".." / ".." 184 | let mut html ← IO.FS.readFile <| ← IO.FS.realPath ( exeDir / "html-template" / "index.html") 185 | for dep in (#[ 186 | "vendor" / "sigma.min.js", 187 | "vendor" / "graphology.min.js", 188 | "vendor" / "graphology-library.min.js" ] : Array FilePath) do 189 | let depContent ← IO.FS.readFile <| ← IO.FS.realPath (exeDir / "html-template" / dep) 190 | html := html.replace s!"" s!"" 191 | -- inline the graph data 192 | -- note: changes in `index.html` might need to be reflected here! 193 | let escapedFile := gexfFile.replace "\n" "" |>.replace "\"" "\\\"" 194 | let toFormatted : String := ", ".intercalate <| (to.map toString).toList 195 | html := html 196 | |>.replace "fetch(\"imports.gexf\").then((res) => res.text()).then(render_gexf)" s!"render_gexf(\"{escapedFile}\")" 197 | |>.replace "

Import Graph

" s!"

Import Graph for {toFormatted}

" 198 | |>.replace "import graph" s!"import graph for {toFormatted}" 199 | IO.FS.writeFile fp html 200 | | some ext => try 201 | _ ← IO.Process.output { cmd := "dot", args := #["-T" ++ ext, "-o", o] } outFiles["dot"]! 202 | catch ex => 203 | IO.eprintln s!"Error occurred while writing out {fp}." 204 | IO.eprintln s!"Make sure you have `graphviz` installed and the file is writable." 205 | throw ex 206 | return 0 207 | -------------------------------------------------------------------------------- /ImportGraph/CurrentModule.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jon Eugster. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jon Eugster 5 | -/ 6 | import Lake.Load.Manifest 7 | 8 | open Lean (Name) 9 | 10 | namespace ImportGraph 11 | 12 | /-- Read the name of the main module from the `lake-manifest`. -/ 13 | def getCurrentModule : IO Name := do 14 | 15 | match (← Lake.Manifest.load? ⟨"lake-manifest.json"⟩) with 16 | | none => 17 | -- TODO: should this be caught? 18 | pure .anonymous 19 | | some manifest => 20 | -- TODO: This assumes that the `package` and the default `lean_lib` 21 | -- have the same name up to capitalisation. 22 | -- Would be better to read the `.defaultTargets` from the 23 | -- `← getRootPackage` from `Lake`, but I can't make that work with the monads involved. 24 | return manifest.name.capitalize 25 | 26 | /-- 27 | Helper which only returns `true` if the `module` is provided and the name `n` lies 28 | inside it. 29 | -/ 30 | def isInModule (module : Option Name) (n : Name) := match module with 31 | | some m => m.isPrefixOf n 32 | | none => false 33 | -------------------------------------------------------------------------------- /ImportGraph/Gexf.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jon Eugster. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jon Eugster 5 | -/ 6 | 7 | import Lean.Data.Name 8 | import Lean.AuxRecursor 9 | import Lean.MonadEnv 10 | import Lean.Meta.Match.MatcherInfo 11 | 12 | open Lean 13 | 14 | namespace ImportGraph 15 | 16 | open Elab Meta in 17 | /-- Filter Lean internal declarations -/ 18 | def isBlackListed (env : Environment) (declName : Name) : Bool := 19 | declName == ``sorryAx 20 | || declName matches .str _ "inj" 21 | || declName matches .str _ "noConfusionType" 22 | || declName.isInternalDetail 23 | || isAuxRecursor env declName 24 | || isNoConfusion env declName 25 | || isRecCore env declName 26 | || isMatcherCore env declName 27 | 28 | /-- Get number of non-blacklisted declarations per file. -/ 29 | def getNumberOfDeclsPerFile (env: Environment) : NameMap Nat := 30 | env.const2ModIdx.fold (fun acc n (idx : ModuleIdx) => 31 | let mod := env.allImportedModuleNames[idx]! 32 | if isBlackListed env n then acc else acc.insert mod ((acc.getD mod 0) + 1) 33 | ) {} 34 | 35 | /-- Gexf template for a node in th graph. -/ 36 | def Gexf.nodeTemplate (n module : Name) (size : Nat) := s!"\n " 37 | 38 | /-- Gexf template for an edge in the graph -/ 39 | def Gexf.edgeTemplate (source target : Name) := s!"\n " 40 | 41 | open Gexf in 42 | /-- Creates a `.gexf` file of the graph, see https://gexf.net/ 43 | 44 | Metadata can be stored in forms of attributes, currently we record the following: 45 | * `decl_count` (Nat): number of declarations in the file 46 | * `in_module` (Bool): whether the file belongs to the main module 47 | (used to strip the first part of the name when displaying). 48 | -/ 49 | def Graph.toGexf (graph : NameMap (Array Name)) (module : Name) (env : Environment) : String := 50 | let sizes : NameMap Nat := getNumberOfDeclsPerFile env 51 | let nodes : String := graph.foldl (fun acc n _ => acc ++ nodeTemplate n module (sizes.getD n 0)) "" 52 | let edges : String := graph.foldl (fun acc n i => acc ++ (i.foldl (fun b j => b ++ edgeTemplate j n) "")) "" 53 | s!" 54 | 55 | 56 | Lean ImportGraph 57 | 58 | 59 | 60 | 61 | 62 | 63 | 64 | {nodes.trim} 65 | 66 | 67 | {edges.trim} 68 | 69 | 70 | 71 | " 72 | -------------------------------------------------------------------------------- /ImportGraph/Imports.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Kim Morrison. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Kim Morrison, Paul Lezeau 5 | -/ 6 | import Lean.Elab.Command 7 | import Lean.Util.SearchPath 8 | import Lean.Server.GoTo 9 | import Lean.Widget.UserWidget 10 | import ImportGraph.RequiredModules 11 | 12 | /-! 13 | # Tools for analyzing imports. 14 | 15 | Provides the commands 16 | 17 | * `#redundant_imports` which lists any transitively redundant imports in the current module. 18 | * `#min_imports` which attempts to construct a minimal set of imports for the declarations 19 | in the current file. 20 | (Must be run at the end of the file. Tactics and macros may result in incorrect output.) 21 | * `#find_home decl` suggests files higher up the import hierarchy to which `decl` could be moved. 22 | -/ 23 | 24 | open Lean 25 | 26 | namespace Lean.Environment 27 | 28 | /-- 29 | Find the imports of a given module. 30 | -/ 31 | def importsOf (env : Environment) (n : Name) : Array Name := 32 | if n = env.header.mainModule then 33 | env.header.imports.map Import.module 34 | else match env.getModuleIdx? n with 35 | | .some idx => env.header.moduleData[idx.toNat]!.imports.map Import.module |>.erase `Init 36 | | .none => #[] 37 | 38 | /-- 39 | Construct the import graph of the current file. 40 | -/ 41 | partial def importGraph (env : Environment) : NameMap (Array Name) := 42 | let main := env.header.mainModule 43 | let imports := env.header.imports.map Import.module 44 | imports.foldl (fun m i => process env i m) (({} : NameMap _).insert main imports) 45 | |>.erase Name.anonymous 46 | where 47 | process (env) (i) (m) : NameMap (Array Name) := 48 | if m.contains i then 49 | m 50 | else 51 | let imports := env.importsOf i 52 | imports.foldr (fun i m => process env i m) (m.insert i imports) 53 | 54 | /-- 55 | Return the redundant imports (i.e. those transitively implied by another import) 56 | amongst a candidate list of imports. 57 | -/ 58 | partial def findRedundantImports (env : Environment) (imports : Array Name) : NameSet := 59 | let run := visit env.importGraph imports 60 | let (_, seen) := imports.foldl (fun ⟨v, s⟩ n => run v s n) ({}, {}) 61 | seen 62 | where 63 | visit (Γ) (targets) (visited) (seen) (n) : NameSet × NameSet := 64 | if visited.contains n then 65 | (visited, seen) 66 | else 67 | let imports := (Γ.find? n).getD #[] 68 | let (visited', seen') := imports.foldl (fun ⟨v, s⟩ t => visit Γ targets v s t) (visited, seen) 69 | (visited'.insert n, 70 | imports.foldl (fun s t => if targets.contains t then s.insert t else s) seen') 71 | 72 | end Lean.Environment 73 | 74 | namespace Lean.NameMap 75 | 76 | /-- 77 | Compute the transitive closure of an import graph. 78 | -/ 79 | partial def transitiveClosure (m : NameMap (Array Name)) : NameMap NameSet := 80 | m.foldl (fun r n i => process r n i) {} 81 | where 82 | process (r : NameMap NameSet) (n : Name) (i : Array Name) : NameMap NameSet := 83 | if r.contains n then 84 | r 85 | else 86 | let r' := i.foldr (fun i r => process r i ((m.find? i).getD #[])) r 87 | let t := i.foldr 88 | (fun j s => ((r'.find? j).getD {}).foldl NameSet.insert s) 89 | (.ofList i.toList) 90 | r'.insert n t 91 | 92 | /-- 93 | Compute the transitive reduction of an import graph. 94 | 95 | Typical usage is `transitiveReduction (← importGraph)`. 96 | -/ 97 | def transitiveReduction (m : NameMap (Array Name)) : NameMap (Array Name) := 98 | let c := transitiveClosure m 99 | m.foldl (fun r n a => 100 | r.insert n (a.foldr (fun i b => b.filter (fun j => ! ((c.find? i).getD {}).contains j)) a)) {} 101 | 102 | /-- Restrict an import graph to only the downstream dependencies of some set of modules. -/ 103 | def downstreamOf (m : NameMap (Array Name)) (targets : NameSet) : NameMap (Array Name) := 104 | let tc := transitiveClosure m 105 | let P (n : Name) := targets.contains n || ((tc.find? n).getD {}).any fun j => targets.contains j 106 | m.foldl (init := {}) fun r n i => 107 | if P n then 108 | r.insert n (i.filter P) 109 | else 110 | r 111 | 112 | /-- Restrict an import graph to only the transitive imports of some set of modules. -/ 113 | def upstreamOf (m : NameMap (Array Name)) (targets : NameSet) : NameMap (Array Name) := 114 | let tc := transitiveClosure m 115 | let P (n : Name) := targets.contains n || targets.any fun t => ((tc.find? t).getD {}).contains n 116 | m.foldl (init := {}) fun r n i => 117 | if P n then 118 | r.insert n (i.filter P) 119 | else 120 | r 121 | 122 | /-- 123 | Filter the list of edges `… → node` inside `graph` by the function `filter`. 124 | 125 | Any such upstream node `source` where `filter source` returns true will be replaced 126 | by all its upstream nodes. This results in a list of all unfiltered nodes 127 | in the `graph` that either had an edge to `node` 128 | or had an indirect edge to `node` going through filtered nodes. 129 | 130 | Will panic if the `node` is not in the `graph`. 131 | -/ 132 | partial 133 | def transitiveFilteredUpstream (node : Name) (graph : NameMap (Array Name)) 134 | (filter : Name → Bool) (replacement : Option Name := none): 135 | List Name := 136 | (graph.get! node).toList.flatMap fun source => 137 | ( if filter source then 138 | -- Add the transitive edges going through the filtered node `source`. 139 | -- If there is a replacement node, add an additional edge `repl → node`. 140 | match replacement with 141 | | none => transitiveFilteredUpstream source graph filter 142 | | some repl => .cons repl <| transitiveFilteredUpstream source graph filter 143 | -- If the node is not filtered, we leave the edge `source → node` intact. 144 | else [source]).eraseDups 145 | 146 | /-- 147 | Filters the `graph` removing all nodes where `filter n` returns false. Additionally, 148 | replace edges from removed nodes by all the transitive edges. 149 | 150 | This means there is a path between two nodes in the filtered graph iff there exists 151 | such a path in the original graph. 152 | 153 | If the optional `(replacement : Name)` is provided, a corresponding node will be 154 | added together with edges to all nodes which had an incoming edge from any 155 | filtered node. 156 | -/ 157 | def filterGraph (graph : NameMap (Array Name)) (filter : Name → Bool) 158 | (replacement : Option Name := none) : NameMap (Array Name) := 159 | -- Create a list of all files imported by any of the filtered files 160 | -- and remove all imports starting with `Mathlib` to avoid loops. 161 | let replImports := graph.toList.flatMap 162 | (fun ⟨n, i⟩ => if filter n then i.toList else []) 163 | |>.eraseDups |>.filter (¬ Name.isPrefixOf `Mathlib ·) |>.toArray 164 | let graph := graph.filterMap (fun node edges => if filter node then none else some <| 165 | -- If the node `node` is not filtered, modify the `edges` going into `node`. 166 | edges.toList.flatMap (fun source => 167 | if filter source then 168 | transitiveFilteredUpstream source graph filter (replacement := replacement) 169 | else [source]) |>.eraseDups.toArray) 170 | -- Add a replacement node if provided. 171 | match replacement with 172 | | none => graph 173 | | some repl => graph.insert repl replImports 174 | 175 | end Lean.NameMap 176 | 177 | /-- 178 | Returns a `List (Name × List Name)` with a key for each module `n` in `amongst`, 179 | whose corresponding value is the list of modules `m` in `amongst` which are transitively imported by `n`, 180 | but no declaration in `n` makes use of a declaration in `m`. 181 | -/ 182 | def unusedTransitiveImports (amongst : List Name) (verbose : Bool := false) : CoreM (List (Name × List Name)) := do 183 | let env ← getEnv 184 | let transitiveImports := env.importGraph.transitiveClosure 185 | let transitivelyRequired ← env.transitivelyRequiredModules' amongst verbose 186 | amongst.mapM fun n => do return (n, 187 | let unused := (transitiveImports.find? n).getD {} \ (transitivelyRequired.find? n |>.getD {}) 188 | amongst.filter (fun m => unused.contains m)) 189 | 190 | def Core.withImportModules (modules : Array Name) {α} (f : CoreM α) : IO α := do 191 | initSearchPath (← findSysroot) 192 | unsafe Lean.withImportModules (modules.map (fun m => {module := m})) {} (trustLevel := 1024) 193 | fun env => Prod.fst <$> Core.CoreM.toIO 194 | (ctx := { fileName := "", fileMap := default }) (s := { env := env }) do f 195 | 196 | /-- 197 | Return the redundant imports (i.e. those transitively implied by another import) 198 | of a specified module (or the current module if `none` is specified). 199 | -/ 200 | def redundantImports (n? : Option Name := none) : CoreM NameSet := do 201 | let env ← getEnv 202 | let imports := env.importsOf (n?.getD (env.header.mainModule)) 203 | return env.findRedundantImports imports 204 | 205 | /-- 206 | List the imports in this file which can be removed 207 | because they are transitively implied by another import. 208 | -/ 209 | elab "#redundant_imports" : command => do 210 | let redundant := (← Elab.Command.liftCoreM do redundantImports) 211 | if redundant.isEmpty then 212 | logInfo "No transitively redundant imports found." 213 | else 214 | logInfo <| "Found the following transitively redundant imports:\n" ++ 215 | m!"{Format.joinSep redundant.toList "\n"}" 216 | 217 | /-- 218 | Return the names of the modules in which constants used in the current file were defined, 219 | with modules already transitively imported removed. 220 | 221 | Note that this will *not* account for tactics and syntax used in the file, 222 | so the results may not suffice as imports. 223 | -/ 224 | def Lean.Environment.minimalRequiredModules (env : Environment) : Array Name := 225 | let required := env.requiredModules.toArray.erase env.header.mainModule 226 | let redundant := findRedundantImports env required 227 | required.filter fun n => ¬ redundant.contains n 228 | 229 | /-- 230 | Try to compute a minimal set of imports for this file, 231 | by analyzing the declarations. 232 | 233 | This must be run at the end of the file, 234 | and is not aware of syntax and tactics, 235 | so the results will likely need to be adjusted by hand. 236 | -/ 237 | elab "#min_imports" : command => do 238 | let imports := (← getEnv).minimalRequiredModules.qsort (·.toString < ·.toString) 239 | |>.toList.map (fun n => "import " ++ n.toString) 240 | logInfo <| Format.joinSep imports "\n" 241 | 242 | -- deprecated since 2024-07-06 243 | elab "#minimize_imports" : command => do 244 | logWarning m!"'#minimize_imports' is deprecated: please use '#min_imports'" 245 | Elab.Command.elabCommand (← `(command| #min_imports)) 246 | 247 | /-- 248 | Find locations as high as possible in the import hierarchy 249 | where the named declaration could live. 250 | -/ 251 | def Lean.Name.findHome (n : Name) (env : Option Environment) : CoreM NameSet := do 252 | let current? := match env with | some env => env.header.mainModule | _ => default 253 | let required := (← n.requiredModules).toArray.erase current? 254 | let imports := (← getEnv).importGraph.transitiveClosure 255 | let mut candidates : NameSet := {} 256 | for (n, i) in imports do 257 | if required.all fun r => n == r || i.contains r then 258 | candidates := candidates.insert n 259 | for c in candidates do 260 | for i in candidates do 261 | if imports.find? i |>.getD {} |>.contains c then 262 | candidates := candidates.erase i 263 | return candidates 264 | 265 | open Server in 266 | /-- Tries to resolve the module `modName` to a source file URI. 267 | This has to be done in the Lean server 268 | since the `Environment` does not keep track of source URIs. -/ 269 | @[server_rpc_method] 270 | def getModuleUri (modName : Name) : RequestM (RequestTask Lsp.DocumentUri) := 271 | RequestM.asTask do 272 | let some uri ← documentUriFromModule? modName 273 | | throw $ RequestError.invalidParams s!"couldn't find URI for module '{modName}'" 274 | return uri 275 | 276 | structure GoToModuleLinkProps where 277 | modName : Name 278 | deriving Server.RpcEncodable 279 | 280 | /-- When clicked, this widget component jumps to the source of the module `modName`, 281 | assuming a source URI can be found for the module. -/ 282 | @[widget_module] 283 | def GoToModuleLink : Widget.Module where 284 | javascript := " 285 | import * as React from 'react' 286 | import { EditorContext, useRpcSession } from '@leanprover/infoview' 287 | 288 | export default function(props) { 289 | const ec = React.useContext(EditorContext) 290 | const rs = useRpcSession() 291 | return React.createElement('a', 292 | { 293 | className: 'link pointer dim', 294 | onClick: async () => { 295 | try { 296 | const uri = await rs.call('getModuleUri', props.modName) 297 | ec.revealPosition({ uri, line: 0, character: 0 }) 298 | } catch {} 299 | } 300 | }, 301 | props.modName) 302 | } 303 | " 304 | 305 | open Elab Command in 306 | /-- 307 | Find locations as high as possible in the import hierarchy 308 | where the named declaration could live. 309 | Using `#find_home!` will forcefully remove the current file. 310 | Note that this works best if used in a file with `import Mathlib`. 311 | 312 | The current file could still be the only suggestion, even using `#find_home! lemma`. 313 | The reason is that `#find_home!` scans the import graph below the current file, 314 | selects all the files containing declarations appearing in `lemma`, excluding 315 | the current file itself and looks for all least upper bounds of such files. 316 | 317 | For a simple example, if `lemma` is in a file importing only `A.lean` and `B.lean` and 318 | uses one lemma from each, then `#find_home! lemma` returns the current file. 319 | -/ 320 | elab "#find_home" bang:"!"? n:ident : command => do 321 | let stx ← getRef 322 | let mut homes : Array MessageData := #[] 323 | let n ← liftCoreM <| realizeGlobalConstNoOverloadWithInfo n 324 | let env? ← bang.mapM fun _ => getEnv 325 | for modName in (← Elab.Command.liftCoreM do n.findHome env?) do 326 | let p : GoToModuleLinkProps := { modName } 327 | homes := homes.push $ .ofWidget 328 | (← liftCoreM $ Widget.WidgetInstance.ofHash 329 | GoToModuleLink.javascriptHash $ 330 | Server.RpcEncodable.rpcEncode p) 331 | (toString modName) 332 | logInfoAt stx[0] m!"{homes}" 333 | 334 | 335 | /-- `#import_diff foo bar ...` computes the new transitive imports that are added to a given file when 336 | modules `foo, bar, ...` are added to the set of imports of the file. More precisely, it computes the 337 | import diff between when `foo, bar, ...` are added to the imports and when `foo, bar, ...` are removed 338 | from the imports. 339 | 340 | Note: the command also works when some of the modules passed as arguments are already present in the file's 341 | imports. -/ 342 | elab "#import_diff" n:ident* : command => do 343 | let name_arr : Array Name := n.map (·.getId) 344 | let sp ← searchPathRef.get 345 | -- First, make sure the files exist. 346 | for name in name_arr do 347 | if (← sp.findWithExt "olean" name).isSome then continue 348 | throwError m!"File {name} cannot be found." 349 | let env ← getEnv 350 | -- Next, check for redundancies: 351 | let current_all_imports := env.allImportedModuleNames 352 | let redundancies := name_arr.filter current_all_imports.contains 353 | unless redundancies.isEmpty do 354 | let out := "\n".intercalate <| redundancies.map Name.toString |>.qsort (· < ·) |>.toList 355 | Lean.logInfo <| m!"The following are already imported (possibly transitively):\n{out}" 356 | -- Now compute the import diffs. 357 | let current_imports := env.imports 358 | let reduced_imports := env.imports.filter (!name_arr.contains ·.module) 359 | let extended_imports := current_imports ++ (name_arr.map ({ module := · })) 360 | let reduced_all_imports := (← Lean.importModules reduced_imports {}).allImportedModuleNames 361 | let extended_all_imports := (← Lean.importModules extended_imports {}).allImportedModuleNames 362 | let import_diff := extended_all_imports.filter (· ∉ reduced_all_imports) 363 | let out := "\n".intercalate <| import_diff.map Name.toString |>.qsort (· < ·) |>.toList 364 | Lean.logInfo s!"Found {import_diff.size} additional imports:\n{out}" 365 | -------------------------------------------------------------------------------- /ImportGraph/Lean/Name.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jon Eugster. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jon Eugster 5 | -/ 6 | import Lean.Data.Name 7 | import Lean.CoreM 8 | import Lean.Meta.Match.MatcherInfo 9 | import Std.Data.HashMap 10 | 11 | /-! 12 | TODO: Some declarations in this file are duplicated from mathlib, but especially `isBlacklisted` 13 | is deemed to specific for upstreaming to Batteries. 14 | -/ 15 | namespace Lean.Name 16 | 17 | open Lean Meta Elab 18 | 19 | namespace ImportGraph 20 | 21 | /-- Note: copied from `Mathlib.Lean.Name` -/ 22 | private def isBlackListed (declName : Name) : CoreM Bool := do 23 | if declName.toString.startsWith "Lean" then return true 24 | let env ← getEnv 25 | pure $ declName.isInternalDetail 26 | || isAuxRecursor env declName 27 | || isNoConfusion env declName 28 | <||> isRec declName <||> isMatcher declName 29 | 30 | /-- 31 | Retrieve all names in the environment satisfying a predicate. 32 | 33 | Note: copied from `Mathlib.Lean.Name` 34 | -/ 35 | def allNames (p : Name → Bool) : CoreM (Array Name) := do 36 | (← getEnv).constants.foldM (init := #[]) fun names n _ => do 37 | if p n && !(← isBlackListed n) then 38 | return names.push n 39 | else 40 | return names 41 | 42 | /-- 43 | Retrieve all names in the environment satisfying a predicate, 44 | gathered together into a `HashMap` according to the module they are defined in. 45 | 46 | Note: copied from `Mathlib.Lean.Name` 47 | -/ 48 | def allNamesByModule (p : Name → Bool) : CoreM (Std.HashMap Name (Array Name)) := do 49 | (← getEnv).constants.foldM (init := ∅) fun names n _ => do 50 | if p n && !(← isBlackListed n) then 51 | let some m ← findModuleOf? n | return names 52 | -- TODO use `modify`/`alter` when available 53 | match names[m]? with 54 | | some others => return names.insert m (others.push n) 55 | | none => return names.insert m #[n] 56 | else 57 | return names 58 | 59 | /-- Returns the very first part of a name: for `ImportGraph.Lean.NameMap` it 60 | returns `ImportGraph`. 61 | -/ 62 | def getModule (name : Name) (s := "") : Name := 63 | match name with 64 | | .anonymous => .mkSimple s 65 | | .num _ _ => panic s!"panic in `getModule`: did not expect numerical name: {name}." 66 | | .str pre s => getModule pre s 67 | -------------------------------------------------------------------------------- /ImportGraph/RequiredModules.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Kim Morrison. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Kim Morrison 5 | -/ 6 | import Lean.CoreM 7 | import Lean.Data.NameMap 8 | import Lean.Environment 9 | import Lean.Util.FoldConsts 10 | 11 | namespace Lean 12 | 13 | /-- Return the name of the module in which a declaration was defined. -/ 14 | def Environment.getModuleFor? (env : Environment) (declName : Name) : Option Name := 15 | match env.getModuleIdxFor? declName with 16 | | none => 17 | if env.constants.map₂.contains declName then 18 | env.header.mainModule 19 | else 20 | none 21 | | some idx => env.header.moduleNames[idx.toNat]! 22 | 23 | open Lean 24 | 25 | /-- 26 | Return the names of the modules in which constants used in the specified declaration were defined. 27 | 28 | Note that this will *not* account for tactics and syntax used in the declaration, 29 | so the results may not suffice as imports. 30 | -/ 31 | def Name.requiredModules (n : Name) : CoreM NameSet := do 32 | let env ← getEnv 33 | let mut requiredModules : NameSet := {} 34 | let ci ← getConstInfo n 35 | for n in ci.getUsedConstantsAsSet do 36 | match env.getModuleFor? n with 37 | | some m => 38 | if ¬ (`Init).isPrefixOf m then 39 | requiredModules := requiredModules.insert m 40 | | none => pure () 41 | return requiredModules 42 | 43 | /-- 44 | Return the names of the constants used in the specified declarations, 45 | and the constants they use transitively. 46 | -/ 47 | def NameSet.transitivelyUsedConstants (s : NameSet) : CoreM NameSet := do 48 | let mut usedConstants : NameSet := {} 49 | let mut toProcess : NameSet := s 50 | while !toProcess.isEmpty do 51 | let current := toProcess.min! 52 | toProcess := toProcess.erase current 53 | usedConstants := usedConstants.insert current 54 | for m in (← getConstInfo current).getUsedConstantsAsSet do 55 | if !usedConstants.contains m then 56 | toProcess := toProcess.insert m 57 | return usedConstants 58 | 59 | /-- 60 | Return the names of the constants used in the specified declaration, 61 | and the constants they use transitively. 62 | -/ 63 | def Name.transitivelyUsedConstants (n : Name) : CoreM NameSet := 64 | NameSet.transitivelyUsedConstants {n} 65 | 66 | /-- 67 | Return the names of the modules in which constants used transitively 68 | in the specified declarations were defined. 69 | 70 | Note that this will *not* account for tactics and syntax used in the declaration, 71 | so the results may not suffice as imports. 72 | -/ 73 | def NameSet.transitivelyRequiredModules (s : NameSet) (env : Environment) : CoreM NameSet := do 74 | let mut requiredModules : NameSet := {} 75 | for m in (← s.transitivelyUsedConstants) do 76 | if let some module := env.getModuleFor? m then 77 | requiredModules := requiredModules.insert module 78 | return requiredModules 79 | 80 | /-- 81 | Return the names of the modules in which constants used transitively 82 | in the specified declaration were defined. 83 | 84 | Note that this will *not* account for tactics and syntax used in the declaration, 85 | so the results may not suffice as imports. 86 | -/ 87 | def Name.transitivelyRequiredModules (n : Name) (env : Environment) : CoreM NameSet := 88 | NameSet.transitivelyRequiredModules {n} env 89 | 90 | /-- 91 | Finds all constants defined in the specified module, 92 | and identifies all modules containing constants which are transitively required by those constants. 93 | -/ 94 | def Environment.transitivelyRequiredModules (env : Environment) (module : Name) : CoreM NameSet := do 95 | let constants := env.constants.map₁.values.map (·.name) 96 | |>.filter (! ·.isInternal) 97 | |>.filter (env.getModuleFor? · = some module) 98 | (NameSet.ofList constants).transitivelyRequiredModules env 99 | 100 | /-- 101 | Computes all the modules transitively required by the specified modules. 102 | Should be equivalent to calling `transitivelyRequiredModules` on each module, but shares more of the work. 103 | -/ 104 | partial def Environment.transitivelyRequiredModules' (env : Environment) (modules : List Name) (verbose : Bool := false) : 105 | CoreM (NameMap NameSet) := do 106 | let N := env.header.moduleNames.size 107 | let mut c2m : NameMap (BitVec N) := {} 108 | let mut pushed : NameSet := {} 109 | let mut result : NameMap NameSet := {} 110 | for m in modules do 111 | if verbose then 112 | IO.println s!"Processing module {m}" 113 | let mut r : BitVec N := 0 114 | for n in env.header.moduleData[(env.header.moduleNames.idxOf? m).getD 0]!.constNames do 115 | if ! n.isInternal then 116 | -- This is messy: Mathlib is big enough that writing a recursive function causes a stack overflow. 117 | -- So we use an explicit stack instead. We visit each constant twice: 118 | -- once to record the constants transitively used by it, 119 | -- and again to record the modules which defined those constants. 120 | let mut stack : List (Name × Option NameSet) := [⟨n, none⟩] 121 | pushed := pushed.insert n 122 | while !stack.isEmpty do 123 | match stack with 124 | | [] => panic! "Stack is empty" 125 | | (c, used?) :: tail => 126 | stack := tail 127 | match used? with 128 | | none => 129 | if !c2m.contains c then 130 | let used := (← getConstInfo c).getUsedConstantsAsSet 131 | stack := ⟨c, some used⟩ :: stack 132 | for u in used do 133 | if !pushed.contains u then 134 | stack := ⟨u, none⟩ :: stack 135 | pushed := pushed.insert u 136 | | some used => 137 | let usedModules : NameSet := 138 | used.foldl (init := {}) (fun s u => if let some m := env.getModuleFor? u then s.insert m else s) 139 | let transitivelyUsed : BitVec N := 140 | used.foldl (init := toBitVec usedModules) (fun s u => s ||| ((c2m.find? u).getD 0)) 141 | c2m := c2m.insert c transitivelyUsed 142 | r := r ||| ((c2m.find? n).getD 0) 143 | result := result.insert m (toNameSet r) 144 | return result 145 | where 146 | toBitVec {N : Nat} (s : NameSet) : BitVec N := 147 | s.foldl (init := 0) (fun b n => b ||| BitVec.twoPow _ ((env.header.moduleNames.idxOf? n).getD 0)) 148 | toNameSet {N : Nat} (b : BitVec N) : NameSet := 149 | env.header.moduleNames.zipIdx.foldl (init := {}) (fun s (n, i) => if b.getLsbD i then s.insert n else s) 150 | 151 | /-- 152 | Return the names of the modules in which constants used in the current file were defined. 153 | 154 | Note that this will *not* account for tactics and syntax used in the file, 155 | so the results may not suffice as imports. 156 | -/ 157 | def Environment.requiredModules (env : Environment) : NameSet := Id.run do 158 | let localConstantInfos := env.constants.map₂ 159 | let mut requiredModules : NameSet := {} 160 | for (_, ci) in localConstantInfos do 161 | for n in ci.getUsedConstantsAsSet do 162 | match env.getModuleFor? n with 163 | | some m => 164 | if ¬ (`Init).isPrefixOf m then 165 | requiredModules := requiredModules.insert m 166 | | none => pure () 167 | return requiredModules 168 | 169 | end Lean 170 | -------------------------------------------------------------------------------- /ImportGraph/UnusedTransitiveImports.lean: -------------------------------------------------------------------------------- 1 | import ImportGraph.Imports 2 | 3 | open Lean 4 | 5 | /-- 6 | `lake exe unused_transitive_imports m1 m2 ...` 7 | 8 | For each specified module `m`, prints those `n` from the argument list which are imported, but transitively unused by `m`. 9 | -/ 10 | def main (args : List String) : IO UInt32 := do 11 | let (flags, args) := args.partition (fun s => s.startsWith "-") 12 | let mut modules := args.map (fun s => s.toName) 13 | Core.withImportModules modules.toArray do 14 | let r ← unusedTransitiveImports modules (verbose := flags.contains "-v" || flags.contains "--verbose") 15 | for (n, u) in r do 16 | IO.println s!"{n}: {u}" 17 | return 0 18 | -------------------------------------------------------------------------------- /ImportGraphTest.lean: -------------------------------------------------------------------------------- 1 | import ImportGraphTest.AnotherFileWithTransitiveImports 2 | import ImportGraphTest.Dot 3 | import ImportGraphTest.FileWithTransitiveImports 4 | import ImportGraphTest.Imports 5 | import ImportGraphTest.Unused 6 | import ImportGraphTest.Used 7 | -------------------------------------------------------------------------------- /ImportGraphTest/.gitignore: -------------------------------------------------------------------------------- 1 | produced.dot -------------------------------------------------------------------------------- /ImportGraphTest/AnotherFileWithTransitiveImports.lean: -------------------------------------------------------------------------------- 1 | import ImportGraph.Imports 2 | import ImportGraphTest.Unused 3 | import ImportGraphTest.FileWithTransitiveImports 4 | 5 | /-- 6 | info: The following are already imported (possibly transitively): 7 | ImportGraphTest.FileWithTransitiveImports 8 | --- 9 | info: Found 2 additional imports: 10 | ImportGraphTest.FileWithTransitiveImports 11 | ImportGraphTest.Used 12 | -/ 13 | #guard_msgs in 14 | #import_diff ImportGraphTest.FileWithTransitiveImports 15 | 16 | /-- 17 | info: The following are already imported (possibly transitively): 18 | ImportGraphTest.FileWithTransitiveImports 19 | ImportGraphTest.Used 20 | --- 21 | info: Found 2 additional imports: 22 | ImportGraphTest.FileWithTransitiveImports 23 | ImportGraphTest.Used 24 | -/ 25 | #guard_msgs in 26 | #import_diff ImportGraphTest.FileWithTransitiveImports ImportGraphTest.Used 27 | 28 | 29 | /-- error: File SomeBogusFilename cannot be found. -/ 30 | #guard_msgs in 31 | #import_diff SomeBogusFilename 32 | -------------------------------------------------------------------------------- /ImportGraphTest/Dot.lean: -------------------------------------------------------------------------------- 1 | import ImportGraph.Cli 2 | 3 | def readFile (path : System.FilePath) : IO String := 4 | IO.FS.readFile path 5 | 6 | def runGraphCommand : IO Unit := do 7 | let _ ← IO.Process.output { 8 | cmd := "lake" 9 | args := #["exe", "graph", "--to", "ImportGraphTest.Used", "ImportGraphTest/produced.dot"] 10 | } 11 | 12 | def compareOutputs (expected : String) (actual : String) : IO Bool := do 13 | let expectedLines := expected.splitOn "\n" |>.filter (·.trim.length > 0) |>.map String.trim 14 | let actualLines := actual.splitOn "\n" |>.filter (·.trim.length > 0) |>.map String.trim 15 | pure (expectedLines == actualLines) 16 | 17 | /-- info: Test passed: The graph command output matches the expected.dot file. -/ 18 | #guard_msgs in 19 | #eval show IO Unit from do 20 | runGraphCommand 21 | let expectedOutput ← readFile "ImportGraphTest/expected.dot" 22 | let actualOutput ← readFile "ImportGraphTest/produced.dot" 23 | let isEqual ← compareOutputs expectedOutput actualOutput 24 | if isEqual then 25 | IO.println "Test passed: The graph command output matches the expected.dot file." 26 | else 27 | IO.println "Test failed: The graph command output does not match the expected.dot file." 28 | IO.println s!"Expected:\n{expectedOutput}" 29 | IO.println s!"Actual:\n{actualOutput}" 30 | -------------------------------------------------------------------------------- /ImportGraphTest/FileWithTransitiveImports.lean: -------------------------------------------------------------------------------- 1 | import ImportGraph.Imports 2 | import ImportGraphTest.Used 3 | 4 | /-- 5 | info: The following are already imported (possibly transitively): ImportGraphTest.Used 6 | --- 7 | info: Found 2 additional imports: 8 | ImportGraphTest.Unused 9 | ImportGraphTest.Used 10 | -/ 11 | #guard_msgs in 12 | #import_diff ImportGraphTest.Used 13 | -------------------------------------------------------------------------------- /ImportGraphTest/Imports.lean: -------------------------------------------------------------------------------- 1 | import ImportGraph.Imports 2 | import ImportGraph.RequiredModules 3 | import ImportGraphTest.Used 4 | 5 | open Lean 6 | 7 | def importTest : CoreM Unit := do 8 | let x ← redundantImports 9 | logInfo s!"{x.toArray}" 10 | 11 | /-- 12 | info: Found the following transitively redundant imports: 13 | ImportGraph.RequiredModules 14 | -/ 15 | #guard_msgs in 16 | #redundant_imports 17 | 18 | /-- info: import ImportGraph.Imports -/ 19 | #guard_msgs in 20 | #min_imports 21 | 22 | /-- info: [ImportGraph.Imports] -/ 23 | #guard_msgs in 24 | #find_home importTest 25 | 26 | open Elab Command 27 | 28 | 29 | /-- 30 | Reports unused transitive imports amongst the specified modules. 31 | -/ 32 | elab "#unused_transitive_imports" names:ident* : command => do 33 | let imports := (names.map Syntax.getId).toList 34 | let unused ← Elab.Command.liftCoreM (unusedTransitiveImports imports) 35 | for (n, u) in unused do 36 | if !u.isEmpty then 37 | logInfo <| s!"Transitively unused imports of {n}:\n{"\n".intercalate (u.map (fun i => s!" {i}"))}" 38 | 39 | /-- 40 | info: Transitively unused imports of Init.System.IO: 41 | Init.Control.StateRef 42 | Init.Control.Reader 43 | -/ 44 | #guard_msgs in 45 | #unused_transitive_imports Init.Control.StateRef Init.System.IO Init.Control.Reader Init.Control.Basic 46 | 47 | /-- 48 | info: Transitively unused imports of ImportGraphTest.Used: 49 | ImportGraphTest.Unused 50 | -/ 51 | #guard_msgs in 52 | #unused_transitive_imports ImportGraphTest.Used ImportGraphTest.Unused Init.Control.Reader 53 | 54 | elab "#transitivelyRequiredModules_test" : command => do 55 | let env ← getEnv 56 | let unused ← liftCoreM <| env.transitivelyRequiredModules `ImportGraph.RequiredModules 57 | logInfo s!"{unused.contains `Init.Data.Option.Lemmas}" 58 | 59 | /-- info: true -/ 60 | #guard_msgs in 61 | #transitivelyRequiredModules_test 62 | 63 | elab "#my_test" : command => do 64 | -- functionality of `#redundant_imports` 65 | let expected := #[`ImportGraph.RequiredModules] 66 | let ri ← liftCoreM redundantImports 67 | if (ri.toArray != expected) then 68 | logError s!"Failed: `redundantImports` returned {ri.toArray} instead of {expected}" 69 | 70 | -- functionality of `#find_home` 71 | let expected := #[`ImportGraph.Imports] 72 | let mi ← liftCoreM <| Lean.Name.findHome `importTest none 73 | if (mi.toArray != expected) then 74 | logError s!"Failed: `findHome` returned {mi.toArray} instead of {expected}" 75 | 76 | -- functionality of `#find_home!` 77 | let expected := #[`ImportGraph.Imports] 78 | let mi! ← liftCoreM <| Lean.Name.findHome `importTest (← getEnv) 79 | if (mi!.toArray != expected) then 80 | logError s!"Failed: `findHome (!)` returned {mi!.toArray} instead of {expected}" 81 | 82 | logInfo s!"{mi.toArray}" 83 | pure () 84 | 85 | /-- info: #[ImportGraph.Imports] -/ 86 | #guard_msgs in 87 | #my_test 88 | -------------------------------------------------------------------------------- /ImportGraphTest/Unused.lean: -------------------------------------------------------------------------------- 1 | /-! 2 | This is a dummy file to test that unused files are shaded in the import graph. 3 | -/ 4 | -------------------------------------------------------------------------------- /ImportGraphTest/Used.lean: -------------------------------------------------------------------------------- 1 | import ImportGraphTest.Unused 2 | -------------------------------------------------------------------------------- /ImportGraphTest/expected.dot: -------------------------------------------------------------------------------- 1 | digraph "import_graph" { 2 | "ImportGraphTest.Unused" [style=filled, fillcolor="#e0e0e0", shape=ellipse]; 3 | "ImportGraphTest.Used" [style=filled, fillcolor="#e0e0e0", shape=house]; 4 | "ImportGraphTest.Unused" -> "ImportGraphTest.Used"; 5 | } -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | APPENDIX: How to apply the Apache License to your work. 179 | 180 | To apply the Apache License to your work, attach the following 181 | boilerplate notice, with the fields enclosed by brackets "[]" 182 | replaced with your own identifying information. (Don't include 183 | the brackets!) The text should be enclosed in the appropriate 184 | comment syntax for the file format. We also recommend that a 185 | file or class name and description of purpose be included on the 186 | same "printed page" as the copyright notice for easier 187 | identification within third-party archives. 188 | 189 | Copyright [yyyy] [name of copyright owner] 190 | 191 | Licensed under the Apache License, Version 2.0 (the "License"); 192 | you may not use this file except in compliance with the License. 193 | You may obtain a copy of the License at 194 | 195 | http://www.apache.org/licenses/LICENSE-2.0 196 | 197 | Unless required by applicable law or agreed to in writing, software 198 | distributed under the License is distributed on an "AS IS" BASIS, 199 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 200 | See the License for the specific language governing permissions and 201 | limitations under the License. 202 | -------------------------------------------------------------------------------- /Main.lean: -------------------------------------------------------------------------------- 1 | import ImportGraph.Cli 2 | 3 | /-! 4 | # `lake exe graph` 5 | 6 | This is a replacement for Lean 3's `leanproject import-graph` tool. 7 | -/ 8 | 9 | open Cli 10 | 11 | /-- Setting up command line options and help text for `lake exe graph`. -/ 12 | def graph : Cmd := `[Cli| 13 | graph VIA importGraphCLI; ["0.0.3"] 14 | "Generate representations of a Lean import graph. \ 15 | By default generates the import graph up to `Mathlib`. \ 16 | If you are working in a downstream project, use `lake exe graph --to MyProject`." 17 | 18 | FLAGS: 19 | "show-transitive"; "Show transitively redundant edges." 20 | "to" : Array ModuleName; "Only show the upstream imports of the specified modules." 21 | "from" : Array ModuleName; "Only show the downstream dependencies of the specified modules." 22 | "exclude-meta"; "Exclude any files starting with `Mathlib.[Tactic|Lean|Util|Mathport]`." 23 | "include-direct"; "Include directly imported files from other libraries" 24 | "include-deps"; "Include used files from other libraries (not including Lean itself and `std`)" 25 | "include-std"; "Include used files from the Lean standard library (implies `--include-deps`)" 26 | "include-lean"; "Include used files from Lean itself (implies `--include-deps` and `--include-std`)" 27 | "mark-package"; "Visually highlight the package containing the first `--to` target (used in combination with some `--include-XXX`)." 28 | 29 | ARGS: 30 | ...outputs : String; "Filename(s) for the output. \ 31 | If none are specified, generates `import_graph.dot`. \ 32 | Automatically chooses the format based on the file extension. \ 33 | Currently supported formats are `.dot`, `.gexf`, `.html`, \ 34 | and if you have `graphviz` installed then any supported output format is allowed." 35 | ] 36 | 37 | 38 | /-- `lake exe graph` -/ 39 | def main (args : List String) : IO UInt32 := 40 | graph.validate args 41 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # importGraph 2 | 3 | A simple tool to create import graphs of lake packages. 4 | 5 | 6 | ## Requirements 7 | 8 | For creating different output formats than `.dot` (for example to create a `.pdf` file), you should have [`graphviz`](https://graphviz.org/) installed. 9 | 10 | ## Usage 11 | 12 | If you are using mathlib, the tool will already be available. If not, see installation notes below. 13 | 14 | Once available in your project, you can create import graphs with 15 | 16 | ```bash 17 | lake exe graph 18 | ``` 19 | 20 | A typical command is 21 | 22 | ```bash 23 | lake exe graph --to MyModule my_graph.pdf 24 | ``` 25 | where `MyModule` follows the same module naming you would use to `import` it in lean. See `lake exe graph --help` for more options. 26 | 27 | You can specify multiple sources and targets e.g. as 28 | ```bash 29 | lake exe graph --from MyModule1,MyModule2 --to MyModule3,MyModule4 my_graph.pdf 30 | ``` 31 | 32 | ### Troubleshoot 33 | 34 | * make sure to `lake build` your project (or the specified `--to` module) before using `lake exe graph`! 35 | 36 | ### Json 37 | 38 | To create a Json file, you can use `.xdot_json` as output type: 39 | 40 | ```bash 41 | lake exe graph my_graph.xdot_json 42 | ``` 43 | 44 | ### HTML 45 | 46 | ``` 47 | lake exe graph my_graph.html 48 | ``` 49 | 50 | creates a stand-alone HTML file visualising the import structure. 51 | 52 | ## Commands 53 | 54 | There are a few commands implemented, which help you analysing the imports of a file. These are accessible by adding `import ImportGraph.Imports` to your lean file. 55 | 56 | * `#redundant_imports`: lists any transitively redundant imports in the current module. 57 | * `#min_imports`: attempts to construct a minimal set of imports for the declarations 58 | in the current file. 59 | (Must be run at the end of the file. Tactics and macros may result in incorrect output.) 60 | * `#find_home decl`: suggests files higher up the import hierarchy to which `decl` could be moved. 61 | 62 | ## Other executables 63 | 64 | `lake exe unused_transitive_imports m1 m2 ...` 65 | 66 | For each specified module `m`, prints those `n` from the argument list which are imported, but transitively unused by `m`. 67 | 68 | ## Installation 69 | 70 | The installation works exactly like for any [Lake package](https://reservoir.lean-lang.org/), 71 | see [Lake docs](https://github.com/leanprover/lean4/tree/master/src/lake#supported-sources). 72 | 73 | *This only relevant if your project does not already require `importGraph` through another lake package (e.g. mathlib). If it does, do not follow these instructions; instead just use the tool with `lake exe graph`!* 74 | 75 | You can import this in any lean projects by the following line to your `lakefile.lean`: 76 | 77 | ```lean 78 | require "leanprover-community" / "importGraph" @ git "main" 79 | ``` 80 | 81 | or, if you have a `lakefile.toml`, it would be 82 | 83 | ```toml 84 | [[require]] 85 | name = "importGraph" 86 | source = "leanprover-community" 87 | rev = "main" 88 | ``` 89 | 90 | Then, you might need to call `lake update -R importGraph` in your project. 91 | 92 | ## Contribution 93 | 94 | Please open PRs/Issues if you have troubles or would like to contribute new features! 95 | 96 | ## Credits 97 | 98 | The main tool has been extracted from [mathlib](https://github.com/leanprover-community/mathlib4), 99 | originally written by Kim Morrison and other mathlib contributors. 100 | 101 | The HTML visualisation has been incorporated from 102 | [a project by Eric Wieser](https://github.com/eric-wieser/mathlib-import-graph). 103 | 104 | ### Maintainers 105 | 106 | Primarily maintained by [Jon Eugster](https://leanprover.zulipchat.com/#narrow/dm/385895-Jon-Eugster), Kim Morrison, and the wider leanprover community. 107 | -------------------------------------------------------------------------------- /html-template/.gitignore: -------------------------------------------------------------------------------- 1 | imports.gexf 2 | -------------------------------------------------------------------------------- /html-template/LICENSE_source: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2021 Eric Wieser 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /html-template/README.md: -------------------------------------------------------------------------------- 1 | # Visualised import graph 2 | 3 | ## Instructions 4 | 5 | To test this, place a file `imports.gexf` inside this directory. You can create such a file with 6 | 7 | ``` 8 | lake exe graph html-template/imports.gexf 9 | ``` 10 | 11 | Then open `index.html` in any browser and you should see the graph. 12 | 13 | ## Development 14 | 15 | Currently `lake exe graph output.html` will use the files here to create a stand-alone 16 | HTML file. It does so by search-replacing the JS-scripts, the `fetch('imports.gexf')` 17 | statement, and the `

` header. 18 | 19 | Therefore any modifications to these lines need to be reflected in `ImportGraph/Cli.lean`! 20 | 21 | # Credits 22 | 23 | This tool has been adapted from its [Lean 3 version](https://github.com/eric-wieser/mathlib-import-graph) written by Eric Wieser, which was published under the [MIT License](./LICENSE_source) 24 | included here. 25 | 26 | Adaptation by Jon Eugster. 27 | -------------------------------------------------------------------------------- /html-template/index.html: -------------------------------------------------------------------------------- 1 | 2 | 3 | 4 | 5 | 6 | 7 | 8 | import graph 9 | 10 | 12 | 13 | 14 | 15 | 16 | 17 | 100 | 101 | 102 |
103 |
104 | 105 |

Import Graph

106 |

Built with Sigma.js. 107 | Node sizes indicate the number of declarations in the file.

108 | 109 |
110 |
111 |
    112 |
    113 |

    Hover over a node to show only the files importing it. Hover over a directory name to highlight only the files in that directory

    114 |
    115 |
    116 |
    117 | 118 | 515 | 516 | 517 | 518 | -------------------------------------------------------------------------------- /html-template/vendor/graphology-LICENSE: -------------------------------------------------------------------------------- 1 | The MIT License (MIT) 2 | 3 | Copyright (c) 2016-2021 Guillaume Plique (Yomguithereal) 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in 13 | all copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN 21 | THE SOFTWARE. 22 | -------------------------------------------------------------------------------- /html-template/vendor/graphology.min.js: -------------------------------------------------------------------------------- 1 | !function(e,t){"object"==typeof exports&&"undefined"!=typeof module?module.exports=t():"function"==typeof define&&define.amd?define(t):(e="undefined"!=typeof globalThis?globalThis:e||self).graphology=t()}(this,(function(){"use strict";function e(t){return e="function"==typeof Symbol&&"symbol"==typeof Symbol.iterator?function(e){return typeof e}:function(e){return e&&"function"==typeof Symbol&&e.constructor===Symbol&&e!==Symbol.prototype?"symbol":typeof e},e(t)}function t(e,t){e.prototype=Object.create(t.prototype),e.prototype.constructor=e,r(e,t)}function n(e){return n=Object.setPrototypeOf?Object.getPrototypeOf:function(e){return e.__proto__||Object.getPrototypeOf(e)},n(e)}function r(e,t){return r=Object.setPrototypeOf||function(e,t){return e.__proto__=t,e},r(e,t)}function i(){if("undefined"==typeof Reflect||!Reflect.construct)return!1;if(Reflect.construct.sham)return!1;if("function"==typeof Proxy)return!0;try{return Boolean.prototype.valueOf.call(Reflect.construct(Boolean,[],(function(){}))),!0}catch(e){return!1}}function o(e,t,n){return o=i()?Reflect.construct:function(e,t,n){var i=[null];i.push.apply(i,t);var o=new(Function.bind.apply(e,i));return n&&r(o,n.prototype),o},o.apply(null,arguments)}function a(e){var t="function"==typeof Map?new Map:void 0;return a=function(e){if(null===e||(i=e,-1===Function.toString.call(i).indexOf("[native code]")))return e;var i;if("function"!=typeof e)throw new TypeError("Super expression must either be null or a function");if(void 0!==t){if(t.has(e))return t.get(e);t.set(e,a)}function a(){return o(e,arguments,n(this).constructor)}return a.prototype=Object.create(e.prototype,{constructor:{value:a,enumerable:!1,writable:!0,configurable:!0}}),r(a,e)},a(e)}function u(e){if(void 0===e)throw new ReferenceError("this hasn't been initialised - super() hasn't been called");return e}var d=function(){for(var e=arguments[0],t=1,n=arguments.length;t0&&a.length>i&&!a.warned){a.warned=!0;var d=new Error("Possible EventEmitter memory leak detected. "+a.length+" "+String(t)+" listeners added. Use emitter.setMaxListeners() to increase limit");d.name="MaxListenersExceededWarning",d.emitter=e,d.type=t,d.count=a.length,u=d,console&&console.warn&&console.warn(u)}return e}function S(){if(!this.fired)return this.target.removeListener(this.type,this.wrapFn),this.fired=!0,0===arguments.length?this.listener.call(this.target):this.listener.apply(this.target,arguments)}function A(e,t,n){var r={fired:!1,wrapFn:void 0,target:e,type:t,listener:n},i=S.bind(r);return i.listener=n,r.wrapFn=i,i}function L(e,t,n){var r=e._events;if(void 0===r)return[];var i=r[t];return void 0===i?[]:"function"==typeof i?n?[i.listener||i]:[i]:n?function(e){for(var t=new Array(e.length),n=0;n0&&(o=t[0]),o instanceof Error)throw o;var a=new Error("Unhandled error."+(o?" ("+o.message+")":""));throw a.context=o,a}var u=i[e];if(void 0===u)return!1;if("function"==typeof u)w(u,this,t);else{var d=u.length,c=D(u,d);for(n=0;n=0;o--)if(n[o]===t||n[o].listener===t){a=n[o].listener,i=o;break}if(i<0)return this;0===i?n.shift():function(e,t){for(;t+1=0;r--)this.removeListener(e,t[r]);return this},_.prototype.listeners=function(e){return L(this,e,!0)},_.prototype.rawListeners=function(e){return L(this,e,!1)},_.listenerCount=function(e,t){return"function"==typeof e.listenerCount?e.listenerCount(t):N.call(e,t)},_.prototype.listenerCount=N,_.prototype.eventNames=function(){return this._eventsCount>0?y(this._events):[]},U.prototype.next=function(){if(this.done)return{done:!0};var e=this._next();return e.done&&(this.done=!0),e},"undefined"!=typeof Symbol&&(U.prototype[Symbol.iterator]=function(){return this}),U.of=function(){var e=arguments,t=e.length,n=0;return new U((function(){return n>=t?{done:!0}:{done:!1,value:e[n++]}}))},U.empty=function(){var e=new U(null);return e.done=!0,e},U.is=function(e){return e instanceof U||"object"==typeof e&&null!==e&&"function"==typeof e.next};var O=U,C=function(e,t){for(var n,r=arguments.length>1?t:1/0,i=r!==1/0?new Array(r):[],o=0;;){if(o===r)return i;if((n=e.next()).done)return o!==t?i.slice(0,o):i;i[o++]=n.value}},K=function(e){function n(t,n){var r;return(r=e.call(this)||this).name="GraphError",r.message=t||"",r.data=n||{},r}return t(n,e),n}(a(Error)),z=function(e){function n(t,r){var i;return(i=e.call(this,t,r)||this).name="InvalidArgumentsGraphError","function"==typeof Error.captureStackTrace&&Error.captureStackTrace(u(i),n.prototype.constructor),i}return t(n,e),n}(K),M=function(e){function n(t,r){var i;return(i=e.call(this,t,r)||this).name="NotFoundGraphError","function"==typeof Error.captureStackTrace&&Error.captureStackTrace(u(i),n.prototype.constructor),i}return t(n,e),n}(K),P=function(e){function n(t,r){var i;return(i=e.call(this,t,r)||this).name="UsageGraphError","function"==typeof Error.captureStackTrace&&Error.captureStackTrace(u(i),n.prototype.constructor),i}return t(n,e),n}(K);function T(e,t){this.key=e,this.attributes=t,this.inDegree=0,this.outDegree=0,this.undirectedDegree=0,this.directedSelfLoops=0,this.undirectedSelfLoops=0,this.in={},this.out={},this.undirected={}}function R(e,t){this.key=e,this.attributes=t,this.inDegree=0,this.outDegree=0,this.directedSelfLoops=0,this.in={},this.out={}}function F(e,t){this.key=e,this.attributes=t,this.undirectedDegree=0,this.undirectedSelfLoops=0,this.undirected={}}function I(e,t,n,r,i,o){this.key=t,this.attributes=o,this.undirected=e,this.source=r,this.target=i,this.generatedKey=n}function W(e,t,n,r,i,o,a){var u,d,c="out",s="in";if(t&&(c=s="undirected"),e.multi){if(void 0===(d=(u=o[c])[i])&&(d=new Set,u[i]=d),d.add(n),r===i&&t)return;void 0===(u=a[s])[r]&&(u[r]=d)}else{if(o[c][i]=n,r===i&&t)return;a[s][r]=n}}function Y(e,t,n){var r=e.multi,i=n.source,o=n.target,a=i.key,u=o.key,d=i[t?"undirected":"out"],c=t?"undirected":"in";if(u in d)if(r){var s=d[u];1===s.size?(delete d[u],delete o[c][a]):s.delete(n)}else delete d[u];r||delete o[c][a]}R.prototype.upgradeToMixed=function(){this.undirectedDegree=0,this.undirectedSelfLoops=0,this.undirected={}},F.prototype.upgradeToMixed=function(){this.inDegree=0,this.outDegree=0,this.directedSelfLoops=0,this.in={},this.out={}};var B=[{name:function(e){return"get".concat(e,"Attribute")},attacher:function(e,t,n){e.prototype[t]=function(e,r){var i;if("mixed"!==this.type&&"mixed"!==n&&n!==this.type)throw new P("Graph.".concat(t,": cannot find this type of edges in your ").concat(this.type," graph."));if(arguments.length>2){if(this.multi)throw new P("Graph.".concat(t,": cannot use a {source,target} combo when asking about an edge's attributes in a MultiGraph since we cannot infer the one you want information about."));var o=""+e,a=""+r;if(r=arguments[2],!(i=c(this,o,a,n)))throw new M("Graph.".concat(t,': could not find an edge for the given path ("').concat(o,'" - "').concat(a,'").'))}else if(e=""+e,!(i=this._edges.get(e)))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" edge in the graph.'));if("mixed"!==n&&i.undirected!==("undirected"===n))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" ').concat(n," edge in the graph."));return i.attributes[r]}}},{name:function(e){return"get".concat(e,"Attributes")},attacher:function(e,t,n){e.prototype[t]=function(e){var r;if("mixed"!==this.type&&"mixed"!==n&&n!==this.type)throw new P("Graph.".concat(t,": cannot find this type of edges in your ").concat(this.type," graph."));if(arguments.length>1){if(this.multi)throw new P("Graph.".concat(t,": cannot use a {source,target} combo when asking about an edge's attributes in a MultiGraph since we cannot infer the one you want information about."));var i=""+e,o=""+arguments[1];if(!(r=c(this,i,o,n)))throw new M("Graph.".concat(t,': could not find an edge for the given path ("').concat(i,'" - "').concat(o,'").'))}else if(e=""+e,!(r=this._edges.get(e)))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" edge in the graph.'));if("mixed"!==n&&r.undirected!==("undirected"===n))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" ').concat(n," edge in the graph."));return r.attributes}}},{name:function(e){return"has".concat(e,"Attribute")},attacher:function(e,t,n){e.prototype[t]=function(e,r){var i;if("mixed"!==this.type&&"mixed"!==n&&n!==this.type)throw new P("Graph.".concat(t,": cannot find this type of edges in your ").concat(this.type," graph."));if(arguments.length>2){if(this.multi)throw new P("Graph.".concat(t,": cannot use a {source,target} combo when asking about an edge's attributes in a MultiGraph since we cannot infer the one you want information about."));var o=""+e,a=""+r;if(r=arguments[2],!(i=c(this,o,a,n)))throw new M("Graph.".concat(t,': could not find an edge for the given path ("').concat(o,'" - "').concat(a,'").'))}else if(e=""+e,!(i=this._edges.get(e)))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" edge in the graph.'));if("mixed"!==n&&i.undirected!==("undirected"===n))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" ').concat(n," edge in the graph."));return i.attributes.hasOwnProperty(r)}}},{name:function(e){return"set".concat(e,"Attribute")},attacher:function(e,t,n){e.prototype[t]=function(e,r,i){var o;if("mixed"!==this.type&&"mixed"!==n&&n!==this.type)throw new P("Graph.".concat(t,": cannot find this type of edges in your ").concat(this.type," graph."));if(arguments.length>3){if(this.multi)throw new P("Graph.".concat(t,": cannot use a {source,target} combo when asking about an edge's attributes in a MultiGraph since we cannot infer the one you want information about."));var a=""+e,u=""+r;if(r=arguments[2],i=arguments[3],!(o=c(this,a,u,n)))throw new M("Graph.".concat(t,': could not find an edge for the given path ("').concat(a,'" - "').concat(u,'").'))}else if(e=""+e,!(o=this._edges.get(e)))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" edge in the graph.'));if("mixed"!==n&&o.undirected!==("undirected"===n))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" ').concat(n," edge in the graph."));return o.attributes[r]=i,this.emit("edgeAttributesUpdated",{key:o.key,type:"set",attributes:o.attributes,name:r}),this}}},{name:function(e){return"update".concat(e,"Attribute")},attacher:function(e,t,n){e.prototype[t]=function(e,r,i){var o;if("mixed"!==this.type&&"mixed"!==n&&n!==this.type)throw new P("Graph.".concat(t,": cannot find this type of edges in your ").concat(this.type," graph."));if(arguments.length>3){if(this.multi)throw new P("Graph.".concat(t,": cannot use a {source,target} combo when asking about an edge's attributes in a MultiGraph since we cannot infer the one you want information about."));var a=""+e,u=""+r;if(r=arguments[2],i=arguments[3],!(o=c(this,a,u,n)))throw new M("Graph.".concat(t,': could not find an edge for the given path ("').concat(a,'" - "').concat(u,'").'))}else if(e=""+e,!(o=this._edges.get(e)))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" edge in the graph.'));if("function"!=typeof i)throw new z("Graph.".concat(t,": updater should be a function."));if("mixed"!==n&&o.undirected!==("undirected"===n))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" ').concat(n," edge in the graph."));return o.attributes[r]=i(o.attributes[r]),this.emit("edgeAttributesUpdated",{key:o.key,type:"set",attributes:o.attributes,name:r}),this}}},{name:function(e){return"remove".concat(e,"Attribute")},attacher:function(e,t,n){e.prototype[t]=function(e,r){var i;if("mixed"!==this.type&&"mixed"!==n&&n!==this.type)throw new P("Graph.".concat(t,": cannot find this type of edges in your ").concat(this.type," graph."));if(arguments.length>2){if(this.multi)throw new P("Graph.".concat(t,": cannot use a {source,target} combo when asking about an edge's attributes in a MultiGraph since we cannot infer the one you want information about."));var o=""+e,a=""+r;if(r=arguments[2],!(i=c(this,o,a,n)))throw new M("Graph.".concat(t,': could not find an edge for the given path ("').concat(o,'" - "').concat(a,'").'))}else if(e=""+e,!(i=this._edges.get(e)))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" edge in the graph.'));if("mixed"!==n&&i.undirected!==("undirected"===n))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" ').concat(n," edge in the graph."));return delete i.attributes[r],this.emit("edgeAttributesUpdated",{key:i.key,type:"remove",attributes:i.attributes,name:r}),this}}},{name:function(e){return"replace".concat(e,"Attributes")},attacher:function(e,t,n){e.prototype[t]=function(e,r){var i;if("mixed"!==this.type&&"mixed"!==n&&n!==this.type)throw new P("Graph.".concat(t,": cannot find this type of edges in your ").concat(this.type," graph."));if(arguments.length>2){if(this.multi)throw new P("Graph.".concat(t,": cannot use a {source,target} combo when asking about an edge's attributes in a MultiGraph since we cannot infer the one you want information about."));var o=""+e,a=""+r;if(r=arguments[2],!(i=c(this,o,a,n)))throw new M("Graph.".concat(t,': could not find an edge for the given path ("').concat(o,'" - "').concat(a,'").'))}else if(e=""+e,!(i=this._edges.get(e)))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" edge in the graph.'));if(!h(r))throw new z("Graph.".concat(t,": provided attributes are not a plain object."));if("mixed"!==n&&i.undirected!==("undirected"===n))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" ').concat(n," edge in the graph."));return i.attributes=r,this.emit("edgeAttributesUpdated",{key:i.key,type:"replace",attributes:i.attributes}),this}}},{name:function(e){return"merge".concat(e,"Attributes")},attacher:function(e,t,n){e.prototype[t]=function(e,r){var i;if("mixed"!==this.type&&"mixed"!==n&&n!==this.type)throw new P("Graph.".concat(t,": cannot find this type of edges in your ").concat(this.type," graph."));if(arguments.length>2){if(this.multi)throw new P("Graph.".concat(t,": cannot use a {source,target} combo when asking about an edge's attributes in a MultiGraph since we cannot infer the one you want information about."));var o=""+e,a=""+r;if(r=arguments[2],!(i=c(this,o,a,n)))throw new M("Graph.".concat(t,': could not find an edge for the given path ("').concat(o,'" - "').concat(a,'").'))}else if(e=""+e,!(i=this._edges.get(e)))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" edge in the graph.'));if(!h(r))throw new z("Graph.".concat(t,": provided attributes are not a plain object."));if("mixed"!==n&&i.undirected!==("undirected"===n))throw new M("Graph.".concat(t,': could not find the "').concat(e,'" ').concat(n," edge in the graph."));return d(i.attributes,r),this.emit("edgeAttributesUpdated",{key:i.key,type:"merge",attributes:i.attributes,data:r}),this}}}];var J=O,q=function(){var e,t=arguments,n=-1;return new J((function r(){if(!e){if(++n>=t.length)return{done:!0};e=t[n]}var i=e.next();return i.done?(e=null,r()):i}))},H=[{name:"edges",type:"mixed"},{name:"inEdges",type:"directed",direction:"in"},{name:"outEdges",type:"directed",direction:"out"},{name:"inboundEdges",type:"mixed",direction:"in"},{name:"outboundEdges",type:"mixed",direction:"out"},{name:"directedEdges",type:"directed"},{name:"undirectedEdges",type:"undirected"}];function Q(e,t){for(var n in t)e.push(t[n].key)}function V(e,t){for(var n in t)t[n].forEach((function(t){return e.push(t.key)}))}function X(e,t,n){for(var r in e)if(r!==n){var i=e[r];t(i.key,i.attributes,i.source.key,i.target.key,i.source.attributes,i.target.attributes,i.undirected,i.generatedKey)}}function Z(e,t,n){for(var r in e)r!==n&&e[r].forEach((function(e){return t(e.key,e.attributes,e.source.key,e.target.key,e.source.attributes,e.target.attributes,e.undirected,e.generatedKey)}))}function $(e,t,n){for(var r in e)if(r!==n){var i=e[r];if(t(i.key,i.attributes,i.source.key,i.target.key,i.source.attributes,i.target.attributes,i.undirected,i.generatedKey))return!0}return!1}function ee(e,t,n){var r,i,o,a,u;for(var d in e)if(d!==n)for(r=e[d].values();!0!==(i=r.next()).done;)if(a=(o=i.value).source,u=o.target,t(o.key,o.attributes,a.key,u.key,a.attributes,u.attributes,o.undirected,o.generatedKey))return!0;return!1}function te(e,t){var n=Object.keys(e),r=n.length,i=null,o=0;return new O((function a(){var u;if(i){var d=i.next();if(d.done)return i=null,o++,a();u=d.value}else{if(o>=r)return{done:!0};var c=n[o];if(c===t)return o++,a();if((u=e[c])instanceof Set)return i=u.values(),a();o++}return{done:!1,value:[u.key,u.attributes,u.source.key,u.target.key,u.source.attributes,u.target.attributes]}}))}function ne(e,t,n){var r=t[n];r&&e.push(r.key)}function re(e,t,n){var r=t[n];r&&r.forEach((function(t){return e.push(t.key)}))}function ie(e,t,n){var r=e[t];if(r){var i=r.source,o=r.target;n(r.key,r.attributes,i.key,o.key,i.attributes,o.attributes,r.undirected,r.generatedKey)}}function oe(e,t,n){var r=e[t];r&&r.forEach((function(e){return n(e.key,e.attributes,e.source.key,e.target.key,e.source.attributes,e.target.attributes,e.undirected,e.generatedKey)}))}function ae(e,t,n){var r=e[t];if(!r)return!1;var i=r.source,o=r.target;return n(r.key,r.attributes,i.key,o.key,i.attributes,o.attributes,r.undirected,r.generatedKey)}function ue(e,t,n){var r=e[t];if(!r)return!1;for(var i,o,a=r.values();!0!==(i=a.next()).done;)if(n((o=i.value).key,o.attributes,o.source.key,o.target.key,o.source.attributes,o.target.attributes,o.undirected,o.generatedKey))return!0;return!1}function de(e,t){var n=e[t];if(n instanceof Set){var r=n.values();return new O((function(){var e=r.next();if(e.done)return e;var t=e.value;return{done:!1,value:[t.key,t.attributes,t.source.key,t.target.key,t.source.attributes,t.target.attributes]}}))}return O.of([n.key,n.attributes,n.source.key,n.target.key,n.source.attributes,n.target.attributes])}function ce(e,t){if(0===e.size)return[];if("mixed"===t||t===e.type)return"function"==typeof Array.from?Array.from(e._edges.keys()):C(e._edges.keys(),e._edges.size);for(var n,r,i="undirected"===t?e.undirectedSize:e.directedSize,o=new Array(i),a="undirected"===t,u=e._edges.values(),d=0;!0!==(n=u.next()).done;)(r=n.value).undirected===a&&(o[d++]=r.key);return o}function se(e,t,n){if(0!==e.size)for(var r,i,o="mixed"!==t&&t!==e.type,a="undirected"===t,u=e._edges.values();!0!==(r=u.next()).done;)if(i=r.value,!o||i.undirected===a){var d=i,c=d.key,s=d.attributes,h=d.source,f=d.target;n(c,s,h.key,f.key,h.attributes,f.attributes,i.undirected,i.generatedKey)}}function he(e,t,n){if(0===e.size)return!1;for(var r,i,o="mixed"!==t&&t!==e.type,a="undirected"===t,u=e._edges.values();!0!==(r=u.next()).done;)if(i=r.value,!o||i.undirected===a){var d=i,c=d.key,s=d.attributes,h=d.source,f=d.target;if(n(c,s,h.key,f.key,h.attributes,f.attributes,i.undirected,i.generatedKey))return!0}return!1}function fe(e,t){if(0===e.size)return O.empty();var n="mixed"!==t&&t!==e.type,r="undirected"===t,i=e._edges.values();return new O((function(){for(var e,t;;){if((e=i.next()).done)return e;if(t=e.value,!n||t.undirected===r)break}return{value:[t.key,t.attributes,t.source.key,t.target.key,t.source.attributes,t.target.attributes],done:!1}}))}function pe(e,t,n,r){var i=[],o=e?V:Q;return"undirected"!==t&&("out"!==n&&o(i,r.in),"in"!==n&&o(i,r.out),!n&&r.directedSelfLoops>0&&i.splice(i.lastIndexOf(r.key),1)),"directed"!==t&&o(i,r.undirected),i}function ge(e,t,n,r,i){var o=e?Z:X;"undirected"!==t&&("out"!==n&&o(r.in,i),"in"!==n&&o(r.out,i,n?null:r.key)),"directed"!==t&&o(r.undirected,i)}function le(e,t,n,r,i){var o=e?ee:$;if("undirected"!==t){if("out"!==n&&o(r.in,i))return!0;if("in"!==n&&o(r.out,i,n?null:r.key))return!0}return!("directed"===t||!o(r.undirected,i))}function ye(e,t,n){var r=O.empty();return"undirected"!==e&&("out"!==t&&void 0!==n.in&&(r=q(r,te(n.in))),"in"!==t&&void 0!==n.out&&(r=q(r,te(n.out,t?null:n.key)))),"directed"!==e&&void 0!==n.undirected&&(r=q(r,te(n.undirected))),r}function ve(e,t,n,r,i){var o=t?re:ne,a=[];return"undirected"!==e&&(void 0!==r.in&&"out"!==n&&o(a,r.in,i),void 0!==r.out&&"in"!==n&&o(a,r.out,i),!n&&r.directedSelfLoops>0&&a.splice(a.lastIndexOf(r.key),1)),"directed"!==e&&void 0!==r.undirected&&o(a,r.undirected,i),a}function be(e,t,n,r,i,o){var a=t?oe:ie;"undirected"!==e&&(void 0!==r.in&&"out"!==n&&a(r.in,i,o),r.key!==i&&void 0!==r.out&&"in"!==n&&a(r.out,i,o)),"directed"!==e&&void 0!==r.undirected&&a(r.undirected,i,o)}function we(e,t,n,r,i,o){var a=t?ue:ae;if("undirected"!==e){if(void 0!==r.in&&"out"!==n&&a(r.in,i,o))return!0;if(r.key!==i&&void 0!==r.out&&"in"!==n&&a(r.out,i,o,n?null:r.key))return!0}return!("directed"===e||void 0===r.undirected||!a(r.undirected,i,o))}function me(e,t,n,r){var i=O.empty();return"undirected"!==e&&(void 0!==n.in&&"out"!==t&&r in n.in&&(i=q(i,de(n.in,r))),void 0!==n.out&&"in"!==t&&r in n.out&&(i=q(i,de(n.out,r)))),"directed"!==e&&void 0!==n.undirected&&r in n.undirected&&(i=q(i,de(n.undirected,r))),i}var _e=[{name:"neighbors",type:"mixed"},{name:"inNeighbors",type:"directed",direction:"in"},{name:"outNeighbors",type:"directed",direction:"out"},{name:"inboundNeighbors",type:"mixed",direction:"in"},{name:"outboundNeighbors",type:"mixed",direction:"out"},{name:"directedNeighbors",type:"directed"},{name:"undirectedNeighbors",type:"undirected"}];function ke(e,t){if(void 0!==t)for(var n in t)e.add(n)}function Ge(e,t,n){if("mixed"!==e){if("undirected"===e)return Object.keys(n.undirected);if("string"==typeof t)return Object.keys(n[t])}var r=new Set;return"undirected"!==e&&("out"!==t&&ke(r,n.in),"in"!==t&&ke(r,n.out)),"directed"!==e&&ke(r,n.undirected),C(r.values(),r.size)}function xe(e,t,n){for(var r in t){var i=t[r];i instanceof Set&&(i=i.values().next().value);var o=i.source,a=i.target,u=o===e?a:o;n(u.key,u.attributes)}}function Ee(e,t,n,r){for(var i in n){var o=n[i];o instanceof Set&&(o=o.values().next().value);var a=o.source,u=o.target,d=a===t?u:a;e.has(d.key)||(e.add(d.key),r(d.key,d.attributes))}}function Se(e,t,n){for(var r in t){var i=t[r];i instanceof Set&&(i=i.values().next().value);var o=i.source,a=i.target,u=o===e?a:o;if(n(u.key,u.attributes))return!0}return!1}function Ae(e,t,n,r){for(var i in n){var o=n[i];o instanceof Set&&(o=o.values().next().value);var a=o.source,u=o.target,d=a===t?u:a;if(!e.has(d.key))if(e.add(d.key),r(d.key,d.attributes))return!0}return!1}function Le(e,t){var n=Object.keys(t),r=n.length,i=0;return new O((function(){if(i>=r)return{done:!0};var o=t[n[i++]];o instanceof Set&&(o=o.values().next().value);var a=o.source,u=o.target,d=a===e?u:a;return{done:!1,value:[d.key,d.attributes]}}))}function Ne(e,t,n){var r=Object.keys(n),i=r.length,o=0;return new O((function a(){if(o>=i)return{done:!0};var u=n[r[o++]];u instanceof Set&&(u=u.values().next().value);var d=u.source,c=u.target,s=d===t?c:d;return e.has(s.key)?a():(e.add(s.key),{done:!1,value:[s.key,s.attributes]})}))}function De(e,t,n,r,i){var o=e._nodes.get(r);if("undirected"!==t){if("out"!==n&&void 0!==o.in)for(var a in o.in)if(a===i)return!0;if("in"!==n&&void 0!==o.out)for(var u in o.out)if(u===i)return!0}if("directed"!==t&&void 0!==o.undirected)for(var d in o.undirected)if(d===i)return!0;return!1}function je(e,t){var n=t.name,r=t.type,i=t.direction,o="forEach"+n[0].toUpperCase()+n.slice(1,-1);e.prototype[o]=function(e,t){if("mixed"===r||"mixed"===this.type||r===this.type){e=""+e;var n=this._nodes.get(e);if(void 0===n)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" node in the graph.'));!function(e,t,n,r){if("mixed"!==e){if("undirected"===e)return xe(n,n.undirected,r);if("string"==typeof t)return xe(n,n[t],r)}var i=new Set;"undirected"!==e&&("out"!==t&&Ee(i,n,n.in,r),"in"!==t&&Ee(i,n,n.out,r)),"directed"!==e&&Ee(i,n,n.undirected,r)}("mixed"===r?this.type:r,i,n,t)}}}function Ue(e,t){var n=t.name,r=t.type,i=t.direction,o="forEach"+n[0].toUpperCase()+n.slice(1,-1)+"Until";e.prototype[o]=function(e,t){if("mixed"===r||"mixed"===this.type||r===this.type){e=""+e;var n=this._nodes.get(e);if(void 0===n)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" node in the graph.'));return function(e,t,n,r){if("mixed"!==e){if("undirected"===e)return Se(n,n.undirected,r);if("string"==typeof t)return Se(n,n[t],r)}var i=new Set;if("undirected"!==e){if("out"!==t&&Ae(i,n,n.in,r))return!0;if("in"!==t&&Ae(i,n,n.out,r))return!0}return!("directed"===e||!Ae(i,n,n.undirected,r))}("mixed"===r?this.type:r,i,n,t)}}}function Oe(e,t){var n=t.name,r=t.type,i=t.direction,o=n.slice(0,-1)+"Entries";e.prototype[o]=function(e){if("mixed"!==r&&"mixed"!==this.type&&r!==this.type)return O.empty();e=""+e;var t=this._nodes.get(e);if(void 0===t)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" node in the graph.'));return function(e,t,n){if("mixed"!==e){if("undirected"===e)return Le(n,n.undirected);if("string"==typeof t)return Le(n,n[t])}var r=O.empty(),i=new Set;return"undirected"!==e&&("out"!==t&&(r=q(r,Ne(i,n,n.in))),"in"!==t&&(r=q(r,Ne(i,n,n.out)))),"directed"!==e&&(r=q(r,Ne(i,n,n.undirected))),r}("mixed"===r?this.type:r,i,t)}}function Ce(e,t,n){for(var r,i,o,a,u,d,c,s=t._nodes.values(),h=t.type;!0!==(r=s.next()).done;){if(i=r.value,"undirected"!==h)for(o in a=i.out)if(d=(u=a[o]).target,c=n(i.key,d.key,i.attributes,d.attributes,u.key,u.attributes,u.undirected,u.generatedKey),e&&c)return!0;if("directed"!==h)for(o in a=i.undirected)if((d=(u=a[o]).target).key!==o&&(d=u.source),c=n(i.key,d.key,i.attributes,d.attributes,u.key,u.attributes,u.undirected,u.generatedKey),e&&c)return!0}return!1}function Ke(e,t,n){for(var r,i,o,a,u,d,c,s,h,f=t._nodes.values(),p=t.type;!0!==(r=f.next()).done;){if(i=r.value,"undirected"!==p)for(o in d=i.out)for(a=d[o].values();!0!==(u=a.next()).done;)if(s=(c=u.value).target,h=n(i.key,s.key,i.attributes,s.attributes,c.key,c.attributes,c.undirected,c.generatedKey),e&&h)return!0;if("directed"!==p)for(o in d=i.undirected)for(a=d[o].values();!0!==(u=a.next()).done;)if((s=(c=u.value).target).key!==o&&(s=c.source),h=n(i.key,s.key,i.attributes,s.attributes,c.key,c.attributes,c.undirected,c.generatedKey),e&&h)return!0}return!1}function ze(e,t){var n={key:e};return f(t.attributes)||(n.attributes=d({},t.attributes)),n}function Me(e,t){var n={source:t.source.key,target:t.target.key};return t.generatedKey||(n.key=e),f(t.attributes)||(n.attributes=d({},t.attributes)),t.undirected&&(n.undirected=!0),n}function Pe(e){return h(e)?"key"in e?!("attributes"in e)||h(e.attributes)&&null!==e.attributes?null:"invalid-attributes":"no-key":"not-object"}function Te(e){return h(e)?"source"in e?"target"in e?!("attributes"in e)||h(e.attributes)&&null!==e.attributes?"undirected"in e&&"boolean"!=typeof e.undirected?"invalid-undirected":null:"invalid-attributes":"no-target":"no-source":"not-object"}var Re=new Set(["directed","undirected","mixed"]),Fe=new Set(["domain","_events","_eventsCount","_maxListeners"]),Ie={allowSelfLoops:!0,edgeKeyGenerator:null,multi:!1,type:"mixed"};function We(e,t,n){var r=new e.NodeDataClass(t,n);return e._nodes.set(t,r),e.emit("nodeAdded",{key:t,attributes:n}),r}function Ye(e,t,n,r,i,o,a,u){if(!r&&"undirected"===e.type)throw new P("Graph.".concat(t,": you cannot add a directed edge to an undirected graph. Use the #.addEdge or #.addUndirectedEdge instead."));if(r&&"directed"===e.type)throw new P("Graph.".concat(t,": you cannot add an undirected edge to a directed graph. Use the #.addEdge or #.addDirectedEdge instead."));if(u&&!h(u))throw new z("Graph.".concat(t,': invalid attributes. Expecting an object but got "').concat(u,'"'));if(o=""+o,a=""+a,u=u||{},!e.allowSelfLoops&&o===a)throw new P("Graph.".concat(t,': source & target are the same ("').concat(o,"\"), thus creating a loop explicitly forbidden by this graph 'allowSelfLoops' option set to false."));var d=e._nodes.get(o),c=e._nodes.get(a);if(!d)throw new M("Graph.".concat(t,': source node "').concat(o,'" not found.'));if(!c)throw new M("Graph.".concat(t,': target node "').concat(a,'" not found.'));var s={key:null,undirected:r,source:o,target:a,attributes:u};if(n&&(i=e._edgeKeyGenerator(s)),i=""+i,e._edges.has(i))throw new P("Graph.".concat(t,': the "').concat(i,'" edge already exists in the graph.'));if(!e.multi&&(r?void 0!==d.undirected[a]:void 0!==d.out[a]))throw new P("Graph.".concat(t,': an edge linking "').concat(o,'" to "').concat(a,"\" already exists. If you really want to add multiple edges linking those nodes, you should create a multi graph by using the 'multi' option."));var f=new I(r,i,n,d,c,u);return e._edges.set(i,f),o===a?r?(d.undirectedSelfLoops++,e._undirectedSelfLoopCount++):(d.directedSelfLoops++,e._directedSelfLoopCount++):r?(d.undirectedDegree++,c.undirectedDegree++):(d.outDegree++,c.inDegree++),W(e,r,f,o,a,d,c),r?e._undirectedSize++:e._directedSize++,s.key=i,e.emit("edgeAdded",s),i}function Be(e,t,n,r,i,o,a,u,c){if(!r&&"undirected"===e.type)throw new P("Graph.".concat(t,": you cannot add a directed edge to an undirected graph. Use the #.addEdge or #.addUndirectedEdge instead."));if(r&&"directed"===e.type)throw new P("Graph.".concat(t,": you cannot add an undirected edge to a directed graph. Use the #.addEdge or #.addDirectedEdge instead."));if(u)if(c){if("function"!=typeof u)throw new z("Graph.".concat(t,': invalid updater function. Expecting a function but got "').concat(u,'"'))}else if(!h(u))throw new z("Graph.".concat(t,': invalid attributes. Expecting an object but got "').concat(u,'"'));var s;if(o=""+o,a=""+a,c&&(s=u,u=void 0),!e.allowSelfLoops&&o===a)throw new P("Graph.".concat(t,': source & target are the same ("').concat(o,"\"), thus creating a loop explicitly forbidden by this graph 'allowSelfLoops' option set to false."));var f,p,g=e._nodes.get(o),l=e._nodes.get(a);if(!n&&(f=e._edges.get(i))){if(f.source.key!==o||f.target.key!==a||r&&(f.source.key!==a||f.target.key!==o))throw new P("Graph.".concat(t,': inconsistency detected when attempting to merge the "').concat(i,'" edge with "').concat(o,'" source & "').concat(a,'" target vs. ("').concat(f.source.key,'", "').concat(f.target.key,'").'));p=f}if(p||e.multi||!g||(p=r?g.undirected[a]:g.out[a]),p){if(c?!s:!u)return p.key;if(c){var y=p.attributes;p.attributes=s(y),e.emit("edgeAttributesUpdated",{type:"replace",key:p.key,attributes:p.attributes})}else d(p.attributes,u),e.emit("edgeAttributesUpdated",{type:"merge",key:p.key,attributes:p.attributes,data:u});return p.key}u=u||{},c&&s&&(u=s(u));var v={key:null,undirected:r,source:o,target:a,attributes:u};if(n&&(i=e._edgeKeyGenerator(v)),i=""+i,e._edges.has(i))throw new P("Graph.".concat(t,': the "').concat(i,'" edge already exists in the graph.'));return g||(g=We(e,o,{}),o===a&&(l=g)),l||(l=We(e,a,{})),f=new I(r,i,n,g,l,u),e._edges.set(i,f),o===a?r?(g.undirectedSelfLoops++,e._undirectedSelfLoopCount++):(g.directedSelfLoops++,e._directedSelfLoopCount++):r?(g.undirectedDegree++,l.undirectedDegree++):(g.outDegree++,l.inDegree++),W(e,r,f,o,a,g,l),r?e._undirectedSize++:e._directedSize++,v.key=i,e.emit("edgeAdded",v),i}var Je=function(e){function n(t){var n;if(n=e.call(this)||this,(t=d({},Ie,t)).edgeKeyGenerator&&"function"!=typeof t.edgeKeyGenerator)throw new z("Graph.constructor: invalid 'edgeKeyGenerator' option. Expecting a function but got \"".concat(t.edgeKeyGenerator,'".'));if("boolean"!=typeof t.multi)throw new z("Graph.constructor: invalid 'multi' option. Expecting a boolean but got \"".concat(t.multi,'".'));if(!Re.has(t.type))throw new z('Graph.constructor: invalid \'type\' option. Should be one of "mixed", "directed" or "undirected" but got "'.concat(t.type,'".'));if("boolean"!=typeof t.allowSelfLoops)throw new z("Graph.constructor: invalid 'allowSelfLoops' option. Expecting a boolean but got \"".concat(t.allowSelfLoops,'".'));var r,i="mixed"===t.type?T:"directed"===t.type?R:F;return p(u(n),"NodeDataClass",i),p(u(n),"_attributes",{}),p(u(n),"_nodes",new Map),p(u(n),"_edges",new Map),p(u(n),"_directedSize",0),p(u(n),"_undirectedSize",0),p(u(n),"_directedSelfLoopCount",0),p(u(n),"_undirectedSelfLoopCount",0),p(u(n),"_edgeKeyGenerator",t.edgeKeyGenerator||(r=0,function(){return r++})),p(u(n),"_options",t),Fe.forEach((function(e){return p(u(n),e,n[e])})),g(u(n),"order",(function(){return n._nodes.size})),g(u(n),"size",(function(){return n._edges.size})),g(u(n),"directedSize",(function(){return n._directedSize})),g(u(n),"undirectedSize",(function(){return n._undirectedSize})),g(u(n),"selfLoopCount",(function(){return n._directedSelfLoopCount+n._undirectedSelfLoopCount})),g(u(n),"directedSelfLoopCount",(function(){return n._directedSelfLoopCount})),g(u(n),"undirectedSelfLoopCount",(function(){return n._undirectedSelfLoopCount})),g(u(n),"multi",n._options.multi),g(u(n),"type",n._options.type),g(u(n),"allowSelfLoops",n._options.allowSelfLoops),g(u(n),"implementation",(function(){return"graphology"})),n}t(n,e);var r=n.prototype;return r.hasNode=function(e){return this._nodes.has(""+e)},r.hasDirectedEdge=function(e,t){if("undirected"===this.type)return!1;if(1===arguments.length){var n=""+e,r=this._edges.get(n);return!!r&&!r.undirected}if(2===arguments.length){e=""+e,t=""+t;var i=this._nodes.get(e);if(!i)return!1;var o=i.out[t];return!!o&&(!this.multi||!!o.size)}throw new z("Graph.hasDirectedEdge: invalid arity (".concat(arguments.length,", instead of 1 or 2). You can either ask for an edge id or for the existence of an edge between a source & a target."))},r.hasUndirectedEdge=function(e,t){if("directed"===this.type)return!1;if(1===arguments.length){var n=""+e,r=this._edges.get(n);return!!r&&r.undirected}if(2===arguments.length){e=""+e,t=""+t;var i=this._nodes.get(e);if(!i)return!1;var o=i.undirected[t];return!!o&&(!this.multi||!!o.size)}throw new z("Graph.hasDirectedEdge: invalid arity (".concat(arguments.length,", instead of 1 or 2). You can either ask for an edge id or for the existence of an edge between a source & a target."))},r.hasEdge=function(e,t){if(1===arguments.length){var n=""+e;return this._edges.has(n)}if(2===arguments.length){e=""+e,t=""+t;var r=this._nodes.get(e);if(!r)return!1;var i=void 0!==r.out&&r.out[t];return i||(i=void 0!==r.undirected&&r.undirected[t]),!!i&&(!this.multi||!!i.size)}throw new z("Graph.hasEdge: invalid arity (".concat(arguments.length,", instead of 1 or 2). You can either ask for an edge id or for the existence of an edge between a source & a target."))},r.directedEdge=function(e,t){if("undirected"!==this.type){if(e=""+e,t=""+t,this.multi)throw new P("Graph.directedEdge: this method is irrelevant with multigraphs since there might be multiple edges between source & target. See #.directedEdges instead.");var n=this._nodes.get(e);if(!n)throw new M('Graph.directedEdge: could not find the "'.concat(e,'" source node in the graph.'));if(!this._nodes.has(t))throw new M('Graph.directedEdge: could not find the "'.concat(t,'" target node in the graph.'));var r=n.out&&n.out[t]||void 0;return r?r.key:void 0}},r.undirectedEdge=function(e,t){if("directed"!==this.type){if(e=""+e,t=""+t,this.multi)throw new P("Graph.undirectedEdge: this method is irrelevant with multigraphs since there might be multiple edges between source & target. See #.undirectedEdges instead.");var n=this._nodes.get(e);if(!n)throw new M('Graph.undirectedEdge: could not find the "'.concat(e,'" source node in the graph.'));if(!this._nodes.has(t))throw new M('Graph.undirectedEdge: could not find the "'.concat(t,'" target node in the graph.'));var r=n.undirected&&n.undirected[t]||void 0;return r?r.key:void 0}},r.edge=function(e,t){if(this.multi)throw new P("Graph.edge: this method is irrelevant with multigraphs since there might be multiple edges between source & target. See #.edges instead.");e=""+e,t=""+t;var n=this._nodes.get(e);if(!n)throw new M('Graph.edge: could not find the "'.concat(e,'" source node in the graph.'));if(!this._nodes.has(t))throw new M('Graph.edge: could not find the "'.concat(t,'" target node in the graph.'));var r=n.out&&n.out[t]||n.undirected&&n.undirected[t]||void 0;if(r)return r.key},r.inDegree=function(e){var t=!(arguments.length>1&&void 0!==arguments[1])||arguments[1];if("boolean"!=typeof t)throw new z('Graph.inDegree: Expecting a boolean but got "'.concat(t,'" for the second parameter (allowing self-loops to be counted).'));e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.inDegree: could not find the "'.concat(e,'" node in the graph.'));if("undirected"===this.type)return 0;var r=t?n.directedSelfLoops:0;return n.inDegree+r},r.outDegree=function(e){var t=!(arguments.length>1&&void 0!==arguments[1])||arguments[1];if("boolean"!=typeof t)throw new z('Graph.outDegree: Expecting a boolean but got "'.concat(t,'" for the second parameter (allowing self-loops to be counted).'));e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.outDegree: could not find the "'.concat(e,'" node in the graph.'));if("undirected"===this.type)return 0;var r=t?n.directedSelfLoops:0;return n.outDegree+r},r.directedDegree=function(e){var t=!(arguments.length>1&&void 0!==arguments[1])||arguments[1];if("boolean"!=typeof t)throw new z('Graph.directedDegree: Expecting a boolean but got "'.concat(t,'" for the second parameter (allowing self-loops to be counted).'));e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.directedDegree: could not find the "'.concat(e,'" node in the graph.'));if("undirected"===this.type)return 0;var r=t?n.directedSelfLoops:0,i=n.inDegree+r,o=n.outDegree+r;return i+o},r.undirectedDegree=function(e){var t=!(arguments.length>1&&void 0!==arguments[1])||arguments[1];if("boolean"!=typeof t)throw new z('Graph.undirectedDegree: Expecting a boolean but got "'.concat(t,'" for the second parameter (allowing self-loops to be counted).'));e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.undirectedDegree: could not find the "'.concat(e,'" node in the graph.'));if("directed"===this.type)return 0;var r=t?n.undirectedSelfLoops:0;return n.undirectedDegree+2*r},r.degree=function(e){var t=!(arguments.length>1&&void 0!==arguments[1])||arguments[1];if("boolean"!=typeof t)throw new z('Graph.degree: Expecting a boolean but got "'.concat(t,'" for the second parameter (allowing self-loops to be counted).'));e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.degree: could not find the "'.concat(e,'" node in the graph.'));var r=0,i=0;return"directed"!==this.type&&(t&&(i=n.undirectedSelfLoops),r+=n.undirectedDegree+2*i),"undirected"!==this.type&&(t&&(i=n.directedSelfLoops),r+=n.inDegree+n.outDegree+2*i),r},r.source=function(e){e=""+e;var t=this._edges.get(e);if(!t)throw new M('Graph.source: could not find the "'.concat(e,'" edge in the graph.'));return t.source.key},r.target=function(e){e=""+e;var t=this._edges.get(e);if(!t)throw new M('Graph.target: could not find the "'.concat(e,'" edge in the graph.'));return t.target.key},r.extremities=function(e){e=""+e;var t=this._edges.get(e);if(!t)throw new M('Graph.extremities: could not find the "'.concat(e,'" edge in the graph.'));return[t.source.key,t.target.key]},r.opposite=function(e,t){e=""+e,t=""+t;var n=this._edges.get(t);if(!n)throw new M('Graph.opposite: could not find the "'.concat(t,'" edge in the graph.'));var r=n.source.key,i=n.target.key;if(e!==r&&e!==i)throw new M('Graph.opposite: the "'.concat(e,'" node is not attached to the "').concat(t,'" edge (').concat(r,", ").concat(i,")."));return e===r?i:r},r.hasExtremity=function(e,t){e=""+e,t=""+t;var n=this._edges.get(e);if(!n)throw new M('Graph.hasExtremity: could not find the "'.concat(e,'" edge in the graph.'));return n.source.key===t||n.target.key===t},r.isUndirected=function(e){e=""+e;var t=this._edges.get(e);if(!t)throw new M('Graph.isUndirected: could not find the "'.concat(e,'" edge in the graph.'));return t.undirected},r.isDirected=function(e){e=""+e;var t=this._edges.get(e);if(!t)throw new M('Graph.isDirected: could not find the "'.concat(e,'" edge in the graph.'));return!t.undirected},r.isSelfLoop=function(e){e=""+e;var t=this._edges.get(e);if(!t)throw new M('Graph.isSelfLoop: could not find the "'.concat(e,'" edge in the graph.'));return t.source===t.target},r.hasGeneratedKey=function(e){e=""+e;var t=this._edges.get(e);if(!t)throw new M('Graph.hasGeneratedKey: could not find the "'.concat(e,'" edge in the graph.'));return t.generatedKey},r.addNode=function(e,t){var n=function(e,t,n){if(n&&!h(n))throw new z('Graph.addNode: invalid attributes. Expecting an object but got "'.concat(n,'"'));if(t=""+t,n=n||{},e._nodes.has(t))throw new P('Graph.addNode: the "'.concat(t,'" node already exist in the graph.'));var r=new e.NodeDataClass(t,n);return e._nodes.set(t,r),e.emit("nodeAdded",{key:t,attributes:n}),r}(this,e,t);return n.key},r.mergeNode=function(e,t){if(t&&!h(t))throw new z('Graph.mergeNode: invalid attributes. Expecting an object but got "'.concat(t,'"'));e=""+e,t=t||{};var n=this._nodes.get(e);return n?(t&&(d(n.attributes,t),this.emit("nodeAttributesUpdated",{type:"merge",key:e,attributes:n.attributes,data:t})),e):(n=new this.NodeDataClass(e,t),this._nodes.set(e,n),this.emit("nodeAdded",{key:e,attributes:t}),e)},r.updateNode=function(e,t){if(t&&"function"!=typeof t)throw new z('Graph.updateNode: invalid updater function. Expecting a function but got "'.concat(t,'"'));e=""+e;var n=this._nodes.get(e);if(n){if(t){var r=n.attributes;n.attributes=t(r),this.emit("nodeAttributesUpdated",{type:"replace",key:e,attributes:n.attributes})}return e}var i=t?t({}):{};return n=new this.NodeDataClass(e,i),this._nodes.set(e,n),this.emit("nodeAdded",{key:e,attributes:i}),e},r.dropNode=function(e){var t=this;e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.dropNode: could not find the "'.concat(e,'" node in the graph.'));this.forEachEdge(e,(function(e){t.dropEdge(e)})),this._nodes.delete(e),this.emit("nodeDropped",{key:e,attributes:n.attributes})},r.dropEdge=function(e){var t;if(arguments.length>1){var n=""+arguments[0],r=""+arguments[1];if(!(t=c(this,n,r,this.type)))throw new M('Graph.dropEdge: could not find the "'.concat(n,'" -> "').concat(r,'" edge in the graph.'))}else if(e=""+e,!(t=this._edges.get(e)))throw new M('Graph.dropEdge: could not find the "'.concat(e,'" edge in the graph.'));this._edges.delete(t.key);var i=t,o=i.source,a=i.target,u=i.attributes,d=t.undirected;return o===a?d?(o.undirectedSelfLoops--,this._undirectedSelfLoopCount--):(o.directedSelfLoops--,this._directedSelfLoopCount--):d?(o.undirectedDegree--,a.undirectedDegree--):(o.outDegree--,a.inDegree--),Y(this,d,t),d?this._undirectedSize--:this._directedSize--,this.emit("edgeDropped",{key:e,attributes:u,source:o.key,target:a.key,undirected:d}),this},r.clear=function(){this._edges.clear(),this._nodes.clear(),this.emit("cleared")},r.clearEdges=function(){this._edges.clear(),this.clearIndex(),this.emit("edgesCleared")},r.getAttribute=function(e){return this._attributes[e]},r.getAttributes=function(){return this._attributes},r.hasAttribute=function(e){return this._attributes.hasOwnProperty(e)},r.setAttribute=function(e,t){return this._attributes[e]=t,this.emit("attributesUpdated",{type:"set",attributes:this._attributes,name:e}),this},r.updateAttribute=function(e,t){if("function"!=typeof t)throw new z("Graph.updateAttribute: updater should be a function.");var n=this._attributes[e];return this._attributes[e]=t(n),this.emit("attributesUpdated",{type:"set",attributes:this._attributes,name:e}),this},r.removeAttribute=function(e){return delete this._attributes[e],this.emit("attributesUpdated",{type:"remove",attributes:this._attributes,name:e}),this},r.replaceAttributes=function(e){if(!h(e))throw new z("Graph.replaceAttributes: provided attributes are not a plain object.");return this._attributes=e,this.emit("attributesUpdated",{type:"replace",attributes:this._attributes}),this},r.mergeAttributes=function(e){if(!h(e))throw new z("Graph.mergeAttributes: provided attributes are not a plain object.");return d(this._attributes,e),this.emit("attributesUpdated",{type:"merge",attributes:this._attributes,data:e}),this},r.getNodeAttribute=function(e,t){e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.getNodeAttribute: could not find the "'.concat(e,'" node in the graph.'));return n.attributes[t]},r.getNodeAttributes=function(e){e=""+e;var t=this._nodes.get(e);if(!t)throw new M('Graph.getNodeAttributes: could not find the "'.concat(e,'" node in the graph.'));return t.attributes},r.hasNodeAttribute=function(e,t){e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.hasNodeAttribute: could not find the "'.concat(e,'" node in the graph.'));return n.attributes.hasOwnProperty(t)},r.setNodeAttribute=function(e,t,n){e=""+e;var r=this._nodes.get(e);if(!r)throw new M('Graph.setNodeAttribute: could not find the "'.concat(e,'" node in the graph.'));if(arguments.length<3)throw new z("Graph.setNodeAttribute: not enough arguments. Either you forgot to pass the attribute's name or value, or you meant to use #.replaceNodeAttributes / #.mergeNodeAttributes instead.");return r.attributes[t]=n,this.emit("nodeAttributesUpdated",{key:e,type:"set",attributes:r.attributes,name:t}),this},r.updateNodeAttribute=function(e,t,n){e=""+e;var r=this._nodes.get(e);if(!r)throw new M('Graph.updateNodeAttribute: could not find the "'.concat(e,'" node in the graph.'));if(arguments.length<3)throw new z("Graph.updateNodeAttribute: not enough arguments. Either you forgot to pass the attribute's name or updater, or you meant to use #.replaceNodeAttributes / #.mergeNodeAttributes instead.");if("function"!=typeof n)throw new z("Graph.updateAttribute: updater should be a function.");var i=r.attributes,o=n(i[t]);return i[t]=o,this.emit("nodeAttributesUpdated",{key:e,type:"set",attributes:r.attributes,name:t}),this},r.removeNodeAttribute=function(e,t){e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.removeNodeAttribute: could not find the "'.concat(e,'" node in the graph.'));return delete n.attributes[t],this.emit("nodeAttributesUpdated",{key:e,type:"remove",attributes:n.attributes,name:t}),this},r.replaceNodeAttributes=function(e,t){e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.replaceNodeAttributes: could not find the "'.concat(e,'" node in the graph.'));if(!h(t))throw new z("Graph.replaceNodeAttributes: provided attributes are not a plain object.");return n.attributes=t,this.emit("nodeAttributesUpdated",{key:e,type:"replace",attributes:n.attributes}),this},r.mergeNodeAttributes=function(e,t){e=""+e;var n=this._nodes.get(e);if(!n)throw new M('Graph.mergeNodeAttributes: could not find the "'.concat(e,'" node in the graph.'));if(!h(t))throw new z("Graph.mergeNodeAttributes: provided attributes are not a plain object.");return d(n.attributes,t),this.emit("nodeAttributesUpdated",{key:e,type:"merge",attributes:n.attributes,data:t}),this},r.updateEachNodeAttributes=function(e,t){if("function"!=typeof e)throw new z("Graph.updateEachNodeAttributes: expecting an updater function.");if(t&&!l(t))throw new z("Graph.updateEachNodeAttributes: invalid hints. Expecting an object having the following shape: {attributes?: [string]}");for(var n,r,i=this._nodes.values();!0!==(n=i.next()).done;)(r=n.value).attributes=e(r.key,r.attributes);this.emit("eachNodeAttributesUpdated",{hints:t||null})},r.updateEachEdgeAttributes=function(e,t){if("function"!=typeof e)throw new z("Graph.updateEachEdgeAttributes: expecting an updater function.");if(t&&!l(t))throw new z("Graph.updateEachEdgeAttributes: invalid hints. Expecting an object having the following shape: {attributes?: [string]}");for(var n,r,i=this._edges.values();!0!==(n=i.next()).done;)(r=n.value).attributes=e(r.key,r.attributes);this.emit("eachEdgeAttributesUpdated",{hints:t||null})},r.forEach=function(e){if("function"!=typeof e)throw new z("Graph.forEach: expecting a callback.");this.multi?Ke(!1,this,e):Ce(!1,this,e)},r.forEachUntil=function(e){if("function"!=typeof e)throw new z("Graph.forEach: expecting a callback.");return this.multi?Ke(!0,this,e):Ce(!0,this,e)},r.adjacency=function(){return this.multi?(o=(e=this)._nodes.values(),a=e.type,u="outer",d=null,new O((function e(){var c;if("outer"===u)return!0===(c=o.next()).done?c:(t=c.value,u="directed",e());if("directed"===u)return"undirected"===a?(u="undirected",e()):(r=t.out,n=Object.keys(t.out),i=0,u="inner-directed",e());if("undirected"===u){if("directed"===a)return u="outer",e();r=t.undirected,n=Object.keys(t.undirected),i=0,u="inner-undirected"}if(!d&&i>=n.length)return u="inner-undirected"===u?"outer":"undirected",e();if(!d){var s=n[i++];return d=r[s].values(),e()}if((c=d.next()).done)return d=null,e();var h=c.value,f=h.target;return"inner-undirected"===u&&f.key===t.key&&(f=h.source),{done:!1,value:[t.key,f.key,t.attributes,f.attributes,h.key,h.attributes]}}))):function(e){var t,n,r,i,o=e._nodes.values(),a=e.type,u="outer";return new O((function e(){var d;if("outer"===u)return!0===(d=o.next()).done?d:(t=d.value,u="directed",e());if("directed"===u)return"undirected"===a?(u="undirected",e()):(r=t.out,n=Object.keys(t.out),i=0,u="inner-directed",e());if("undirected"===u){if("directed"===a)return u="outer",e();r=t.undirected,n=Object.keys(t.undirected),i=0,u="inner-undirected"}if(i>=n.length)return u="inner-undirected"===u?"outer":"undirected",e();var c=n[i++],s=r[c],h=s.target;return"inner-undirected"===u&&h.key===t.key&&(h=s.source),{done:!1,value:[t.key,h.key,t.attributes,h.attributes,s.key,s.attributes]}}))}(this);var e,t,n,r,i,o,a,u,d},r.nodes=function(){return"function"==typeof Array.from?Array.from(this._nodes.keys()):C(this._nodes.keys(),this._nodes.size)},r.forEachNode=function(e){if("function"!=typeof e)throw new z("Graph.forEachNode: expecting a callback.");this._nodes.forEach((function(t,n){e(n,t.attributes)}))},r.forEachNodeUntil=function(e){if("function"!=typeof e)throw new z("Graph.forEachNode: expecting a callback.");for(var t,n,r=this._nodes.values();!0!==(t=r.next()).done;)if(e((n=t.value).key,n.attributes))return!0;return!1},r.nodeEntries=function(){var e=this._nodes.values();return new O((function(){var t=e.next();if(t.done)return t;var n=t.value;return{value:[n.key,n.attributes],done:!1}}))},r.exportNode=function(e){e=""+e;var t=this._nodes.get(e);if(!t)throw new M('Graph.exportNode: could not find the "'.concat(e,'" node in the graph.'));return ze(e,t)},r.exportEdge=function(e){e=""+e;var t=this._edges.get(e);if(!t)throw new M('Graph.exportEdge: could not find the "'.concat(e,'" edge in the graph.'));return Me(e,t)},r.export=function(){var e=new Array(this._nodes.size),t=0;this._nodes.forEach((function(n,r){e[t++]=ze(r,n)}));var n=new Array(this._edges.size);return t=0,this._edges.forEach((function(e,r){n[t++]=Me(r,e)})),{attributes:this.getAttributes(),nodes:e,edges:n,options:{type:this.type,multi:this.multi,allowSelfLoops:this.allowSelfLoops}}},r.importNode=function(e){var t=arguments.length>1&&void 0!==arguments[1]&&arguments[1],n=Pe(e);if(n){if("not-object"===n)throw new z('Graph.importNode: invalid serialized node. A serialized node should be a plain object with at least a "key" property.');if("no-key"===n)throw new z("Graph.importNode: no key provided.");if("invalid-attributes"===n)throw new z("Graph.importNode: invalid attributes. Attributes should be a plain object, null or omitted.")}var r=e.key,i=e.attributes,o=void 0===i?{}:i;return t?this.mergeNode(r,o):this.addNode(r,o),this},r.importEdge=function(e){var t=arguments.length>1&&void 0!==arguments[1]&&arguments[1],n=Te(e);if(n){if("not-object"===n)throw new z('Graph.importEdge: invalid serialized edge. A serialized edge should be a plain object with at least a "source" & "target" property.');if("no-source"===n)throw new z("Graph.importEdge: missing souce.");if("no-target"===n)throw new z("Graph.importEdge: missing target.");if("invalid-attributes"===n)throw new z("Graph.importEdge: invalid attributes. Attributes should be a plain object, null or omitted.");if("invalid-undirected"===n)throw new z("Graph.importEdge: invalid undirected. Undirected should be boolean or omitted.")}var r=e.source,i=e.target,o=e.attributes,a=void 0===o?{}:o,u=e.undirected,d=void 0!==u&&u;return"key"in e?(t?d?this.mergeUndirectedEdgeWithKey:this.mergeDirectedEdgeWithKey:d?this.addUndirectedEdgeWithKey:this.addDirectedEdgeWithKey).call(this,e.key,r,i,a):(t?d?this.mergeUndirectedEdge:this.mergeDirectedEdge:d?this.addUndirectedEdge:this.addDirectedEdge).call(this,r,i,a),this},r.import=function(e){var t,n,r,i=arguments.length>1&&void 0!==arguments[1]&&arguments[1];if(s(e))return this.import(e.export(),i),this;if(!h(e))throw new z("Graph.import: invalid argument. Expecting a serialized graph or, alternatively, a Graph instance.");if(e.attributes){if(!h(e.attributes))throw new z("Graph.import: invalid attributes. Expecting a plain object.");i?this.mergeAttributes(e.attributes):this.replaceAttributes(e.attributes)}if(e.nodes){if(r=e.nodes,!Array.isArray(r))throw new z("Graph.import: invalid nodes. Expecting an array.");for(t=0,n=r.length;tn)){var a=new Set;a.add(t.undirected[o]),t.undirected[o]=a,e._nodes.get(o).undirected[n]=a}}))),this;var e},r.clearIndex=function(){return this._nodes.forEach((function(e){void 0!==e.in&&(e.in={},e.out={}),void 0!==e.undirected&&(e.undirected={})})),this},r.toJSON=function(){return this.export()},r.toString=function(){return"[object Graph]"},r.inspect=function(){var e=this,t={};this._nodes.forEach((function(e,n){t[n]=e.attributes}));var n={},r={};this._edges.forEach((function(t,i){var o=t.undirected?"--":"->",a="",u="(".concat(t.source.key,")").concat(o,"(").concat(t.target.key,")");t.generatedKey?e.multi&&(void 0===r[u]?r[u]=0:r[u]++,a+="".concat(r[u],". ")):a+="[".concat(i,"]: "),n[a+=u]=t.attributes}));var i={};for(var o in this)this.hasOwnProperty(o)&&!Fe.has(o)&&"function"!=typeof this[o]&&(i[o]=this[o]);return i.attributes=this._attributes,i.nodes=t,i.edges=n,p(i,"constructor",this.constructor),i},n}(v.exports.EventEmitter);"undefined"!=typeof Symbol&&(Je.prototype[Symbol.for("nodejs.util.inspect.custom")]=Je.prototype.inspect),[{name:function(e){return"".concat(e,"Edge")},generateKey:!0},{name:function(e){return"".concat(e,"DirectedEdge")},generateKey:!0,type:"directed"},{name:function(e){return"".concat(e,"UndirectedEdge")},generateKey:!0,type:"undirected"},{name:function(e){return"".concat(e,"EdgeWithKey")}},{name:function(e){return"".concat(e,"DirectedEdgeWithKey")},type:"directed"},{name:function(e){return"".concat(e,"UndirectedEdgeWithKey")},type:"undirected"}].forEach((function(e){["add","merge","update"].forEach((function(t){var n=e.name(t),r="add"===t?Ye:Be;e.generateKey?Je.prototype[n]=function(i,o,a){return r(this,n,!0,"undirected"===(e.type||this.type),null,i,o,a,"update"===t)}:Je.prototype[n]=function(i,o,a,u){return r(this,n,!1,"undirected"===(e.type||this.type),i,o,a,u,"update"===t)}}))})),"undefined"!=typeof Symbol&&(Je.prototype[Symbol.iterator]=Je.prototype.adjacency),function(e){B.forEach((function(t){var n=t.name,r=t.attacher;r(e,n("Edge"),"mixed"),r(e,n("DirectedEdge"),"directed"),r(e,n("UndirectedEdge"),"undirected")}))}(Je),function(e){H.forEach((function(t){!function(e,t){var n=t.name,r=t.type,i=t.direction;e.prototype[n]=function(e,t){if("mixed"!==r&&"mixed"!==this.type&&r!==this.type)return[];if(!arguments.length)return ce(this,r);if(1===arguments.length){e=""+e;var o=this._nodes.get(e);if(void 0===o)throw new M("Graph.".concat(n,': could not find the "').concat(e,'" node in the graph.'));return pe(this.multi,"mixed"===r?this.type:r,i,o)}if(2===arguments.length){e=""+e,t=""+t;var a=this._nodes.get(e);if(!a)throw new M("Graph.".concat(n,': could not find the "').concat(e,'" source node in the graph.'));if(!this._nodes.has(t))throw new M("Graph.".concat(n,': could not find the "').concat(t,'" target node in the graph.'));return ve(r,this.multi,i,a,t)}throw new z("Graph.".concat(n,": too many arguments (expecting 0, 1 or 2 and got ").concat(arguments.length,")."))}}(e,t),function(e,t){var n=t.name,r=t.type,i=t.direction,o="forEach"+n[0].toUpperCase()+n.slice(1,-1);e.prototype[o]=function(e,t,n){if("mixed"===r||"mixed"===this.type||r===this.type){if(1===arguments.length)return se(this,r,n=e);if(2===arguments.length){e=""+e,n=t;var a=this._nodes.get(e);if(void 0===a)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" node in the graph.'));return ge(this.multi,"mixed"===r?this.type:r,i,a,n)}if(3===arguments.length){e=""+e,t=""+t;var u=this._nodes.get(e);if(!u)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" source node in the graph.'));if(!this._nodes.has(t))throw new M("Graph.".concat(o,': could not find the "').concat(t,'" target node in the graph.'));return be(r,this.multi,i,u,t,n)}throw new z("Graph.".concat(o,": too many arguments (expecting 1, 2 or 3 and got ").concat(arguments.length,")."))}}}(e,t),function(e,t){var n=t.name,r=t.type,i=t.direction,o="forEach"+n[0].toUpperCase()+n.slice(1,-1)+"Until";e.prototype[o]=function(e,t,n){if("mixed"!==r&&"mixed"!==this.type&&r!==this.type)return!1;if(1===arguments.length)return he(this,r,n=e);if(2===arguments.length){e=""+e,n=t;var a=this._nodes.get(e);if(void 0===a)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" node in the graph.'));return le(this.multi,"mixed"===r?this.type:r,i,a,n)}if(3===arguments.length){e=""+e,t=""+t;var u=this._nodes.get(e);if(!u)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" source node in the graph.'));if(!this._nodes.has(t))throw new M("Graph.".concat(o,': could not find the "').concat(t,'" target node in the graph.'));return we(r,this.multi,i,u,t,n)}throw new z("Graph.".concat(o,": too many arguments (expecting 1, 2 or 3 and got ").concat(arguments.length,")."))}}(e,t),function(e,t){var n=t.name,r=t.type,i=t.direction,o=n.slice(0,-1)+"Entries";e.prototype[o]=function(e,t){if("mixed"!==r&&"mixed"!==this.type&&r!==this.type)return O.empty();if(!arguments.length)return fe(this,r);if(1===arguments.length){e=""+e;var n=this._nodes.get(e);if(!n)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" node in the graph.'));return ye(r,i,n)}if(2===arguments.length){e=""+e,t=""+t;var a=this._nodes.get(e);if(!a)throw new M("Graph.".concat(o,': could not find the "').concat(e,'" source node in the graph.'));if(!this._nodes.has(t))throw new M("Graph.".concat(o,': could not find the "').concat(t,'" target node in the graph.'));return me(r,i,a,t)}throw new z("Graph.".concat(o,": too many arguments (expecting 0, 1 or 2 and got ").concat(arguments.length,")."))}}(e,t)}))}(Je),function(e){_e.forEach((function(t){!function(e,t){var n=t.name,r=t.type,i=t.direction;e.prototype[n]=function(e){if("mixed"!==r&&"mixed"!==this.type&&r!==this.type)return[];if(2===arguments.length){var t=""+arguments[0],o=""+arguments[1];if(!this._nodes.has(t))throw new M("Graph.".concat(n,': could not find the "').concat(t,'" node in the graph.'));if(!this._nodes.has(o))throw new M("Graph.".concat(n,': could not find the "').concat(o,'" node in the graph.'));return De(this,r,i,t,o)}if(1===arguments.length){e=""+e;var a=this._nodes.get(e);if(void 0===a)throw new M("Graph.".concat(n,': could not find the "').concat(e,'" node in the graph.'));return Ge("mixed"===r?this.type:r,i,a)}throw new z("Graph.".concat(n,": invalid number of arguments (expecting 1 or 2 and got ").concat(arguments.length,")."))}}(e,t),je(e,t),Ue(e,t),Oe(e,t)}))}(Je);var qe=function(e){function n(t){var n=d({type:"directed"},t);if("multi"in n&&!1!==n.multi)throw new z("DirectedGraph.from: inconsistent indication that the graph should be multi in given options!");if("directed"!==n.type)throw new z('DirectedGraph.from: inconsistent "'+n.type+'" type in given options!');return e.call(this,n)||this}return t(n,e),n}(Je),He=function(e){function n(t){var n=d({type:"undirected"},t);if("multi"in n&&!1!==n.multi)throw new z("UndirectedGraph.from: inconsistent indication that the graph should be multi in given options!");if("undirected"!==n.type)throw new z('UndirectedGraph.from: inconsistent "'+n.type+'" type in given options!');return e.call(this,n)||this}return t(n,e),n}(Je),Qe=function(e){function n(t){var n=d({multi:!0},t);if("multi"in n&&!0!==n.multi)throw new z("MultiGraph.from: inconsistent indication that the graph should be simple in given options!");return e.call(this,n)||this}return t(n,e),n}(Je),Ve=function(e){function n(t){var n=d({type:"directed",multi:!0},t);if("multi"in n&&!0!==n.multi)throw new z("MultiDirectedGraph.from: inconsistent indication that the graph should be simple in given options!");if("directed"!==n.type)throw new z('MultiDirectedGraph.from: inconsistent "'+n.type+'" type in given options!');return e.call(this,n)||this}return t(n,e),n}(Je),Xe=function(e){function n(t){var n=d({type:"undirected",multi:!0},t);if("multi"in n&&!0!==n.multi)throw new z("MultiUndirectedGraph.from: inconsistent indication that the graph should be simple in given options!");if("undirected"!==n.type)throw new z('MultiUndirectedGraph.from: inconsistent "'+n.type+'" type in given options!');return e.call(this,n)||this}return t(n,e),n}(Je);function Ze(e){e.from=function(t,n){var r=d({},t.options,n),i=new e(r);return i.import(t),i}}return Ze(Je),Ze(qe),Ze(He),Ze(Qe),Ze(Ve),Ze(Xe),Je.Graph=Je,Je.DirectedGraph=qe,Je.UndirectedGraph=He,Je.MultiGraph=Qe,Je.MultiDirectedGraph=Ve,Je.MultiUndirectedGraph=Xe,Je.InvalidArgumentsGraphError=z,Je.NotFoundGraphError=M,Je.UsageGraphError=P,Je})); 2 | //# sourceMappingURL=graphology.umd.min.js.map 3 | -------------------------------------------------------------------------------- /html-template/vendor/sigma-LICENSE: -------------------------------------------------------------------------------- 1 | Copyright (C) 2013-2024, Alexis Jacomy, Guillaume Plique, Benoît Simard https://www.sigmajs.org 2 | 3 | Permission is hereby granted, free of charge, to any person obtaining a copy of this software and associated documentation files (the "Software"), 4 | to deal in the Software without restriction, including without limitation the rights to use, copy, modify, merge, publish, distribute, sublicense, 5 | and/or sell copies of the Software, and to permit persons to whom the Software is furnished to do so, subject to the following conditions: 6 | 7 | The above copyright notice and this permission notice shall be included in all copies or substantial portions of the Software. 8 | 9 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 10 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 11 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS 12 | IN THE SOFTWARE. 13 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover/lean4-cli", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "leanprover", 8 | "rev": "b62fd39acc32da6fb8bb160c82d1bbc3cb3c186e", 9 | "name": "Cli", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "main", 12 | "inherited": false, 13 | "configFile": "lakefile.toml"}], 14 | "name": "importGraph", 15 | "lakeDir": ".lake"} 16 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "importGraph" 2 | description = "Tools to analyse and visualise the import structure of Lean packages and their files." 3 | license = "Apache-2.0" 4 | licenseFiles = ["LICENSE"] 5 | 6 | defaultTargets = ["ImportGraph", "graph"] 7 | testRunner = "ImportGraphTest" 8 | 9 | [[require]] 10 | name = "Cli" 11 | scope = "leanprover" 12 | rev = "main" 13 | 14 | [[lean_lib]] 15 | name = "ImportGraph" 16 | 17 | # `lake exe graph` constructs import graphs in `.dot` or graphical formats. 18 | [[lean_exe]] 19 | name = "graph" 20 | root = "Main" 21 | # Enables the use of the Lean interpreter by the executable (e.g., `runFrontend`) 22 | # at the expense of increased binary size on Linux. 23 | # Remove this line if you do not need such functionality. 24 | supportInterpreter = true 25 | 26 | # `lake exe unused_transitive_imports` prints unused transitive imports from amongst a given list of modules. 27 | [[lean_exe]] 28 | name = "unused_transitive_imports" 29 | root = "ImportGraph.UnusedTransitiveImports" 30 | supportInterpreter = true 31 | 32 | [[lean_lib]] 33 | name = "ImportGraphTest" 34 | 35 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.24.0-rc1 2 | --------------------------------------------------------------------------------