From 193176aae29bb45d1942faeeb024783851e61b51 Mon Sep 17 00:00:00 2001 From: Jon Eugster <eugster.jon@gmail.com> Date: Sun, 1 Sep 2024 18:03:23 +0200 Subject: [PATCH] feat: incorporate Eric's html visualisation --- ImportGraph/Cli.lean | 55 +- ImportGraph/Gexf.lean | 77 +++ Main.lean | 2 +- README.md | 22 +- html-template/.gitignore | 1 + html-template/LICENSE | 21 + html-template/README.md | 25 + html-template/index.html | 509 ++++++++++++++++++ .../vendor/graphology-library.min.js | 2 + html-template/vendor/graphology.min.js | 2 + html-template/vendor/sigma.min.js | 1 + 11 files changed, 706 insertions(+), 11 deletions(-) create mode 100644 ImportGraph/Gexf.lean create mode 100644 html-template/.gitignore create mode 100644 html-template/LICENSE create mode 100644 html-template/README.md create mode 100644 html-template/index.html create mode 100644 html-template/vendor/graphology-library.min.js create mode 100644 html-template/vendor/graphology.min.js create mode 100644 html-template/vendor/sigma.min.js diff --git a/ImportGraph/Cli.lean b/ImportGraph/Cli.lean index 6fcf231..5933634 100644 --- a/ImportGraph/Cli.lean +++ b/ImportGraph/Cli.lean @@ -8,6 +8,7 @@ import Batteries.Lean.IO.Process import ImportGraph.CurrentModule import ImportGraph.Imports import ImportGraph.Lean.Name +import ImportGraph.Gexf open Cli @@ -60,6 +61,18 @@ open Lean Core System open IO.FS IO.Process Name in /-- Implementation of the import graph command line program. -/ def importGraphCLI (args : Cli.Parsed) : IO UInt32 := do + -- file extensions that should be created + let extensions : Array String := match args.variableArgsAs! String with + | #[] => #["dot"] + | outputs => outputs.foldl (fun acc (o : String) => + match FilePath.extension o with + | none => if acc.contains "dot" then acc else acc.push "dot" + | some "gexf" => if acc.contains "gexf" then acc else acc.push "gexf" + | some "html" => if acc.contains "gexf" then acc else acc.push "gexf" + -- currently all other formats are handled by passing the `.dot` file to + -- graphviz + | some _ => if acc.contains "dot" then acc else acc.push "dot" ) #[] + let to ← match args.flag? "to" with | some to => pure <| to.as! ModuleName | none => getCurrentModule @@ -67,14 +80,15 @@ def importGraphCLI (args : Cli.Parsed) : IO UInt32 := do | some fr => some <| fr.as! ModuleName | none => none searchPathRef.set compile_time_search_path% - let dotFile ← try unsafe withImportModules #[{module := to}] {} (trustLevel := 1024) fun env => do + + let outFiles ← try unsafe withImportModules #[{module := to}] {} (trustLevel := 1024) fun env => do let p := ImportGraph.getModule to + let ctx := { options := {}, fileName := "<input>", fileMap := default } + let state := { env } let mut graph := env.importGraph let unused ← match args.flag? "to" with | some _ => - let ctx := { options := {}, fileName := "<input>", fileMap := default } - let state := { env } let used ← Prod.fst <$> (CoreM.toIO (env.transitivelyRequiredModules to) ctx state) pure <| graph.fold (fun acc n _ => if used.contains n then acc else acc.insert n) NameSet.empty | none => pure NameSet.empty @@ -112,7 +126,16 @@ def importGraphCLI (args : Cli.Parsed) : IO UInt32 := do let markedModule : Option Name := if args.hasFlag "mark-module" then p else none - return asDotGraph graph (unused := unused) (markedModule := markedModule) (directDeps := directDeps) + -- Create all output files that are requested + let mut outFiles : HashMap String String := {} + if extensions.contains "dot" then + let dotFile := asDotGraph graph (unused := unused) (markedModule := markedModule) (directDeps := directDeps) + outFiles := outFiles.insert "dot" dotFile + if extensions.contains "gexf" then + let (out, _) ← CoreM.toIO (Graph.toGexf graph p) ctx state + outFiles := outFiles.insert "gexf" out + return outFiles + catch err => -- TODO: try to build `to` first, so this doesn't happen throw <| IO.userError <| s!"{err}\nIf the error above says `unknown package`, " ++ @@ -120,14 +143,32 @@ def importGraphCLI (args : Cli.Parsed) : IO UInt32 := do throw err match args.variableArgsAs! String with - | #[] => writeFile "import_graph.dot" dotFile + | #[] => writeFile "import_graph.dot" (outFiles.find! "dot") | outputs => for o in outputs do let fp : FilePath := o match fp.extension with | none - | "dot" => writeFile fp dotFile + | "dot" => writeFile fp (outFiles.find! "dot") + | "gexf" => IO.FS.writeFile fp (outFiles.find! "gexf") + | "html" => + -- use `html-template/index.html` and insert any dependencies to make it + -- a stand-alone HTML file. + let gexFile := (outFiles.find! "gexf") + let mut html ← IO.FS.readFile "html-template/index.html" + for dep in #[ + "vendor/sigma.min.js", + "vendor/graphology.min.js", + "vendor/graphology-library.min.js" ] do + let depContent ← IO.FS.readFile ("html-template" / dep) + html := html.replace s!"<script src=\"{dep}\"></script>" + s!"<script>{depContent}</script>" + html := html.replace "fetch(\"imports.gexf\").then((res) => res.text()).then(render_gexf)" + s!"render_gexf(\"{gexFile.replace "\n" ""|>.replace "\"" "\\\""}\")" + |>.replace "<h1>Import Graph</h1>" + s!"<h1>Import Graph for {to}</h1>" + IO.FS.writeFile fp html | some ext => try - _ ← runCmdWithInput "dot" #["-T" ++ ext, "-o", o] dotFile + _ ← runCmdWithInput "dot" #["-T" ++ ext, "-o", o] (outFiles.find! "dot") catch ex => IO.eprintln s!"Error occurred while writing out {fp}." IO.eprintln s!"Make sure you have `graphviz` installed and the file is writable." diff --git a/ImportGraph/Gexf.lean b/ImportGraph/Gexf.lean new file mode 100644 index 0000000..0270493 --- /dev/null +++ b/ImportGraph/Gexf.lean @@ -0,0 +1,77 @@ +/- +Copyright (c) 2024 Jon Eugster. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jon Eugster +-/ + +import Lean +import Batteries.Lean.NameMap +import Batteries.Tactic.OpenPrivate + +open Lean + +namespace ImportGraph + +open Elab Meta in +/-- Filter Lean internal declarations -/ +def isBlackListed {m} [Monad m] [MonadEnv m] (declName : Name) : m Bool := do + if declName == ``sorryAx then return true + if declName matches .str _ "inj" then return true + if declName matches .str _ "noConfusionType" then return true + let env ← getEnv + pure <| declName.isInternalDetail + || isAuxRecursor env declName + || isNoConfusion env declName + <||> isRec declName <||> isMatcher declName + +/-- Get all declarations in the specified file. -/ +def getNumberOfDeclsInFile (module : Name) : CoreM (NameSet) := do + let env ← getEnv + match env.moduleIdxForModule? module with + | none => return {} + | some modIdx => + let decls := env.const2ModIdx + let declsIn ← decls.foldM (fun acc n idx => do + if idx == modIdx && (! (← isBlackListed n)) then return acc.insert n else return acc) ({} : NameSet) + return declsIn + +/-- Gexf template for a node in th graph. -/ +def Gexf.nodeTemplate (n module : Name) (size : Nat) := s!"<node id=\"{n}\" label=\"{n}\"><attvalues><attvalue for=\"0\" value=\"{size}\" /><attvalue for=\"1\" value=\"{module.isPrefixOf n}\" /></attvalues></node>\n " + +/-- Gexf template for an edge in the graph -/ +def Gexf.edgeTemplate (source target : Name) := s!"<edge source=\"{source}\" target=\"{target}\" id=\"{source}--{target}\" />\n " + +open Gexf in +/-- Creates a `.gexf` file of the graph, see https://gexf.net/ + +Metadata can be stored in forms of attributes, currently we record the following: +* `decl_count` (Nat): number of declarations in the file +* `in_module` (Bool): whether the file belongs to the main module + (used to strip the first part of the name when displaying). +-/ +def Graph.toGexf (graph : NameMap (Array Name)) (module : Name) : CoreM String := do + let sizes : NameMap Nat ← graph.foldM (fun acc n _ => do + pure <| acc.insert n (← getNumberOfDeclsInFile n).size ) {} + -- graph.fold (fun acc _ i => i.foldl (fun acc₂ j => acc₂.insert j ((acc₂.findD j 0) + 1)) acc) {} + + let nodes : String := graph.fold (fun acc n _ => acc ++ nodeTemplate n module (sizes.findD n 0)) "" + let edges : String := graph.fold (fun acc n i => acc ++ (i.foldl (fun b j => b ++ edgeTemplate j n) "")) "" + return s!"<?xml version='1.0' encoding='utf-8'?> + <gexf xmlns=\"http://www.gexf.net/1.2draft\" xmlns:xsi=\"http://www.w3.org/2001/XMLSchema-instance\" xsi:schemaLocation=\"http://www.gexf.net/1.2draft http://www.gexf.net/1.2draft/gexf.xsd\" version=\"1.2\"> + <meta> + <creator>Lean ImportGraph</creator> + </meta> + <graph defaultedgetype=\"directed\" mode=\"static\" name=\"\"> + <attributes mode=\"static\" class=\"node\"> + <attribute id=\"0\" title=\"decl_count\" type=\"long\" /> + <attribute id=\"1\" title=\"in_module\" type=\"boolean\" /> + </attributes> + <nodes> + {nodes.trim} + </nodes> + <edges> + {edges.trim} + </edges> + </graph> + </gexf> + " diff --git a/Main.lean b/Main.lean index c84f5d6..bd15e26 100644 --- a/Main.lean +++ b/Main.lean @@ -29,7 +29,7 @@ def graph : Cmd := `[Cli| ...outputs : String; "Filename(s) for the output. \ If none are specified, generates `import_graph.dot`. \ Automatically chooses the format based on the file extension. \ - Currently `.dot` is supported, \ + Currently supported formats are `.dot`, `.gexf`, `.html`, \ and if you have `graphviz` installed then any supported output format is allowed." ] diff --git a/README.md b/README.md index d890376..7d43271 100644 --- a/README.md +++ b/README.md @@ -26,12 +26,20 @@ where `MyModule` follows the same module naming you would use to `import` it in ### Json -To create a Json file, you can use `.xdot_json` as output type: +To create a Json file, you can use `.xdot_json` or `.json` as output type: ``` lake exe graph my_graph.xdot_json ``` +### HTML + +``` +lake exe graph my_graph.html +``` + +creates a stand-alone HTML file visualising the import structure. + ## Installation The installation works exactly like for any [Lake package](https://reservoir.lean-lang.org/). @@ -65,10 +73,18 @@ There are a few commands implemented, which help you analysing the imports of a (Must be run at the end of the file. Tactics and macros may result in incorrect output.) * `#find_home decl`: suggests files higher up the import hierarchy to which `decl` could be moved. +## Contribution + +Please open PRs/Issues if you have troubles or would like to contribute new features! + ## Credits -This code has been extracted from [mathlib](https://github.com/leanprover-community/mathlib4) and has mainly been written by Kim Morrison and a few other mathlib contributors. +The main tool has been extracted from [mathlib](https://github.com/leanprover-community/mathlib4), +originally written by Kim Morrison and other mathlib contributors. + +The HTML visualisation has been incorporated from +[a project by Eric Wieser](https://github.com/eric-wieser/mathlib-import-graph). ### Maintainers -For issues, questions, or feature requests, please reach out to [Jon Eugster](https://leanprover.zulipchat.com/#narrow/dm/385895-Jon-Eugster). +Primarily maintained by [Jon Eugster](https://leanprover.zulipchat.com/#narrow/dm/385895-Jon-Eugster), Kim Morrison, and the wider leanprover community. diff --git a/html-template/.gitignore b/html-template/.gitignore new file mode 100644 index 0000000..0219bf0 --- /dev/null +++ b/html-template/.gitignore @@ -0,0 +1 @@ +imports.gexf \ No newline at end of file diff --git a/html-template/LICENSE b/html-template/LICENSE new file mode 100644 index 0000000..9c0b6d7 --- /dev/null +++ b/html-template/LICENSE @@ -0,0 +1,21 @@ +MIT License + +Copyright (c) 2021 Eric Wieser + +Permission is hereby granted, free of charge, to any person obtaining a copy +of this software and associated documentation files (the "Software"), to deal +in the Software without restriction, including without limitation the rights +to use, copy, modify, merge, publish, distribute, sublicense, 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: + +The above copyright notice and this permission notice shall be included in all +copies or substantial portions of the Software. + +THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR +IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, +FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE +AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER +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 IN THE +SOFTWARE. diff --git a/html-template/README.md b/html-template/README.md new file mode 100644 index 0000000..3a90d1f --- /dev/null +++ b/html-template/README.md @@ -0,0 +1,25 @@ +# Visualised import graph + +## Instructions + +To test this, place a file `imports.gexf` inside this directory. You can create such a file with + +``` +lake exe graph html-template/imports.gexf +``` + +Then open `index.html` in any browser and you should see the graph. + +## Development + +Currently `lake exe graph output.html` will use the files here to create a stand-alone +HTML file. It does so by search-replacing the JS-scripts, the `fetch('imports.gexf')` +statement, and the `<h1>` header. + +Therefore any modifications to these lines need to be reflected in `ImportGraph/Cli.lean`! + +# Credits + +This tool has been adapted from it's [Lean 3 version](https://github.com/eric-wieser/mathlib-import-graph) written by Eric Wieser. + +Adaptation by Jon Eugster. diff --git a/html-template/index.html b/html-template/index.html new file mode 100644 index 0000000..e1e30fb --- /dev/null +++ b/html-template/index.html @@ -0,0 +1,509 @@ + +<!doctype html> +<html class="no-js" lang="en"> +<head> + <meta charset="utf-8"> + + <title>Mathlib import graph</title> + + <!-- NOTE: currently the Lean script searches for these imports to replace them (exact search)! + If you modify it, modify the Lean script (ImportGraph/Cli.lean) accordingly! --> + <script src="vendor/sigma.min.js"></script> + <script src="vendor/graphology.min.js"></script> + <script src="vendor/graphology-library.min.js"></script> + + <link rel="stylesheet" href="style.css" /> + <style> + html, body { + height: 100%; width: 100%; + margin: 0; + padding: 0; + position: relative; + } + .sigma { + position: absolute; + width: 100%; + height: 100%; + top: 0; + left: 0; + overflow: hidden; + } + p, label, .summary { + text-align: center; + } + .summary { + position: absolute; + top: 0; + left: 0; + right: 0; + } + .key-wrapper { + position: absolute; + bottom: 0; + left: 0; + right: 0; + } + #key { + list-style: none; + columns: 10em auto; + } + .summary, .key-wrapper { + text-shadow: + 0 0 5px white, + -1px -1px 0 white, + 0 -1px 0 white, + 1px -1px 0 white, + 1px 0 0 white, + 1px 1px 0 white, + 0 1px 0 white, + -1px 1px 0 white, + -1px 0 0 white; + } + + ul > li { + white-space: nowrap; + } + ul > li:hover { + font-weight: bold; + } + + #key > li > span { + width: 1ch; + height: 1ch; + display: inline-block; + margin-right: 0.25em; + border: 1px solid gray; + margin-bottom: -1px + } + + .status.active .placeholder { + visibility: hidden; + } + .status { + position: relative; + } + .status .message { + position: absolute; + top: 0; + bottom: 0; + width: 100%; + text-align: center; + } + .pause-checkbox { + position: absolute; + right: 0; + top: 0; + padding: .5em; + } + </style> +</head> +<body> + <div class="sigma" id="sigma-example"></div> + <div class="summary"> + <!-- Note `<h1>Import Graph</h1>` is searched for by `lake exe graph`, keep it updated there. --> + <h1>Import Graph</h1> + <p>Built with <a href="https://www.sigmajs.org/">Sigma.js</a>. + Node sizes indicate the number of declarations in the file.</p> + <label class="pause-checkbox"><input type="checkbox" id="pause" checked/> Pause graph layout</label> + </div> + <div class="key-wrapper"> + <ul id="key"></ul> + <div class="status" id="statusWrapper"> + <p class="placeholder">Hover over a node to show only the files it imports. Hover over a directory name to highlight only the files in that directory</p> + <div class="message" id="statusMessage"></div> + </div> + </div> + +<script type="text/javascript"> + +var params = new URLSearchParams(window.location.search); +var docs_url = params.get("docs_url") || "https://leanprover-community.github.io/mathlib4_docs/"; + +// From stackoverflow +/** + * Converts an HSL color value to RGB. Conversion formula + * adapted from http://en.wikipedia.org/wiki/HSL_color_space. + * Assumes h, s, and l are contained in the set [0, 1] and + * returns r, g, and b in the set [0, 255]. + * + * @param {number} h The hue + * @param {number} s The saturation + * @param {number} l The lightness + * @return {Array} The RGB representation + */ +function hslToRgb(h, s, l){ + var r, g, b; + + if(s == 0){ + r = g = b = l; // achromatic + } else{ + var hue2rgb = function hue2rgb(p, q, t){ + if(t < 0) t += 1; + if(t > 1) t -= 1; + if(t < 1/6) return p + (q - p) * 6 * t; + if(t < 1/2) return q; + if(t < 2/3) return p + (q - p) * (2/3 - t) * 6; + return p; + } + + var q = l < 0.5 ? l * (1 + s) : l + s - l * s; + var p = 2 * l - q; + r = hue2rgb(p, q, h + 1/3); + g = hue2rgb(p, q, h); + b = hue2rgb(p, q, h - 1/3); + } + + return [Math.round(r * 255), Math.round(g * 255), Math.round(b * 255)]; +} + +var common_prefix = function(a, b) { + for (var i = 0; i < Math.min(a.length, b.length); i++) { + if (a[i] != b[i]) break; + } + return i; +} + +let container = document.getElementById('sigma-example'); +let key = document.getElementById('key'); +let pause = document.getElementById('pause'); + +var render_gexf = (gexf) => { + let graph = graphologyLibrary.gexf.parse(graphology.Graph, gexf); + let rev_graph = graphologyLibrary.operators.reverse(graph); + + // interaction state + var state = { + hoveredNode: undefined, + hoveredDescendants: undefined, + hoveredPath: undefined, + defaultNode: params.get("highlight"), + }; + var default_hover = !!params.get("default_hover"); + + // unpack the node labels into data + for (let [node, node_data] of graph.nodeEntries()) { + let path = node.split('.'); + node_data.proj = path; + if (graph.getNodeAttribute(node, 'in_module')) { + node_data.path = path.slice(1); + node_data.label = node_data.path.join('.') + } else { + node_data.path = path; + } + } + + // represent the directory tree as a nested map + var tree = new Map(); + for (let [node, node_data] of graph.nodeEntries()) { + var at = tree; + for (let p of node_data.path) { + var curr = at.get(p); + if (!curr) { + curr = new Map(); + at.set(p, curr); + } + at = curr; + } + } + + // return [f, depth] where f is in [0, 1) and corresponds to some measure + // of order within the heirarchy. + var get_frac = function (path, max_depth=Infinity) { + var pos = 0; + var width = 1; + var at = tree; + for (let [i, p] of path.entries()) { + var key_arr = Array.from(at.keys()).sort(); + width /= key_arr.length; + var ind = key_arr.indexOf(p); + if (ind == -1) { + pos = NaN; + break; + } + pos += width * ind; + if (i >= max_depth) { + break; + } + at = at.get(p); + } + return [pos, path.length]; + }; + + var get_color = function (node_data, s=1, l=0.5) { + var [h, d] = get_frac(node_data.path, 0); + var [r, g, b] = hslToRgb(h, s, l); + return 'rgb(' + [r, g, b].join(',') + ')' + }; + + // build the key + while(key.firstChild) key.removeChild(key.lastChild); + var keyElems = new Map(); + for (let top_level of Array.from(tree.keys()).sort()) { + var li = document.createElement('li'); + var sp = document.createElement('span'); + sp.style.backgroundColor = get_color({path: [top_level]}); + sp.style.borderColor = get_color({path: [top_level]}, 1, 0.4); + li.appendChild(sp); + li.appendChild(document.createTextNode(top_level)); + key.appendChild(li); + keyElems.set(top_level, li); + } + + // get max decl count + let max_decls = 0 + for (let [node, node_data] of graph.nodeEntries()) { + if (node_data.decl_count > max_decls) { + max_decls = node_data.decl_count + } + } + + // color the nodes + for (let [node, node_data] of graph.nodeEntries()) { + graph.setNodeAttribute(node, 'color', get_color(node_data)); + graph.setNodeAttribute(node, 'dark_color', get_color(node_data, 1, 0.4)); + graph.setNodeAttribute(node, 'dim_color', get_color(node_data, 0.5, 7/8)); + graph.setNodeAttribute(node, 'very_dim_color', get_color(node_data, 0.5, 15/16)); + } + + //size the nodes + for (let [node, node_data] of graph.nodeEntries()) { + node_data.size = 2 + Math.pow(node_data.decl_count, 1/2) * 20 / Math.pow(max_decls, 1/2); + } + + // set up the edge weights, to make the layout somewhat ok + for (let [key, attrs, s, t, sdata, tdata] of graph.edgeEntries()) { + var ncommon = common_prefix(sdata.path, tdata.path); + if (sdata.proj != tdata.proj) { + attrs.weight = 1 + } + else if (ncommon == 0) { + attrs.weight = 2; + } + else if (ncommon == 1) { + attrs.weight = 4; + } + else { + attrs.weight = 8; + } + // attrs.weight = (sdata.proj == tdata.proj) * Math.pow(5, ncommon) + 1; + } + + // pick a somewhat arbitrary initial layout + let sorted_nodes = Array.from(graph.nodes()).sort(); + for (let [i, node] of sorted_nodes.entries()) { + let node_data = graph.getNodeAttributes(node); + // var [theta, r] = get_frac(node_data.path); + // r = 1; + // graphologyLibrary.traversal.bfsFromNode(graph, node, function(node, attr, depth) { + // r = Math.max(r, depth + 1); + // }); + var theta = i / sorted_nodes.length; + var r = 1; + node_data.x = 0.1*r * Math.cos(theta * 2 * Math.PI); + node_data.y = 0.1*r * Math.sin(theta * 2 * Math.PI); + } + + var force_atlas_settings = { + barnesHutOptimize: false, + strongGravityMode: false, + gravity: 0.05, + scalingRatio: 10, + slowDown: 10, + edgeWeightInfluence: 1, + }; + + // run the optimizer a little + graphologyLibrary.layoutForceAtlas2.assign(graph, { + iterations: 10, + settings: force_atlas_settings, + }); + + // now send it to a background task + var worker = new graphologyLibrary.FA2Layout(graph, { + settings: force_atlas_settings, + }); + + if (!pause.checked) { + worker.start(); + } + + pause.addEventListener('change', (event) => { + if (event.currentTarget.checked) { + worker.stop(); + } else { + worker.start(); + } + }); + + window.graph = graph; + if (window.renderer) { + window.renderer.kill(); + delete window.renderer; + } + window.renderer = new Sigma.Sigma(graph, container, { + defaultEdgeType: "arrow", + defaultEdgeColor: "#c0c0c0", + minArrowSize: 30, + zIndex: true, + nodeReducer: (node, data) => { + var res = { ...data }; + + if (state.hoveredNode) { + if (state.hoveredNode === node) { + res.highlighted = true; + } + if (state.hoveredDescendants.has(node)) { + res.zIndex = 2; + } + else if (state.hoveredAncestors.has(node)) { + res.color = res.dim_color; + } + else { + res.label = undefined; + res.color = res.very_dim_color; + } + } + else if (state.hoveredPath) { + if (data.path[0] == state.hoveredPath) { + res.zIndex = 2; + } + else { + res.label = undefined; + res.color = res.dim_color; + } + } + else if (state.defaultNode) { + if (state.defaultNode === node) { + res.highlighted = true; + } + } + + return res; + }, + edgeReducer: (edge, data) => { + var res = { ...data }; + + if (state.hoveredNode) { + if (state.hoveredDescendants.has(graph.source(edge))) { + res.size = 1.5; + } + else if (state.hoveredNode == graph.target(edge)) { + res.color = graph.getNodeAttributes(state.hoveredNode).dark_color; + } + else { + res.hidden = true; + } + } + else if (state.hoveredPath) { + let src = graph.getNodeAttributes(graph.source(edge)); + let target = graph.getNodeAttributes(graph.target(edge)); + if (src.path[0] == state.hoveredPath && target.path[0] == state.hoveredPath) { + res.color = get_color({path: [state.hoveredPath]}, 1, 0.4); + } + else if (src.path[0] == state.hoveredPath) { + res.color = "#808080"; + } + else { + res.color = "#E0E0E0" + } + } + + return res; + }, + }); + + let statusElem = document.getElementById('statusWrapper'); + let statusMsgElem = document.getElementById('statusMessage'); + let setStatus = function(html) { + if (html) { + statusElem.classList.add('active'); + statusMsgElem.innerHTML = html; + } + else { + statusElem.classList.remove('active'); + statusMsgElem.innerHTML = ''; + } + } + + for (let [top_level, li] of keyElems) { + li.addEventListener('mouseenter', (event) => { + setHoveredNode(undefined); + state.hoveredPath = top_level; + let nfiles = 0; + let ndecls = 0; + for (let [n, node_data] of graph.nodeEntries()) { + if (node_data.path[0] == top_level) { + nfiles += 1; + ndecls += node_data.decl_count; + } + } + setStatus("Highlighting " + nfiles + " files consisting of "+ndecls+" declarations"); + renderer.refresh(); + }); + li.addEventListener('mouseleave', (event) => { + state.hoveredPath = undefined; + setHoveredNode(default_hover ? state.defaultNode : undefined); + setStatus(null); + renderer.refresh(); + }); + } + + // Bind graph interactions: + let setHoveredNode = function(node) { + if (node) { + state.hoveredNode = node; + var descendants = new Set(); + var ancestors = new Set(); + let ndecls = graph.getNodeAttribute(node, 'decl_count') + // let ndecls_trans = 0; + graphologyLibrary.traversal.bfsFromNode(graph, node, function(node2, attr, depth) { + descendants.add(node2); + // ndecls_trans += graph.getNodeAttribute(node2, 'decl_count'); + }); + graphologyLibrary.traversal.bfsFromNode(rev_graph, node, function(node, attr, depth) { + ancestors.add(node); + }); + state.hoveredDescendants = descendants; + state.hoveredAncestors = ancestors; + setStatus(`imported by ${descendants.size - 1} files; importing ${ancestors.size - 1} files; ${ndecls} declarations in this file.`); + container.style.cursor = "pointer"; + } else { + state.hoveredNode = undefined; + state.hoveredDescendants = undefined; + state.hoveredAncestors = undefined; + setStatus(null); + container.style.cursor = "inherit"; + } + + renderer.refresh(); + } + renderer.on("enterNode", ({ node }) => { setHoveredNode(node); }); + renderer.on("leaveNode", () => { setHoveredNode(default_hover ? state.defaultNode : undefined); }); + renderer.on("clickNode", ({ node }) => { + let path = graph.getNodeAttribute(node, 'path'); + window.open(docs_url + path.join('/') + ".html"); + }); + + setHoveredNode(default_hover ? state.defaultNode : undefined); +}; + +// NOTE: currently the Lean script searches for this line to replace it (exact search)! +// If you modify it, modify the Lean script (ImportGraph/Cli.lean) accordingly! +fetch("imports.gexf").then((res) => res.text()).then(render_gexf) + +container.addEventListener("dragover", ev => { + event.preventDefault(); +}); +container.addEventListener('drop', ev => { + ev.preventDefault(); + for (const item of ev.dataTransfer.items) { + if (item.kind == "file") { + item.getAsFile().text().then(render_gexf); + break; + } + } +}); +</script> + +</body> +</html> diff --git a/html-template/vendor/graphology-library.min.js b/html-template/vendor/graphology-library.min.js new file mode 100644 index 0000000..6fc8588 --- /dev/null +++ b/html-template/vendor/graphology-library.min.js @@ -0,0 +1,2 @@ +!function(t,e){"object"==typeof exports&&"undefined"!=typeof module?e(exports):"function"==typeof define&&define.amd?define(["exports"],e):e((t="undefined"!=typeof globalThis?globalThis:t||self).graphologyLibrary={})}(this,(function(t){"use strict";var e="undefined"!=typeof globalThis?globalThis:"undefined"!=typeof window?window:"undefined"!=typeof global?global:"undefined"!=typeof self?self:{},r={},n={};var i=function(){this.__data__=[],this.size=0};var o=function(t,e){return t===e||t!=t&&e!=e},a=o;var s=function(t,e){for(var r=t.length;r--;)if(a(t[r][0],e))return r;return-1},u=s,h=Array.prototype.splice;var c=s;var l=s;var d=s;var f=i,g=function(t){var e=this.__data__,r=u(e,t);return!(r<0)&&(r==e.length-1?e.pop():h.call(e,r,1),--this.size,!0)},p=function(t){var e=this.__data__,r=c(e,t);return r<0?void 0:e[r][1]},y=function(t){return l(this.__data__,t)>-1},v=function(t,e){var r=this.__data__,n=d(r,t);return n<0?(++this.size,r.push([t,e])):r[n][1]=e,this};function m(t){var e=-1,r=null==t?0:t.length;for(this.clear();++e<r;){var n=t[e];this.set(n[0],n[1])}}m.prototype.clear=f,m.prototype.delete=g,m.prototype.get=p,m.prototype.has=y,m.prototype.set=v;var b=m,w=b;var E=function(){this.__data__=new w,this.size=0};var A=function(t){var e=this.__data__,r=e.delete(t);return this.size=e.size,r};var x=function(t){return this.__data__.get(t)};var z=function(t){return this.__data__.has(t)},M="object"==typeof e&&e&&e.Object===Object&&e,S=M,N="object"==typeof self&&self&&self.Object===Object&&self,j=S||N||Function("return this")(),k=j.Symbol,D=k,O=Object.prototype,_=O.hasOwnProperty,C=O.toString,W=D?D.toStringTag:void 0;var P=function(t){var e=_.call(t,W),r=t[W];try{t[W]=void 0;var n=!0}catch(t){}var i=C.call(t);return n&&(e?t[W]=r:delete t[W]),i},I=Object.prototype.toString;var U=P,T=function(t){return I.call(t)},L=k?k.toStringTag:void 0;var F=function(t){return null==t?void 0===t?"[object Undefined]":"[object Null]":L&&L in Object(t)?U(t):T(t)};var q=function(t){var e=typeof t;return null!=t&&("object"==e||"function"==e)},R=F,B=q;var G,H=function(t){if(!B(t))return!1;var e=R(t);return"[object Function]"==e||"[object GeneratorFunction]"==e||"[object AsyncFunction]"==e||"[object Proxy]"==e},K=j["__core-js_shared__"],V=(G=/[^.]+$/.exec(K&&K.keys&&K.keys.IE_PROTO||""))?"Symbol(src)_1."+G:"";var $=function(t){return!!V&&V in t},J=Function.prototype.toString;var Q=function(t){if(null!=t){try{return J.call(t)}catch(t){}try{return t+""}catch(t){}}return""},X=H,Y=$,Z=q,tt=Q,et=/^\[object .+?Constructor\]$/,rt=Function.prototype,nt=Object.prototype,it=rt.toString,ot=nt.hasOwnProperty,at=RegExp("^"+it.call(ot).replace(/[\\^$.*+?()[\]{}|]/g,"\\$&").replace(/hasOwnProperty|(function).*?(?=\\\()| for .+?(?=\\\])/g,"$1.*?")+"$");var st=function(t){return!(!Z(t)||Y(t))&&(X(t)?at:et).test(tt(t))},ut=function(t,e){return null==t?void 0:t[e]};var ht=function(t,e){var r=ut(t,e);return st(r)?r:void 0},ct=ht(j,"Map"),lt=ht(Object,"create"),dt=lt;var ft=function(){this.__data__=dt?dt(null):{},this.size=0};var gt=function(t){var e=this.has(t)&&delete this.__data__[t];return this.size-=e?1:0,e},pt=lt,yt=Object.prototype.hasOwnProperty;var vt=function(t){var e=this.__data__;if(pt){var r=e[t];return"__lodash_hash_undefined__"===r?void 0:r}return yt.call(e,t)?e[t]:void 0},mt=lt,bt=Object.prototype.hasOwnProperty;var wt=lt;var Et=ft,At=gt,xt=vt,zt=function(t){var e=this.__data__;return mt?void 0!==e[t]:bt.call(e,t)},Mt=function(t,e){var r=this.__data__;return this.size+=this.has(t)?0:1,r[t]=wt&&void 0===e?"__lodash_hash_undefined__":e,this};function St(t){var e=-1,r=null==t?0:t.length;for(this.clear();++e<r;){var n=t[e];this.set(n[0],n[1])}}St.prototype.clear=Et,St.prototype.delete=At,St.prototype.get=xt,St.prototype.has=zt,St.prototype.set=Mt;var Nt=St,jt=b,kt=ct;var Dt=function(t){var e=typeof t;return"string"==e||"number"==e||"symbol"==e||"boolean"==e?"__proto__"!==t:null===t};var Ot=function(t,e){var r=t.__data__;return Dt(e)?r["string"==typeof e?"string":"hash"]:r.map},_t=Ot;var Ct=Ot;var Wt=Ot;var Pt=Ot;var It=function(){this.size=0,this.__data__={hash:new Nt,map:new(kt||jt),string:new Nt}},Ut=function(t){var e=_t(this,t).delete(t);return this.size-=e?1:0,e},Tt=function(t){return Ct(this,t).get(t)},Lt=function(t){return Wt(this,t).has(t)},Ft=function(t,e){var r=Pt(this,t),n=r.size;return r.set(t,e),this.size+=r.size==n?0:1,this};function qt(t){var e=-1,r=null==t?0:t.length;for(this.clear();++e<r;){var n=t[e];this.set(n[0],n[1])}}qt.prototype.clear=It,qt.prototype.delete=Ut,qt.prototype.get=Tt,qt.prototype.has=Lt,qt.prototype.set=Ft;var Rt=qt,Bt=b,Gt=ct,Ht=Rt;var Kt=b,Vt=E,$t=A,Jt=x,Qt=z,Xt=function(t,e){var r=this.__data__;if(r instanceof Bt){var n=r.__data__;if(!Gt||n.length<199)return n.push([t,e]),this.size=++r.size,this;r=this.__data__=new Ht(n)}return r.set(t,e),this.size=r.size,this};function Yt(t){var e=this.__data__=new Kt(t);this.size=e.size}Yt.prototype.clear=Vt,Yt.prototype.delete=$t,Yt.prototype.get=Jt,Yt.prototype.has=Qt,Yt.prototype.set=Xt;var Zt=Yt;var te=Rt,ee=function(t){return this.__data__.set(t,"__lodash_hash_undefined__"),this},re=function(t){return this.__data__.has(t)};function ne(t){var e=-1,r=null==t?0:t.length;for(this.__data__=new te;++e<r;)this.add(t[e])}ne.prototype.add=ne.prototype.push=ee,ne.prototype.has=re;var ie=ne,oe=function(t,e){for(var r=-1,n=null==t?0:t.length;++r<n;)if(e(t[r],r,t))return!0;return!1},ae=function(t,e){return t.has(e)};var se=function(t,e,r,n,i,o){var a=1&r,s=t.length,u=e.length;if(s!=u&&!(a&&u>s))return!1;var h=o.get(t),c=o.get(e);if(h&&c)return h==e&&c==t;var l=-1,d=!0,f=2&r?new ie:void 0;for(o.set(t,e),o.set(e,t);++l<s;){var g=t[l],p=e[l];if(n)var y=a?n(p,g,l,e,t,o):n(g,p,l,t,e,o);if(void 0!==y){if(y)continue;d=!1;break}if(f){if(!oe(e,(function(t,e){if(!ae(f,e)&&(g===t||i(g,t,r,n,o)))return f.push(e)}))){d=!1;break}}else if(g!==p&&!i(g,p,r,n,o)){d=!1;break}}return o.delete(t),o.delete(e),d},ue=j.Uint8Array;var he=function(t){var e=-1,r=Array(t.size);return t.forEach((function(t){r[++e]=t})),r},ce=ue,le=o,de=se,fe=function(t){var e=-1,r=Array(t.size);return t.forEach((function(t,n){r[++e]=[n,t]})),r},ge=he,pe=k?k.prototype:void 0,ye=pe?pe.valueOf:void 0;var ve=function(t,e,r,n,i,o,a){switch(r){case"[object DataView]":if(t.byteLength!=e.byteLength||t.byteOffset!=e.byteOffset)return!1;t=t.buffer,e=e.buffer;case"[object ArrayBuffer]":return!(t.byteLength!=e.byteLength||!o(new ce(t),new ce(e)));case"[object Boolean]":case"[object Date]":case"[object Number]":return le(+t,+e);case"[object Error]":return t.name==e.name&&t.message==e.message;case"[object RegExp]":case"[object String]":return t==e+"";case"[object Map]":var s=fe;case"[object Set]":var u=1&n;if(s||(s=ge),t.size!=e.size&&!u)return!1;var h=a.get(t);if(h)return h==e;n|=2,a.set(t,e);var c=de(s(t),s(e),n,i,o,a);return a.delete(t),c;case"[object Symbol]":if(ye)return ye.call(t)==ye.call(e)}return!1};var me=function(t,e){for(var r=-1,n=e.length,i=t.length;++r<n;)t[i+r]=e[r];return t},be=Array.isArray,we=me,Ee=be;var Ae=function(t,e,r){var n=e(t);return Ee(t)?n:we(n,r(t))};var xe=function(t,e){for(var r=-1,n=null==t?0:t.length,i=0,o=[];++r<n;){var a=t[r];e(a,r,t)&&(o[i++]=a)}return o},ze=function(){return[]},Me=Object.prototype.propertyIsEnumerable,Se=Object.getOwnPropertySymbols,Ne=Se?function(t){return null==t?[]:(t=Object(t),xe(Se(t),(function(e){return Me.call(t,e)})))}:ze;var je=function(t,e){for(var r=-1,n=Array(t);++r<t;)n[r]=e(r);return n};var ke=function(t){return null!=t&&"object"==typeof t},De=F,Oe=ke;var _e=function(t){return Oe(t)&&"[object Arguments]"==De(t)},Ce=ke,We=Object.prototype,Pe=We.hasOwnProperty,Ie=We.propertyIsEnumerable,Ue=_e(function(){return arguments}())?_e:function(t){return Ce(t)&&Pe.call(t,"callee")&&!Ie.call(t,"callee")},Te=Ue,Le={exports:{}};var Fe=function(){return!1};!function(t,e){var r=j,n=Fe,i=e&&!e.nodeType&&e,o=i&&t&&!t.nodeType&&t,a=o&&o.exports===i?r.Buffer:void 0,s=(a?a.isBuffer:void 0)||n;t.exports=s}(Le,Le.exports);var qe=/^(?:0|[1-9]\d*)$/;var Re=function(t,e){var r=typeof t;return!!(e=null==e?9007199254740991:e)&&("number"==r||"symbol"!=r&&qe.test(t))&&t>-1&&t%1==0&&t<e};var Be=function(t){return"number"==typeof t&&t>-1&&t%1==0&&t<=9007199254740991},Ge=F,He=Be,Ke=ke,Ve={};Ve["[object Float32Array]"]=Ve["[object Float64Array]"]=Ve["[object Int8Array]"]=Ve["[object Int16Array]"]=Ve["[object Int32Array]"]=Ve["[object Uint8Array]"]=Ve["[object Uint8ClampedArray]"]=Ve["[object Uint16Array]"]=Ve["[object Uint32Array]"]=!0,Ve["[object Arguments]"]=Ve["[object Array]"]=Ve["[object ArrayBuffer]"]=Ve["[object Boolean]"]=Ve["[object DataView]"]=Ve["[object Date]"]=Ve["[object Error]"]=Ve["[object Function]"]=Ve["[object Map]"]=Ve["[object Number]"]=Ve["[object Object]"]=Ve["[object RegExp]"]=Ve["[object Set]"]=Ve["[object String]"]=Ve["[object WeakMap]"]=!1;var $e=function(t){return Ke(t)&&He(t.length)&&!!Ve[Ge(t)]};var Je=function(t){return function(e){return t(e)}},Qe={exports:{}};!function(t,e){var r=M,n=e&&!e.nodeType&&e,i=n&&t&&!t.nodeType&&t,o=i&&i.exports===n&&r.process,a=function(){try{var t=i&&i.require&&i.require("util").types;return t||o&&o.binding&&o.binding("util")}catch(t){}}();t.exports=a}(Qe,Qe.exports);var Xe=$e,Ye=Je,Ze=Qe.exports,tr=Ze&&Ze.isTypedArray,er=tr?Ye(tr):Xe,rr=je,nr=Te,ir=be,or=Le.exports,ar=Re,sr=er,ur=Object.prototype.hasOwnProperty;var hr=function(t,e){var r=ir(t),n=!r&&nr(t),i=!r&&!n&&or(t),o=!r&&!n&&!i&&sr(t),a=r||n||i||o,s=a?rr(t.length,String):[],u=s.length;for(var h in t)!e&&!ur.call(t,h)||a&&("length"==h||i&&("offset"==h||"parent"==h)||o&&("buffer"==h||"byteLength"==h||"byteOffset"==h)||ar(h,u))||s.push(h);return s},cr=Object.prototype;var lr=function(t){var e=t&&t.constructor;return t===("function"==typeof e&&e.prototype||cr)};var dr=function(t,e){return function(r){return t(e(r))}},fr=dr(Object.keys,Object),gr=lr,pr=fr,yr=Object.prototype.hasOwnProperty;var vr=H,mr=Be;var br=function(t){return null!=t&&mr(t.length)&&!vr(t)},wr=hr,Er=function(t){if(!gr(t))return pr(t);var e=[];for(var r in Object(t))yr.call(t,r)&&"constructor"!=r&&e.push(r);return e},Ar=br;var xr=Ae,zr=Ne,Mr=function(t){return Ar(t)?wr(t):Er(t)};var Sr=function(t){return xr(t,Mr,zr)},Nr=Object.prototype.hasOwnProperty;var jr=function(t,e,r,n,i,o){var a=1&r,s=Sr(t),u=s.length;if(u!=Sr(e).length&&!a)return!1;for(var h=u;h--;){var c=s[h];if(!(a?c in e:Nr.call(e,c)))return!1}var l=o.get(t),d=o.get(e);if(l&&d)return l==e&&d==t;var f=!0;o.set(t,e),o.set(e,t);for(var g=a;++h<u;){var p=t[c=s[h]],y=e[c];if(n)var v=a?n(y,p,c,e,t,o):n(p,y,c,t,e,o);if(!(void 0===v?p===y||i(p,y,r,n,o):v)){f=!1;break}g||(g="constructor"==c)}if(f&&!g){var m=t.constructor,b=e.constructor;m==b||!("constructor"in t)||!("constructor"in e)||"function"==typeof m&&m instanceof m&&"function"==typeof b&&b instanceof b||(f=!1)}return o.delete(t),o.delete(e),f},kr=ht(j,"DataView"),Dr=ct,Or=ht(j,"Promise"),_r=ht(j,"Set"),Cr=ht(j,"WeakMap"),Wr=F,Pr=Q,Ir="[object Map]",Ur="[object Promise]",Tr="[object Set]",Lr="[object WeakMap]",Fr="[object DataView]",qr=Pr(kr),Rr=Pr(Dr),Br=Pr(Or),Gr=Pr(_r),Hr=Pr(Cr),Kr=Wr;(kr&&Kr(new kr(new ArrayBuffer(1)))!=Fr||Dr&&Kr(new Dr)!=Ir||Or&&Kr(Or.resolve())!=Ur||_r&&Kr(new _r)!=Tr||Cr&&Kr(new Cr)!=Lr)&&(Kr=function(t){var e=Wr(t),r="[object Object]"==e?t.constructor:void 0,n=r?Pr(r):"";if(n)switch(n){case qr:return Fr;case Rr:return Ir;case Br:return Ur;case Gr:return Tr;case Hr:return Lr}return e});var Vr=Zt,$r=se,Jr=ve,Qr=jr,Xr=Kr,Yr=be,Zr=Le.exports,tn=er,en="[object Arguments]",rn="[object Array]",nn="[object Object]",on=Object.prototype.hasOwnProperty;var an=function(t,e,r,n,i,o){var a=Yr(t),s=Yr(e),u=a?rn:Xr(t),h=s?rn:Xr(e),c=(u=u==en?nn:u)==nn,l=(h=h==en?nn:h)==nn,d=u==h;if(d&&Zr(t)){if(!Zr(e))return!1;a=!0,c=!1}if(d&&!c)return o||(o=new Vr),a||tn(t)?$r(t,e,r,n,i,o):Jr(t,e,u,r,n,i,o);if(!(1&r)){var f=c&&on.call(t,"__wrapped__"),g=l&&on.call(e,"__wrapped__");if(f||g){var p=f?t.value():t,y=g?e.value():e;return o||(o=new Vr),i(p,y,r,n,o)}}return!!d&&(o||(o=new Vr),Qr(t,e,r,n,i,o))},sn=ke;var un=function t(e,r,n,i,o){return e===r||(null==e||null==r||!sn(e)&&!sn(r)?e!=e&&r!=r:an(e,r,n,i,t,o))},hn=un;var cn=function(t,e){return hn(t,e)};n.isGraph=function(t){return null!==t&&"object"==typeof t&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode&&"boolean"==typeof t.multi},n.isGraphConstructor=function(t){return null!==t&&"function"==typeof t&&"object"==typeof t.prototype&&"function"==typeof t.prototype.addUndirectedEdgeWithKey&&"function"==typeof t.prototype.dropNode},n.sameNodes=function(t,e){if(t.order!==e.order)return!1;var r,n=t.nodes(),i=t.order;for(r=0;r<i;r++)if(!e.hasNode(n[r]))return!1;return!0},n.sameEdges=function(t,e){if(t.directedSize!==e.directedSize||t.undirectedSize!==e.undirectedSize)return!1},n.sameNodesDeep=function(t,e){if(t.order!==e.order)return!1;var r,n=t.nodes(),i=t.order;for(r=0;r<i;r++){if(!e.hasNode(n[r]))return!1;if(!cn(t.getNodeAttributes(n[r]),e.getNodeAttributes(n[r])))return!1}return!0};var ln=n;var dn=function(t,e,r){switch(r.length){case 0:return t.call(e);case 1:return t.call(e,r[0]);case 2:return t.call(e,r[0],r[1]);case 3:return t.call(e,r[0],r[1],r[2])}return t.apply(e,r)};var fn=function(t){return t},gn=dn,pn=Math.max;var yn=function(t,e,r){return e=pn(void 0===e?t.length-1:e,0),function(){for(var n=arguments,i=-1,o=pn(n.length-e,0),a=Array(o);++i<o;)a[i]=n[e+i];i=-1;for(var s=Array(e+1);++i<e;)s[i]=n[i];return s[e]=r(a),gn(t,this,s)}};var vn=function(t){return function(){return t}},mn=ht,bn=function(){try{var t=mn(Object,"defineProperty");return t({},"",{}),t}catch(t){}}(),wn=vn,En=bn,An=En?function(t,e){return En(t,"toString",{configurable:!0,enumerable:!1,value:wn(e),writable:!0})}:fn,xn=Date.now;var zn=function(t){var e=0,r=0;return function(){var n=xn(),i=16-(n-r);if(r=n,i>0){if(++e>=800)return arguments[0]}else e=0;return t.apply(void 0,arguments)}},Mn=zn(An),Sn=fn,Nn=yn,jn=Mn;var kn=function(t,e){return jn(Nn(t,e,Sn),t+"")},Dn=bn;var On=function(t,e,r){"__proto__"==e&&Dn?Dn(t,e,{configurable:!0,enumerable:!0,value:r,writable:!0}):t[e]=r},_n=On,Cn=o;var Wn=function(t,e,r){(void 0!==r&&!Cn(t[e],r)||void 0===r&&!(e in t))&&_n(t,e,r)};var Pn=function(t){return function(e,r,n){for(var i=-1,o=Object(e),a=n(e),s=a.length;s--;){var u=a[t?s:++i];if(!1===r(o[u],u,o))break}return e}}(),In={exports:{}};!function(t,e){var r=j,n=e&&!e.nodeType&&e,i=n&&t&&!t.nodeType&&t,o=i&&i.exports===n?r.Buffer:void 0,a=o?o.allocUnsafe:void 0;t.exports=function(t,e){if(e)return t.slice();var r=t.length,n=a?a(r):new t.constructor(r);return t.copy(n),n}}(In,In.exports);var Un=ue;var Tn=function(t){var e=new t.constructor(t.byteLength);return new Un(e).set(new Un(t)),e};var Ln=function(t,e){var r=e?Tn(t.buffer):t.buffer;return new t.constructor(r,t.byteOffset,t.length)};var Fn=function(t,e){var r=-1,n=t.length;for(e||(e=Array(n));++r<n;)e[r]=t[r];return e},qn=q,Rn=Object.create,Bn=function(){function t(){}return function(e){if(!qn(e))return{};if(Rn)return Rn(e);t.prototype=e;var r=new t;return t.prototype=void 0,r}}(),Gn=dr(Object.getPrototypeOf,Object),Hn=Bn,Kn=Gn,Vn=lr;var $n=function(t){return"function"!=typeof t.constructor||Vn(t)?{}:Hn(Kn(t))},Jn=br,Qn=ke;var Xn=function(t){return Qn(t)&&Jn(t)},Yn=F,Zn=Gn,ti=ke,ei=Function.prototype,ri=Object.prototype,ni=ei.toString,ii=ri.hasOwnProperty,oi=ni.call(Object);var ai=function(t){if(!ti(t)||"[object Object]"!=Yn(t))return!1;var e=Zn(t);if(null===e)return!0;var r=ii.call(e,"constructor")&&e.constructor;return"function"==typeof r&&r instanceof r&&ni.call(r)==oi};var si=function(t,e){if(("constructor"!==e||"function"!=typeof t[e])&&"__proto__"!=e)return t[e]},ui=On,hi=o,ci=Object.prototype.hasOwnProperty;var li=function(t,e,r){var n=t[e];ci.call(t,e)&&hi(n,r)&&(void 0!==r||e in t)||ui(t,e,r)},di=On;var fi=function(t,e,r,n){var i=!r;r||(r={});for(var o=-1,a=e.length;++o<a;){var s=e[o],u=n?n(r[s],t[s],s,r,t):void 0;void 0===u&&(u=t[s]),i?di(r,s,u):li(r,s,u)}return r};var gi=q,pi=lr,yi=function(t){var e=[];if(null!=t)for(var r in Object(t))e.push(r);return e},vi=Object.prototype.hasOwnProperty;var mi=hr,bi=function(t){if(!gi(t))return yi(t);var e=pi(t),r=[];for(var n in t)("constructor"!=n||!e&&vi.call(t,n))&&r.push(n);return r},wi=br;var Ei=function(t){return wi(t)?mi(t,!0):bi(t)},Ai=fi,xi=Ei;var zi=Wn,Mi=In.exports,Si=Ln,Ni=Fn,ji=$n,ki=Te,Di=be,Oi=Xn,_i=Le.exports,Ci=H,Wi=q,Pi=ai,Ii=er,Ui=si,Ti=function(t){return Ai(t,xi(t))};var Li=Zt,Fi=Wn,qi=Pn,Ri=function(t,e,r,n,i,o,a){var s=Ui(t,r),u=Ui(e,r),h=a.get(u);if(h)zi(t,r,h);else{var c=o?o(s,u,r+"",t,e,a):void 0,l=void 0===c;if(l){var d=Di(u),f=!d&&_i(u),g=!d&&!f&&Ii(u);c=u,d||f||g?Di(s)?c=s:Oi(s)?c=Ni(s):f?(l=!1,c=Mi(u,!0)):g?(l=!1,c=Si(u,!0)):c=[]:Pi(u)||ki(u)?(c=s,ki(s)?c=Ti(s):Wi(s)&&!Ci(s)||(c=ji(u))):l=!1}l&&(a.set(u,c),i(c,u,n,o,a),a.delete(u)),zi(t,r,c)}},Bi=q,Gi=Ei,Hi=si;var Ki=function t(e,r,n,i,o){e!==r&&qi(r,(function(a,s){if(o||(o=new Li),Bi(a))Ri(e,r,s,n,t,i,o);else{var u=i?i(Hi(e,s),a,s+"",e,r,o):void 0;void 0===u&&(u=a),Fi(e,s,u)}}),Gi)},Vi=Ki,$i=q;var Ji=function t(e,r,n,i,o,a){return $i(e)&&$i(r)&&(a.set(r,e),Vi(e,r,void 0,t,a),a.delete(r)),e},Qi=o,Xi=br,Yi=Re,Zi=q;var to=function(t,e,r){if(!Zi(r))return!1;var n=typeof e;return!!("number"==n?Xi(r)&&Yi(e,r.length):"string"==n&&e in r)&&Qi(r[e],t)},eo=kn,ro=to;var no=Ki,io=function(t){return eo((function(e,r){var n=-1,i=r.length,o=i>1?r[i-1]:void 0,a=i>2?r[2]:void 0;for(o=t.length>3&&"function"==typeof o?(i--,o):void 0,a&&ro(r[0],r[1],a)&&(o=i<3?void 0:o,i=1),e=Object(e);++n<i;){var s=r[n];s&&t(e,s,n,o)}return e}))}((function(t,e,r,n){no(t,e,r,n)})),oo=dn,ao=Ji,so=io,uo=kn((function(t){return t.push(void 0,ao),oo(so,void 0,t)})),ho=function(t){return null!==t&&"object"==typeof t&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode&&"boolean"==typeof t.multi},co=ho,lo=function(t){if(!co(t))throw new Error("graphology-utils/infer-type: expecting a valid graphology instance.");var e=t.type;return"mixed"!==e?e:0===t.directedSize&&0===t.undirectedSize||t.directedSize>0&&t.undirectedSize>0?"mixed":t.directedSize>0?"directed":"undirected"};function fo(t){Object.defineProperty(this,"_next",{writable:!1,enumerable:!1,value:t}),this.done=!1}fo.prototype.next=function(){if(this.done)return{done:!0};var t=this._next();return t.done&&(this.done=!0),t},"undefined"!=typeof Symbol&&(fo.prototype[Symbol.iterator]=function(){return this}),fo.of=function(){var t=arguments,e=t.length,r=0;return new fo((function(){return r>=e?{done:!0}:{done:!1,value:t[r++]}}))},fo.empty=function(){var t=new fo(null);return t.done=!0,t},fo.is=function(t){return t instanceof fo||"object"==typeof t&&null!==t&&"function"==typeof t.next};var go=fo,po={};!function(t){var e=Math.pow(2,8)-1,r=Math.pow(2,16)-1,n=Math.pow(2,32)-1,i=Math.pow(2,7)-1,o=Math.pow(2,15)-1,a=Math.pow(2,31)-1;t.getPointerArray=function(t){var i=t-1;if(i<=e)return Uint8Array;if(i<=r)return Uint16Array;if(i<=n)return Uint32Array;throw new Error("mnemonist: Pointer Array of size > 4294967295 is not supported.")},t.getSignedPointerArray=function(t){var e=t-1;return e<=i?Int8Array:e<=o?Int16Array:e<=a?Int32Array:Float64Array},t.getNumberType=function(t){return t===(0|t)?-1===Math.sign(t)?t<=127&&t>=-128?Int8Array:t<=32767&&t>=-32768?Int16Array:Int32Array:t<=255?Uint8Array:t<=65535?Uint16Array:Uint32Array:Float64Array};var s={Uint8Array:1,Int8Array:2,Uint16Array:3,Int16Array:4,Uint32Array:5,Int32Array:6,Float32Array:7,Float64Array:8};t.getMinimalRepresentation=function(e,r){var n,i,o,a,u,h=null,c=0;for(a=0,u=e.length;a<u;a++)o=r?r(e[a]):e[a],i=t.getNumberType(o),(n=s[i.name])>c&&(c=n,h=i);return h},t.isTypedArray=function(t){return"undefined"!=typeof ArrayBuffer&&ArrayBuffer.isView(t)},t.concat=function(){var t,e,r,n=0;for(t=0,r=arguments.length;t<r;t++)n+=arguments[t].length;var i=new arguments[0].constructor(n);for(t=0,e=0;t<r;t++)i.set(arguments[t],e),e+=arguments[t].length;return i},t.indices=function(e){for(var r=new(t.getPointerArray(e))(e),n=0;n<e;n++)r[n]=n;return r}}(po);var yo=go,vo=po.getPointerArray;function mo(t,e){arguments.length<2&&(e=t,t=Array);var r=vo(e);this.size=0,this.length=e,this.dense=new r(e),this.sparse=new r(e),this.vals=new t(e)}mo.prototype.clear=function(){this.size=0},mo.prototype.has=function(t){var e=this.sparse[t];return e<this.size&&this.dense[e]===t},mo.prototype.get=function(t){var e=this.sparse[t];if(e<this.size&&this.dense[e]===t)return this.vals[e]},mo.prototype.set=function(t,e){var r=this.sparse[t];return r<this.size&&this.dense[r]===t?(this.vals[r]=e,this):(this.dense[this.size]=t,this.sparse[t]=this.size,this.vals[this.size]=e,this.size++,this)},mo.prototype.delete=function(t){var e=this.sparse[t];return!(e>=this.size||this.dense[e]!==t)&&(e=this.dense[this.size-1],this.dense[this.sparse[t]]=e,this.sparse[e]=this.sparse[t],this.size--,!0)},mo.prototype.forEach=function(t,e){e=arguments.length>1?e:this;for(var r=0;r<this.size;r++)t.call(e,this.vals[r],this.dense[r])},mo.prototype.keys=function(){var t=this.size,e=this.dense,r=0;return new yo((function(){if(r<t){var n=e[r];return r++,{value:n}}return{done:!0}}))},mo.prototype.values=function(){var t=this.size,e=this.vals,r=0;return new yo((function(){if(r<t){var n=e[r];return r++,{value:n}}return{done:!0}}))},mo.prototype.entries=function(){var t=this.size,e=this.dense,r=this.vals,n=0;return new yo((function(){if(n<t){var i=[e[n],r[n]];return n++,{value:i}}return{done:!0}}))},"undefined"!=typeof Symbol&&(mo.prototype[Symbol.iterator]=mo.prototype.entries),mo.prototype.inspect=function(){for(var t=new Map,e=0;e<this.size;e++)t.set(this.dense[e],this.vals[e]);return Object.defineProperty(t,"constructor",{value:mo,enumerable:!1}),t.length=this.length,this.vals.constructor!==Array&&(t.type=this.vals.constructor.name),t},"undefined"!=typeof Symbol&&(mo.prototype[Symbol.for("nodejs.util.inspect.custom")]=mo.prototype.inspect);var bo=mo,wo=go,Eo=po.getPointerArray;function Ao(t){var e=Eo(t);this.start=0,this.size=0,this.capacity=t,this.dense=new e(t),this.sparse=new e(t)}Ao.prototype.clear=function(){this.start=0,this.size=0},Ao.prototype.has=function(t){if(0===this.size)return!1;var e=this.sparse[t];return(e<this.capacity&&e>=this.start&&e<this.start+this.size||e<(this.start+this.size)%this.capacity)&&this.dense[e]===t},Ao.prototype.enqueue=function(t){var e=this.sparse[t];if(0!==this.size&&((e<this.capacity&&e>=this.start&&e<this.start+this.size||e<(this.start+this.size)%this.capacity)&&this.dense[e]===t))return this;return e=(this.start+this.size)%this.capacity,this.dense[e]=t,this.sparse[t]=e,this.size++,this},Ao.prototype.dequeue=function(){if(0!==this.size){var t=this.start;this.size--,this.start++,this.start===this.capacity&&(this.start=0);var e=this.dense[t];return this.sparse[e]=this.capacity,e}},Ao.prototype.forEach=function(t,e){e=arguments.length>1?e:this;for(var r=this.capacity,n=this.size,i=this.start,o=0;o<n;)t.call(e,this.dense[i],o,this),o++,++i===r&&(i=0)},Ao.prototype.values=function(){var t=this.dense,e=this.capacity,r=this.size,n=this.start,i=0;return new wo((function(){if(i>=r)return{done:!0};var o=t[n];return n++,i++,n===e&&(n=0),{value:o,done:!1}}))},"undefined"!=typeof Symbol&&(Ao.prototype[Symbol.iterator]=Ao.prototype.values),Ao.prototype.inspect=function(){var t=[];return this.forEach((function(e){t.push(e)})),Object.defineProperty(t,"constructor",{value:Ao,enumerable:!1}),t.capacity=this.capacity,t},"undefined"!=typeof Symbol&&(Ao.prototype[Symbol.for("nodejs.util.inspect.custom")]=Ao.prototype.inspect);var xo=Ao;function zo(t){return function(e,r){return e+Math.floor(t()*(r-e+1))}}var Mo=zo(Math.random);Mo.createRandom=zo;var So=Mo.createRandom;function No(t){var e=So(t);return function(t){return"number"!=typeof t&&(t=t.length),e(0,t-1)}}var jo=No(Math.random);jo.createRandomIndex=No;var ko=jo,Do={},Oo=po,_o=Symbol.for("nodejs.util.inspect.custom"),Co={weight:"weight"},Wo=1;function Po(t,e){var r=(e=e||{}).attributes||{},n=!0===e.keepDendrogram,i="number"==typeof e.resolution?e.resolution:Wo,o=!0===e.weighted,a=r.weight||Co.weight,s=2*(t.size-t.selfLoopCount),u=Oo.getPointerArray(s),h=Oo.getPointerArray(t.order+1),c=o?Float64Array:Oo.getPointerArray(2*t.size);this.C=t.order,this.M=0,this.E=s,this.U=0,this.resolution=i,this.level=0,this.graph=t,this.nodes=new Array(t.order),this.keepDendrogram=n,this.neighborhood=new h(s),this.weights=new c(s),this.loops=new c(t.order),this.starts=new u(t.order+1),this.belongings=new h(t.order),this.dendrogram=[],this.mapping=null,this.counts=new h(t.order),this.unused=new h(t.order),this.totalWeights=new c(t.order);var l,d={},f=0,g=0,p=this;t.forEachNode((function(e){p.nodes[f]=e,d[e]=f,g+=t.undirectedDegree(e,!1),p.starts[f]=g,p.belongings[f]=f,p.counts[f]=1,f++})),t.forEachEdge((function(t,e,r,n){if(l=function(t){if(!o)return 1;var e=t[a];return"number"!=typeof e||isNaN(e)?1:e}(e),r=d[r],n=d[n],p.M+=l,r===n)p.totalWeights[r]+=2*l,p.loops[r]=2*l;else{p.totalWeights[r]+=l,p.totalWeights[n]+=l;var i=--p.starts[r],s=--p.starts[n];p.neighborhood[i]=n,p.neighborhood[s]=r,p.weights[i]=l,p.weights[s]=l}})),this.starts[f]=this.E,this.keepDendrogram?this.dendrogram.push(this.belongings.slice()):this.mapping=this.belongings.slice()}function Io(t,e){var r=(e=e||{}).attributes||{},n=!0===e.keepDendrogram,i="number"==typeof e.resolution?e.resolution:Wo,o=!0===e.weighted,a=r.weight||Co.weight,s=2*(t.size-t.selfLoopCount),u=Oo.getPointerArray(s),h=Oo.getPointerArray(t.order+1),c=o?Float64Array:Oo.getPointerArray(2*t.size);this.C=t.order,this.M=0,this.E=s,this.U=0,this.resolution=i,this.level=0,this.graph=t,this.nodes=new Array(t.order),this.keepDendrogram=n,this.neighborhood=new h(s),this.weights=new c(s),this.loops=new c(t.order),this.starts=new u(t.order+1),this.offsets=new u(t.order),this.belongings=new h(t.order),this.dendrogram=[],this.counts=new h(t.order),this.unused=new h(t.order),this.totalInWeights=new c(t.order),this.totalOutWeights=new c(t.order);var l,d={},f=0,g=0,p=this;t.forEachNode((function(e){p.nodes[f]=e,d[e]=f,g+=t.outDegree(e,!1),p.starts[f]=g,g+=t.inDegree(e,!1),p.offsets[f]=g,p.belongings[f]=f,p.counts[f]=1,f++})),t.forEachEdge((function(t,e,r,n){if(l=function(t){if(!o)return 1;var e=t[a];return"number"!=typeof e||isNaN(e)?1:e}(e),r=d[r],n=d[n],p.M+=l,r===n)p.loops[r]+=l,p.totalInWeights[r]+=l,p.totalOutWeights[r]+=l;else{p.totalOutWeights[r]+=l,p.totalInWeights[n]+=l;var i=--p.starts[r],s=--p.offsets[n];p.neighborhood[i]=n,p.neighborhood[s]=r,p.weights[i]=l,p.weights[s]=l}})),this.starts[f]=this.E,this.keepDendrogram?this.dendrogram.push(this.belongings.slice()):this.mapping=this.belongings.slice()}Po.prototype.isolate=function(t,e){var r=this.belongings[t];if(1===this.counts[r])return r;var n=this.unused[--this.U],i=this.loops[t];return this.totalWeights[r]-=e+i,this.totalWeights[n]+=e+i,this.belongings[t]=n,this.counts[r]--,this.counts[n]++,n},Po.prototype.move=function(t,e,r){var n=this.belongings[t],i=this.loops[t];this.totalWeights[n]-=e+i,this.totalWeights[r]+=e+i,this.belongings[t]=r;var o=1==this.counts[n]--;this.counts[r]++,o&&(this.unused[this.U++]=n)},Po.prototype.computeNodeDegree=function(t){var e,r,n=0;for(e=this.starts[t],r=this.starts[t+1];e<r;e++)n+=this.weights[e];return n},Po.prototype.expensiveIsolate=function(t){var e=this.computeNodeDegree(t);return this.isolate(t,e)},Po.prototype.expensiveMove=function(t,e){var r=this.computeNodeDegree(t);this.move(t,r,e)},Po.prototype.zoomOut=function(){var t,e,r,n,i,o,a,s,u,h,c,l=new Array(this.C-this.U),d={},f=this.nodes.length,g=0,p=0;for(t=0,r=this.C;t<r;t++)(o=this.belongings[t])in d||(d[o]=g,l[g]={adj:{},totalWeights:this.totalWeights[o],internalWeights:0},g++),this.belongings[t]=d[o];if(this.keepDendrogram){for(h=this.dendrogram[this.level],c=new(Oo.getPointerArray(g))(f),t=0;t<f;t++)c[t]=this.belongings[h[t]];this.dendrogram.push(c)}else for(t=0;t<f;t++)this.mapping[t]=this.belongings[this.mapping[t]];for(t=0,r=this.C;t<r;t++)for(u=(s=l[o=this.belongings[t]]).adj,s.internalWeights+=this.loops[t],e=this.starts[t],n=this.starts[t+1];e<n;e++)i=this.neighborhood[e],o!==(a=this.belongings[i])?(a in u||(u[a]=0),u[a]+=this.weights[e]):s.internalWeights+=this.weights[e];for(this.C=g,i=0,o=0;o<g;o++)for(a in u=(s=l[o]).adj,o=+o,this.totalWeights[o]=s.totalWeights,this.loops[o]=s.internalWeights,this.counts[o]=1,this.starts[o]=i,this.belongings[o]=o,u)this.neighborhood[i]=+a,this.weights[i]=u[a],p++,i++;return this.starts[g]=p,this.E=p,this.U=0,this.level++,d},Po.prototype.modularity=function(){var t,e,r,n,i=0,o=2*this.M,a=new Float64Array(this.C);for(e=0;e<this.C;e++)for(a[t=this.belongings[e]]+=this.loops[e],r=this.starts[e],n=this.starts[e+1];r<n;r++)t===this.belongings[this.neighborhood[r]]&&(a[t]+=this.weights[r]);for(e=0;e<this.C;e++)i+=a[e]/o-Math.pow(this.totalWeights[e]/o,2)*this.resolution;return i},Po.prototype.delta=function(t,e,r,n){var i=this.M;return r/i-this.totalWeights[n]*(e+=this.loops[t])*this.resolution/(2*i*i)},Po.prototype.deltaWithOwnCommunity=function(t,e,r,n){var i=this.M;return r/i-(this.totalWeights[n]-(e+=this.loops[t]))*e*this.resolution/(2*i*i)},Po.prototype.fastDelta=function(t,e,r,n){var i=this.M,o=this.totalWeights[n];return r-(e+=this.loops[t])*o*this.resolution/(2*i)},Po.prototype.fastDeltaWithOwnCommunity=function(t,e,r,n){var i=this.M,o=this.totalWeights[n];return r-(e+=this.loops[t])*(o-e)*this.resolution/(2*i)},Po.prototype.bounds=function(t){return[this.starts[t],this.starts[t+1]]},Po.prototype.project=function(){var t=this,e={};return t.nodes.slice(0,this.C).forEach((function(r,n){e[r]=Array.from(t.neighborhood.slice(t.starts[n],t.starts[n+1])).map((function(e){return t.nodes[e]}))})),e},Po.prototype.collect=function(t){arguments.length<1&&(t=this.level);var e,r,n={},i=this.keepDendrogram?this.dendrogram[t]:this.mapping;for(e=0,r=i.length;e<r;e++)n[this.nodes[e]]=i[e];return n},Po.prototype.assign=function(t,e){arguments.length<2&&(e=this.level);var r,n,i=this.keepDendrogram?this.dendrogram[e]:this.mapping;for(r=0,n=i.length;r<n;r++)this.graph.setNodeAttribute(this.nodes[r],t,i[r])},Po.prototype[_o]=function(){var t={};Object.defineProperty(t,"constructor",{value:Po,enumerable:!1}),t.C=this.C,t.M=this.M,t.E=this.E,t.U=this.U,t.resolution=this.resolution,t.level=this.level,t.nodes=this.nodes,t.starts=this.starts.slice(0,t.C+1);var e=this;return["neighborhood","weights"].forEach((function(r){t[r]=e[r].slice(0,t.E)})),["counts","loops","belongings","totalWeights"].forEach((function(r){t[r]=e[r].slice(0,t.C)})),t.unused=this.unused.slice(0,this.U),this.keepDendrogram?t.dendrogram=this.dendrogram:t.mapping=this.mapping,t},Io.prototype.bounds=Po.prototype.bounds,Io.prototype.inBounds=function(t){return[this.offsets[t],this.starts[t+1]]},Io.prototype.outBounds=function(t){return[this.starts[t],this.offsets[t]]},Io.prototype.project=Po.prototype.project,Io.prototype.projectIn=function(){var t=this,e={};return t.nodes.slice(0,this.C).forEach((function(r,n){e[r]=Array.from(t.neighborhood.slice(t.offsets[n],t.starts[n+1])).map((function(e){return t.nodes[e]}))})),e},Io.prototype.projectOut=function(){var t=this,e={};return t.nodes.slice(0,this.C).forEach((function(r,n){e[r]=Array.from(t.neighborhood.slice(t.starts[n],t.offsets[n])).map((function(e){return t.nodes[e]}))})),e},Io.prototype.isolate=function(t,e,r){var n=this.belongings[t];if(1===this.counts[n])return n;var i=this.unused[--this.U],o=this.loops[t];return this.totalInWeights[n]-=e+o,this.totalInWeights[i]+=e+o,this.totalOutWeights[n]-=r+o,this.totalOutWeights[i]+=r+o,this.belongings[t]=i,this.counts[n]--,this.counts[i]++,i},Io.prototype.move=function(t,e,r,n){var i=this.belongings[t],o=this.loops[t];this.totalInWeights[i]-=e+o,this.totalInWeights[n]+=e+o,this.totalOutWeights[i]-=r+o,this.totalOutWeights[n]+=r+o,this.belongings[t]=n;var a=1==this.counts[i]--;this.counts[n]++,a&&(this.unused[this.U++]=i)},Io.prototype.computeNodeInDegree=function(t){var e,r,n=0;for(e=this.offsets[t],r=this.starts[t+1];e<r;e++)n+=this.weights[e];return n},Io.prototype.computeNodeOutDegree=function(t){var e,r,n=0;for(e=this.starts[t],r=this.offsets[t];e<r;e++)n+=this.weights[e];return n},Io.prototype.expensiveMove=function(t,e){var r=this.computeNodeInDegree(t),n=this.computeNodeOutDegree(t);this.move(t,r,n,e)},Io.prototype.zoomOut=function(){var t,e,r,n,i,o,a,s,u,h,c,l,d,f,g,p=new Array(this.C-this.U),y={},v=this.nodes.length,m=0,b=0;for(t=0,r=this.C;t<r;t++)(o=this.belongings[t])in y||(y[o]=m,p[m]={inAdj:{},outAdj:{},totalInWeights:this.totalInWeights[o],totalOutWeights:this.totalOutWeights[o],internalWeights:0},m++),this.belongings[t]=y[o];if(this.keepDendrogram){for(f=this.dendrogram[this.level],g=new(Oo.getPointerArray(m))(v),t=0;t<v;t++)g[t]=this.belongings[f[t]];this.dendrogram.push(g)}else for(t=0;t<v;t++)this.mapping[t]=this.belongings[this.mapping[t]];for(t=0,r=this.C;t<r;t++)for(o=this.belongings[t],u=this.offsets[t],l=(s=p[o]).inAdj,d=s.outAdj,s.internalWeights+=this.loops[t],e=this.starts[t],n=this.starts[t+1];e<n;e++)i=this.neighborhood[e],c=(h=e<u)?d:l,o!==(a=this.belongings[i])?(a in c||(c[a]=0),c[a]+=this.weights[e]):h&&(s.internalWeights+=this.weights[e]);for(this.C=m,i=0,o=0;o<m;o++){for(a in l=(s=p[o]).inAdj,d=s.outAdj,o=+o,this.totalInWeights[o]=s.totalInWeights,this.totalOutWeights[o]=s.totalOutWeights,this.loops[o]=s.internalWeights,this.counts[o]=1,this.starts[o]=i,this.belongings[o]=o,d)this.neighborhood[i]=+a,this.weights[i]=d[a],b++,i++;for(a in this.offsets[o]=i,l)this.neighborhood[i]=+a,this.weights[i]=l[a],b++,i++}return this.starts[m]=b,this.E=b,this.U=0,this.level++,y},Io.prototype.modularity=function(){var t,e,r,n,i=0,o=this.M,a=new Float64Array(this.C);for(e=0;e<this.C;e++)for(a[t=this.belongings[e]]+=this.loops[e],r=this.starts[e],n=this.offsets[e];r<n;r++)t===this.belongings[this.neighborhood[r]]&&(a[t]+=this.weights[r]);for(e=0;e<this.C;e++)i+=a[e]/o-this.totalInWeights[e]*this.totalOutWeights[e]/Math.pow(o,2)*this.resolution;return i},Io.prototype.delta=function(t,e,r,n,i){var o=this.M,a=this.totalInWeights[i],s=this.totalOutWeights[i],u=this.loops[t];return n/o-((r+=u)*a+(e+=u)*s)*this.resolution/(o*o)},Io.prototype.deltaWithOwnCommunity=function(t,e,r,n,i){var o=this.M,a=this.totalInWeights[i],s=this.totalOutWeights[i],u=this.loops[t];return n/o-((r+=u)*(a-(e+=u))+e*(s-r))*this.resolution/(o*o)},Io.prototype.collect=Po.prototype.collect,Io.prototype.assign=Po.prototype.assign,Io.prototype[_o]=function(){var t={};Object.defineProperty(t,"constructor",{value:Io,enumerable:!1}),t.C=this.C,t.M=this.M,t.E=this.E,t.U=this.U,t.resolution=this.resolution,t.level=this.level,t.nodes=this.nodes,t.starts=this.starts.slice(0,t.C+1);var e=this;return["neighborhood","weights"].forEach((function(r){t[r]=e[r].slice(0,t.E)})),["counts","offsets","loops","belongings","totalInWeights","totalOutWeights"].forEach((function(r){t[r]=e[r].slice(0,t.C)})),t.unused=this.unused.slice(0,this.U),this.keepDendrogram?t.dendrogram=this.dendrogram:t.mapping=this.mapping,t},Do.UndirectedLouvainIndex=Po,Do.DirectedLouvainIndex=Io;var Uo=uo,To=ho,Lo=lo,Fo=bo,qo=xo,Ro=ko.createRandomIndex,Bo=Do,Go=Bo.UndirectedLouvainIndex,Ho=Bo.DirectedLouvainIndex,Ko={attributes:{community:"community",weight:"weight"},fastLocalMoves:!0,randomWalk:!0,resolution:1,rng:Math.random,weighted:!1};function Vo(t,e,r){var n=t.get(e);void 0===n&&(n=0),n+=r,t.set(e,n)}function $o(t,e,r,n,i){return Math.abs(n-i)<1e-10?t!==e&&r>t:n>i}function Jo(t,e,r){var n,i,o,a,s,u,h,c,l,d,f,g,p,y,v,m,b,w,E,A=new Go(e,{attributes:{weight:r.attributes.weight},keepDendrogram:t,resolution:r.resolution,weighted:r.weighted}),x=Ro(r.rng),z=!0,M=!0,S=new Fo(Float64Array,A.C),N=0,j=0,k=[];for(r.fastLocalMoves&&(o=new qo(A.C));z;){if(g=A.C,z=!1,M=!0,r.fastLocalMoves){for(E=0,c=r.randomWalk?x(g):0,l=0;l<g;l++,c++)d=c%g,o.enqueue(d);for(;0!==o.size;){for(d=o.dequeue(),j++,p=0,S.clear(),n=A.belongings[d],a=A.starts[d],s=A.starts[d+1];a<s;a++)f=A.neighborhood[a],p+=u=A.weights[a],Vo(S,i=A.belongings[f],u);for(m=A.fastDeltaWithOwnCommunity(d,p,S.get(n)||0,n),v=n,h=0;h<S.size;h++)(i=S.dense[h])!==n&&(y=S.vals[h],N++,$o(v,n,i,b=A.fastDelta(d,p,y,i),m)&&(m=b,v=i));if(m<0){if((v=A.isolate(d,p))===n)continue}else{if(v===n)continue;A.move(d,p,v)}for(z=!0,E++,a=A.starts[d],s=A.starts[d+1];a<s;a++)f=A.neighborhood[a],(i=A.belongings[f])!==v&&o.enqueue(f)}k.push(E)}else for(w=[],k.push(w);M;){for(M=!1,E=0,c=r.randomWalk?x(g):0,l=0;l<g;l++,c++){for(d=c%g,j++,p=0,S.clear(),n=A.belongings[d],a=A.starts[d],s=A.starts[d+1];a<s;a++)f=A.neighborhood[a],p+=u=A.weights[a],Vo(S,i=A.belongings[f],u);for(m=A.fastDeltaWithOwnCommunity(d,p,S.get(n)||0,n),v=n,h=0;h<S.size;h++)(i=S.dense[h])!==n&&(y=S.vals[h],N++,$o(v,n,i,b=A.fastDelta(d,p,y,i),m)&&(m=b,v=i));if(m<0){if((v=A.isolate(d,p))===n)continue}else{if(v===n)continue;A.move(d,p,v)}M=!0,E++}w.push(E),z=M||z}z&&A.zoomOut()}return{index:A,deltaComputations:N,nodesVisited:j,moves:k}}function Qo(t,e,r){var n,i,o,a,s,u,h,c,l,d,f,g,p,y,v,m,b,w,E,A,x,z,M=new Ho(e,{attributes:{weight:r.attributes.weight},keepDendrogram:t,resolution:r.resolution,weighted:r.weighted}),S=Ro(r.rng),N=!0,j=!0,k=new Fo(Float64Array,M.C),D=0,O=0,_=[];for(r.fastLocalMoves&&(o=new qo(M.C));N;){if(y=M.C,N=!1,j=!0,r.fastLocalMoves){for(z=0,d=r.randomWalk?S(y):0,f=0;f<y;f++,d++)g=d%y,o.enqueue(g);for(;0!==o.size;){for(g=o.dequeue(),O++,v=0,m=0,k.clear(),n=M.belongings[g],a=M.starts[g],s=M.starts[g+1],u=M.offsets[g];a<s;a++)h=a<u,p=M.neighborhood[a],c=M.weights[a],h?m+=c:v+=c,Vo(k,i=M.belongings[p],c);for(E=M.deltaWithOwnCommunity(g,v,m,k.get(n)||0,n),w=n,l=0;l<k.size;l++)(i=k.dense[l])!==n&&(b=k.vals[l],D++,$o(w,n,i,A=M.delta(g,v,m,b,i),E)&&(E=A,w=i));if(E<0){if((w=M.isolate(g,v,m))===n)continue}else{if(w===n)continue;M.move(g,v,m,w)}for(N=!0,z++,a=M.starts[g],s=M.starts[g+1];a<s;a++)p=M.neighborhood[a],(i=M.belongings[p])!==w&&o.enqueue(p)}_.push(z)}else for(x=[],_.push(x);j;){for(j=!1,z=0,d=r.randomWalk?S(y):0,f=0;f<y;f++,d++){for(g=d%y,O++,v=0,m=0,k.clear(),n=M.belongings[g],a=M.starts[g],s=M.starts[g+1],u=M.offsets[g];a<s;a++)h=a<u,p=M.neighborhood[a],c=M.weights[a],h?m+=c:v+=c,Vo(k,i=M.belongings[p],c);for(E=M.deltaWithOwnCommunity(g,v,m,k.get(n)||0,n),w=n,l=0;l<k.size;l++)(i=k.dense[l])!==n&&(b=k.vals[l],D++,$o(w,n,i,A=M.delta(g,v,m,b,i),E)&&(E=A,w=i));if(E<0){if((w=M.isolate(g,v,m))===n)continue}else{if(w===n)continue;M.move(g,v,m,w)}j=!0,z++}x.push(z),N=j||N}N&&M.zoomOut()}return{index:M,deltaComputations:D,nodesVisited:O,moves:_}}function Xo(t,e,r,n){if(!To(r))throw new Error("graphology-communities-louvain: the given graph is not a valid graphology instance.");var i=Lo(r);if("mixed"===i)throw new Error("graphology-communities-louvain: cannot run the algorithm on a true mixed graph.");n=Uo({},n,Ko);var o=0;if(0===r.size){if(t)return void r.forEachNode((function(t){r.setNodeAttribute(t,n.attributes.communities,o++)}));var a={};return r.forEachNode((function(t){a[t]=o++})),e?{communities:a,count:r.order,deltaComputations:0,dendrogram:null,level:0,modularity:NaN,moves:null,nodesVisited:0,resolution:n.resolution}:a}var s=("undirected"===i?Jo:Qo)(e,r,n),u=s.index;if(!e)return t?void u.assign(n.attributes.community):u.collect();var h={count:u.C,deltaComputations:s.deltaComputations,dendrogram:u.dendrogram,level:u.level,modularity:u.modularity(),moves:s.moves,nodesVisited:s.nodesVisited,resolution:n.resolution};return t?(u.assign(n.attributes.community),h):(h.communities=u.collect(),h)}var Yo=Xo.bind(null,!1,!1);Yo.assign=Xo.bind(null,!0,!1),Yo.detailed=Xo.bind(null,!1,!0),Yo.defaults=Ko;var Zo=Yo,ta={},ea=function(t,e){var r=e.length;if(0!==r){var n=t.length;t.length+=r;for(var i=0;i<r;i++)t[n+i]=e[i]}},ra=function(t){return null!==t&&"object"==typeof t&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode&&"boolean"==typeof t.multi},na=ea;ta.connectedComponents=function(t){if(!ra(t))throw new Error("graphology-components: the given graph is not a valid graphology instance.");if(!t.order)return[];if(!t.size)return t.nodes().map((function(t){return[t]}));var e,r,n,i,o,a=new Set,s=[],u=[],h=t.nodes();for(r=0,n=h.length;r<n;r++)if(i=h[r],!a.has(i)){for(e=[],u.push(i);0!==u.length;)o=u.pop(),a.has(o)||(a.add(o),e.push(o),na(u,t.neighbors(o)));s.push(e)}return s},ta.largestConnectedComponent=function(t){if(!ra(t))throw new Error("graphology-components: the given graph is not a valid graphology instance.");if(!t.order)return[];if(!t.size)return[t.nodes()[0]];var e,r,n,i,o,a,s=t.order,u=new Set,h=[],c=[],l=t.nodes();for(n=0,i=l.length;n<i;n++)if(o=l[n],!u.has(o)){for(r=[],c.push(o);0!==c.length;)a=c.pop(),u.has(a)||(u.add(a),r.push(a),na(c,t.neighbors(a)));if(r.length>h.length&&(h=r),e=s-u.size,h.length>e)return h}return h},ta.stronglyConnectedComponents=function(t){if(!ra(t))throw new Error("graphology-components: the given graph is not a valid graphology instance.");if(!t.order)return[];if("undirected"===t.type)throw new Error("graphology-components: the given graph is undirected");var e,r,n=t.nodes(),i=[];if(!t.size){for(e=0,r=n.length;e<r;e++)i.push([n[e]]);return i}var o,a,s,u=1,h=[],c=[],l=new Map,d=new Set,f=function(e){var r,n,s=t.outNeighbors(e).concat(t.undirectedNeighbors(e));l.set(e,u++),h.push(e),c.push(e);for(var g=0,p=s.length;g<p;g++)if(r=s[g],l.has(r)){if(n=l.get(r),!d.has(r))for(;l.get(h[h.length-1])>n;)h.pop()}else f(r);if(l.get(h[h.length-1])===l.get(e)){o=[];do{a=c.pop(),o.push(a),d.add(a)}while(a!==e);i.push(o),h.pop()}};for(e=0,r=n.length;e<r;e++)s=n[e],d.has(s)||f(s);return i};var ia=ta,oa={},aa={},sa=function(t){return null!==t&&"function"==typeof t&&"object"==typeof t.prototype&&"function"==typeof t.prototype.addUndirectedEdgeWithKey&&"function"==typeof t.prototype.dropNode},ua=sa,ha=sa,ca=function(t,e){if(!ha(t))throw new Error("graphology-generators/classic/empty: invalid Graph constructor.");var r,n=new t;for(r=0;r<e;r++)n.addNode(r);return n},la=function(t,e){var r,n,i,o;if(0!==e.length)for(t.mergeNode(e[0]),i=1,o=e.length;i<o;i++)r=e[i-1],n=e[i],t.mergeEdge(r,n)},da=Math.ceil,fa=Math.max;var ga=/\s/;var pa=function(t){for(var e=t.length;e--&&ga.test(t.charAt(e)););return e},ya=/^\s+/;var va=F,ma=ke;var ba=function(t){return t?t.slice(0,pa(t)+1).replace(ya,""):t},wa=q,Ea=function(t){return"symbol"==typeof t||ma(t)&&"[object Symbol]"==va(t)},Aa=/^[-+]0x[0-9a-f]+$/i,xa=/^0b[01]+$/i,za=/^0o[0-7]+$/i,Ma=parseInt;var Sa=function(t){if("number"==typeof t)return t;if(Ea(t))return NaN;if(wa(t)){var e="function"==typeof t.valueOf?t.valueOf():t;t=wa(e)?e+"":e}if("string"!=typeof t)return 0===t?t:+t;t=ba(t);var r=xa.test(t);return r||za.test(t)?Ma(t.slice(2),r?2:8):Aa.test(t)?NaN:+t},Na=1/0;var ja=function(t,e,r,n){for(var i=-1,o=fa(da((e-t)/(r||1)),0),a=Array(o);o--;)a[n?o:++i]=t,t+=r;return a},ka=to,Da=function(t){return t?(t=Sa(t))===Na||t===-1/0?17976931348623157e292*(t<0?-1:1):t==t?t:0:0===t?t:0};var Oa=function(t){return function(e,r,n){return n&&"number"!=typeof n&&ka(e,r,n)&&(r=n=void 0),e=Da(e),void 0===r?(r=e,e=0):r=Da(r),n=void 0===n?e<r?1:-1:Da(n),ja(e,r,n,t)}}(),_a=sa,Ca=la,Wa=Oa,Pa=sa,Ia=la,Ua=Oa;aa.complete=function(t,e){if(!ua(t))throw new Error("graphology-generators/classic/complete: invalid Graph constructor.");var r,n,i=new t;for(r=0;r<e;r++)i.addNode(r);for(r=0;r<e;r++)for(n=r+1;n<e;n++)"directed"!==i.type&&i.addUndirectedEdge(r,n),"undirected"!==i.type&&(i.addDirectedEdge(r,n),i.addDirectedEdge(n,r));return i},aa.empty=ca,aa.ladder=function(t,e){if(!_a(t))throw new Error("graphology-generators/classic/ladder: invalid Graph constructor.");var r=new t;Ca(r,Wa(e)),Ca(r,Wa(e,2*e));for(var n=0;n<e;n++)r.addEdge(n,n+e);return r},aa.path=function(t,e){if(!Pa(t))throw new Error("graphology-generators/classic/path: invalid Graph constructor.");var r=new t;return Ia(r,Ua(e)),r};var Ta={},La=sa,Fa=ca,qa=sa,Ra=ca;Ta.caveman=function(t,e,r){if(!La(t))throw new Error("graphology-generators/community/caveman: invalid Graph constructor.");var n,i,o,a=e*r,s=Fa(t,a);if(r<2)return s;for(n=0;n<a;n+=r)for(i=n;i<n+r;i++)for(o=i+1;o<n+r;o++)s.addEdge(i,o);return s},Ta.connectedCaveman=function(t,e,r){if(!qa(t))throw new Error("graphology-generators/community/connected-caveman: invalid Graph constructor.");var n,i,o,a=e*r,s=Ra(t,a);if(r<2)return s;for(n=0;n<a;n+=r){for(i=n;i<n+r;i++)for(o=i+1;o<n+r;o++)i===n&&i===o-1||s.addEdge(i,o);n>0&&s.addEdge(n,(n-1)%a)}return s.addEdge(0,a-1),s};var Ba={},Ga=sa,Ha=ho;function Ka(t,e){return 2*e/(t*(t-1))}function Va(t,e){return e/(t*(t-1))}function $a(t,e){var r=t*(t-1);return e/(r+r/2)}function Ja(t){var e,r,n=t.nodes(),i=0;for(e=0,r=n.length;e<r;e++)i+=t.outNeighbors(n[e]).length,i+=t.undirectedNeighbors(n[e]).length/2;return i}function Qa(t,e,r){var n,i;if(arguments.length>3){if(i=arguments[3],"number"!=typeof(n=r))throw new Error("graphology-metrics/density: given order is not a number.");if("number"!=typeof i)throw new Error("graphology-metrics/density: given size is not a number.")}else{if(!Ha(r))throw new Error("graphology-metrics/density: given graph is not a valid graphology instance.");n=r.order,i=r.size,r.multi&&!1===e&&(i=Ja(r))}return n<2?0:(null===t&&(t=r.type),null===e&&(e=r.multi),("undirected"===t?Ka:"directed"===t?Va:$a)(n,i))}var Xa=Qa.bind(null,null,null);Xa.directedDensity=Qa.bind(null,"directed",!1),Xa.undirectedDensity=Qa.bind(null,"undirected",!1),Xa.mixedDensity=Qa.bind(null,"mixed",!1),Xa.multiDirectedDensity=Qa.bind(null,"directed",!0),Xa.multiUndirectedDensity=Qa.bind(null,"undirected",!0),Xa.multiMixedDensity=Qa.bind(null,"mixed",!0);var Ya=Xa,Za=sa,ts=Ya;function es(t,e){if(!Za(t))throw new Error("graphology-generators/random/erdos-renyi: invalid Graph constructor.");var r,n,i=e.order,o=e.probability,a=e.rng||Math.random,s=new t;"number"==typeof e.approximateSize&&(o=(0,ts[s.type+"Density"])(i,e.approximateSize));if("number"!=typeof i||i<=0)throw new Error("graphology-generators/random/erdos-renyi: invalid `order`. Should be a positive number.");if("number"!=typeof o||o<0||o>1)throw new Error("graphology-generators/random/erdos-renyi: invalid `probability`. Should be a number between 0 and 1. Or maybe you gave an `approximateSize` exceeding the graph's density.");if("function"!=typeof a)throw new Error("graphology-generators/random/erdos-renyi: invalid `rng`. Should be a function.");for(r=0;r<i;r++)s.addNode(r);if(o<=0)return s;for(r=0;r<i;r++)for(n=r+1;n<i;n++)"directed"!==s.type&&a()<o&&s.addUndirectedEdge(r,n),"undirected"!==s.type&&(a()<o&&s.addDirectedEdge(r,n),a()<o&&s.addDirectedEdge(n,r));return s}es.sparse=function(t,e){if(!Za(t))throw new Error("graphology-generators/random/erdos-renyi: invalid Graph constructor.");var r=e.order,n=e.probability,i=e.rng||Math.random,o=new t;if("number"==typeof e.approximateSize&&(n=(0,ts[o.type+"Density"])(r,e.approximateSize)),"number"!=typeof r||r<=0)throw new Error("graphology-generators/random/erdos-renyi: invalid `order`. Should be a positive number.");if("number"!=typeof n||n<0||n>1)throw new Error("graphology-generators/random/erdos-renyi: invalid `probability`. Should be a number between 0 and 1. Or maybe you gave an `approximateSize` exceeding the graph's density.");if("function"!=typeof i)throw new Error("graphology-generators/random/erdos-renyi: invalid `rng`. Should be a function.");for(var a=0;a<r;a++)o.addNode(a);if(n<=0)return o;var s,u=-1,h=Math.log(1-n);if("undirected"!==o.type)for(s=0;s<r;){for(s===(u+=1+(Math.log(1-i())/h|0))&&u++;s<r&&r<=u;)++s===(u-=r)&&u++;s<r&&o.addDirectedEdge(s,u)}if(u=-1,"directed"!==o.type)for(s=1;s<r;){for(u+=1+(Math.log(1-i())/h|0);u>=s&&s<r;)u-=s,s++;s<r&&o.addUndirectedEdge(s,u)}return o};var rs=es,ns=sa;Ba.clusters=function(t,e){if(!Ga(t))throw new Error("graphology-generators/random/clusters: invalid Graph constructor.");var r="clusterDensity"in(e=e||{})?e.clusterDensity:.5,n=e.rng||Math.random,i=e.order,o=e.size,a=e.clusters;if("number"!=typeof r||r>1||r<0)throw new Error("graphology-generators/random/clusters: `clusterDensity` option should be a number between 0 and 1.");if("function"!=typeof n)throw new Error("graphology-generators/random/clusters: `rng` option should be a function.");if("number"!=typeof i||i<=0)throw new Error("graphology-generators/random/clusters: `order` option should be a positive number.");if("number"!=typeof o||o<=0)throw new Error("graphology-generators/random/clusters: `size` option should be a positive number.");if("number"!=typeof a||a<=0)throw new Error("graphology-generators/random/clusters: `clusters` option should be a positive number.");var s=new t;if(!i)return s;var u,h,c,l,d,f,g=new Array(a);for(c=0;c<a;c++)g[c]=[];for(c=0;c<i;c++)u=n()*a|0,s.addNode(c,{cluster:u}),g[u].push(c);if(!o)return s;for(c=0;c<o;c++){if(n()<1-r){l=n()*i|0;do{d=n()*i|0}while(l===d)}else{if(!(f=(h=g[u=n()*a|0]).length)||f<2)continue;l=h[n()*f|0];do{d=h[n()*f|0]}while(l===d)}s.multi?s.addEdge(l,d):s.mergeEdge(l,d)}return s},Ba.erdosRenyi=rs,Ba.girvanNewman=function(t,e){if(!ns(t))throw new Error("graphology-generators/random/girvan-newman: invalid Graph constructor.");var r=e.zOut,n=e.rng||Math.random;if("number"!=typeof r)throw new Error("graphology-generators/random/girvan-newman: invalid `zOut`. Should be a number.");if("function"!=typeof n)throw new Error("graphology-generators/random/girvan-newman: invalid `rng`. Should be a function.");var i,o,a,s=r/96,u=(16-96*s)/31,h=new t;for(o=0;o<128;o++)h.addNode(o);for(o=0;o<128;o++)for(a=o+1;a<128;a++)i=n(),o%4==a%4?i<u&&h.addEdge(o,a):i<s&&h.addEdge(o,a);return h};var is={},os=sa,as=function(t,e){if(0!==e.length){var r,n,i,o=e[0];for(t.mergeNode(o),n=1,i=e.length;n<i;n++)r=e[n],t.mergeEdge(o,r)}},ss=[["Andre","Beverley","Carol","Diane","Fernando"],["Beverley","Andre","Ed","Garth"],["Carol","Andre","Diane","Fernando"],["Diane","Andre","Beverley","Carol","Ed","Fernando","Garth"],["Ed","Beverley","Diane","Garth"],["Fernando","Andre","Carol","Diane","Garth","Heather"],["Garth","Beverley","Diane","Ed","Fernando","Heather"],["Heather","Fernando","Garth","Ike"],["Ike","Heather","Jane"],["Jane","Ike"]];is.krackhardtKite=function(t){if(!os(t))throw new Error("graphology-generators/social/krackhardt-kite: invalid Graph constructor.");var e,r,n=new t;for(e=0,r=ss.length;e<r;e++)as(n,ss[e]);return n};var us={},hs=sa,cs=[["Acciaiuoli","Medici"],["Castellani","Peruzzi"],["Castellani","Strozzi"],["Castellani","Barbadori"],["Medici","Barbadori"],["Medici","Ridolfi"],["Medici","Tornabuoni"],["Medici","Albizzi"],["Medici","Salviati"],["Salviati","Pazzi"],["Peruzzi","Strozzi"],["Peruzzi","Bischeri"],["Strozzi","Ridolfi"],["Strozzi","Bischeri"],["Ridolfi","Tornabuoni"],["Tornabuoni","Guadagni"],["Albizzi","Ginori"],["Albizzi","Guadagni"],["Bischeri","Guadagni"],["Guadagni","Lamberteschi"]],ls=sa,ds=["0111111110111100010101000000000100","1011000100000100010101000000001000","1101000111000100000000000001100010","1110000100001100000000000000000000","1000001000100000000000000000000000","1000001000100000100000000000000000","1000110000000000100000000000000000","1111000000000000000000000000000000","1010000000000000000000000000001011","0010000000000000000000000000000001","1000110000000000000000000000000000","1000000000000000000000000000000000","1001000000000000000000000000000000","1111000000000000000000000000000001","0000000000000000000000000000000011","0000000000000000000000000000000011","0000011000000000000000000000000000","1100000000000000000000000000000000","0000000000000000000000000000000011","1100000000000000000000000000000001","0000000000000000000000000000000011","1100000000000000000000000000000000","0000000000000000000000000000000011","0000000000000000000000000101010011","0000000000000000000000000101000100","0000000000000000000000011000000100","0000000000000000000000000000010001","0010000000000000000000011000000001","0010000000000000000000000000000101","0000000000000000000000010010000011","0100000010000000000000000000000011","1000000000000000000000001100100011","0010000010000011001010110000011101","0000000011000111001110110011111110"],fs=new Set([0,1,2,3,4,5,6,7,8,10,11,12,13,16,17,19,21]);us.florentineFamilies=function(t){if(!hs(t))throw new Error("graphology-generators/social/florentine-families: invalid Graph constructor.");var e,r,n,i=new t;for(r=0,n=cs.length;r<n;r++)e=cs[r],i.mergeEdge(e[0],e[1]);return i},us.karateClub=function(t){if(!ls(t))throw new Error("graphology-generators/social/karate: invalid Graph constructor.");for(var e,r,n,i,o,a,s=new t,u=0;u<34;u++)e=fs.has(u)?"Mr. Hi":"Officer",s.addNode(u,{club:e});for(n=0,o=ds.length;n<o;n++)for(i=n+1,a=(r=ds[n].split("")).length;i<a;i++)+r[i]&&s.addEdgeWithKey(n+"->"+i,n,i);return s},oa.classic=aa,oa.community=Ta,oa.random=Ba,oa.small=is,oa.social=us;var gs=oa,ps=uo,ys=ho,vs={attributes:{authority:"authority",hub:"hub",weight:"weight"},maxIterations:100,normalize:!0,tolerance:1e-8};function ms(t,e){var r,n,i=Object.create(null);for(r=0,n=t.length;r<n;r++)i[t[r]]=e;return i}function bs(t){var e=0;for(var r in t)e+=t[r];return e}function ws(t,e,r){if(!ys(e))throw new Error("graphology-hits: the given graph is not a valid graphology instance.");if(e.multi)throw new Error("graphology-hits: the HITS algorithm does not work with MultiGraphs.");r=ps(r,vs);var n,i,o,a,s,u,h,c,l,d,f,g,p,y=e.order,v=e.size,m=e.nodes(),b=e.edges(),w=ms(m,1/y),E={},A=!1;for(f=0;f<v;f++)E[s=b[f]]=e.getEdgeAttribute(s,r.attributes.weight)||1;for(u=0;u<r.maxIterations;u++){for(n=w,w=ms(m,0),i=ms(m,0),c=0,h=0,f=0;f<y;f++)for(o=m[f],g=0,p=(b=e.outEdges(o).concat(e.undirectedEdges(o))).length;g<p;g++)s=b[g],i[a=e.opposite(o,s)]+=n[o]*E[s],i[a]>h&&(h=i[a]);for(f=0;f<y;f++)for(o=m[f],g=0,p=(b=e.outEdges(o).concat(e.undirectedEdges(o))).length;g<p;g++)s=b[g],a=e.opposite(o,s),w[o]+=i[a]*E[s],w[a]>c&&(c=w[a]);for(o in d=1/c,w)w[o]*=d;for(o in d=1/h,i)i[o]*=d;for(o in l=0,w)l+=Math.abs(w[o]-n[o]);if(l<r.tolerance){A=!0;break}}if(r.normalize){for(o in d=1/bs(i),i)i[o]*=d;for(o in d=1/bs(w),w)w[o]*=d}if(t)for(f=0;f<y;f++)o=m[f],e.setNodeAttribute(o,r.attributes.authority,i[o]),e.setNodeAttribute(o,r.attributes.hub,w[o]);return{converged:A,hubs:w,authorities:i}}var Es=ws.bind(null,!1);Es.assign=ws.bind(null,!0);var As=Es,xs={},zs=function(t){return null!==t&&"object"==typeof t&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode&&"boolean"==typeof t.multi};function Ms(t){return function(e,r){return e+Math.floor(t()*(r-e+1))}}var Ss=Ms(Math.random);Ss.createRandom=Ms;var Ns=Ss.createRandom;function js(t){var e=Ns(t);return function(t){for(var r=t.length,n=r-1,i=-1;++i<r;){var o=e(i,n),a=t[o];t[o]=t[i],t[i]=a}}}var ks=js(Math.random);ks.createShuffleInPlace=js;var Ds=uo,Os=zs,_s=ks,Cs={attributes:{x:"x",y:"y"},center:0,hierarchyAttributes:[],rng:Math.random,scale:1};function Ws(t,e,r,n,i){this.wrappedCircle=i||null,this.children={},this.countChildren=0,this.id=t||null,this.next=null,this.previous=null,this.x=e||null,this.y=r||null,this.r=i?1010101:n||999}function Ps(t,e,r){for(var n in e.children){var i=e.children[n];i.hasChildren()?Ps(t,i,r):r[i.id]={x:i.x,y:i.y}}}function Is(t,e){var r=t.r-e.r,n=e.x-t.x,i=e.y-t.y;return r<0||r*r<n*n+i*i}function Us(t,e){var r=t.r-e.r+1e-6,n=e.x-t.x,i=e.y-t.y;return r>0&&r*r>n*n+i*i}function Ts(t,e){for(var r=0;r<e.length;++r)if(!Us(t,e[r]))return!1;return!0}function Ls(t,e){var r=t.x,n=t.y,i=t.r,o=e.x,a=e.y,s=e.r,u=o-r,h=a-n,c=s-i,l=Math.sqrt(u*u+h*h);return new Ws(null,(r+o+u/l*c)/2,(n+a+h/l*c)/2,(l+i+s)/2)}function Fs(t,e,r){var n=t.x,i=t.y,o=t.r,a=e.x,s=e.y,u=e.r,h=r.x,c=r.y,l=r.r,d=n-a,f=n-h,g=i-s,p=i-c,y=u-o,v=l-o,m=n*n+i*i-o*o,b=m-a*a-s*s+u*u,w=m-h*h-c*c+l*l,E=f*g-d*p,A=(g*w-p*b)/(2*E)-n,x=(p*y-g*v)/E,z=(f*b-d*w)/(2*E)-i,M=(d*v-f*y)/E,S=x*x+M*M-1,N=2*(o+A*x+z*M),j=A*A+z*z-o*o,k=-(S?(N+Math.sqrt(N*N-4*S*j))/(2*S):j/N);return new Ws(null,n+A+x*k,i+z+M*k,k)}function qs(t){switch(t.length){case 1:return new Ws(null,(e=t[0]).x,e.y,e.r);case 2:return Ls(t[0],t[1]);case 3:return Fs(t[0],t[1],t[2]);default:throw new Error("graphology-layout/circlepack: Invalid basis length "+t.length)}var e}function Rs(t,e){var r,n;if(Ts(e,t))return[e];for(r=0;r<t.length;++r)if(Is(e,t[r])&&Ts(Ls(t[r],e),t))return[t[r],e];for(r=0;r<t.length-1;++r)for(n=r+1;n<t.length;++n)if(Is(Ls(t[r],t[n]),e)&&Is(Ls(t[r],e),t[n])&&Is(Ls(t[n],e),t[r])&&Ts(Fs(t[r],t[n],e),t))return[t[r],t[n],e];throw new Error("graphology-layout/circlepack: extendBasis failure !")}function Bs(t){var e=t.wrappedCircle,r=t.next.wrappedCircle,n=e.r+r.r,i=(e.x*r.r+r.x*e.r)/n,o=(e.y*r.r+r.y*e.r)/n;return i*i+o*o}function Gs(t,e,r){var n,i,o,a,s=t.x-e.x,u=t.y-e.y,h=s*s+u*u;h?(i=e.r+r.r,i*=i,a=t.r+r.r,i>(a*=a)?(n=(h+a-i)/(2*h),o=Math.sqrt(Math.max(0,a/h-n*n)),r.x=t.x-n*s-o*u,r.y=t.y-n*u+o*s):(n=(h+i-a)/(2*h),o=Math.sqrt(Math.max(0,i/h-n*n)),r.x=e.x+n*s-o*u,r.y=e.y+n*u+o*s)):(r.x=e.x+r.r,r.y=e.y)}function Hs(t,e){var r=t.r+e.r-1e-6,n=e.x-t.x,i=e.y-t.y;return r>0&&r*r>n*n+i*i}function Ks(t,e){var r,n,i,o,a,s,u,h,c,l,d=t.length;if(0===d)return 0;if((r=t[0]).x=0,r.y=0,d<=1)return r.r;if(n=t[1],r.x=-n.r,n.x=r.r,n.y=0,d<=2)return r.r+n.r;Gs(n,r,i=t[2]),r=new Ws(null,null,null,null,r),n=new Ws(null,null,null,null,n),i=new Ws(null,null,null,null,i),r.next=i.previous=n,n.next=r.previous=i,i.next=n.previous=r;t:for(s=3;s<d;++s){i=t[s],Gs(r.wrappedCircle,n.wrappedCircle,i),i=new Ws(null,null,null,null,i),u=n.next,h=r.previous,c=n.wrappedCircle.r,l=r.wrappedCircle.r;do{if(c<=l){if(Hs(u.wrappedCircle,i.wrappedCircle)){n=u,r.next=n,n.previous=r,--s;continue t}c+=u.wrappedCircle.r,u=u.next}else{if(Hs(h.wrappedCircle,i.wrappedCircle)){(r=h).next=n,n.previous=r,--s;continue t}l+=h.wrappedCircle.r,h=h.previous}}while(u!==h.next);for(i.previous=r,i.next=n,r.next=n.previous=n=i,o=Bs(r);(i=i.next)!==n;)(a=Bs(i))<o&&(r=i,o=a);n=r.next}r=[n.wrappedCircle],i=n;for(var f=1e4;(i=i.next)!==n&&0!=--f;)r.push(i.wrappedCircle);for(i=function(t,e){var r,n,i=0,o=t.slice(),a=t.length,s=[];for(e(o);i<a;)r=o[i],n&&Us(n,r)?++i:(n=qs(s=Rs(s,r)),i=0);return n}(r,e),s=0;s<d;++s)(r=t[s]).x-=i.x,r.y-=i.y;return i.r}function Vs(t,e){var r=0;if(t.hasChildren()){for(var n in t.children){var i=t.children[n];i.hasChildren()&&(i.r=Vs(i,e))}r=Ks(Object.values(t.children),e)}return r}function $s(t,e){for(var r in Vs(t,e),t.children){t.children[r].applyPositionToChildren()}}function Js(t,e,r){if(!Os(e))throw new Error("graphology-layout/circlepack: the given graph is not a valid graphology instance.");r=Ds(r,Cs);var n={},i={},o=e.nodes(),a=r.center,s=r.hierarchyAttributes,u=_s.createShuffleInPlace(r.rng),h=r.scale,c=new Ws;e.forEachNode((function(t,e){var r=new Ws(t,null,null,e.size?e.size:1),n=c;s.forEach((function(t){var r=e[t];n=n.getChild(r)})),n.addChild(t,r)})),$s(c,u),Ps(e,c,n);var l,d,f,g=o.length;for(f=0;f<g;f++){var p=o[f];l=a+h*n[p].x,d=a+h*n[p].y,i[p]={x:l,y:d},t&&(e.setNodeAttribute(p,r.attributes.x,l),e.setNodeAttribute(p,r.attributes.y,d))}return i}Ws.prototype.hasChildren=function(){return this.countChildren>0},Ws.prototype.addChild=function(t,e){this.children[t]=e,++this.countChildren},Ws.prototype.getChild=function(t){if(!this.children.hasOwnProperty(t)){var e=new Ws;this.children[t]=e,++this.countChildren}return this.children[t]},Ws.prototype.applyPositionToChildren=function(){if(this.hasChildren()){var t=this;for(var e in t.children){var r=t.children[e];r.x+=t.x,r.y+=t.y,r.applyPositionToChildren()}}};var Qs=Js.bind(null,!1);Qs.assign=Js.bind(null,!0);var Xs=Qs,Ys=uo,Zs=zs,tu={attributes:{x:"x",y:"y"},center:.5,scale:1};function eu(t,e,r){if(!Zs(e))throw new Error("graphology-layout/random: the given graph is not a valid graphology instance.");r=Ys(r,tu);var n,i,o,a,s={},u=e.nodes(),h=r.center,c=r.scale,l=2*Math.PI,d=u.length;for(a=0;a<d;a++)n=u[a],i=c*Math.cos(a*l/d),o=c*Math.sin(a*l/d),.5!==h&&(i+=h-.5*c,o+=h-.5*c),s[n]={x:i,y:o},t&&(e.setNodeAttribute(n,r.attributes.x,i),e.setNodeAttribute(n,r.attributes.y,o));return s}var ru=eu.bind(null,!1);ru.assign=eu.bind(null,!0);var nu=ru,iu=uo,ou=zs,au={attributes:{x:"x",y:"y"},center:.5,rng:Math.random,scale:1};function su(t,e,r){if(!ou(e))throw new Error("graphology-layout/random: the given graph is not a valid graphology instance.");r=iu(r,au);var n,i,o,a,s={},u=e.nodes(),h=r.center,c=r.rng,l=r.scale,d=u.length;for(a=0;a<d;a++)n=u[a],i=c()*l,o=c()*l,.5!==h&&(i+=h-.5*l,o+=h-.5*l),s[n]={x:i,y:o},t&&(e.setNodeAttribute(n,r.attributes.x,i),e.setNodeAttribute(n,r.attributes.y,o));return s}var uu=su.bind(null,!1);uu.assign=su.bind(null,!0);var hu=Xs,cu=nu,lu=uu;xs.circlepack=hu,xs.circular=cu,xs.random=lu;var du=xs,fu=10,gu={};gu.assign=function(t){t=t||{};var e,r,n,i=Array.prototype.slice.call(arguments).slice(1);for(e=0,n=i.length;e<n;e++)if(i[e])for(r in i[e])t[r]=i[e][r];return t},gu.validateSettings=function(t){return"linLogMode"in t&&"boolean"!=typeof t.linLogMode?{message:"the `linLogMode` setting should be a boolean."}:"outboundAttractionDistribution"in t&&"boolean"!=typeof t.outboundAttractionDistribution?{message:"the `outboundAttractionDistribution` setting should be a boolean."}:"adjustSizes"in t&&"boolean"!=typeof t.adjustSizes?{message:"the `adjustSizes` setting should be a boolean."}:"edgeWeightInfluence"in t&&"number"!=typeof t.edgeWeightInfluence&&t.edgeWeightInfluence<0?{message:"the `edgeWeightInfluence` setting should be a number >= 0."}:"scalingRatio"in t&&"number"!=typeof t.scalingRatio&&t.scalingRatio<0?{message:"the `scalingRatio` setting should be a number >= 0."}:"strongGravityMode"in t&&"boolean"!=typeof t.strongGravityMode?{message:"the `strongGravityMode` setting should be a boolean."}:"gravity"in t&&"number"!=typeof t.gravity&&t.gravity<0?{message:"the `gravity` setting should be a number >= 0."}:"slowDown"in t&&"number"!=typeof t.slowDown&&t.slowDown<0?{message:"the `slowDown` setting should be a number >= 0."}:"barnesHutOptimize"in t&&"boolean"!=typeof t.barnesHutOptimize?{message:"the `barnesHutOptimize` setting should be a boolean."}:"barnesHutTheta"in t&&"number"!=typeof t.barnesHutTheta&&t.barnesHutTheta<0?{message:"the `barnesHutTheta` setting should be a number >= 0."}:null},gu.graphToByteArrays=function(t){var e,r=t.order,n=t.size,i={},o=new Float32Array(10*r),a=new Float32Array(3*n);return e=0,t.forEachNode((function(r,n){i[r]=e,o[e]=n.x,o[e+1]=n.y,o[e+2]=0,o[e+3]=0,o[e+4]=0,o[e+5]=0,o[e+6]=1+t.degree(r),o[e+7]=1,o[e+8]=n.size||1,o[e+9]=n.fixed?1:0,e+=10})),e=0,t.forEachEdge((function(t,r,n,o){a[e]=i[n],a[e+1]=i[o],a[e+2]=r.weight||0,e+=3})),{nodes:o,edges:a}},gu.assignLayoutChanges=function(t,e){var r=0;t.updateEachNodeAttributes((function(t,n){return n.x=e[r],n.y=e[r+1],r+=10,n}),{attributes:["x","y"]})},gu.collectLayoutChanges=function(t,e){for(var r=t.nodes(),n={},i=0,o=0,a=e.length;i<a;i+=10)n[r[o]]={x:e[i],y:e[i+1]},o++;return n},gu.createWorker=function(t){var e=window.URL||window.webkitURL,r=t.toString(),n=e.createObjectURL(new Blob(["("+r+").call(this);"],{type:"text/javascript"})),i=new Worker(n);return e.revokeObjectURL(n),i};var pu={linLogMode:!1,outboundAttractionDistribution:!1,adjustSizes:!1,edgeWeightInfluence:0,scalingRatio:1,strongGravityMode:!1,gravity:1,slowDown:1,barnesHutOptimize:!1,barnesHutTheta:.5},yu=ho,vu=function(t,e,r){var n,i,o,a,s,u,h,c,l,d,f,g,p,y,v,m,b,w,E,A,x,z,M,S=e.length,N=r.length,j=t.adjustSizes,k=t.barnesHutTheta*t.barnesHutTheta,D=[];for(o=0;o<S;o+=fu)e[o+4]=e[o+2],e[o+5]=e[o+3],e[o+2]=0,e[o+3]=0;if(t.outboundAttractionDistribution){for(f=0,o=0;o<S;o+=fu)f+=e[o+6];f/=S/fu}if(t.barnesHutOptimize){var O,_,C,W=1/0,P=-1/0,I=1/0,U=-1/0;for(o=0;o<S;o+=fu)W=Math.min(W,e[o+0]),P=Math.max(P,e[o+0]),I=Math.min(I,e[o+1]),U=Math.max(U,e[o+1]);var T=P-W,L=U-I;for(T>L?U=(I-=(T-L)/2)+T:P=(W-=(L-T)/2)+L,D[0]=-1,D[1]=(W+P)/2,D[2]=(I+U)/2,D[3]=Math.max(P-W,U-I),D[4]=-1,D[5]=-1,D[6]=0,D[7]=0,D[8]=0,n=1,o=0;o<S;o+=fu)for(i=0,C=3;;){if(!(D[i+5]>=0)){if(D[i+0]<0){D[i+0]=o;break}if(D[i+5]=9*n,c=D[i+3]/2,D[(l=D[i+5])+0]=-1,D[l+1]=D[i+1]-c,D[l+2]=D[i+2]-c,D[l+3]=c,D[l+4]=l+9,D[l+5]=-1,D[l+6]=0,D[l+7]=0,D[l+8]=0,D[(l+=9)+0]=-1,D[l+1]=D[i+1]-c,D[l+2]=D[i+2]+c,D[l+3]=c,D[l+4]=l+9,D[l+5]=-1,D[l+6]=0,D[l+7]=0,D[l+8]=0,D[(l+=9)+0]=-1,D[l+1]=D[i+1]+c,D[l+2]=D[i+2]-c,D[l+3]=c,D[l+4]=l+9,D[l+5]=-1,D[l+6]=0,D[l+7]=0,D[l+8]=0,D[(l+=9)+0]=-1,D[l+1]=D[i+1]+c,D[l+2]=D[i+2]+c,D[l+3]=c,D[l+4]=D[i+4],D[l+5]=-1,D[l+6]=0,D[l+7]=0,D[l+8]=0,n+=4,O=e[D[i+0]+0]<D[i+1]?e[D[i+0]+1]<D[i+2]?D[i+5]:D[i+5]+9:e[D[i+0]+1]<D[i+2]?D[i+5]+18:D[i+5]+27,D[i+6]=e[D[i+0]+6],D[i+7]=e[D[i+0]+0],D[i+8]=e[D[i+0]+1],D[O+0]=D[i+0],D[i+0]=-1,O===(_=e[o+0]<D[i+1]?e[o+1]<D[i+2]?D[i+5]:D[i+5]+9:e[o+1]<D[i+2]?D[i+5]+18:D[i+5]+27)){if(C--){i=O;continue}C=3;break}D[_+0]=o;break}O=e[o+0]<D[i+1]?e[o+1]<D[i+2]?D[i+5]:D[i+5]+9:e[o+1]<D[i+2]?D[i+5]+18:D[i+5]+27,D[i+7]=(D[i+7]*D[i+6]+e[o+0]*e[o+6])/(D[i+6]+e[o+6]),D[i+8]=(D[i+8]*D[i+6]+e[o+1]*e[o+6])/(D[i+6]+e[o+6]),D[i+6]+=e[o+6],i=O}}if(t.barnesHutOptimize){for(g=t.scalingRatio,o=0;o<S;o+=fu)for(i=0;;)if(D[i+5]>=0){if(m=Math.pow(e[o+0]-D[i+7],2)+Math.pow(e[o+1]-D[i+8],2),4*(d=D[i+3])*d/m<k){if(p=e[o+0]-D[i+7],y=e[o+1]-D[i+8],!0===j?m>0?(b=g*e[o+6]*D[i+6]/m,e[o+2]+=p*b,e[o+3]+=y*b):m<0&&(b=-g*e[o+6]*D[i+6]/Math.sqrt(m),e[o+2]+=p*b,e[o+3]+=y*b):m>0&&(b=g*e[o+6]*D[i+6]/m,e[o+2]+=p*b,e[o+3]+=y*b),(i=D[i+4])<0)break;continue}i=D[i+5]}else if((u=D[i+0])>=0&&u!==o&&(m=(p=e[o+0]-e[u+0])*p+(y=e[o+1]-e[u+1])*y,!0===j?m>0?(b=g*e[o+6]*e[u+6]/m,e[o+2]+=p*b,e[o+3]+=y*b):m<0&&(b=-g*e[o+6]*e[u+6]/Math.sqrt(m),e[o+2]+=p*b,e[o+3]+=y*b):m>0&&(b=g*e[o+6]*e[u+6]/m,e[o+2]+=p*b,e[o+3]+=y*b)),(i=D[i+4])<0)break}else for(g=t.scalingRatio,a=0;a<S;a+=fu)for(s=0;s<a;s+=fu)p=e[a+0]-e[s+0],y=e[a+1]-e[s+1],!0===j?(m=Math.sqrt(p*p+y*y)-e[a+8]-e[s+8])>0?(b=g*e[a+6]*e[s+6]/m/m,e[a+2]+=p*b,e[a+3]+=y*b,e[s+2]+=p*b,e[s+3]+=y*b):m<0&&(b=100*g*e[a+6]*e[s+6],e[a+2]+=p*b,e[a+3]+=y*b,e[s+2]-=p*b,e[s+3]-=y*b):(m=Math.sqrt(p*p+y*y))>0&&(b=g*e[a+6]*e[s+6]/m/m,e[a+2]+=p*b,e[a+3]+=y*b,e[s+2]-=p*b,e[s+3]-=y*b);for(l=t.gravity/t.scalingRatio,g=t.scalingRatio,o=0;o<S;o+=fu)b=0,p=e[o+0],y=e[o+1],m=Math.sqrt(Math.pow(p,2)+Math.pow(y,2)),t.strongGravityMode?m>0&&(b=g*e[o+6]*l):m>0&&(b=g*e[o+6]*l/m),e[o+2]-=p*b,e[o+3]-=y*b;for(g=1*(t.outboundAttractionDistribution?f:1),h=0;h<N;h+=3)a=r[h+0],s=r[h+1],c=r[h+2],v=Math.pow(c,t.edgeWeightInfluence),p=e[a+0]-e[s+0],y=e[a+1]-e[s+1],!0===j?(m=Math.sqrt(Math.pow(p,2)+Math.pow(y,2)-e[a+8]-e[s+8]),t.linLogMode?t.outboundAttractionDistribution?m>0&&(b=-g*v*Math.log(1+m)/m/e[a+6]):m>0&&(b=-g*v*Math.log(1+m)/m):t.outboundAttractionDistribution?m>0&&(b=-g*v/e[a+6]):m>0&&(b=-g*v)):(m=Math.sqrt(Math.pow(p,2)+Math.pow(y,2)),t.linLogMode?t.outboundAttractionDistribution?m>0&&(b=-g*v*Math.log(1+m)/m/e[a+6]):m>0&&(b=-g*v*Math.log(1+m)/m):t.outboundAttractionDistribution?(m=1,b=-g*v/e[a+6]):(m=1,b=-g*v)),m>0&&(e[a+2]+=p*b,e[a+3]+=y*b,e[s+2]-=p*b,e[s+3]-=y*b);if(!0===j)for(o=0;o<S;o+=fu)1!==e[o+9]&&((w=Math.sqrt(Math.pow(e[o+2],2)+Math.pow(e[o+3],2)))>10&&(e[o+2]=10*e[o+2]/w,e[o+3]=10*e[o+3]/w),E=e[o+6]*Math.sqrt((e[o+4]-e[o+2])*(e[o+4]-e[o+2])+(e[o+5]-e[o+3])*(e[o+5]-e[o+3])),A=Math.sqrt((e[o+4]+e[o+2])*(e[o+4]+e[o+2])+(e[o+5]+e[o+3])*(e[o+5]+e[o+3]))/2,x=.1*Math.log(1+A)/(1+Math.sqrt(E)),z=e[o+0]+e[o+2]*(x/t.slowDown),e[o+0]=z,M=e[o+1]+e[o+3]*(x/t.slowDown),e[o+1]=M);else for(o=0;o<S;o+=fu)1!==e[o+9]&&(E=e[o+6]*Math.sqrt((e[o+4]-e[o+2])*(e[o+4]-e[o+2])+(e[o+5]-e[o+3])*(e[o+5]-e[o+3])),A=Math.sqrt((e[o+4]+e[o+2])*(e[o+4]+e[o+2])+(e[o+5]+e[o+3])*(e[o+5]+e[o+3]))/2,x=e[o+7]*Math.log(1+A)/(1+Math.sqrt(E)),e[o+7]=Math.min(1,Math.sqrt(x*(Math.pow(e[o+2],2)+Math.pow(e[o+3],2))/(1+Math.sqrt(E)))),z=e[o+0]+e[o+2]*(x/t.slowDown),e[o+0]=z,M=e[o+1]+e[o+3]*(x/t.slowDown),e[o+1]=M);return{}},mu=gu,bu=pu;function wu(t,e,r){if(!yu(e))throw new Error("graphology-layout-forceatlas2: the given graph is not a valid graphology instance.");"number"==typeof r&&(r={iterations:r});var n=r.iterations;if("number"!=typeof n)throw new Error("graphology-layout-forceatlas2: invalid number of iterations.");if(n<=0)throw new Error("graphology-layout-forceatlas2: you should provide a positive number of iterations.");var i=mu.assign({},bu,r.settings),o=mu.validateSettings(i);if(o)throw new Error("graphology-layout-forceatlas2: "+o.message);var a,s=mu.graphToByteArrays(e);for(a=0;a<n;a++)vu(i,s.nodes,s.edges);if(!t)return mu.collectLayoutChanges(e,s.nodes);mu.assignLayoutChanges(e,s.nodes)}var Eu=wu.bind(null,!1);Eu.assign=wu.bind(null,!0),Eu.inferSettings=function(t){var e="number"==typeof t?t:t.order;return{barnesHutOptimize:e>2e3,strongGravityMode:!0,gravity:.05,scalingRatio:10,slowDown:1+Math.log(e)}};var Au=Eu;function xu(){return.01*(.5-Math.random())}var zu={};zu.validateSettings=function(t){return"gridSize"in t&&"number"!=typeof t.gridSize||t.gridSize<=0?{message:"the `gridSize` setting should be a positive number."}:"margin"in t&&"number"!=typeof t.margin||t.margin<0?{message:"the `margin` setting should be 0 or a positive number."}:"expansion"in t&&"number"!=typeof t.expansion||t.expansion<=0?{message:"the `expansion` setting should be a positive number."}:"ratio"in t&&"number"!=typeof t.ratio||t.ratio<=0?{message:"the `ratio` setting should be a positive number."}:"speed"in t&&"number"!=typeof t.speed||t.speed<=0?{message:"the `speed` setting should be a positive number."}:null},zu.graphToByteArray=function(t,e){var r=t.order,n=new Float32Array(3*r),i=0;return t.forEachNode((function(t,r){"function"==typeof e&&(r=e(t,r)),n[i]=r.x,n[i+1]=r.y,n[i+2]=r.size||1,i+=3})),n},zu.assignLayoutChanges=function(t,e,r){var n=0;t.forEachNode((function(i){var o={x:e[n],y:e[n+1]};"function"==typeof r&&(o=r(i,o)),t.mergeNodeAttributes(i,o),n+=3}))},zu.collectLayoutChanges=function(t,e,r){var n={},i=0;return t.forEachNode((function(t){var o={x:e[i],y:e[i+1]};"function"==typeof r&&(o=r(t,o)),n[t]=o,i+=3})),n},zu.createWorker=function(t){var e=window.URL||window.webkitURL,r=t.toString(),n=e.createObjectURL(new Blob(["("+r+").call(this);"],{type:"text/javascript"})),i=new Worker(n);return e.revokeObjectURL(n),i};var Mu=function(t){return null!==t&&"object"==typeof t&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode&&"boolean"==typeof t.multi},Su=function(t,e){var r,n,i,o,a,s,u=t.margin,h=t.ratio,c=t.expansion,l=t.gridSize,d=t.speed,f=!0,g=e.length,p=g/3|0,y=new Float32Array(p),v=new Float32Array(p),m=1/0,b=1/0,w=-1/0,E=-1/0;for(r=0;r<g;r+=3)i=e[r+0],o=e[r+1],s=e[r+2]*h+u,m=Math.min(m,i-s),w=Math.max(w,i+s),b=Math.min(b,o-s),E=Math.max(E,o+s);var A=w-m,x=E-b,z=(m+w)/2,M=(b+E)/2;m=z-c*A/2,w=z+c*A/2,b=M-c*x/2,E=M+c*x/2;var S,N,j,k,D,O,_,C,W,P,I,U,T=new Array(l*l),L=T.length;for(S=0;S<L;S++)T[S]=[];for(r=0;r<g;r+=3)for(i=e[r+0],o=e[r+1],N=i-(s=e[r+2]*h+u),j=i+s,k=o-s,D=o+s,O=Math.floor(l*(N-m)/(w-m)),_=Math.floor(l*(j-m)/(w-m)),C=Math.floor(l*(k-b)/(E-b)),W=Math.floor(l*(D-b)/(E-b)),P=O;P<=_;P++)for(I=C;I<=W;I++)T[P*l+I].push(r);var F,q,R,B,G,H,K,V,$,J,Q,X,Y=new Set;for(S=0;S<L;S++)for(r=0,a=(U=T[S]).length;r<a;r++)for(R=e[(F=U[r])+0],G=e[F+1],K=e[F+2],n=r+1;n<a;n++)q=U[n],$=F+"§"+q,L>1&&Y.has($)||(L>1&&Y.add($),B=e[q+0],H=e[q+1],V=e[q+2],J=B-R,Q=H-G,(X=Math.sqrt(J*J+Q*Q))<K*h+u+(V*h+u)&&(f=!1,q=q/3|0,X>0?(y[q]+=J/X*(1+K),v[q]+=Q/X*(1+K)):(y[q]+=A*xu(),v[q]+=x*xu())));for(r=0,n=0;r<g;r+=3,n++)e[r+0]+=.1*y[n]*d,e[r+1]+=.1*v[n]*d;return{converged:f}},Nu=zu,ju={gridSize:20,margin:5,expansion:1.1,ratio:1,speed:3};function ku(t,e,r){if(!Mu(e))throw new Error("graphology-layout-noverlap: the given graph is not a valid graphology instance.");var n=(r="number"==typeof r?{maxIterations:r}:r||{}).maxIterations||500;if("number"!=typeof n)throw new Error("graphology-layout-noverlap: invalid number of maximum iterations.");if(n<=0)throw new Error("graphology-layout-noverlap: you should provide a positive number of maximum iterations.");var i=Object.assign({},ju,r.settings),o=Nu.validateSettings(i);if(o)throw new Error("graphology-layout-noverlap: "+o.message);var a,s=Nu.graphToByteArray(e,r.inputReducer),u=!1;for(a=0;a<n&&!u;a++)u=Su(i,s).converged;if(!t)return Nu.collectLayoutChanges(e,s,r.outputReducer);Nu.assignLayoutChanges(e,s,r.outputReducer)}var Du=ku.bind(null,!1);Du.assign=ku.bind(null,!0);var Ou=Du,_u={},Cu={},Wu={},Pu={},Iu="undefined"!=typeof ArrayBuffer,Uu="undefined"!=typeof Symbol;function Tu(t,e){var r,n,i,o,a;if(!t)throw new Error("obliterator/forEach: invalid iterable.");if("function"!=typeof e)throw new Error("obliterator/forEach: expecting a callback.");if(Array.isArray(t)||Iu&&ArrayBuffer.isView(t)||"string"==typeof t||"[object Arguments]"===t.toString())for(i=0,o=t.length;i<o;i++)e(t[i],i);else if("function"!=typeof t.forEach)if(Uu&&Symbol.iterator in t&&"function"!=typeof t.next&&(t=t[Symbol.iterator]()),"function"!=typeof t.next)for(n in t)t.hasOwnProperty(n)&&e(t[n],n);else for(r=t,i=0;!0!==(a=r.next()).done;)e(a.value,i),i++;else t.forEach(e)}Tu.forEachWithNullKeys=function(t,e){var r,n,i,o,a;if(!t)throw new Error("obliterator/forEachWithNullKeys: invalid iterable.");if("function"!=typeof e)throw new Error("obliterator/forEachWithNullKeys: expecting a callback.");if(Array.isArray(t)||Iu&&ArrayBuffer.isView(t)||"string"==typeof t||"[object Arguments]"===t.toString())for(i=0,o=t.length;i<o;i++)e(t[i],null);else if(t instanceof Set)t.forEach((function(t){e(t,null)}));else if("function"!=typeof t.forEach)if(Uu&&Symbol.iterator in t&&"function"!=typeof t.next&&(t=t[Symbol.iterator]()),"function"!=typeof t.next)for(n in t)t.hasOwnProperty(n)&&e(t[n],n);else for(r=t,i=0;!0!==(a=r.next()).done;)e(a.value,null),i++;else t.forEach(e)};var Lu=Tu,Fu=Lu,qu=po;function Ru(t){return"number"==typeof t.length?t.length:"number"==typeof t.size?t.size:void 0}Pu.isArrayLike=function(t){return Array.isArray(t)||qu.isTypedArray(t)},Pu.guessLength=Ru,Pu.toArray=function(t){var e=Ru(t),r="number"==typeof e?new Array(e):[],n=0;return Fu(t,(function(t){r[n++]=t})),r},Pu.toArrayWithIndices=function(t){var e=Ru(t),r="number"==typeof e?qu.getPointerArray(e):Array,n="number"==typeof e?new Array(e):[],i="number"==typeof e?new r(e):[],o=0;return Fu(t,(function(t){n[o]=t,i[o]=o++})),[n,i]};var Bu=Pu,Gu=go;function Hu(t,e){if(arguments.length<2)throw new Error("mnemonist/fixed-deque: expecting an Array class and a capacity.");if("number"!=typeof e||e<=0)throw new Error("mnemonist/fixed-deque: `capacity` should be a positive number.");this.ArrayClass=t,this.capacity=e,this.items=new t(this.capacity),this.clear()}Hu.prototype.clear=function(){this.start=0,this.size=0},Hu.prototype.push=function(t){if(this.size===this.capacity)throw new Error("mnemonist/fixed-deque.push: deque capacity ("+this.capacity+") exceeded!");var e=(this.start+this.size)%this.capacity;return this.items[e]=t,++this.size},Hu.prototype.unshift=function(t){if(this.size===this.capacity)throw new Error("mnemonist/fixed-deque.unshift: deque capacity ("+this.capacity+") exceeded!");var e=this.start-1;return 0===this.start&&(e=this.capacity-1),this.items[e]=t,this.start=e,++this.size},Hu.prototype.pop=function(){if(0===this.size)return;const t=(this.start+this.size-1)%this.capacity;return this.size--,this.items[t]},Hu.prototype.shift=function(){if(0!==this.size){var t=this.start;return this.size--,this.start++,this.start===this.capacity&&(this.start=0),this.items[t]}},Hu.prototype.peekFirst=function(){if(0!==this.size)return this.items[this.start]},Hu.prototype.peekLast=function(){if(0!==this.size){var t=this.start+this.size-1;return t>this.capacity&&(t-=this.capacity),this.items[t]}},Hu.prototype.get=function(t){if(0!==this.size)return(t=this.start+t)>this.capacity&&(t-=this.capacity),this.items[t]},Hu.prototype.forEach=function(t,e){e=arguments.length>1?e:this;for(var r=this.capacity,n=this.size,i=this.start,o=0;o<n;)t.call(e,this.items[i],o,this),o++,++i===r&&(i=0)},Hu.prototype.toArray=function(){var t=this.start+this.size;if(t<this.capacity)return this.items.slice(this.start,t);for(var e=new this.ArrayClass(this.size),r=this.capacity,n=this.size,i=this.start,o=0;o<n;)e[o]=this.items[i],o++,++i===r&&(i=0);return e},Hu.prototype.values=function(){var t=this.items,e=this.capacity,r=this.size,n=this.start,i=0;return new Gu((function(){if(i>=r)return{done:!0};var o=t[n];return n++,i++,n===e&&(n=0),{value:o,done:!1}}))},Hu.prototype.entries=function(){var t=this.items,e=this.capacity,r=this.size,n=this.start,i=0;return new Gu((function(){if(i>=r)return{done:!0};var o=t[n];return++n===e&&(n=0),{value:[i++,o],done:!1}}))},"undefined"!=typeof Symbol&&(Hu.prototype[Symbol.iterator]=Hu.prototype.values),Hu.prototype.inspect=function(){var t=this.toArray();return t.type=this.ArrayClass.name,t.capacity=this.capacity,Object.defineProperty(t,"constructor",{value:Hu,enumerable:!1}),t},"undefined"!=typeof Symbol&&(Hu.prototype[Symbol.for("nodejs.util.inspect.custom")]=Hu.prototype.inspect),Hu.from=function(t,e,r){if(arguments.length<3&&"number"!=typeof(r=Bu.guessLength(t)))throw new Error("mnemonist/fixed-deque.from: could not guess iterable length. Please provide desired capacity as last argument.");var n=new Hu(e,r);if(Bu.isArrayLike(t)){var i,o;for(i=0,o=t.length;i<o;i++)n.items[i]=t[i];return n.size=o,n}return Bu.forEach(t,(function(t){n.push(t)})),n};var Ku=Hu,Vu=go,$u=Pu;function Ju(t,e){if(arguments.length<2)throw new Error("mnemonist/fixed-stack: expecting an Array class and a capacity.");if("number"!=typeof e||e<=0)throw new Error("mnemonist/fixed-stack: `capacity` should be a positive number.");this.capacity=e,this.ArrayClass=t,this.items=new this.ArrayClass(this.capacity),this.clear()}Ju.prototype.clear=function(){this.size=0},Ju.prototype.push=function(t){if(this.size===this.capacity)throw new Error("mnemonist/fixed-stack.push: stack capacity ("+this.capacity+") exceeded!");return this.items[this.size++]=t,this.size},Ju.prototype.pop=function(){if(0!==this.size)return this.items[--this.size]},Ju.prototype.peek=function(){return this.items[this.size-1]},Ju.prototype.forEach=function(t,e){e=arguments.length>1?e:this;for(var r=0,n=this.items.length;r<n;r++)t.call(e,this.items[n-r-1],r,this)},Ju.prototype.toArray=function(){for(var t=new this.ArrayClass(this.size),e=this.size-1,r=this.size;r--;)t[r]=this.items[e-r];return t},Ju.prototype.values=function(){var t=this.items,e=this.size,r=0;return new Vu((function(){if(r>=e)return{done:!0};var n=t[e-r-1];return r++,{value:n,done:!1}}))},Ju.prototype.entries=function(){var t=this.items,e=this.size,r=0;return new Vu((function(){if(r>=e)return{done:!0};var n=t[e-r-1];return{value:[r++,n],done:!1}}))},"undefined"!=typeof Symbol&&(Ju.prototype[Symbol.iterator]=Ju.prototype.values),Ju.prototype.toString=function(){return this.toArray().join(",")},Ju.prototype.toJSON=function(){return this.toArray()},Ju.prototype.inspect=function(){var t=this.toArray();return t.type=this.ArrayClass.name,t.capacity=this.capacity,Object.defineProperty(t,"constructor",{value:Ju,enumerable:!1}),t},"undefined"!=typeof Symbol&&(Ju.prototype[Symbol.for("nodejs.util.inspect.custom")]=Ju.prototype.inspect),Ju.from=function(t,e,r){if(arguments.length<3&&"number"!=typeof(r=$u.guessLength(t)))throw new Error("mnemonist/fixed-stack.from: could not guess iterable length. Please provide desired capacity as last argument.");var n=new Ju(e,r);if($u.isArrayLike(t)){var i,o;for(i=0,o=t.length;i<o;i++)n.items[i]=t[i];return n.size=o,n}return $u.forEach(t,(function(t){n.push(t)})),n};var Qu=Ju,Xu={};Xu.DEFAULT_COMPARATOR=function(t,e){return t<e?-1:t>e?1:0},Xu.DEFAULT_REVERSE_COMPARATOR=function(t,e){return t<e?1:t>e?-1:0},Xu.reverseComparator=function(t){return function(e,r){return t(r,e)}},Xu.createTupleComparator=function(t){return 2===t?function(t,e){return t[0]<e[0]?-1:t[0]>e[0]?1:t[1]<e[1]?-1:t[1]>e[1]?1:0}:function(e,r){for(var n=0;n<t;){if(e[n]<r[n])return-1;if(e[n]>r[n])return 1;n++}return 0}};var Yu=Lu,Zu=Xu,th=Pu,eh=Zu.DEFAULT_COMPARATOR,rh=Zu.reverseComparator;function nh(t,e,r,n){for(var i,o,a=e[n];n>r&&t(a,o=e[i=n-1>>1])<0;)e[n]=o,n=i;e[n]=a}function ih(t,e,r){for(var n,i=e.length,o=r,a=e[r],s=2*r+1;s<i;)(n=s+1)<i&&t(e[s],e[n])>=0&&(s=n),e[r]=e[s],s=2*(r=s)+1;e[r]=a,nh(t,e,o,r)}function oh(t,e,r){e.push(r),nh(t,e,0,e.length-1)}function ah(t,e){var r=e.pop();if(0!==e.length){var n=e[0];return e[0]=r,ih(t,e,0),n}return r}function sh(t,e,r){if(0===e.length)throw new Error("mnemonist/heap.replace: cannot pop an empty heap.");var n=e[0];return e[0]=r,ih(t,e,0),n}function uh(t,e,r){var n;return 0!==e.length&&t(e[0],r)<0&&(n=e[0],e[0]=r,r=n,ih(t,e,0)),r}function hh(t,e){for(var r=e.length>>1;--r>=0;)ih(t,e,r)}function ch(t,e){for(var r=e.length,n=0,i=new Array(r);n<r;)i[n++]=ah(t,e);return i}function lh(t){if(this.clear(),this.comparator=t||eh,"function"!=typeof this.comparator)throw new Error("mnemonist/Heap.constructor: given comparator should be a function.")}function dh(t){if(this.clear(),this.comparator=t||eh,"function"!=typeof this.comparator)throw new Error("mnemonist/MaxHeap.constructor: given comparator should be a function.");this.comparator=rh(this.comparator)}lh.prototype.clear=function(){this.items=[],this.size=0},lh.prototype.push=function(t){return oh(this.comparator,this.items,t),++this.size},lh.prototype.peek=function(){return this.items[0]},lh.prototype.pop=function(){return 0!==this.size&&this.size--,ah(this.comparator,this.items)},lh.prototype.replace=function(t){return sh(this.comparator,this.items,t)},lh.prototype.pushpop=function(t){return uh(this.comparator,this.items,t)},lh.prototype.consume=function(){return this.size=0,ch(this.comparator,this.items)},lh.prototype.toArray=function(){return ch(this.comparator,this.items.slice())},lh.prototype.inspect=function(){var t=this.toArray();return Object.defineProperty(t,"constructor",{value:lh,enumerable:!1}),t},"undefined"!=typeof Symbol&&(lh.prototype[Symbol.for("nodejs.util.inspect.custom")]=lh.prototype.inspect),dh.prototype=lh.prototype,lh.from=function(t,e){var r,n=new lh(e);return r=th.isArrayLike(t)?t.slice():th.toArray(t),hh(n.comparator,r),n.items=r,n.size=r.length,n},dh.from=function(t,e){var r,n=new dh(e);return r=th.isArrayLike(t)?t.slice():th.toArray(t),hh(n.comparator,r),n.items=r,n.size=r.length,n},lh.siftUp=ih,lh.siftDown=nh,lh.push=oh,lh.pop=ah,lh.replace=sh,lh.pushpop=uh,lh.heapify=hh,lh.consume=ch,lh.nsmallest=function(t,e,r){2===arguments.length&&(r=e,e=t,t=eh);var n,i,o,a,s=rh(t),u=1/0;if(1===e){if(th.isArrayLike(r)){for(n=0,i=r.length;n<i;n++)o=r[n],(u===1/0||t(o,u)<0)&&(u=o);return(a=new r.constructor(1))[0]=u,a}return Yu(r,(function(e){(u===1/0||t(e,u)<0)&&(u=e)})),[u]}if(th.isArrayLike(r)){if(e>=r.length)return r.slice().sort(t);for(a=r.slice(0,e),hh(s,a),n=e,i=r.length;n<i;n++)s(r[n],a[0])>0&&sh(s,a,r[n]);return a.sort(t)}var h=th.guessLength(r);return null!==h&&h<e&&(e=h),a=new Array(e),n=0,Yu(r,(function(t){n<e?a[n]=t:(n===e&&hh(s,a),s(t,a[0])>0&&sh(s,a,t)),n++})),a.length>n&&(a.length=n),a.sort(t)},lh.nlargest=function(t,e,r){2===arguments.length&&(r=e,e=t,t=eh);var n,i,o,a,s=rh(t),u=-1/0;if(1===e){if(th.isArrayLike(r)){for(n=0,i=r.length;n<i;n++)o=r[n],(u===-1/0||t(o,u)>0)&&(u=o);return(a=new r.constructor(1))[0]=u,a}return Yu(r,(function(e){(u===-1/0||t(e,u)>0)&&(u=e)})),[u]}if(th.isArrayLike(r)){if(e>=r.length)return r.slice().sort(s);for(a=r.slice(0,e),hh(t,a),n=e,i=r.length;n<i;n++)t(r[n],a[0])>0&&sh(t,a,r[n]);return a.sort(s)}var h=th.guessLength(r);return null!==h&&h<e&&(e=h),a=new Array(e),n=0,Yu(r,(function(r){n<e?a[n]=r:(n===e&&hh(t,a),t(r,a[0])>0&&sh(t,a,r)),n++})),a.length>n&&(a.length=n),a.sort(s)},lh.MinHeap=lh,lh.MaxHeap=dh;var fh=lh,gh={},ph=po;function yh(t){var e=t.directedSize+2*t.undirectedSize,r=ph.getPointerArray(e),n=ph.getPointerArray(t.order);this.graph=t,this.neighborhood=new n(e),this.starts=new r(t.order+1),this.nodes=t.nodes();var i,o,a,s,u,h,c={},l=0;for(i=0,o=t.order;i<o;i++)c[this.nodes[i]]=i;for(i=0,o=t.order;i<o;i++)for(u=this.nodes[i],h=t.outboundNeighbors(u),this.starts[i]=l,a=0,s=h.length;a<s;a++)this.neighborhood[l++]=c[h[a]];this.starts[i]=e}function vh(t,e){var r=t.directedSize+2*t.undirectedSize,n=ph.getPointerArray(r),i=ph.getPointerArray(t.order);e=e||"weight",this.graph=t,this.neighborhood=new i(r),this.weights=new Float64Array(r),this.starts=new n(t.order+1),this.nodes=t.nodes();var o,a,s,u,h,c,l,d,f,g={},p=0;for(o=0,a=t.order;o<a;o++)g[this.nodes[o]]=o;for(o=0,a=t.order;o<a;o++)for(h=this.nodes[o],l=t.outboundEdges(h),this.starts[o]=p,s=0,u=l.length;s<u;s++)d=l[s],c=t.opposite(h,d),"number"!=typeof(f=t.getEdgeAttribute(d,e))&&(f=1),this.neighborhood[p]=g[c],this.weights[p++]=f;this.starts[o]=r}yh.prototype.bounds=function(t){return[this.starts[t],this.starts[t+1]]},yh.prototype.project=function(){var t=this,e={};return t.nodes.forEach((function(r,n){e[r]=Array.from(t.neighborhood.slice(t.starts[n],t.starts[n+1])).map((function(e){return t.nodes[e]}))})),e},yh.prototype.collect=function(t){var e,r,n={};for(e=0,r=t.length;e<r;e++)n[this.nodes[e]]=t[e];return n},yh.prototype.assign=function(t,e){var r,n;for(r=0,n=e.length;r<n;r++)this.graph.setNodeAttribute(this.nodes[r],t,e[r])},gh.OutboundNeighborhoodIndex=yh,vh.prototype.bounds=yh.prototype.bounds,vh.prototype.project=yh.prototype.project,vh.prototype.collect=yh.prototype.collect,vh.prototype.assign=yh.prototype.assign,gh.WeightedOutboundNeighborhoodIndex=vh;var mh=Ku,bh=Qu,wh=fh,Eh=po,Ah=gh,xh=Ah.OutboundNeighborhoodIndex,zh=Ah.WeightedOutboundNeighborhoodIndex;function Mh(t,e){return t[0]>e[0]?1:t[0]<e[0]?-1:t[1]>e[1]?1:t[1]<e[1]?-1:t[2]>e[2]?1:t[2]<e[2]?-1:t[3]>e[3]?1:t[3]<e[3]?-1:0}Wu.createUnweightedIndexedBrandes=function(t){var e=new xh(t),r=e.neighborhood,n=e.starts,i=t.order,o=new bh(Eh.getPointerArray(i),i),a=new Uint32Array(i),s=new Array(i),u=new Int32Array(i),h=new mh(Uint32Array,i),c=function(t){var e,c,l,d,f,g,p;for(g=0;g<i;g++)s[g]=[],a[g]=0,u[g]=-1;for(a[t]=1,u[t]=0,h.push(t);0!==h.size;)for(g=h.shift(),o.push(g),e=u[g],c=a[g],l=n[g],d=n[g+1],f=l;f<d;f++)p=r[f],-1===u[p]&&(h.push(p),u[p]=e+1),u[p]===e+1&&(a[p]+=c,s[p].push(g));return[o,s,a]};return c.index=e,c},Wu.createDijkstraIndexedBrandes=function(t,e){var r=new zh(t,e),n=r.neighborhood,i=r.weights,o=r.starts,a=t.order,s=new bh(Eh.getPointerArray(a),a),u=new Uint32Array(a),h=new Array(a),c=new Float64Array(a),l=new Float64Array(a),d=new wh(Mh),f=function(t){var e,r,f,g,p,y,v,m,b,w=0;for(m=0;m<a;m++)h[m]=[],u[m]=0,c[m]=-1,l[m]=-1;for(u[t]=1,l[t]=0,d.push([0,w++,t,t]);0!==d.size;)if(g=(f=d.pop())[0],p=f[2],m=f[3],-1===c[m])for(s.push(m),c[m]=g,u[m]+=u[p],e=o[m],r=o[m+1],v=e;v<r;v++)b=n[v],y=g+i[v],-1===c[b]&&(-1===l[b]||y<l[b])?(l[b]=y,d.push([y,w++,m,b]),u[b]=0,h[b]=[m]):y===l[b]&&(u[b]+=u[m],h[b].push(m));return[s,h,u]};return f.index=r,f};var Sh=kn,Nh=o,jh=to,kh=Ei,Dh=Object.prototype,Oh=Dh.hasOwnProperty,_h=Sh((function(t,e){t=Object(t);var r=-1,n=e.length,i=n>2?e[2]:void 0;for(i&&jh(e[0],e[1],i)&&(n=1);++r<n;)for(var o=e[r],a=kh(o),s=-1,u=a.length;++s<u;){var h=a[s],c=t[h];(void 0===c||Nh(c,Dh[h])&&!Oh.call(t,h))&&(t[h]=o[h])}return t})),Ch=ho,Wh=Wu,Ph=_h,Ih=Wh.createUnweightedIndexedBrandes,Uh=Wh.createDijkstraIndexedBrandes,Th={attributes:{centrality:"betweennessCentrality",weight:"weight"},normalized:!0,weighted:!1};function Lh(t,e,r){if(!Ch(e))throw new Error("graphology-centrality/beetweenness-centrality: the given graph is not a valid graphology instance.");var n,i,o,a,s,u,h,c,l,d,f=(r=Ph({},r,Th)).attributes.weight,g=r.attributes.centrality,p=r.normalized,y=r.weighted?Uh(e,f):Ih(e),v=e.order,m=new Float64Array(v),b=new Float64Array(v);for(u=0;u<v;u++){for(i=(n=y(u))[0],o=n[1],a=n[2],h=i.size;h--;)m[i.items[i.size-h]]=0;for(;0!==i.size;){for(s=(1+m[d=i.pop()])/a[d],h=0,c=o[d].length;h<c;h++)m[l=o[d][h]]+=a[l]*s;d!==u&&(b[d]+=m[d])}}var w=null;if(null!==(w=p?v<=2?null:1/((v-1)*(v-2)):"undirected"===e.type?.5:null))for(u=0;u<v;u++)b[u]*=w;return t?y.index.assign(g,b):y.index.collect(b)}var Fh=Lh.bind(null,!1);Fh.assign=Lh.bind(null,!0);var qh=Fh,Rh=ho;function Bh(t,e,r,n){var i=e+"Centrality";if(!Rh(r))throw new Error("graphology-centrality/"+i+": the given graph is not a valid graphology instance.");if("degree"!==e&&"undirected"===r.type)throw new Error("graphology-centrality/"+i+": cannot compute "+e+" centrality on an undirected graph.");var o=((n=n||{}).attributes||{}).centrality||i,a=r.order,s=r.nodes(),u=r[e].bind(r),h={};if(0===a)return t?void 0:h;var c,l,d,f=1/(a-1);for(d=0;d<a;d++)l=u(c=s[d])*f,t?r.setNodeAttribute(c,o,l):h[c]=l;return t?void 0:h}var Gh=Bh.bind(null,!1,"degree"),Hh=Bh.bind(null,!1,"inDegree"),Kh=Bh.bind(null,!1,"outDegree");Gh.assign=Bh.bind(null,!0,"degree"),Hh.assign=Bh.bind(null,!0,"inDegree"),Kh.assign=Bh.bind(null,!0,"outDegree"),Gh.degreeCentrality=Gh,Gh.inDegreeCentrality=Hh,Gh.outDegreeCentrality=Kh;var Vh=Gh;Cu.betweenness=qh,Cu.degree=Vh;var $h=go,Jh=Lu;function Qh(){this.clear()}Qh.prototype.clear=function(){this.items=[],this.offset=0,this.size=0},Qh.prototype.enqueue=function(t){return this.items.push(t),++this.size},Qh.prototype.dequeue=function(){if(this.size){var t=this.items[this.offset];return 2*++this.offset>=this.items.length&&(this.items=this.items.slice(this.offset),this.offset=0),this.size--,t}},Qh.prototype.peek=function(){if(this.size)return this.items[this.offset]},Qh.prototype.forEach=function(t,e){e=arguments.length>1?e:this;for(var r=this.offset,n=0,i=this.items.length;r<i;r++,n++)t.call(e,this.items[r],n,this)},Qh.prototype.toArray=function(){return this.items.slice(this.offset)},Qh.prototype.values=function(){var t=this.items,e=this.offset;return new $h((function(){if(e>=t.length)return{done:!0};var r=t[e];return e++,{value:r,done:!1}}))},Qh.prototype.entries=function(){var t=this.items,e=this.offset,r=0;return new $h((function(){if(e>=t.length)return{done:!0};var n=t[e];return e++,{value:[r++,n],done:!1}}))},"undefined"!=typeof Symbol&&(Qh.prototype[Symbol.iterator]=Qh.prototype.values),Qh.prototype.toString=function(){return this.toArray().join(",")},Qh.prototype.toJSON=function(){return this.toArray()},Qh.prototype.inspect=function(){var t=this.toArray();return Object.defineProperty(t,"constructor",{value:Qh,enumerable:!1}),t},"undefined"!=typeof Symbol&&(Qh.prototype[Symbol.for("nodejs.util.inspect.custom")]=Qh.prototype.inspect),Qh.from=function(t){var e=new Qh;return Jh(t,(function(t){e.enqueue(t)})),e},Qh.of=function(){return Qh.from(arguments)};var Xh=ho,Yh=Qh,Zh=ea;function tc(t,e,r){if(!Xh(t))throw new Error("graphology-shortest-path: invalid graphology instance.");if(arguments.length<3)throw new Error("graphology-shortest-path: invalid number of arguments. Expecting at least 3.");if(!t.hasNode(e))throw new Error('graphology-shortest-path: the "'+e+'" source node does not exist in the given graph.');if(!t.hasNode(r))throw new Error('graphology-shortest-path: the "'+r+'" target node does not exist in the given graph.');if((e=""+e)===(r=""+r))return[e];var n=t.inboundNeighbors.bind(t),i=t.outboundNeighbors.bind(t),o={},a={};o[e]=null,a[r]=null;var s,u,h,c,l,d,f,g,p=[e],y=[r],v=!1;t:for(;p.length&&y.length;)if(p.length<=y.length){for(s=p,p=[],l=0,f=s.length;l<f;l++)for(d=0,g=(h=i(u=s[l])).length;d<g;d++)if((c=h[d])in o||(p.push(c),o[c]=u),c in a){v=!0;break t}}else for(s=y,y=[],l=0,f=s.length;l<f;l++)for(d=0,g=(h=n(u=s[l])).length;d<g;d++)if((c=h[d])in a||(y.push(c),a[c]=u),c in o){v=!0;break t}if(!v)return null;for(var m=[];c;)m.unshift(c),c=o[c];for(c=a[m[m.length-1]];c;)m.push(c),c=a[c];return m.length?m:null}function ec(t,e){if(!Xh(t))throw new Error("graphology-shortest-path: invalid graphology instance.");if(arguments.length<2)throw new Error("graphology-shortest-path: invalid number of arguments. Expecting at least 2.");if(!t.hasNode(e))throw new Error('graphology-shortest-path: the "'+e+'" source node does not exist in the given graph.');var r,n,i,o,a,s,u={},h={};for(u[e=""+e]=!0,h[e]=[e];Object.keys(u).length;)for(i in r=u,u={},r)for(a=0,s=(n=t.outboundNeighbors(i)).length;a<s;a++)h[o=n[a]]||(h[o]=h[i].concat(o),u[o]=!0);return h}function rc(t,e,r){if(!Xh(e))throw new Error("graphology-shortest-path: invalid graphology instance.");if(!e.hasNode(r))throw new Error('graphology-shortest-path: the "'+r+'" source node does not exist in the given graph.');r=""+r;var n=new Set,i={},o=0;i[r]=0;for(var a,s,u,h=[r];0!==h.length;){var c=[];for(a=0,s=h.length;a<s;a++)u=h[a],n.has(u)||(n.add(u),Zh(c,e[t](u)),i[u]=o);o++,h=c}return i}var nc=rc.bind(null,"outboundNeighbors"),ic=rc.bind(null,"neighbors");function oc(t,e,r){return arguments.length<3?ec(t,e):tc(t,e,r)}oc.bidirectional=tc,oc.singleSource=ec,oc.singleSourceLength=nc,oc.undirectedSingleSourceLength=ic,oc.brandes=function(t,e){e=""+e;var r,n,i,o,a,s,u,h,c,l=[],d={},f={},g=t.nodes();for(s=0,h=g.length;s<h;s++)d[o=g[s]]=[],f[o]=0;var p={};f[e]=1,p[e]=0;for(var y=Yh.of(e);y.size;)for(o=y.dequeue(),l.push(o),r=p[o],n=f[o],u=0,c=(i=t.outboundNeighbors(o)).length;u<c;u++)(a=i[u])in p||(y.enqueue(a),p[a]=r+1),p[a]===r+1&&(f[a]+=n,d[a].push(o));return[l,d,f]};var ac=oc,sc=ho,uc=ac.singleSourceLength,hc=function(t,e){if(!sc(t))throw new Error("graphology-metrics/eccentricity: given graph is not a valid graphology instance.");if(0===t.size)return 1/0;var r,n,i=-1/0,o=uc(t,e),a=0;for(r in o)(n=o[r])>i&&(i=n),a++;return a<t.order?1/0:i},cc=ho,lc=hc,dc=ho;function fc(t,e){if(!dc(t))throw new Error("graphology-metrics/extent: the given graph is not a valid graphology instance.");var r,n,i,o=[].concat(e),a={};for(i=0;i<o.length;i++)n=o[i],a[n]=[1/0,-1/0];return t.forEachNode((function(t,e){for(i=0;i<o.length;i++)n=o[i],(r=e[n])<a[n][0]&&(a[n][0]=r),r>a[n][1]&&(a[n][1]=r)})),"string"==typeof e?a[e]:a}var gc=fc;gc.nodeExtent=fc,gc.edgeExtent=function(t,e){if(!dc(t))throw new Error("graphology-metrics/extent: the given graph is not a valid graphology instance.");var r,n,i,o=[].concat(e),a={};for(i=0;i<o.length;i++)n=o[i],a[n]=[1/0,-1/0];return t.forEachEdge((function(t,e){for(i=0;i<o.length;i++)n=o[i],(r=e[n])<a[n][0]&&(a[n][0]=r),r>a[n][1]&&(a[n][1]=r)})),"string"==typeof e?a[e]:a};var pc=gc,yc={},vc=ho;var mc=Xu,bc=fh,wc=mc.DEFAULT_COMPARATOR,Ec=mc.reverseComparator;function Ac(t,e,r,n){for(var i,o=r,a=n,s=e[n],u=2*n+1;u<o;)(i=u+1)<o&&t(e[u],e[i])>=0&&(u=i),e[n]=e[u],u=2*(n=u)+1;e[n]=s,bc.siftDown(t,e,a,n)}function xc(t,e,r,n){for(var i,o,a=n,s=new t(n);a>0;)i=r[--a],0!==a&&(o=r[0],r[0]=i,Ac(e,r,--n,0),i=o),s[a]=i;return s}function zc(t,e,r){if(2===arguments.length&&(r=e,e=null),this.ArrayClass=t,this.capacity=r,this.items=new t(r),this.clear(),this.comparator=e||wc,"number"!=typeof r&&r<=0)throw new Error("mnemonist/FixedReverseHeap.constructor: capacity should be a number > 0.");if("function"!=typeof this.comparator)throw new Error("mnemonist/FixedReverseHeap.constructor: given comparator should be a function.");this.comparator=Ec(this.comparator)}zc.prototype.clear=function(){this.size=0},zc.prototype.push=function(t){return this.size<this.capacity?(this.items[this.size]=t,bc.siftDown(this.comparator,this.items,0,this.size),this.size++):this.comparator(t,this.items[0])>0&&bc.replace(this.comparator,this.items,t),this.size},zc.prototype.peek=function(){return this.items[0]},zc.prototype.consume=function(){var t=xc(this.ArrayClass,this.comparator,this.items,this.size);return this.size=0,t},zc.prototype.toArray=function(){return xc(this.ArrayClass,this.comparator,this.items.slice(0,this.size),this.size)},zc.prototype.inspect=function(){var t=this.toArray();return Object.defineProperty(t,"constructor",{value:zc,enumerable:!1}),t},"undefined"!=typeof Symbol&&(zc.prototype[Symbol.for("nodejs.util.inspect.custom")]=zc.prototype.inspect);var Mc=zc,Sc={},Nc=new Float64Array(64),jc=new Float64Array(64);Sc.inplaceQuickSort=function(t,e,r){var n,i,o,a,s;for(Nc[0]=e,jc[0]=r,i=0;i>=0;)if((o=Nc[i])<(a=jc[i]-1)){for(n=t[o];o<a;){for(;t[a]>=n&&o<a;)a--;for(o<a&&(t[o++]=t[a]);t[o]<=n&&o<a;)o++;o<a&&(t[a--]=t[o])}t[o]=n,Nc[i+1]=o+1,jc[i+1]=jc[i],jc[i++]=o,jc[i]-Nc[i]>jc[i-1]-Nc[i-1]&&(s=Nc[i],Nc[i]=Nc[i-1],Nc[i-1]=s,s=jc[i],jc[i]=jc[i-1],jc[i-1]=s)}else i--;return t},Sc.inplaceQuickSortIndices=function(t,e,r,n){var i,o,a,s,u,h;for(Nc[0]=r,jc[0]=n,o=0;o>=0;)if((a=Nc[o])<(s=jc[o]-1)){for(i=t[u=e[a]];a<s;){for(;t[e[s]]>=i&&a<s;)s--;for(a<s&&(e[a++]=e[s]);t[e[a]]<=i&&a<s;)a++;a<s&&(e[s--]=e[a])}e[a]=u,Nc[o+1]=a+1,jc[o+1]=jc[o],jc[o++]=a,jc[o]-Nc[o]>jc[o-1]-Nc[o-1]&&(h=Nc[o],Nc[o]=Nc[o-1],Nc[o-1]=h,h=jc[o],jc[o]=jc[o-1],jc[o-1]=h)}else o--;return e};var kc=Pu,Dc=po,Oc=Xu.createTupleComparator,_c=Mc,Cc=Sc.inplaceQuickSortIndices;function Wc(t,e,r,n){var i,o,a=0;for(i=0;i<t;i++)a+=(o=e[i][r]-n[i])*o;return a}function Pc(t,e,r,n){for(var i,o,a,s,u,h,c,l,d=n.length,f=Dc.getPointerArray(d+1),g=new f(d),p=new f(d),y=new f(d),v=[[0,0,r.length,-1,0]],m=0;0!==v.length;)l=(i=v.pop())[0],h=i[1],c=i[2],o=i[3],a=i[4],Cc(e[l],r,h,c),u=r[s=h+((d=c-h)>>>1)],g[m]=u,o>-1&&(0===a?p[o]=m+1:y[o]=m+1),l=(l+1)%t,s!==h&&s!==c-1&&v.push([l,s+1,c,m,1]),s!==h&&v.push([l,h,s,m,0]),m++;return{axes:e,labels:n,pivots:g,lefts:p,rights:y}}function Ic(t,e){this.dimensions=t,this.visited=0,this.axes=e.axes,this.labels=e.labels,this.pivots=e.pivots,this.lefts=e.lefts,this.rights=e.rights,this.size=this.labels.length}Ic.prototype.nearestNeighbor=function(t){var e=1/0,r=null,n=this.dimensions,i=this.axes,o=this.pivots,a=this.lefts,s=this.rights,u=0;return function h(c,l){u++;var d=a[l],f=s[l],g=o[l],p=Wc(n,i,g,t);if(!(p<e&&(r=g,e=p,0===p))){var y=i[c][g]-t[c];c=(c+1)%n,y>0?0!==d&&h(c,d-1):0!==f&&h(c,f-1),y*y<e&&(y>0?0!==f&&h(c,f-1):0!==d&&h(c,d-1))}}(0,0),this.visited=u,this.labels[r]};var Uc=Oc(3),Tc=Oc(2);Ic.prototype.kNearestNeighbors=function(t,e){if(t<=0)throw new Error("mnemonist/kd-tree.kNearestNeighbors: k should be a positive number.");if(1===(t=Math.min(t,this.size)))return[this.nearestNeighbor(e)];var r=new _c(Array,Uc,t),n=this.dimensions,i=this.axes,o=this.pivots,a=this.lefts,s=this.rights,u=0;!function h(c,l){var d=a[l],f=s[l],g=o[l],p=Wc(n,i,g,e);r.push([p,u++,g]);var y=e[c],v=i[c][g],m=y-v;c=(c+1)%n,y<v?0!==d&&h(c,d-1):0!==f&&h(c,f-1),(m*m<r.peek()[0]||r.size<t)&&(y<v?0!==f&&h(c,f-1):0!==d&&h(c,d-1))}(0,0),this.visited=u;for(var h=r.consume(),c=0;c<h.length;c++)h[c]=this.labels[h[c][2]];return h},Ic.prototype.linearKNearestNeighbors=function(t,e){if(t<=0)throw new Error("mnemonist/kd-tree.kNearestNeighbors: k should be a positive number.");t=Math.min(t,this.size);var r,n,i,o=new _c(Array,Tc,t);for(r=0,n=this.size;r<n;r++)i=Wc(this.dimensions,this.axes,this.pivots[r],e),o.push([i,r]);var a=o.consume();for(r=0;r<a.length;r++)a[r]=this.labels[this.pivots[a[r][1]]];return a},Ic.prototype.inspect=function(){var t,e,r,n=new Map;for(n.dimensions=this.dimensions,Object.defineProperty(n,"constructor",{value:Ic,enumerable:!1}),t=0;t<this.size;t++){for(r=new Array(this.dimensions),e=0;e<this.dimensions;e++)r[e]=this.axes[e][t];n.set(this.labels[t],r)}return n},"undefined"!=typeof Symbol&&(Ic.prototype[Symbol.for("nodejs.util.inspect.custom")]=Ic.prototype.inspect),Ic.from=function(t,e){var r=function(t,e){var r,n,i,o,a=e.length,s=new Array(t),u=new Array(a),h=new(Dc.getPointerArray(a))(a),c=!0;for(n=0;n<t;n++){for(r=new Float64Array(a),i=0;i<a;i++)o=e[i],r[i]=o[1][n],c&&(u[i]=o[0],h[i]=i);c=!1,s[n]=r}return{axes:s,ids:h,labels:u}}(e,kc.toArray(t));return new Ic(e,Pc(e,r.axes,r.ids,r.labels))},Ic.fromAxes=function(t,e){return e||(e=Dc.indices(t[0].length)),new Ic(t.length,Pc(t.length,t,Dc.indices(e.length),e))};var Lc=Ic,Fc={};!function(t){t.intersection=function(){if(arguments.length<2)throw new Error("mnemonist/Set.intersection: needs at least two arguments.");var t,e,r=new Set,n=1/0,i=null,o=arguments.length;for(e=0;e<o;e++){if(0===(t=arguments[e]).size)return r;t.size<n&&(n=t.size,i=t)}for(var a,s,u,h,c=i.values();!(a=c.next()).done;){for(s=a.value,u=!0,e=0;e<o;e++)if((h=arguments[e])!==i&&!h.has(s)){u=!1;break}u&&r.add(s)}return r},t.union=function(){if(arguments.length<2)throw new Error("mnemonist/Set.union: needs at least two arguments.");var t,e,r,n=new Set,i=arguments.length;for(t=0;t<i;t++)for(e=arguments[t].values();!(r=e.next()).done;)n.add(r.value);return n},t.difference=function(t,e){if(!t.size)return new Set;if(!e.size)return new Set(t);for(var r,n=new Set,i=t.values();!(r=i.next()).done;)e.has(r.value)||n.add(r.value);return n},t.symmetricDifference=function(t,e){for(var r,n=new Set,i=t.values();!(r=i.next()).done;)e.has(r.value)||n.add(r.value);for(i=e.values();!(r=i.next()).done;)t.has(r.value)||n.add(r.value);return n},t.isSubset=function(t,e){var r,n=t.values();if(t===e)return!0;if(t.size>e.size)return!1;for(;!(r=n.next()).done;)if(!e.has(r.value))return!1;return!0},t.isSuperset=function(e,r){return t.isSubset(r,e)},t.add=function(t,e){for(var r,n=e.values();!(r=n.next()).done;)t.add(r.value)},t.subtract=function(t,e){for(var r,n=e.values();!(r=n.next()).done;)t.delete(r.value)},t.intersect=function(t,e){for(var r,n=t.values();!(r=n.next()).done;)e.has(r.value)||t.delete(r.value)},t.disjunct=function(t,e){for(var r,n=t.values(),i=[];!(r=n.next()).done;)e.has(r.value)&&i.push(r.value);for(n=e.values();!(r=n.next()).done;)t.has(r.value)||t.add(r.value);for(var o=0,a=i.length;o<a;o++)t.delete(i[o])},t.intersectionSize=function(t,e){var r;if(t.size>e.size&&(r=t,t=e,e=r),0===t.size)return 0;if(t===e)return t.size;for(var n,i=t.values(),o=0;!(n=i.next()).done;)e.has(n.value)&&o++;return o},t.unionSize=function(e,r){var n=t.intersectionSize(e,r);return e.size+r.size-n},t.jaccard=function(e,r){var n=t.intersectionSize(e,r);return 0===n?0:n/(e.size+r.size-n)},t.overlap=function(e,r){var n=t.intersectionSize(e,r);return 0===n?0:n/Math.min(e.size,r.size)}}(Fc);var qc=ho,Rc=Lc,Bc=Fc.intersectionSize,Gc=ho,Hc=ac.undirectedSingleSourceLength;yc.edgeUniformity=function(t){if(!vc(t))throw new Error("graphology-metrics/layout-quality/edge-uniformity: given graph is not a valid graphology instance.");if(0===t.size)return 0;var e,r=0,n=0,i=new Float64Array(t.size);t.forEachEdge((function(t,e,o,a,s,u){var h,c,l=(h=s,c=u,Math.sqrt(Math.pow(h.x-c.x,2)+Math.pow(h.y-c.y,2)));i[n++]=l,r+=l}));var o=r/t.size,a=0;for(n=0,e=t.size;n<e;n++)a+=Math.pow(o-i[n],2);var s=a/(t.size*Math.pow(o,2));return Math.sqrt(s)},yc.neighborhoodPreservation=function(t){if(!qc(t))throw new Error("graphology-metrics/layout-quality/neighborhood-preservation: given graph is not a valid graphology instance.");if(0===t.order)throw new Error("graphology-metrics/layout-quality/neighborhood-preservation: cannot compute stress for a null graph.");if(0===t.size)return 0;var e=[new Float64Array(t.order),new Float64Array(t.order)],r=0;t.forEachNode((function(t,n){e[0][r]=n.x,e[1][r++]=n.y}));var n=Rc.fromAxes(e,t.nodes()),i=0;return t.forEachNode((function(e,r){var o=new Set(t.neighbors(e));if(0!==o.size&&o.size!==t.order-1){var a=n.kNearestNeighbors(o.size+1,[r.x,r.y]);a=new Set(a.slice(1));var s=Bc(o,a);i+=s/a.size}else i+=1})),i/t.order},yc.stress=function(t){if(!Gc(t))throw new Error("graphology-metrics/layout-quality/stress: given graph is not a valid graphology instance.");if(0===t.order)throw new Error("graphology-metrics/layout-quality/stress: cannot compute stress for a null graph.");var e,r,n,i,o,a,s,u=new Array(t.order),h=new Array(t.order),c=0,l=4*t.order;t.forEachNode((function(t,e){u[c]=t,h[c++]=e}));var d,f,g=0;for(c=0,r=t.order;c<r;c++)for(o=Hc(t,u[c]),n=h[c],e=c+1;e<r;e++)i=h[e],void 0===(a=o[u[e]])&&(a=l),d=n,f=i,s=Math.sqrt(Math.pow(d.x-f.x,2)+Math.pow(d.y-f.y,2)),g+=1/(a*a)*Math.pow(s-a,2);return g};var Kc=uo,Vc=ho,$c=lo,Jc={attributes:{community:"community",weight:"weight"},communities:null,resolution:1,weighted:!0};function Qc(t,e){return function(r){if(!r)return 0;if(!t)return 1;var n=r[e];return"number"!=typeof n&&(n=1),n}}function Xc(t,e){var r,n,i,o,a,s,u=e.resolution,h=function(t,e){var r=new Array(t.order),n=new Float64Array(t.order),i=0,o={},a=Qc(e.weighted,e.attributes.weight),s=0;return t.forEachNode((function(t,n){o[t]=s,r[s++]=e.communities?e.communities[t]:n[e.attributes.community]})),t.forEachUndirectedEdge((function(t,e,r,s){var u=a(e);i+=u,n[o[r]]+=u,r!==s&&(n[o[s]]+=u)})),{getWeight:a,communities:r,weightedDegrees:n,M:i}}(t,e),c=h.communities,l=h.weightedDegrees,d=h.M,f=t.nodes(),g=0,p=2*d;for(r=0,i=t.order;r<i;r++)for(n=0;n<i;n++)c[r]===c[n]&&(s=t.undirectedEdge(f[r],f[n]),o=h.getWeight(s),a=l[r]*l[n],g+=r===n&&void 0!==s?2*(o-a/p*u):o-a/p*u);return g/p}function Yc(t,e){var r,n,i,o,a=e.resolution,s=function(t,e){var r=new Array(t.order),n=new Float64Array(t.order),i=new Float64Array(t.order),o=0,a={},s=Qc(e.weighted,e.attributes.weight),u=0;return t.forEachNode((function(t,n){a[t]=u,r[u++]=e.communities?e.communities[t]:n[e.attributes.community]})),t.forEachDirectedEdge((function(t,e,r,u){var h=s(e);o+=h,i[a[r]]+=h,n[a[u]]+=h})),{getWeight:s,communities:r,weightedInDegrees:n,weightedOutDegrees:i,M:o}}(t,e),u=s.communities,h=s.weightedInDegrees,c=s.weightedOutDegrees,l=s.M,d=t.nodes(),f=0;for(r=0,i=t.order;r<i;r++)for(n=0;n<i;n++)u[r]===u[n]&&(o=t.directedEdge(d[r],d[n]),f+=s.getWeight(o)-h[r]*c[n]/l*a);return f/l}function Zc(t,e){var r=e.resolution,n=function(t,e){var r={},n={},i={};return e.communities&&(r=e.communities),t.forEachNode((function(t,o){var a;if(e.communities?a=r[t]:(a=o[e.attributes.community],r[t]=a),void 0===a)throw new Error('graphology-metrics/modularity: the "'+t+'" node is not in the partition.');n[a]=0,i[a]=0})),{communities:r,totalWeights:n,internalWeights:i}}(t,e),i=0,o=n.totalWeights,a=n.internalWeights,s=n.communities,u=Qc(e.weighted,e.attributes.weight);t.forEachUndirectedEdge((function(t,e,r,n){var h=u(e);i+=h;var c=s[r],l=s[n];o[c]+=h,o[l]+=h,c===l&&(a[c]+=2*h)}));var h=0,c=2*i;for(var l in a)h+=a[l]/c-Math.pow(o[l]/c,2)*r;return h}function tl(t,e){var r=e.resolution,n=function(t,e){var r={},n={},i={},o={};return e.communities&&(r=e.communities),t.forEachNode((function(t,a){var s;if(e.communities?s=r[t]:(s=a[e.attributes.community],r[t]=s),void 0===s)throw new Error('graphology-metrics/modularity: the "'+t+'" node is not in the partition.');n[s]=0,i[s]=0,o[s]=0})),{communities:r,totalInWeights:n,totalOutWeights:i,internalWeights:o}}(t,e),i=0,o=n.totalInWeights,a=n.totalOutWeights,s=n.internalWeights,u=n.communities,h=Qc(e.weighted,e.attributes.weight);t.forEachDirectedEdge((function(t,e,r,n){var c=h(e);i+=c;var l=u[r],d=u[n];a[l]+=c,o[d]+=c,l===d&&(s[l]+=c)}));var c=0;for(var l in s)c+=s[l]/i-o[l]*a[l]/Math.pow(i,2)*r;return c}function el(t,e){if(!Vc(t))throw new Error("graphology-metrics/modularity: given graph is not a valid graphology instance.");if(0===t.size)throw new Error("graphology-metrics/modularity: cannot compute modularity of an empty graph.");if(t.multi)throw new Error("graphology-metrics/modularity: cannot compute modularity of a multi graph. Cast it to a simple one beforehand.");var r=$c(t);if("mixed"===r)throw new Error("graphology-metrics/modularity: cannot compute modularity of a mixed graph.");return e=Kc({},e||{},Jc),"directed"===r?tl(t,e):Zc(t,e)}var rl=el;rl.sparse=el,rl.dense=function(t,e){if(!Vc(t))throw new Error("graphology-metrics/modularity: given graph is not a valid graphology instance.");if(0===t.size)throw new Error("graphology-metrics/modularity: cannot compute modularity of an empty graph.");if(t.multi)throw new Error("graphology-metrics/modularity: cannot compute modularity of a multi graph. Cast it to a simple one beforehand.");var r=$c(t);if("mixed"===r)throw new Error("graphology-metrics/modularity: cannot compute modularity of a mixed graph.");return e=Kc({},e||{},Jc),"directed"===r?Yc(t,e):Xc(t,e)},rl.undirectedDelta=function(t,e,r,n){return n/(2*t)-e*r/(t*t*2)},rl.directedDelta=function(t,e,r,n,i,o){return o/t-(i*e+n*r)/(t*t)};var nl=rl,il=ho,ol="weight";function al(t,e,r,n,i){if(!il(n))throw new Error("graphology-metrics/"+t+": the given graph is not a valid graphology instance.");if("edges"!==r&&"undirected"===n.type)throw new Error("graphology-metrics/"+t+": cannot compute "+t+" on an undirected graph.");var o=null;5===arguments.length&&"object"!=typeof arguments[4]?o=arguments[4]:6===arguments.length&&(o=arguments[4],i=arguments[5]);var a,s,u,h,c,l=(i=i||{}).attributes||{},d=l.weight||ol,f=l.weightedDegree||t;if(o){for(s=0,h=0,c=(a=n[r](o)).length;h<c;h++)"number"==typeof(u=n.getEdgeAttribute(a[h],d))&&(s+=u);return e?void n.setNodeAttribute(o,f,s):s}var g,p,y,v=n.nodes(),m={};for(h=0,c=v.length;h<c;h++){for(g=v[h],s=0,p=0,y=(a=n[r](g)).length;p<y;p++)"number"==typeof(u=n.getEdgeAttribute(a[p],d))&&(s+=u);e?n.setNodeAttribute(g,f,s):m[g]=s}if(!e)return m}var sl=al.bind(null,"weightedDegree",!1,"edges"),ul=al.bind(null,"weightedInDegree",!1,"inEdges"),hl=al.bind(null,"weightedOutDegree",!1,"outEdges");sl.assign=al.bind(null,"weightedDegree",!0,"edges"),ul.assign=al.bind(null,"weightedInDegree",!0,"inEdges"),hl.assign=al.bind(null,"weightedOutDegree",!0,"outEdges"),sl.weightedDegree=sl,sl.weightedInDegree=ul,sl.weightedOutDegree=hl;var cl=sl,ll=ho;_u.centrality=Cu,_u.density=Ya,_u.diameter=function(t){if(!cc(t))throw new Error("graphology-metrics/diameter: given graph is not a valid graphology instance.");if(0===t.size)return 1/0;for(var e=-1/0,r=0,n=t.nodes(),i=0,o=n.length;i<o&&((r=lc(t,n[i]))>e&&(e=r),e!==1/0);i++);return e},_u.eccentricity=hc,_u.extent=pc,_u.layoutQuality=yc,_u.modularity=nl,_u.weightedDegree=cl,_u.weightedSize=function(t,e){if(!ll(t))throw new Error("graphology-metrics/weighted-size: the given graph is not a valid graphology instance.");e=e||"weight";var r,n,i,o=t.edges(),a=0;for(n=0,i=o.length;n<i;n++)"number"==typeof(r=t.getEdgeAttribute(o[n],e))&&(a+=r);return a};var dl=_u,fl={},gl={addEdge:function(t,e,r,n,i,o){e?null==r?t.addUndirectedEdge(n,i,o):t.addUndirectedEdgeWithKey(r,n,i,o):null==r?t.addDirectedEdge(n,i,o):t.addDirectedEdgeWithKey(r,n,i,o)},copyEdge:function(t,e,r,n,i,o){o=Object.assign({},o),e?null==r?t.addUndirectedEdge(n,i,o):t.addUndirectedEdgeWithKey(r,n,i,o):null==r?t.addDirectedEdge(n,i,o):t.addDirectedEdgeWithKey(r,n,i,o)},mergeEdge:function(t,e,r,n,i,o){e?null==r?t.mergeUndirectedEdge(n,i,o):t.mergeUndirectedEdgeWithKey(r,n,i,o):null==r?t.mergeDirectedEdge(n,i,o):t.mergeDirectedEdgeWithKey(r,n,i,o)},updateEdge:function(t,e,r,n,i,o){e?null==r?t.updateUndirectedEdge(n,i,o):t.updateUndirectedEdgeWithKey(r,n,i,o):null==r?t.updateDirectedEdge(n,i,o):t.updateDirectedEdgeWithKey(r,n,i,o)}},pl=ho,yl=gl.copyEdge,vl={copyNode:function(t,e,r){r=Object.assign({},r),t.addNode(e,r)}},ml=ho,bl=vl.copyNode,wl=gl.copyEdge,El=ho,Al=ho,xl=ho,zl=ho;fl.reverse=function(t){if(!pl(t))throw new Error("graphology-operators/reverse: invalid graph.");var e=t.emptyCopy();return t.forEachUndirectedEdge((function(t,r,n,i,o,a,s,u){yl(e,!0,u?null:t,n,i,r)})),t.forEachDirectedEdge((function(t,r,n,i,o,a,s,u){yl(e,!1,u?null:t,i,n,r)})),e},fl.subgraph=function(t,e){if(!ml(t))throw new Error("graphology-operators/subgraph: invalid graph instance.");var r=t.nullCopy(),n=e;if(Array.isArray(e)){if(0===e.length)return r;e=new Set(e)}if(e instanceof Set){if(0===e.size)return r;n=function(t){return e.has(t)};var i=e;e=new Set,i.forEach((function(t){e.add(""+t)}))}if("function"!=typeof n)throw new Error("graphology-operators/subgraph: invalid nodes. Expecting an array or a set or a filtering function.");if("function"==typeof e){if(t.forEachNode((function(t,e){n(t,e)&&bl(r,t,e)})),0===r.order)return r}else e.forEach((function(e){if(!t.hasNode(e))throw new Error('graphology-operators/subgraph: the "'+e+'" node was not found in the graph.');bl(r,e,t.getNodeAttributes(e))}));return t.forEachEdge((function(t,e,i,o,a,s,u,h){n(i,a)&&(o===i||n(o,s))&&wl(r,u,h?null:t,i,o,e)})),r},fl.toDirected=function(t,e){if(!El(t))throw new Error("graphology-operators/to-directed: expecting a valid graphology instance.");"function"==typeof e&&(e={mergeEdge:e});var r="function"==typeof(e=e||{}).mergeEdge?e.mergeEdge:null;if("directed"===t.type)return t.copy();var n=t.emptyCopy({type:"directed"});return t.forEachDirectedEdge((function(t,e,r,i){n.addDirectedEdge(r,i,Object.assign({},e))})),t.forEachUndirectedEdge((function(e,i,o,a){var s="mixed"===t.type&&n.edge(o,a),u="mixed"===t.type&&n.edge(a,o);s?n.replaceEdgeAttributes(s,r(n.getEdgeAttributes(s),i)):n.addDirectedEdge(o,a,Object.assign({},i)),u?n.replaceEdgeAttributes(u,r(n.getEdgeAttributes(u),i)):n.addDirectedEdge(a,o,Object.assign({},i))})),n},fl.toSimple=function(t){if(!Al(t))throw new Error("graphology-operators/to-simple: expecting a valid graphology instance.");if(!t.multi)return t.copy();var e=t.emptyCopy({multi:!1});return t.forEachDirectedEdge((function(r,n,i,o){e.hasDirectedEdge(i,o)||e.importEdge(t.exportEdge(r))})),t.forEachUndirectedEdge((function(r,n,i,o){e.hasUndirectedEdge(i,o)||e.importEdge(t.exportEdge(r))})),e},fl.toUndirected=function(t,e){if(!xl(t))throw new Error("graphology-operators/to-undirected: expecting a valid graphology instance.");"function"==typeof e&&(e={mergeEdge:e});var r="function"==typeof(e=e||{}).mergeEdge?e.mergeEdge:null;if("undirected"===t.type)return t.copy();var n=t.emptyCopy({type:"undirected"});return t.forEachUndirectedEdge((function(t,e,r,i){n.addUndirectedEdge(r,i,Object.assign({},e))})),t.forEachDirectedEdge((function(t,e,i,o){var a=n.edge(i,o);a?r&&n.replaceEdgeAttributes(a,r(n.getEdgeAttributes(a),e)):n.addUndirectedEdge(i,o,Object.assign({},e))})),n},fl.union=function(t,e){if(!zl(t)||!zl(e))throw new Error("graphology-operators/union: invalid graph.");if(t.multi!==e.multi)throw new Error("graphology-operators/union: both graph should be simple or multi.");var r=t.copy();return r.import(e,!0),r};var Ml=fl,Sl=function(t){return null!==t&&"object"==typeof t&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode&&"boolean"==typeof t.multi},Nl=_h,jl={attributes:{pagerank:"pagerank",weight:"weight"},alpha:.85,maxIterations:100,tolerance:1e-6,weighted:!1};function kl(t,e,r){if(!Sl(e))throw new Error("graphology-pagerank: the given graph is not a valid graphology instance.");if(e.multi)throw new Error("graphology-pagerank: the pagerank algorithm does not work with MultiGraphs.");var n,i,o,a,s,u,h,c,l,d,f,g,p,y=(r=Nl(r,jl)).attributes.pagerank,v=r.attributes.weight,m=r.alpha,b=r.maxIterations,w=r.tolerance,E=r.weighted,A=e.order,x=1/A,z={},M=[],S=e.nodes(),N={},j={},k=0;for(d=0;d<A;d++){if(z[c=S[d]]=x,E)for(u=0,f=0,p=(n=e.undirectedEdges().concat(e.outEdges(c))).length;f<p;f++)l=n[f],u+=e.getEdgeAttribute(l,v)||1;else u=e.undirectedDegree(c)+e.outDegree(c);j[c]=u,0===u&&M.push(c)}for(n=e.edges(),d=0,g=e.size;d<g;d++)l=n[d],u=j[c=e.source(l)],i=E&&e.getEdgeAttribute(l,v)||1,N[l]=i/u;for(;k<b;){for(h in a=z,z={},a)z[h]=0;for(o=0,d=0,g=M.length;d<g;d++)o+=a[M[d]];for(o*=m,d=0;d<A;d++){for(c=S[d],f=0,p=(n=e.undirectedEdges(c).concat(e.outEdges(c))).length;f<p;f++)l=n[f],z[e.opposite(c,l)]+=m*a[c]*N[l];z[c]+=o*x+(1-m)*x}for(c in s=0,z)s+=Math.abs(z[c]-a[c]);if(s<A*w){if(t)for(c in z)e.setNodeAttribute(c,y,z[c]);return z}k++}throw Error("graphology-pagerank: failed to converge.")}var Dl=kl.bind(null,!1);Dl.assign=kl.bind(null,!0);var Ol=Dl,_l=ho,Cl=fh,Wl="weight";function Pl(t,e){return t[0]>e[0]?1:t[0]<e[0]?-1:t[1]>e[1]?1:t[1]<e[1]?-1:t[2]>e[2]?1:t[2]<e[2]?-1:0}function Il(t,e){return t[0]>e[0]?1:t[0]<e[0]?-1:t[1]>e[1]?1:t[1]<e[1]?-1:t[2]>e[2]?1:t[2]<e[2]?-1:t[3]>e[3]?1:t[3]<e[3]?-1:0}var Ul={bidirectional:function(t,e,r,n){return function(t,e,r,n){if(e=""+e,r=""+r,!_l(t))throw new Error("graphology-shortest-path/dijkstra: invalid graphology instance.");if(e&&!t.hasNode(e))throw new Error('graphology-shortest-path/dijkstra: the "'+e+'" source node does not exist in the given graph.');if(r&&!t.hasNode(r))throw new Error('graphology-shortest-path/dijkstra: the "'+r+'" target node does not exist in the given graph.');if(n=n||Wl,e===r)return[0,[e]];var i=[{},{}],o=[{},{}],a=[new Cl(Pl),new Cl(Pl)],s=[{},{}];o[0][e]=[e],o[1][r]=[r],s[0][e]=0,s[1][r]=0;var u,h,c,l,d,f,g,p,y,v,m,b=[],w=1/0,E=0,A=1;for(a[0].push([0,E++,e]),a[1].push([0,E++,r]);a[0].size&&a[1].size;)if(l=(u=a[A=1-A].pop())[0],!((d=u[2])in i[A])){if(i[A][d]=l,d in i[1-A])return[w,b];for(p=0,y=(h=1===A?t.inboundEdges(d):t.outboundEdges(d)).length;p<y;p++){if(g=h[p],f=t.opposite(d,g),c=i[A][d]+(v=g,m=void 0,"number"!=typeof(m=t.getEdgeAttribute(v,n))||isNaN(m)?1:m),f in i[A]&&c<i[A][f])throw Error("graphology-shortest-path/dijkstra: contradictory paths found. Do some of your edges have a negative weight?");(!(f in s[A])||c<s[A][f])&&(s[A][f]=c,a[A].push([c,E++,f]),o[A][f]=o[A][d].concat(f),f in s[0]&&f in s[1]&&(l=s[0][f]+s[1][f],(0===b.length||w>l)&&(w=l,b=o[0][f].concat(o[1][f].slice(0,-1).reverse()))))}}return[1/0,null]}(t,e,r,n)[1]},singleSource:function(t,e,r){var n={};return function(t,e,r,n,i,o){if(!_l(t))throw new Error("graphology-shortest-path/dijkstra: invalid graphology instance.");if(i&&!t.hasNode(i))throw new Error('graphology-shortest-path/dijkstra: the "'+i+'" target node does not exist in the given graph.');r=r||Wl;var a,s,u,h,c,l,d,f,g,p,y,v,m,b={},w={},E=new Cl(Pl),A=0;for(f=0,p=e.length;f<p;f++)w[h=e[f]]=0,E.push([0,A++,h]),o&&(o[h]=[h]);for(;E.size;)if(d=(s=E.pop())[0],!((h=s[2])in b)){if(b[h]=d,h===i)break;for(g=0,y=(a=t.outboundEdges(h)).length;g<y;g++)if(l=a[g],c=t.opposite(h,l),v=l,m=void 0,u=("number"!=typeof(m=t.getEdgeAttribute(v,r))||isNaN(m)?1:m)+b[h],!(n&&u>n)){if(c in b&&u<b[c])throw Error("graphology-shortest-path/dijkstra: contradictory paths found. Do some of your edges have a negative weight?");(!(c in w)||u<w[c])&&(w[c]=u,E.push([u,A++,c]),o&&(o[c]=o[h].concat(c)))}}}(t,[e],r,0,null,n),n},brandes:function(t,e,r){e=""+e,r=r||Wl;var n,i,o,a,s,u,h,c,l,d,f=[],g={},p={},y=t.nodes();for(l=0,d=y.length;l<d;l++)g[u=y[l]]=[],p[u]=0;var v={};p[e]=1;var m={};m[e]=0;var b=0,w=new Cl(Il);for(w.push([0,b++,e,e]);w.size;)if(a=(i=w.pop())[0],o=i[2],!((u=i[3])in v))for(p[u]+=p[o],f.push(u),v[u]=a,l=0,d=(n=t.outboundEdges(u)).length;l<d;l++)c=n[l],h=t.opposite(u,c),s=a+(t.getEdgeAttribute(c,r)||1),h in v||h in m&&!(s<m[h])?s===m[h]&&(p[h]+=p[u],g[h].push(u)):(m[h]=s,w.push([s,b++,u,h]),p[h]=0,g[h]=[u]);return[f,g,p]}},Tl=ac;Tl.dijkstra=Ul,Tl.unweighted=Tl;var Ll=Tl,Fl={},ql=function(t){return null!==t&&"object"==typeof t&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode&&"boolean"==typeof t.multi};function Rl(){this.set=new Set,this.stack=[],this.size=0}function Bl(){this.set=new Set,this.stack=[],this.size=0}function Gl(t,e){var r=[];return t.forEachOutboundEdge(e,(function(t,n,i,o){r.push([t,e===i?o:i])})),r}function Hl(t,e){var r,n={};t.forEachOutboundEdge(e,(function(t,i,o,a){(r=e===o?a:o)in n||(n[r]=[]),n[r].push(t)}));var i=[];for(r in n)i.push([n[r],r]);return i}Rl.prototype.has=function(t){return this.set.has(t)},Rl.prototype.push=function(t){this.stack.push(t),this.set.add(t),this.size++},Rl.prototype.pop=function(){this.set.delete(this.stack.pop()),this.size--},Rl.prototype.path=function(t){return this.stack.concat(t)},Rl.of=function(t){var e=new Rl;return e.push(t),e},Bl.prototype.has=function(t){return this.set.has(t)},Bl.prototype.push=function(t){this.stack.push(t),this.set.add(t[1]),this.size++},Bl.prototype.pop=function(){this.set.delete(this.stack.pop()[1]),this.size--},Bl.prototype.path=function(t){return this.stack.slice(1).map((function(t){return t[0]})).concat([t[0]])},Bl.of=function(t){var e=new Bl;return e.push([null,t]),e},Fl.allSimplePaths=function(t,e,r){if(!ql(t))throw new Error("graphology-simple-path.allSimplePaths: expecting a graphology instance.");if(!t.hasNode(e))throw new Error('graphology-simple-path.allSimplePaths: expecting: could not find source node "'+e+'" in the graph.');if(!t.hasNode(r))throw new Error('graphology-simple-path.allSimplePaths: expecting: could not find target node "'+r+'" in the graph.');for(var n,i,o=(e=""+e)===(r=""+r),a=[t.outboundNeighbors(e)],s=Rl.of(o?"§SOURCE§":e),u=[];0!==a.length;)if(i=a[a.length-1].pop()){if(s.has(i))continue;i===r&&(n=s.path(i),o&&(n[0]=e),u.push(n)),s.push(i),s.has(r)?s.pop():a.push(t.outboundNeighbors(i))}else a.pop(),s.pop();return u},Fl.allSimpleEdgePaths=function(t,e,r){if(!ql(t))throw new Error("graphology-simple-path.allSimpleEdgePaths: expecting a graphology instance.");if(t.multi)throw new Error("graphology-simple-path.allSimpleEdgePaths: not implemented for multi graphs.");if(!t.hasNode(e))throw new Error('graphology-simple-path.allSimpleEdgePaths: expecting: could not find source node "'+e+'" in the graph.');if(!t.hasNode(r))throw new Error('graphology-simple-path.allSimpleEdgePaths: expecting: could not find target node "'+r+'" in the graph.');for(var n,i,o,a=(e=""+e)===(r=""+r),s=[Gl(t,e)],u=Bl.of(a?"§SOURCE§":e),h=[];0!==s.length;)if(i=s[s.length-1].pop()){if(o=i[1],u.has(o))continue;o===r&&(n=u.path(i),h.push(n)),u.push(i),u.has(r)?u.pop():s.push(Gl(t,o))}else s.pop(),u.pop();return h},Fl.allSimpleEdgeGroupPaths=function(t,e,r){if(!ql(t))throw new Error("graphology-simple-path.allSimpleEdgeGroupPaths: expecting a graphology instance.");if(!t.hasNode(e))throw new Error('graphology-simple-path.allSimpleEdgeGroupPaths: expecting: could not find source node "'+e+'" in the graph.');if(!t.hasNode(r))throw new Error('graphology-simple-path.allSimpleEdgeGroupPaths: expecting: could not find target node "'+r+'" in the graph.');for(var n,i,o,a=(e=""+e)===(r=""+r),s=[Hl(t,e)],u=Bl.of(a?"§SOURCE§":e),h=[];0!==s.length;)if(i=s[s.length-1].pop()){if(o=i[1],u.has(o))continue;o===r&&(n=u.path(i),h.push(n)),u.push(i),u.has(r)?u.pop():s.push(Hl(t,o))}else s.pop(),u.pop();return h};var Kl=Fl,Vl={},$l={},Jl={};Jl.TraversalRecord=function(t,e,r){this.node=t,this.attributes=e,this.depth=r};var Ql=ho,Xl=Ku,Yl=Jl.TraversalRecord;$l.bfs=function(t,e){if(!Ql(t))throw new Error("graphology-traversal/bfs: expecting a graphology instance.");if("function"!=typeof e)throw new Error("graphology-traversal/bfs: given callback is not a function.");if(0!==t.order){var r,n,i=new Set,o=new Xl(Array,t.order);t.forEachNode((function(s,u){if(!i.has(s))for(i.add(s),o.push(new Yl(s,u,0));0!==o.size;)r=o.shift(),n=r.depth,e(r.node,r.attributes,n),t.forEachOutboundNeighbor(r.node,a)}))}function a(t,e){i.has(t)||(i.add(t),o.push(new Yl(t,e,n+1)))}},$l.bfsFromNode=function(t,e,r){if(!Ql(t))throw new Error("graphology-traversal/dfs: expecting a graphology instance.");if("function"!=typeof r)throw new Error("graphology-traversal/dfs: given callback is not a function.");if(0!==t.order){e=""+e;var n,i,o=new Set,a=new Xl(Array,t.order);for(o.add(e),a.push(new Yl(e,t.getNodeAttributes(e),0));0!==a.size;)i=a.shift(),n=i.depth,r(i.node,i.attributes,n),t.forEachOutboundNeighbor(i.node,s)}function s(t,e){o.has(t)||(o.add(t),a.push(new Yl(t,e,n+1)))}};var Zl={},td=ho,ed=Jl.TraversalRecord;Zl.dfs=function(t,e){if(!td(t))throw new Error("graphology-traversal/dfs: expecting a graphology instance.");if("function"!=typeof e)throw new Error("graphology-traversal/dfs: given callback is not a function.");if(0!==t.order){var r,n,i=new Set,o=[];t.forEachNode((function(s,u){if(!i.has(s))for(i.add(s),o.push(new ed(s,u,0));0!==o.length;)n=o.pop(),r=n.depth,e(n.node,n.attributes,r),t.forEachOutboundNeighbor(n.node,a)}))}function a(t,e){i.has(t)||(i.add(t),o.push(new ed(t,e,r+1)))}},Zl.dfsFromNode=function(t,e,r){if(!td(t))throw new Error("graphology-traversal/dfs: expecting a graphology instance.");if("function"!=typeof r)throw new Error("graphology-traversal/dfs: given callback is not a function.");if(0!==t.order){e=""+e;var n,i,o=new Set,a=[];for(o.add(e),a.push(new ed(e,t.getNodeAttributes(e),0));0!==a.length;)i=a.pop(),n=i.depth,r(i.node,i.attributes,n),t.forEachOutboundNeighbor(i.node,s)}function s(t,e){o.has(t)||(o.add(t),a.push(new ed(t,e,n+1)))}},function(t){var e,r=$l,n=Zl;for(e in r)t[e]=r[e];for(e in n)t[e]=n[e]}(Vl);var rd=Vl,nd={},id=function(t){return null!==t&&"function"==typeof t&&"object"==typeof t.prototype&&"function"==typeof t.prototype.addUndirectedEdgeWithKey&&"function"==typeof t.prototype.dropNode},od=gl.copyEdge,ad=gl.copyEdge,sd=ho,ud=gl,hd=vl;nd.addEdge=ud.addEdge,nd.copyEdge=ud.copyEdge,nd.mergeEdge=ud.mergeEdge,nd.updateEdge=ud.updateEdge,nd.copyNode=hd.updateNode,nd.inferType=lo,nd.isGraph=ho,nd.isGraphConstructor=id,nd.mergeClique=function(t,e){var r,n,i,o,a;if(0!==e.length)for(i=0,a=e.length;i<a;i++)for(r=e[i],o=i+1;o<a;o++)n=e[o],t.mergeEdge(r,n)},nd.mergeCycle=function(t,e){var r,n,i,o;if(0!==e.length&&(t.mergeNode(e[0]),1!==e.length)){for(i=1,o=e.length;i<o;i++)r=e[i-1],n=e[i],t.mergeEdge(r,n);t.mergeEdge(n,e[0])}},nd.mergePath=function(t,e){var r,n,i,o;if(0!==e.length)for(t.mergeNode(e[0]),i=1,o=e.length;i<o;i++)r=e[i-1],n=e[i],t.mergeEdge(r,n)},nd.mergeStar=function(t,e){if(0!==e.length){var r,n,i,o=e[0];for(t.mergeNode(o),n=1,i=e.length;n<i;n++)r=e[n],t.mergeEdge(o,r)}},nd.renameGraphKeys=function(t,e,r){void 0===e&&(e={}),void 0===r&&(r={});var n,i,o=t.nullCopy();return t.forEachNode((function(t,r){var n=e[t];void 0===n&&(n=t),o.addNode(n,r)})),t.forEach((function(t,a,s,u,h,c,l){t!==n&&(n=t,void 0===(i=e[t])&&(i=t));var d=e[a];void 0===d&&(d=a);var f=r[h];void 0===f&&(f=h),od(o,l,f,i,d,c)})),o},nd.updateGraphKeys=function(t,e,r){var n,i,o=t.nullCopy();return t.forEachNode((function(t,r){var n=e?e(t,r):t;o.addNode(n,r)})),t.forEach((function(t,a,s,u,h,c,l,d){t!==n&&(n=t,i=e?e(t,s):t);var f=e?e(a,u):a,g=r?r(h,c,t,a,s,u,l,d):h;ad(o,l,g,i,f,c)})),o},nd.memoizedForEach=function(t,e,r){if(!sd(t))throw new Error("graphology-utils/memoized-for-each: expecting a valid graphology instance.");var n,i;t.forEach((function(t,o,a,s,u,h,c,l){n!==t&&(n=t,i=e(t,a)),r(i,t,o,a,s,u,h,c,l)}))};var cd=nd,ld=function(){var t,e,r={};!function(){var t=10;r.exports=function(e,r,n){var i,o,a,s,u,h,c,l,d,f,g,p,y,v,m,b,w,E,A,x,z,M,S,N=r.length,j=n.length,k=e.adjustSizes,D=e.barnesHutTheta*e.barnesHutTheta,O=[];for(a=0;a<N;a+=t)r[a+4]=r[a+2],r[a+5]=r[a+3],r[a+2]=0,r[a+3]=0;if(e.outboundAttractionDistribution){for(g=0,a=0;a<N;a+=t)g+=r[a+6];g/=N/t}if(e.barnesHutOptimize){var _,C,W,P=1/0,I=-1/0,U=1/0,T=-1/0;for(a=0;a<N;a+=t)P=Math.min(P,r[a+0]),I=Math.max(I,r[a+0]),U=Math.min(U,r[a+1]),T=Math.max(T,r[a+1]);var L=I-P,F=T-U;for(L>F?T=(U-=(L-F)/2)+L:I=(P-=(F-L)/2)+F,O[0]=-1,O[1]=(P+I)/2,O[2]=(U+T)/2,O[3]=Math.max(I-P,T-U),O[4]=-1,O[5]=-1,O[6]=0,O[7]=0,O[8]=0,i=1,a=0;a<N;a+=t)for(o=0,W=3;;){if(!(O[o+5]>=0)){if(O[o+0]<0){O[o+0]=a;break}if(O[o+5]=9*i,l=O[o+3]/2,O[(d=O[o+5])+0]=-1,O[d+1]=O[o+1]-l,O[d+2]=O[o+2]-l,O[d+3]=l,O[d+4]=d+9,O[d+5]=-1,O[d+6]=0,O[d+7]=0,O[d+8]=0,O[(d+=9)+0]=-1,O[d+1]=O[o+1]-l,O[d+2]=O[o+2]+l,O[d+3]=l,O[d+4]=d+9,O[d+5]=-1,O[d+6]=0,O[d+7]=0,O[d+8]=0,O[(d+=9)+0]=-1,O[d+1]=O[o+1]+l,O[d+2]=O[o+2]-l,O[d+3]=l,O[d+4]=d+9,O[d+5]=-1,O[d+6]=0,O[d+7]=0,O[d+8]=0,O[(d+=9)+0]=-1,O[d+1]=O[o+1]+l,O[d+2]=O[o+2]+l,O[d+3]=l,O[d+4]=O[o+4],O[d+5]=-1,O[d+6]=0,O[d+7]=0,O[d+8]=0,i+=4,_=r[O[o+0]+0]<O[o+1]?r[O[o+0]+1]<O[o+2]?O[o+5]:O[o+5]+9:r[O[o+0]+1]<O[o+2]?O[o+5]+18:O[o+5]+27,O[o+6]=r[O[o+0]+6],O[o+7]=r[O[o+0]+0],O[o+8]=r[O[o+0]+1],O[_+0]=O[o+0],O[o+0]=-1,_===(C=r[a+0]<O[o+1]?r[a+1]<O[o+2]?O[o+5]:O[o+5]+9:r[a+1]<O[o+2]?O[o+5]+18:O[o+5]+27)){if(W--){o=_;continue}W=3;break}O[C+0]=a;break}_=r[a+0]<O[o+1]?r[a+1]<O[o+2]?O[o+5]:O[o+5]+9:r[a+1]<O[o+2]?O[o+5]+18:O[o+5]+27,O[o+7]=(O[o+7]*O[o+6]+r[a+0]*r[a+6])/(O[o+6]+r[a+6]),O[o+8]=(O[o+8]*O[o+6]+r[a+1]*r[a+6])/(O[o+6]+r[a+6]),O[o+6]+=r[a+6],o=_}}if(e.barnesHutOptimize){for(p=e.scalingRatio,a=0;a<N;a+=t)for(o=0;;)if(O[o+5]>=0){if(b=Math.pow(r[a+0]-O[o+7],2)+Math.pow(r[a+1]-O[o+8],2),4*(f=O[o+3])*f/b<D){if(y=r[a+0]-O[o+7],v=r[a+1]-O[o+8],!0===k?b>0?(w=p*r[a+6]*O[o+6]/b,r[a+2]+=y*w,r[a+3]+=v*w):b<0&&(w=-p*r[a+6]*O[o+6]/Math.sqrt(b),r[a+2]+=y*w,r[a+3]+=v*w):b>0&&(w=p*r[a+6]*O[o+6]/b,r[a+2]+=y*w,r[a+3]+=v*w),(o=O[o+4])<0)break;continue}o=O[o+5]}else if((h=O[o+0])>=0&&h!==a&&(b=(y=r[a+0]-r[h+0])*y+(v=r[a+1]-r[h+1])*v,!0===k?b>0?(w=p*r[a+6]*r[h+6]/b,r[a+2]+=y*w,r[a+3]+=v*w):b<0&&(w=-p*r[a+6]*r[h+6]/Math.sqrt(b),r[a+2]+=y*w,r[a+3]+=v*w):b>0&&(w=p*r[a+6]*r[h+6]/b,r[a+2]+=y*w,r[a+3]+=v*w)),(o=O[o+4])<0)break}else for(p=e.scalingRatio,s=0;s<N;s+=t)for(u=0;u<s;u+=t)y=r[s+0]-r[u+0],v=r[s+1]-r[u+1],!0===k?(b=Math.sqrt(y*y+v*v)-r[s+8]-r[u+8])>0?(w=p*r[s+6]*r[u+6]/b/b,r[s+2]+=y*w,r[s+3]+=v*w,r[u+2]+=y*w,r[u+3]+=v*w):b<0&&(w=100*p*r[s+6]*r[u+6],r[s+2]+=y*w,r[s+3]+=v*w,r[u+2]-=y*w,r[u+3]-=v*w):(b=Math.sqrt(y*y+v*v))>0&&(w=p*r[s+6]*r[u+6]/b/b,r[s+2]+=y*w,r[s+3]+=v*w,r[u+2]-=y*w,r[u+3]-=v*w);for(d=e.gravity/e.scalingRatio,p=e.scalingRatio,a=0;a<N;a+=t)w=0,y=r[a+0],v=r[a+1],b=Math.sqrt(Math.pow(y,2)+Math.pow(v,2)),e.strongGravityMode?b>0&&(w=p*r[a+6]*d):b>0&&(w=p*r[a+6]*d/b),r[a+2]-=y*w,r[a+3]-=v*w;for(p=1*(e.outboundAttractionDistribution?g:1),c=0;c<j;c+=3)s=n[c+0],u=n[c+1],l=n[c+2],m=Math.pow(l,e.edgeWeightInfluence),y=r[s+0]-r[u+0],v=r[s+1]-r[u+1],!0===k?(b=Math.sqrt(Math.pow(y,2)+Math.pow(v,2)-r[s+8]-r[u+8]),e.linLogMode?e.outboundAttractionDistribution?b>0&&(w=-p*m*Math.log(1+b)/b/r[s+6]):b>0&&(w=-p*m*Math.log(1+b)/b):e.outboundAttractionDistribution?b>0&&(w=-p*m/r[s+6]):b>0&&(w=-p*m)):(b=Math.sqrt(Math.pow(y,2)+Math.pow(v,2)),e.linLogMode?e.outboundAttractionDistribution?b>0&&(w=-p*m*Math.log(1+b)/b/r[s+6]):b>0&&(w=-p*m*Math.log(1+b)/b):e.outboundAttractionDistribution?(b=1,w=-p*m/r[s+6]):(b=1,w=-p*m)),b>0&&(r[s+2]+=y*w,r[s+3]+=v*w,r[u+2]-=y*w,r[u+3]-=v*w);if(!0===k)for(a=0;a<N;a+=t)1!==r[a+9]&&((E=Math.sqrt(Math.pow(r[a+2],2)+Math.pow(r[a+3],2)))>10&&(r[a+2]=10*r[a+2]/E,r[a+3]=10*r[a+3]/E),A=r[a+6]*Math.sqrt((r[a+4]-r[a+2])*(r[a+4]-r[a+2])+(r[a+5]-r[a+3])*(r[a+5]-r[a+3])),x=Math.sqrt((r[a+4]+r[a+2])*(r[a+4]+r[a+2])+(r[a+5]+r[a+3])*(r[a+5]+r[a+3]))/2,z=.1*Math.log(1+x)/(1+Math.sqrt(A)),M=r[a+0]+r[a+2]*(z/e.slowDown),r[a+0]=M,S=r[a+1]+r[a+3]*(z/e.slowDown),r[a+1]=S);else for(a=0;a<N;a+=t)1!==r[a+9]&&(A=r[a+6]*Math.sqrt((r[a+4]-r[a+2])*(r[a+4]-r[a+2])+(r[a+5]-r[a+3])*(r[a+5]-r[a+3])),x=Math.sqrt((r[a+4]+r[a+2])*(r[a+4]+r[a+2])+(r[a+5]+r[a+3])*(r[a+5]+r[a+3]))/2,z=r[a+7]*Math.log(1+x)/(1+Math.sqrt(A)),r[a+7]=Math.min(1,Math.sqrt(z*(Math.pow(r[a+2],2)+Math.pow(r[a+3],2))/(1+Math.sqrt(A)))),M=r[a+0]+r[a+2]*(z/e.slowDown),r[a+0]=M,S=r[a+1]+r[a+3]*(z/e.slowDown),r[a+1]=S);return{}}}();var n=r.exports;self.addEventListener("message",(function(r){var i=r.data;t=new Float32Array(i.nodes),i.edges&&(e=new Float32Array(i.edges)),n(i.settings,t,e),self.postMessage({nodes:t.buffer},[t.buffer])}))},dd=ho,fd=gu,gd=pu;function pd(t,e){if(e=e||{},!dd(t))throw new Error("graphology-layout-forceatlas2/worker: the given graph is not a valid graphology instance.");var r=fd.assign({},gd,e.settings),n=fd.validateSettings(r);if(n)throw new Error("graphology-layout-forceatlas2/worker: "+n.message);this.worker=null,this.graph=t,this.settings=r,this.matrices=null,this.running=!1,this.killed=!1,this.handleMessage=this.handleMessage.bind(this);var i=void 0,o=this;this.handleGraphUpdate=function(){o.worker&&o.worker.terminate(),i&&clearTimeout(i),i=setTimeout((function(){i=void 0,o.spawnWorker()}),0)},t.on("nodeAdded",this.handleGraphUpdate),t.on("edgeAdded",this.handleGraphUpdate),t.on("nodeDropped",this.handleGraphUpdate),t.on("edgeDropped",this.handleGraphUpdate),this.spawnWorker()}pd.prototype.spawnWorker=function(){this.worker&&this.worker.terminate(),this.worker=fd.createWorker(ld),this.worker.addEventListener("message",this.handleMessage),this.running&&(this.running=!1,this.start())},pd.prototype.handleMessage=function(t){if(this.running){var e=new Float32Array(t.data.nodes);fd.assignLayoutChanges(this.graph,e),this.matrices.nodes=e,this.askForIterations()}},pd.prototype.askForIterations=function(t){var e=this.matrices,r={settings:this.settings,nodes:e.nodes.buffer},n=[e.nodes.buffer];return t&&(r.edges=e.edges.buffer,n.push(e.edges.buffer)),this.worker.postMessage(r,n),this},pd.prototype.start=function(){if(this.killed)throw new Error("graphology-layout-forceatlas2/worker.start: layout was killed.");return this.running||(this.matrices=fd.graphToByteArrays(this.graph),this.running=!0,this.askForIterations(!0)),this},pd.prototype.stop=function(){return this.running=!1,this},pd.prototype.kill=function(){if(this.killed)return this;this.running=!1,this.killed=!0,this.matrices=null,this.worker.terminate(),this.graph.removeListener("nodeAdded",this.handleGraphUpdate),this.graph.removeListener("edgeAdded",this.handleGraphUpdate),this.graph.removeListener("nodeDropped",this.handleGraphUpdate),this.graph.removeListener("edgeDropped",this.handleGraphUpdate)};var yd=pd,vd={},md={cast:function(t,e){switch(t){case"boolean":e="true"===e;break;case"integer":case"long":case"float":case"double":e=+e;break;case"liststring":e=e?e.split("|"):[]}return e}},bd=/["'<>&\s]/g;md.sanitizeTagName=function(t){return t.replace(bd,"").trim()};var wd=id,Ed=gl.mergeEdge,Ad=md.cast;function xd(t,e){var r=t.getElementsByTagName("viz:"+e)[0];return r||(r=t.getElementsByTagNameNS("viz",e)[0]),r||(r=t.getElementsByTagName(e)[0]),r}function zd(t){for(var e,r,n,i,o={},a={},s=0,u=t.length;s<u;s++)o[n=(e=t[s]).getAttribute("id")||e.getAttribute("for")]={id:n,type:e.getAttribute("type")||"string",title:(i=+n,i==i&&e.getAttribute("title")||n)},(r=e.getElementsByTagName("default")[0])&&(a[o[n].title]=Ad(o[n].type,r.textContent));return[o,a]}function Md(t,e,r){var n={},i=r.getAttribute("label"),o=r.getAttribute("weight");i&&(n.label=i),o&&(n.weight=+o);for(var a,s,u,h=r.getElementsByTagName("attvalue"),c=0,l=h.length;c<l;c++)n[t[s=(a=h[c]).getAttribute("id")||a.getAttribute("for")].title]=Ad(t[s].type,a.getAttribute("value"));for(u in e)u in n||(n[u]=e[u]);var d,f,g,p=xd(r,"color");return p&&(n.color=function(t){var e=t.getAttribute("a"),r=t.getAttribute("r"),n=t.getAttribute("g"),i=t.getAttribute("b");return e?"rgba("+r+","+n+","+i+","+e+")":"rgb("+r+","+n+","+i+")"}(p)),(p=xd(r,"size"))&&(n.size=+p.getAttribute("value")),(p=xd(r,"position"))&&(d=p.getAttribute("x"),f=p.getAttribute("y"),g=p.getAttribute("z"),d&&(n.x=+d),f&&(n.y=+f),g&&(n.z=+g)),(p=xd(r,"shape"))&&(n.shape=p.getAttribute("value")),(p=xd(r,"thickness"))&&(n.thickness=+p.getAttribute("value")),n}var Sd=function(t,e){return function(r,n){var i,o,a,s,u,h,c,l,d,f=n;if(!wd(r))throw new Error("graphology-gexf/parser: invalid Graph constructor.");if("string"==typeof n&&(f=(new t).parseFromString(n,"application/xml")),!(f instanceof e))throw new Error("graphology-gexf/parser: source should either be a XML document or a string.");var g=f.getElementsByTagName("graph")[0],p=f.getElementsByTagName("meta")[0],y=p&&p.childNodes||[],v=f.getElementsByTagName("node"),m=f.getElementsByTagName("edge"),b=f.getElementsByTagName("attributes"),w=[],E=[];for(l=0,d=b.length;l<d;l++)"node"===(i=b[l]).getAttribute("class")?w=i.getElementsByTagName("attribute"):"edge"===i.getAttribute("class")&&(E=i.getElementsByTagName("attribute"));var A=g.getAttribute("defaultedgetype")||"undirected";"mutual"===A&&(A="undirected");var x=(o=zd(w))[0],z=o[1],M=(o=zd(E))[0],S=o[1],N=new r({type:m[0]?m[0].getAttribute("type")||A:"mixed"}),j=function(t){for(var e,r={},n=0,i=t.length;n<i;n++)"#text"!==(e=t[n]).nodeName&&e.textContent.trim()&&(r[e.tagName.toLowerCase()]=e.textContent);return r}(y),k=p&&p.getAttribute("lastmodifieddate");for(N.replaceAttributes(j),k&&N.setAttribute("lastModifiedDate",k),l=0,d=v.length;l<d;l++)i=v[l],N.addNode(i.getAttribute("id"),Md(x,z,i));for(l=0,d=m.length;l<d;l++)u=(i=m[l]).getAttribute("id"),a=i.getAttribute("type")||A,h=i.getAttribute("source"),c=i.getAttribute("target"),s=Md(M,S,i),a!==N.type&&"mixed"!==N.type&&N.upgradeToMixed(),!N.multi&&("directed"===a&&N.hasDirectedEdge(h,c)||N.hasUndirectedEdge(h,c))&&N.upgradeToMulti(),Ed(N,"directed"!==a,u||null,h,c,s);return N}}(DOMParser,Document);function Nd(t){return"number"!=typeof t&&!t}function jd(t){if("string"==typeof t)return t;if("number"==typeof t)return t+"";if("function"==typeof t)return t();if(t instanceof kd)return t.toString();throw Error("Bad Parameter")}function kd(t,e){if(!(this instanceof kd))return new kd;this.name_regex=/[_:A-Za-z][-._:A-Za-z0-9]*/,this.indent=!!t,this.indentString=this.indent&&"string"==typeof t?t:" ",this.output="",this.stack=[],this.tags=0,this.attributes=0,this.attribute=0,this.texts=0,this.comment=0,this.dtd=0,this.root="",this.pi=0,this.cdata=0,this.started_write=!1,this.writer,this.writer_encoding="UTF-8",this.writer="function"==typeof e?e:function(t,e){this.output+=t}}kd.prototype={toString:function(){return this.flush(),this.output},indenter:function(){if(this.indent){this.write("\n");for(var t=1;t<this.tags;t++)this.write(this.indentString)}},write:function(){for(var t=0;t<arguments.length;t++)this.writer(arguments[t],this.writer_encoding)},flush:function(){for(var t=this.tags;t>0;t--)this.endElement();this.tags=0},startDocument:function(t,e,r){return this.tags||this.attributes||(this.startPI("xml"),this.startAttribute("version"),this.text("string"==typeof t?t:"1.0"),this.endAttribute(),"string"==typeof e&&(this.startAttribute("encoding"),this.text(e),this.endAttribute(),this.writer_encoding=e),r&&(this.startAttribute("standalone"),this.text("yes"),this.endAttribute()),this.endPI(),this.indent||this.write("\n")),this},endDocument:function(){return this.attributes&&this.endAttributes(),this},writeElement:function(t,e){return this.startElement(t).text(e).endElement()},writeElementNS:function(t,e,r,n){return n||(n=r),this.startElementNS(t,e,r).text(n).endElement()},startElement:function(t){if(!(t=jd(t)).match(this.name_regex))throw Error("Invalid Parameter");if(0===this.tags&&this.root&&this.root!==t)throw Error("Invalid Parameter");return this.attributes&&this.endAttributes(),++this.tags,this.texts=0,this.stack.length>0&&(this.stack[this.stack.length-1].containsTag=!0),this.stack.push({name:t,tags:this.tags}),this.started_write&&this.indenter(),this.write("<",t),this.startAttributes(),this.started_write=!0,this},startElementNS:function(t,e,r){if(t=jd(t),e=jd(e),!t.match(this.name_regex))throw Error("Invalid Parameter");if(!e.match(this.name_regex))throw Error("Invalid Parameter");return this.attributes&&this.endAttributes(),++this.tags,this.texts=0,this.stack.length>0&&(this.stack[this.stack.length-1].containsTag=!0),this.stack.push({name:t+":"+e,tags:this.tags}),this.started_write&&this.indenter(),this.write("<",t+":"+e),this.startAttributes(),this.started_write=!0,this},endElement:function(){if(!this.tags)return this;var t=this.stack.pop();return this.attributes>0?(this.attribute&&(this.texts&&this.endAttribute(),this.endAttribute()),this.write("/"),this.endAttributes()):(t.containsTag&&this.indenter(),this.write("</",t.name,">")),--this.tags,this.texts=0,this},writeAttribute:function(t,e){return"function"==typeof e&&(e=e()),Nd(e)?this:this.startAttribute(t).text(e).endAttribute()},writeAttributeNS:function(t,e,r,n){return n||(n=r),"function"==typeof n&&(n=n()),Nd(n)?this:this.startAttributeNS(t,e,r).text(n).endAttribute()},startAttributes:function(){return this.attributes=1,this},endAttributes:function(){return this.attributes?(this.attribute&&this.endAttribute(),this.attributes=0,this.attribute=0,this.texts=0,this.write(">"),this):this},startAttribute:function(t){if(!(t=jd(t)).match(this.name_regex))throw Error("Invalid Parameter");return this.attributes||this.pi?(this.attribute||(this.attribute=1,this.write(" ",t,'="')),this):this},startAttributeNS:function(t,e,r){if(t=jd(t),e=jd(e),!t.match(this.name_regex))throw Error("Invalid Parameter");if(!e.match(this.name_regex))throw Error("Invalid Parameter");return this.attributes||this.pi?(this.attribute||(this.attribute=1,this.write(" ",t+":"+e,'="')),this):this},endAttribute:function(){return this.attribute?(this.attribute=0,this.texts=0,this.write('"'),this):this},text:function(t){return t=jd(t),this.tags||this.comment||this.pi||this.cdata?this.attributes&&this.attribute?(++this.texts,this.write(t.replace(/&/g,"&").replace(/</g,"<").replace(/"/g,""").replace(/\t/g,"	").replace(/\n/g,"
").replace(/\r/g,"
")),this):(this.attributes&&!this.attribute&&this.endAttributes(),this.comment||this.cdata?this.write(t):this.write(t.replace(/&/g,"&").replace(/</g,"<").replace(/>/g,">")),++this.texts,this.started_write=!0,this):this},writeComment:function(t){return this.startComment().text(t).endComment()},startComment:function(){return this.comment||(this.attributes&&this.endAttributes(),this.indenter(),this.write("\x3c!--"),this.comment=1,this.started_write=!0),this},endComment:function(){return this.comment?(this.write("--\x3e"),this.comment=0,this):this},writeDocType:function(t,e,r,n){return this.startDocType(t,e,r,n).endDocType()},startDocType:function(t,e,r,n){if(this.dtd||this.tags)return this;if(t=jd(t),e=e?jd(e):e,r=r?jd(r):r,n=n?jd(n):n,!t.match(this.name_regex))throw Error("Invalid Parameter");if(e&&!e.match(/^[\w\-][\w\s\-\/\+\:\.]*/))throw Error("Invalid Parameter");if(r&&!r.match(/^[\w\.][\w\-\/\\\:\.]*/))throw Error("Invalid Parameter");if(n&&!n.match(/[\w\s\<\>\+\.\!\#\-\?\*\,\(\)\|]*/))throw Error("Invalid Parameter");return e=e?' PUBLIC "'+e+'"':r?" SYSTEM":"",r=r?' "'+r+'"':"",n=n?" ["+n+"]":"",this.started_write&&this.indenter(),this.write("<!DOCTYPE ",t,e,r,n),this.root=t,this.dtd=1,this.started_write=!0,this},endDocType:function(){return this.dtd?(this.write(">"),this):this},writePI:function(t,e){return this.startPI(t).text(e).endPI()},startPI:function(t){if(!(t=jd(t)).match(this.name_regex))throw Error("Invalid Parameter");return this.pi||(this.attributes&&this.endAttributes(),this.started_write&&this.indenter(),this.write("<?",t),this.pi=1,this.started_write=!0),this},endPI:function(){return this.pi?(this.write("?>"),this.pi=0,this):this},writeCData:function(t){return this.startCData().text(t).endCData()},startCData:function(){return this.cdata||(this.attributes&&this.endAttributes(),this.indenter(),this.write("<![CDATA["),this.cdata=1,this.started_write=!0),this},endCData:function(){return this.cdata?(this.write("]]>"),this.cdata=0,this):this},writeRaw:function(t){return t=jd(t),this.tags||this.comment||this.pi||this.cdata?this.attributes&&this.attribute?(++this.texts,this.write(t.replace("&","&").replace('"',""")),this):(this.attributes&&!this.attribute&&this.endAttributes(),++this.texts,this.write(t),this.started_write=!0,this):this}};var Dd=ho,Od=lo,_d=kd,Cd=md.sanitizeTagName,Wd=new Set(["color","size","x","y","z","shape","thickness"]),Pd=/^\s*rgba?\s*\(/i,Id=/^\s*rgba?\s*\(\s*([0-9]*)\s*,\s*([0-9]*)\s*,\s*([0-9]*)\s*(?:,\s*([.0-9]*))?\)\s*$/;function Ud(t){if(!t||"string"!=typeof t)return{};if("#"===t[0])return 3===(t=t.slice(1)).length?{r:parseInt(t[0]+t[0],16),g:parseInt(t[1]+t[1],16),b:parseInt(t[2]+t[2],16)}:{r:parseInt(t[0]+t[1],16),g:parseInt(t[2]+t[3],16),b:parseInt(t[4]+t[5],16)};if(Pd.test(t)){var e={};return t=t.match(Id),e.r=+t[1],e.g=+t[2],e.b=+t[3],t[4]&&(e.a=+t[4]),e}return{}}function Td(t,e,r){var n,i={};for(n in r)"label"===n?i.label=r.label:"edge"===t&&"weight"===n?i.weight=r.weight:Wd.has(n)?(i.viz=i.viz||{},i.viz[n]=r[n]):(i.attributes=i.attributes||{},i.attributes[n]=r[n]);return i}function Ld(t){return null==t||""===t||t!=t}function Fd(t,e){return"liststring"===t&&Array.isArray(e)?e.join("|"):""+e}function qd(t){for(var e,r,n,i,o,a={},s=0,u=t.length;s<u;s++)if(e=t[s].attributes)for(n in e)i=e[n],o=void 0,"empty"!==(r=Ld(i)?"empty":Array.isArray(i)?"liststring":"boolean"==typeof i?"boolean":"object"==typeof i?"string":"number"==typeof i?i===(0|i)?(o=i)<=2147483647&&o>=-2147483647?"integer":"long":"double":"string")&&(a[n]?"integer"===a[n]&&"long"===r?a[n]=r:a[n]!==r&&(a[n]="string"):a[n]=r);return a}function Rd(t,e,r){var n;if(Object.keys(e).length){for(n in t.startElement("attributes"),t.writeAttribute("class",r),e)t.startElement("attribute"),t.writeAttribute("id",n),t.writeAttribute("title",n),t.writeAttribute("type",e[n]),t.endElement();t.endElement()}}function Bd(t,e,r,n){var i,o,a,s,u,h,c,l,d,f,g,p=!Object.keys(r).length;for(t.startElement(e+"s"),f=0,g=n.length;f<g;f++){if(h=(i=n[f]).attributes,l=i.viz,t.startElement(e),t.writeAttribute("id",i.key),"edge"===e&&((u=i.undirected?"undirected":"directed")!==t.defaultEdgeType&&t.writeAttribute("type",u),t.writeAttribute("source",i.source),t.writeAttribute("target",i.target),("number"==typeof(c=i.weight)&&!isNaN(c)||"string"==typeof c)&&t.writeAttribute("weight",i.weight)),i.label&&t.writeAttribute("label",i.label),!p&&h){for(o in t.startElement("attvalues"),r)if(o in h){if(Ld(s=h[o]))continue;t.startElement("attvalue"),t.writeAttribute("for",o),t.writeAttribute("value",Fd(r[o],s)),t.endElement()}t.endElement()}if(l){if(l.color){for(d in a=Ud(l.color),t.startElementNS("viz","color"),a)t.writeAttribute(d,a[d]);t.endElement()}"size"in l&&(t.startElementNS("viz","size"),t.writeAttribute("value",l.size),t.endElement()),("x"in l||"y"in l||"z"in l)&&(t.startElementNS("viz","position"),"x"in l&&t.writeAttribute("x",l.x),"y"in l&&t.writeAttribute("y",l.y),"z"in l&&t.writeAttribute("z",l.z),t.endElement()),l.shape&&(t.startElementNS("viz","shape"),t.writeAttribute("value",l.shape),t.endElement()),"thickness"in l&&(t.startElementNS("viz","thickness"),t.writeAttribute("value",l.thickness),t.endElement())}t.endElement()}t.endElement()}var Gd={encoding:"UTF-8",pretty:!0,formatNode:Td.bind(null,"node"),formatEdge:Td.bind(null,"edge")};vd.parse=Sd,vd.write=function(t,e){if(!Dd(t))throw new Error("graphology-gexf/writer: invalid graphology instance.");var r=!1!==(e=e||{}).pretty&&" ",n=e.formatNode||Gd.formatNode,i=e.formatEdge||Gd.formatEdge,o=new _d(r);o.startDocument("1.0",e.encoding||Gd.encoding),o.startElement("gexf"),o.writeAttribute("version","1.2"),o.writeAttribute("xmlns","http://www.gexf.net/1.2draft"),o.writeAttribute("xmlns:viz","http:///www.gexf.net/1.1draft/viz"),o.startElement("meta");var a,s=t.getAttributes();for(var u in s.lastModifiedDate&&o.writeAttribute("lastmodifieddate",s.lastModifiedDate),s)"lastModifiedDate"!==u&&(a=Cd(u))&&o.writeElement(a,s[u]);o.endElement(),o.startElement("graph");var h=Od(t);o.defaultEdgeType="mixed"===h?"directed":h,o.writeAttribute("defaultedgetype",o.defaultEdgeType);var c=function(t,e){var r=new Array(t.order),n=0;return t.forEachNode((function(t,i){var o=e(t,i);o.key=t,r[n++]=o})),r}(t,n),l=function(t,e){var r=new Array(t.size),n=0;return t.forEachEdge((function(t,i,o,a,s,u,h){var c=e(t,i);c.key=t,c.source=o,c.target=a,c.undirected=h,r[n++]=c})),r}(t,i),d=qd(c);Rd(o,d,"node");var f=qd(l);return Rd(o,f,"edge"),Bd(o,"node",d,c),Bd(o,"edge",f,l),o.toString()};var Hd={},Kd={};function Vd(t){return("0"+(0|t).toString(16)).slice(-2)}Kd.DEFAULT_FORMATTER=function(t){var e,r,n,i;return"number"==typeof t.r&&"number"==typeof t.g&&"number"==typeof t.b?((e=function(t){var e={};for(var r in t)"r"!==r&&"g"!==r&&"b"!==r&&(e[r]=t[r]);return e}(t)).color=(r=t.r,n=t.g,i=t.b,"#"+Vd(r)+Vd(n)+Vd(i)),e):t};var $d=id,Jd=Kd.DEFAULT_FORMATTER;function Qd(t){return+t}var Xd={boolean:function(t){return"true"===t.toLowerCase()},int:Qd,long:Qd,float:Qd,double:Qd,string:function(t){return t}};function Yd(t,e,r){var n,i,o,a,s,u=r.getElementsByTagName("data"),h={};for(i=0,o=u.length;i<o;i++)void 0===(s=t[a=(n=u[i]).getAttribute("key")])?h[a]=n.textContent:h[s.name]=s.cast(n.textContent);for(a in e)a in h||(h[a]=e[a]);return h}var Zd=function(t,e){return function(r,n){var i=n;if(!$d(r))throw new Error("graphology-graphml/parser: invalid Graph constructor.");if("string"==typeof n&&(i=(new t).parseFromString(n,"application/xml")),!(i instanceof e))throw new Error("graphology-gexf/parser: source should either be a XML document or a string.");var o=i.getElementsByTagName("graph")[0],a=function(t){for(var e,r=t.childNodes,n=[],i=0,o=r.length;i<o;i++)if(1===(e=r[i]).nodeType){if("data"!==e.tagName.toLowerCase())break;n.push(e)}return n}(o),s=i.getElementsByTagName("key"),u=i.getElementsByTagName("node"),h=i.getElementsByTagName("edge"),c=o.getAttribute("edgedefault")||"undirected",l=function(t){var e,r,n,i,o,a,s,u,h,c={graph:{},node:{},edge:{}},l={graph:{},node:{},edge:{}};for(e=0,r=t.length;e<r;e++)n=(s=t[e]).getAttribute("for")||"node",i=s.getAttribute("id"),o=s.getAttribute("attr.name"),a=s.getAttribute("attr.type")||"string",h=void 0,0!==(u=s.getElementsByTagName("default")).length&&(h=u[0].textContent),c[n][i]={name:o,cast:Xd[a]},void 0!==h&&(l[n][o]=h);return{models:c,defaults:l}}(s),d=new r({type:c}),f=o.getAttribute("id");f&&d.setAttribute("id",f);var g=i.createElement("graph");a.forEach((function(t){g.appendChild(t)}));var p,y,v,m,b,w,E,A,x,z=Yd(l.models.graph,l.defaults.graph,g);for(d.mergeAttributes(z),p=0,y=u.length;p<y;p++)m=(v=u[p]).getAttribute("id"),b=Yd(l.models.node,l.defaults.node,v),b=Jd(b),d.addNode(m,b);for(p=0,y=h.length;p<y;p++)m=(w=h[p]).getAttribute("id"),E=w.getAttribute("source"),A=w.getAttribute("target"),x="true"===w.getAttribute("directed")?"directed":c,b=Yd(l.models.edge,l.defaults.edge,w),b=Jd(b),"mixed"!==!d.type&&x!==d.type&&d.upgradeToMixed(),d.multi||("undirected"===x?d.hasUndirectedEdge(E,A)&&d.upgradeToMulti():d.hasDirectedEdge(E,A)&&d.upgradeToMulti()),"undirected"===x?m?d.addUndirectedEdgeWithKey(m,E,A,b):d.addUndirectedEdge(E,A,b):m?d.addDirectedEdgeWithKey(m,E,A,b):d.addDirectedEdge(E,A,b);return d}},tf=Zd(DOMParser,Document);Hd.parse=tf;var ef=r.assertions=ln,rf=r.communitiesLouvain=Zo,nf=r.components=ia,of=r.generators=gs,af=r.hits=As,sf=r.layout=du,uf=r.layoutForceAtlas2=Au,hf=r.layoutNoverlap=Ou,cf=r.metrics=dl,lf=r.operators=Ml,df=r.pagerank=Ol,ff=r.shortestPath=Ll,gf=r.simplePath=Kl,pf=r.traversal=rd,yf=r.utils=cd,vf=r.FA2Layout=yd,mf=r.gexf=vd,bf=r.graphml=Hd;t.FA2Layout=vf,t.assertions=ef,t.communitiesLouvain=rf,t.components=nf,t.default=r,t.generators=of,t.gexf=mf,t.graphml=bf,t.hits=af,t.layout=sf,t.layoutForceAtlas2=uf,t.layoutNoverlap=hf,t.metrics=cf,t.operators=lf,t.pagerank=df,t.shortestPath=ff,t.simplePath=gf,t.traversal=pf,t.utils=yf,Object.defineProperty(t,"__esModule",{value:!0})})); +//# sourceMappingURL=graphology-library.min.js.map diff --git a/html-template/vendor/graphology.min.js b/html-template/vendor/graphology.min.js new file mode 100644 index 0000000..2ef09f0 --- /dev/null +++ b/html-template/vendor/graphology.min.js @@ -0,0 +1,2 @@ +!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;t<n;t++)if(arguments[t])for(var r in arguments[t])e[r]=arguments[t][r];return e};function c(e,t,n,r){var i=e._nodes.get(t),o=null;return i?o="mixed"===r?i.out&&i.out[n]||i.undirected&&i.undirected[n]:"directed"===r?i.out&&i.out[n]:i.undirected&&i.undirected[n]:o}function s(t){return null!==t&&"object"===e(t)&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode}function h(t){return"object"===e(t)&&null!==t&&t.constructor===Object}function f(e){var t;for(t in e)return!1;return!0}function p(e,t,n){Object.defineProperty(e,t,{enumerable:!1,configurable:!1,writable:!0,value:n})}function g(e,t,n){var r={enumerable:!0,configurable:!0};"function"==typeof n?r.get=n:(r.value=n,r.writable=!1),Object.defineProperty(e,t,r)}function l(e){return!!h(e)&&!(e.attributes&&!Array.isArray(e.attributes))}"function"==typeof Object.assign&&(d=Object.assign);var y,v={exports:{}},b="object"==typeof Reflect?Reflect:null,w=b&&"function"==typeof b.apply?b.apply:function(e,t,n){return Function.prototype.apply.call(e,t,n)};y=b&&"function"==typeof b.ownKeys?b.ownKeys:Object.getOwnPropertySymbols?function(e){return Object.getOwnPropertyNames(e).concat(Object.getOwnPropertySymbols(e))}:function(e){return Object.getOwnPropertyNames(e)};var m=Number.isNaN||function(e){return e!=e};function _(){_.init.call(this)}v.exports=_,v.exports.once=function(e,t){return new Promise((function(n,r){function i(n){e.removeListener(t,o),r(n)}function o(){"function"==typeof e.removeListener&&e.removeListener("error",i),n([].slice.call(arguments))}j(e,t,o,{once:!0}),"error"!==t&&function(e,t,n){"function"==typeof e.on&&j(e,"error",t,n)}(e,i,{once:!0})}))},_.EventEmitter=_,_.prototype._events=void 0,_.prototype._eventsCount=0,_.prototype._maxListeners=void 0;var k=10;function G(e){if("function"!=typeof e)throw new TypeError('The "listener" argument must be of type Function. Received type '+typeof e)}function x(e){return void 0===e._maxListeners?_.defaultMaxListeners:e._maxListeners}function E(e,t,n,r){var i,o,a,u;if(G(n),void 0===(o=e._events)?(o=e._events=Object.create(null),e._eventsCount=0):(void 0!==o.newListener&&(e.emit("newListener",t,n.listener?n.listener:n),o=e._events),a=o[t]),void 0===a)a=o[t]=n,++e._eventsCount;else if("function"==typeof a?a=o[t]=r?[n,a]:[a,n]:r?a.unshift(n):a.push(n),(i=x(e))>0&&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;n<t.length;++n)t[n]=e[n].listener||e[n];return t}(i):D(i,i.length)}function N(e){var t=this._events;if(void 0!==t){var n=t[e];if("function"==typeof n)return 1;if(void 0!==n)return n.length}return 0}function D(e,t){for(var n=new Array(t),r=0;r<t;++r)n[r]=e[r];return n}function j(e,t,n,r){if("function"==typeof e.on)r.once?e.once(t,n):e.on(t,n);else{if("function"!=typeof e.addEventListener)throw new TypeError('The "emitter" argument must be of type EventEmitter. Received type '+typeof e);e.addEventListener(t,(function i(o){r.once&&e.removeEventListener(t,i),n(o)}))}}function U(e){Object.defineProperty(this,"_next",{writable:!1,enumerable:!1,value:e}),this.done=!1}Object.defineProperty(_,"defaultMaxListeners",{enumerable:!0,get:function(){return k},set:function(e){if("number"!=typeof e||e<0||m(e))throw new RangeError('The value of "defaultMaxListeners" is out of range. It must be a non-negative number. Received '+e+".");k=e}}),_.init=function(){void 0!==this._events&&this._events!==Object.getPrototypeOf(this)._events||(this._events=Object.create(null),this._eventsCount=0),this._maxListeners=this._maxListeners||void 0},_.prototype.setMaxListeners=function(e){if("number"!=typeof e||e<0||m(e))throw new RangeError('The value of "n" is out of range. It must be a non-negative number. Received '+e+".");return this._maxListeners=e,this},_.prototype.getMaxListeners=function(){return x(this)},_.prototype.emit=function(e){for(var t=[],n=1;n<arguments.length;n++)t.push(arguments[n]);var r="error"===e,i=this._events;if(void 0!==i)r=r&&void 0===i.error;else if(!r)return!1;if(r){var o;if(t.length>0&&(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<d;++n)w(c[n],this,t)}return!0},_.prototype.addListener=function(e,t){return E(this,e,t,!1)},_.prototype.on=_.prototype.addListener,_.prototype.prependListener=function(e,t){return E(this,e,t,!0)},_.prototype.once=function(e,t){return G(t),this.on(e,A(this,e,t)),this},_.prototype.prependOnceListener=function(e,t){return G(t),this.prependListener(e,A(this,e,t)),this},_.prototype.removeListener=function(e,t){var n,r,i,o,a;if(G(t),void 0===(r=this._events))return this;if(void 0===(n=r[e]))return this;if(n===t||n.listener===t)0==--this._eventsCount?this._events=Object.create(null):(delete r[e],r.removeListener&&this.emit("removeListener",e,n.listener||t));else if("function"!=typeof n){for(i=-1,o=n.length-1;o>=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<e.length;t++)e[t]=e[t+1];e.pop()}(n,i),1===n.length&&(r[e]=n[0]),void 0!==r.removeListener&&this.emit("removeListener",e,a||t)}return this},_.prototype.off=_.prototype.removeListener,_.prototype.removeAllListeners=function(e){var t,n,r;if(void 0===(n=this._events))return this;if(void 0===n.removeListener)return 0===arguments.length?(this._events=Object.create(null),this._eventsCount=0):void 0!==n[e]&&(0==--this._eventsCount?this._events=Object.create(null):delete n[e]),this;if(0===arguments.length){var i,o=Object.keys(n);for(r=0;r<o.length;++r)"removeListener"!==(i=o[r])&&this.removeAllListeners(i);return this.removeAllListeners("removeListener"),this._events=Object.create(null),this._eventsCount=0,this}if("function"==typeof(t=n[e]))this.removeListener(e,t);else if(void 0!==t)for(r=t.length-1;r>=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;t<n;t++)this.importNode(r[t],i)}if(e.edges){if(r=e.edges,!Array.isArray(r))throw new z("Graph.import: invalid edges. Expecting an array.");for(t=0,n=r.length;t<n;t++)this.importEdge(r[t],i)}return this},r.nullCopy=function(e){return new n(d({},this._options,e))},r.emptyCopy=function(e){var t=this.nullCopy(e);return this._nodes.forEach((function(e,n){var r=d({},e.attributes);e=new t.NodeDataClass(n,r),t._nodes.set(n,e)})),t},r.copy=function(){var e=this.emptyCopy();return this.forEachEdge((function(t,n,r,i,o,a,u,d){Ye(e,"copy",d,u,t,r,i,n)})),e},r.upgradeToMixed=function(){return"mixed"===this.type||(this._nodes.forEach((function(e){return e.upgradeToMixed()})),this._options.type="mixed",g(this,"type",this._options.type),p(this,"NodeDataClass",T)),this},r.upgradeToMulti=function(){return this.multi||(this._options.multi=!0,g(this,"multi",!0),(e=this)._nodes.forEach((function(t,n){if(t.out)for(var r in t.out){var i=new Set;i.add(t.out[r]),t.out[r]=i,e._nodes.get(r).in[n]=i}if(t.undirected)for(var o in t.undirected)if(!(o>n)){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})); +//# sourceMappingURL=graphology.umd.min.js.map diff --git a/html-template/vendor/sigma.min.js b/html-template/vendor/sigma.min.js new file mode 100644 index 0000000..3e6579f --- /dev/null +++ b/html-template/vendor/sigma.min.js @@ -0,0 +1 @@ +!function(t,e){"object"==typeof exports&&"object"==typeof module?module.exports=e():"function"==typeof define&&define.amd?define([],e):"object"==typeof exports?exports.Sigma=e():t.Sigma=e()}(self,(function(){return(()=>{var t={796:t=>{t.exports=function(t,e){var r=e.length;if(0!==r){var i=t.length;t.length+=r;for(var o=0;o<r;o++)t[i+o]=e[o]}}},187:t=>{"use strict";var e,r="object"==typeof Reflect?Reflect:null,i=r&&"function"==typeof r.apply?r.apply:function(t,e,r){return Function.prototype.apply.call(t,e,r)};e=r&&"function"==typeof r.ownKeys?r.ownKeys:Object.getOwnPropertySymbols?function(t){return Object.getOwnPropertyNames(t).concat(Object.getOwnPropertySymbols(t))}:function(t){return Object.getOwnPropertyNames(t)};var o=Number.isNaN||function(t){return t!=t};function n(){n.init.call(this)}t.exports=n,t.exports.once=function(t,e){return new Promise((function(r,i){function o(r){t.removeListener(e,n),i(r)}function n(){"function"==typeof t.removeListener&&t.removeListener("error",o),r([].slice.call(arguments))}g(t,e,n,{once:!0}),"error"!==e&&function(t,e,r){"function"==typeof t.on&&g(t,"error",e,{once:!0})}(t,o)}))},n.EventEmitter=n,n.prototype._events=void 0,n.prototype._eventsCount=0,n.prototype._maxListeners=void 0;var a=10;function s(t){if("function"!=typeof t)throw new TypeError('The "listener" argument must be of type Function. Received type '+typeof t)}function h(t){return void 0===t._maxListeners?n.defaultMaxListeners:t._maxListeners}function c(t,e,r,i){var o,n,a,c;if(s(r),void 0===(n=t._events)?(n=t._events=Object.create(null),t._eventsCount=0):(void 0!==n.newListener&&(t.emit("newListener",e,r.listener?r.listener:r),n=t._events),a=n[e]),void 0===a)a=n[e]=r,++t._eventsCount;else if("function"==typeof a?a=n[e]=i?[r,a]:[a,r]:i?a.unshift(r):a.push(r),(o=h(t))>0&&a.length>o&&!a.warned){a.warned=!0;var l=new Error("Possible EventEmitter memory leak detected. "+a.length+" "+String(e)+" listeners added. Use emitter.setMaxListeners() to increase limit");l.name="MaxListenersExceededWarning",l.emitter=t,l.type=e,l.count=a.length,c=l,console&&console.warn&&console.warn(c)}return t}function l(){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 u(t,e,r){var i={fired:!1,wrapFn:void 0,target:t,type:e,listener:r},o=l.bind(i);return o.listener=r,i.wrapFn=o,o}function d(t,e,r){var i=t._events;if(void 0===i)return[];var o=i[e];return void 0===o?[]:"function"==typeof o?r?[o.listener||o]:[o]:r?function(t){for(var e=new Array(t.length),r=0;r<e.length;++r)e[r]=t[r].listener||t[r];return e}(o):p(o,o.length)}function f(t){var e=this._events;if(void 0!==e){var r=e[t];if("function"==typeof r)return 1;if(void 0!==r)return r.length}return 0}function p(t,e){for(var r=new Array(e),i=0;i<e;++i)r[i]=t[i];return r}function g(t,e,r,i){if("function"==typeof t.on)i.once?t.once(e,r):t.on(e,r);else{if("function"!=typeof t.addEventListener)throw new TypeError('The "emitter" argument must be of type EventEmitter. Received type '+typeof t);t.addEventListener(e,(function o(n){i.once&&t.removeEventListener(e,o),r(n)}))}}Object.defineProperty(n,"defaultMaxListeners",{enumerable:!0,get:function(){return a},set:function(t){if("number"!=typeof t||t<0||o(t))throw new RangeError('The value of "defaultMaxListeners" is out of range. It must be a non-negative number. Received '+t+".");a=t}}),n.init=function(){void 0!==this._events&&this._events!==Object.getPrototypeOf(this)._events||(this._events=Object.create(null),this._eventsCount=0),this._maxListeners=this._maxListeners||void 0},n.prototype.setMaxListeners=function(t){if("number"!=typeof t||t<0||o(t))throw new RangeError('The value of "n" is out of range. It must be a non-negative number. Received '+t+".");return this._maxListeners=t,this},n.prototype.getMaxListeners=function(){return h(this)},n.prototype.emit=function(t){for(var e=[],r=1;r<arguments.length;r++)e.push(arguments[r]);var o="error"===t,n=this._events;if(void 0!==n)o=o&&void 0===n.error;else if(!o)return!1;if(o){var a;if(e.length>0&&(a=e[0]),a instanceof Error)throw a;var s=new Error("Unhandled error."+(a?" ("+a.message+")":""));throw s.context=a,s}var h=n[t];if(void 0===h)return!1;if("function"==typeof h)i(h,this,e);else{var c=h.length,l=p(h,c);for(r=0;r<c;++r)i(l[r],this,e)}return!0},n.prototype.addListener=function(t,e){return c(this,t,e,!1)},n.prototype.on=n.prototype.addListener,n.prototype.prependListener=function(t,e){return c(this,t,e,!0)},n.prototype.once=function(t,e){return s(e),this.on(t,u(this,t,e)),this},n.prototype.prependOnceListener=function(t,e){return s(e),this.prependListener(t,u(this,t,e)),this},n.prototype.removeListener=function(t,e){var r,i,o,n,a;if(s(e),void 0===(i=this._events))return this;if(void 0===(r=i[t]))return this;if(r===e||r.listener===e)0==--this._eventsCount?this._events=Object.create(null):(delete i[t],i.removeListener&&this.emit("removeListener",t,r.listener||e));else if("function"!=typeof r){for(o=-1,n=r.length-1;n>=0;n--)if(r[n]===e||r[n].listener===e){a=r[n].listener,o=n;break}if(o<0)return this;0===o?r.shift():function(t,e){for(;e+1<t.length;e++)t[e]=t[e+1];t.pop()}(r,o),1===r.length&&(i[t]=r[0]),void 0!==i.removeListener&&this.emit("removeListener",t,a||e)}return this},n.prototype.off=n.prototype.removeListener,n.prototype.removeAllListeners=function(t){var e,r,i;if(void 0===(r=this._events))return this;if(void 0===r.removeListener)return 0===arguments.length?(this._events=Object.create(null),this._eventsCount=0):void 0!==r[t]&&(0==--this._eventsCount?this._events=Object.create(null):delete r[t]),this;if(0===arguments.length){var o,n=Object.keys(r);for(i=0;i<n.length;++i)"removeListener"!==(o=n[i])&&this.removeAllListeners(o);return this.removeAllListeners("removeListener"),this._events=Object.create(null),this._eventsCount=0,this}if("function"==typeof(e=r[t]))this.removeListener(t,e);else if(void 0!==e)for(i=e.length-1;i>=0;i--)this.removeListener(t,e[i]);return this},n.prototype.listeners=function(t){return d(this,t,!0)},n.prototype.rawListeners=function(t){return d(this,t,!1)},n.listenerCount=function(t,e){return"function"==typeof t.listenerCount?t.listenerCount(e):f.call(t,e)},n.prototype.listenerCount=f,n.prototype.eventNames=function(){return this._eventsCount>0?e(this._events):[]}},886:(t,e,r)=>{var i=r(186);function o(t,e){if(!i(t))throw new Error("graphology-metrics/extent: the given graph is not a valid graphology instance.");var r,o,n,a=[].concat(e),s={};for(n=0;n<a.length;n++)o=a[n],s[o]=[1/0,-1/0];return t.forEachNode((function(t,e){for(n=0;n<a.length;n++)o=a[n],(r=e[o])<s[o][0]&&(s[o][0]=r),r>s[o][1]&&(s[o][1]=r)})),"string"==typeof e?s[e]:s}var n=o;n.nodeExtent=o,n.edgeExtent=function(t,e){if(!i(t))throw new Error("graphology-metrics/extent: the given graph is not a valid graphology instance.");var r,o,n,a=[].concat(e),s={};for(n=0;n<a.length;n++)o=a[n],s[o]=[1/0,-1/0];return t.forEachEdge((function(t,e){for(n=0;n<a.length;n++)o=a[n],(r=e[o])<s[o][0]&&(s[o][0]=r),r>s[o][1]&&(s[o][1]=r)})),"string"==typeof e?s[e]:s},t.exports=n},186:t=>{t.exports=function(t){return null!==t&&"object"==typeof t&&"function"==typeof t.addUndirectedEdgeWithKey&&"function"==typeof t.dropNode&&"boolean"==typeof t.multi}},973:(t,e,r)=>{"use strict";r.r(e),r.d(e,{default:()=>i});const i="precision mediump float;\n\nvarying vec4 v_color;\n\nvoid main(void) {\n gl_FragColor = v_color;\n}\n"},912:(t,e,r)=>{"use strict";r.r(e),r.d(e,{default:()=>i});const i="attribute vec2 a_position;\nattribute vec2 a_normal;\nattribute float a_thickness;\nattribute float a_radius;\nattribute vec4 a_color;\nattribute vec3 a_barycentric;\n\nuniform mat3 u_matrix;\nuniform float u_scale;\nuniform float u_cameraRatio;\nuniform float u_viewportRatio;\nuniform float u_thicknessRatio;\n\nvarying vec4 v_color;\n\nconst float arrowHeadLengthThicknessRatio = 2.5;\nconst float arrowHeadWidthLengthRatio = 0.66;\nconst float minThickness = 0.8;\nconst float bias = 255.0 / 254.0;\n\nvoid main() {\n\n // Computing thickness in screen space:\n float thickness = a_thickness * u_thicknessRatio * u_scale * u_viewportRatio / 2.0;\n thickness = max(thickness, minThickness * u_viewportRatio);\n\n float nodeRadius = a_radius * u_thicknessRatio * u_viewportRatio * u_cameraRatio;\n float arrowHeadLength = thickness * 2.0 * arrowHeadLengthThicknessRatio * u_cameraRatio;\n float arrowHeadHalfWidth = arrowHeadWidthLengthRatio * arrowHeadLength / 2.0;\n\n float da = a_barycentric.x;\n float db = a_barycentric.y;\n float dc = a_barycentric.z;\n\n vec2 delta = vec2(\n da * ((nodeRadius) * a_normal.y)\n + db * ((nodeRadius + arrowHeadLength) * a_normal.y + arrowHeadHalfWidth * a_normal.x)\n + dc * ((nodeRadius + arrowHeadLength) * a_normal.y - arrowHeadHalfWidth * a_normal.x),\n\n da * (-(nodeRadius) * a_normal.x)\n + db * (-(nodeRadius + arrowHeadLength) * a_normal.x + arrowHeadHalfWidth * a_normal.y)\n + dc * (-(nodeRadius + arrowHeadLength) * a_normal.x - arrowHeadHalfWidth * a_normal.y)\n );\n\n vec2 position = (u_matrix * vec3(a_position + delta, 1)).xy;\n\n gl_Position = vec4(position, 0, 1);\n\n // Extract the color:\n v_color = a_color;\n v_color.a *= bias;\n}\n"},620:(t,e,r)=>{"use strict";r.r(e),r.d(e,{default:()=>i});const i="attribute vec2 a_position;\nattribute vec2 a_normal;\nattribute float a_thickness;\nattribute vec4 a_color;\nattribute float a_radius;\n\nuniform mat3 u_matrix;\nuniform float u_scale;\nuniform float u_cameraRatio;\nuniform float u_viewportRatio;\nuniform float u_thicknessRatio;\n\nvarying vec4 v_color;\nvarying vec2 v_normal;\nvarying float v_thickness;\n\nconst float arrowHeadLengthThicknessRatio = 2.5;\nconst float minThickness = 0.8;\nconst float bias = 255.0 / 254.0;\n\nvoid main() {\n\n // Computing thickness in screen space:\n float thickness = a_thickness * u_thicknessRatio * u_scale * u_viewportRatio / 2.0;\n thickness = max(thickness, minThickness * u_viewportRatio);\n\n float direction = sign(a_radius);\n float nodeRadius = direction * a_radius * u_thicknessRatio * u_viewportRatio;\n float arrowHeadLength = thickness * 2.0 * arrowHeadLengthThicknessRatio;\n\n vec2 arrowHeadVector = vec2(-direction * a_normal.y, direction * a_normal.x);\n\n // Add normal vector to the position in screen space, but correct thickness first:\n vec2 position = a_position + a_normal * thickness * u_cameraRatio;\n // Add vector that corrects the arrow head length:\n position = position + arrowHeadVector * (arrowHeadLength + nodeRadius) * u_cameraRatio;\n // Apply camera\n position = (u_matrix * vec3(position, 1)).xy;\n\n gl_Position = vec4(position, 0, 1);\n\n v_normal = a_normal;\n v_thickness = thickness;\n\n // Extract the color:\n v_color = a_color;\n v_color.a *= bias;\n}\n"},498:(t,e,r)=>{"use strict";r.r(e),r.d(e,{default:()=>i});const i="precision mediump float;\n\nvarying vec4 v_color;\nvarying vec2 v_normal;\nvarying float v_thickness;\n\nconst float feather = 0.001;\nconst vec4 color0 = vec4(0.0, 0.0, 0.0, 0.0);\n\nvoid main(void) {\n float dist = length(v_normal) * v_thickness;\n\n float t = smoothstep(\n v_thickness - feather,\n v_thickness,\n dist\n );\n\n gl_FragColor = mix(v_color, color0, t);\n}\n"},223:(t,e,r)=>{"use strict";r.r(e),r.d(e,{default:()=>i});const i="attribute vec2 a_position;\nattribute vec2 a_normal;\nattribute float a_thickness;\nattribute vec4 a_color;\n\nuniform mat3 u_matrix;\nuniform float u_scale;\nuniform float u_cameraRatio;\nuniform float u_viewportRatio;\nuniform float u_thicknessRatio;\n\nvarying vec4 v_color;\nvarying vec2 v_normal;\nvarying float v_thickness;\n\nconst float minThickness = 0.8;\nconst float bias = 255.0 / 254.0;\n\nvoid main() {\n\n // Computing thickness in screen space:\n float thickness = a_thickness * u_thicknessRatio * u_scale * u_viewportRatio / 2.0;\n thickness = max(thickness, minThickness * u_viewportRatio);\n\n // Add normal vector to the position in screen space, but correct thickness first:\n vec2 position = (u_matrix * vec3(a_position + a_normal * thickness * u_cameraRatio, 1)).xy;\n\n gl_Position = vec4(position, 0, 1);\n\n v_normal = a_normal;\n v_thickness = thickness;\n\n // Extract the color:\n v_color = a_color;\n v_color.a *= bias;\n}\n"},262:(t,e,r)=>{"use strict";r.r(e),r.d(e,{default:()=>i});const i="precision mediump float;\n\nvarying vec4 v_color;\nvarying float v_border;\n\nconst float radius = 0.5;\n\nvoid main(void) {\n vec4 color0 = vec4(0.0, 0.0, 0.0, 0.0);\n vec2 m = gl_PointCoord - vec2(0.5, 0.5);\n float dist = radius - length(m);\n\n float t = 0.0;\n if (dist > v_border)\n t = 1.0;\n else if (dist > 0.0)\n t = dist / v_border;\n\n gl_FragColor = mix(color0, v_color, t);\n}\n"},106:(t,e,r)=>{"use strict";r.r(e),r.d(e,{default:()=>i});const i="attribute vec2 a_position;\nattribute float a_size;\nattribute vec4 a_color;\n\nuniform float u_ratio;\nuniform float u_scale;\nuniform mat3 u_matrix;\n\nvarying vec4 v_color;\nvarying float v_border;\n\nconst float bias = 255.0 / 254.0;\n\nvoid main() {\n\n gl_Position = vec4(\n (u_matrix * vec3(a_position, 1)).xy,\n 0,\n 1\n );\n\n // Multiply the point size twice:\n // - x SCALING_RATIO to correct the canvas scaling\n // - x 2 to correct the formulae\n gl_PointSize = a_size * u_ratio * u_scale * 2.0;\n\n v_border = (1.0 / u_ratio) * (0.5 / a_size);\n\n // Extract the color:\n v_color = a_color;\n v_color.a *= bias;\n}\n"},764:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)}),n=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0});var a=r(187),s=r(751),h=n(r(358)),c=r(928),l=1.5,u=function(t){function e(){var e=t.call(this)||this;return e.x=.5,e.y=.5,e.angle=0,e.ratio=1,e.nextFrame=null,e.previousState=null,e.enabled=!0,e.previousState=e.getState(),e}return o(e,t),e.from=function(t){return(new e).setState(t)},e.prototype.enable=function(){return this.enabled=!0,this},e.prototype.disable=function(){return this.enabled=!1,this},e.prototype.getState=function(){return{x:this.x,y:this.y,angle:this.angle,ratio:this.ratio}},e.prototype.hasState=function(t){return this.x===t.x&&this.y===t.y&&this.ratio===t.ratio&&this.angle===t.angle},e.prototype.getPreviousState=function(){var t=this.previousState;return t?{x:t.x,y:t.y,angle:t.angle,ratio:t.ratio}:null},e.prototype.isAnimated=function(){return!!this.nextFrame},e.prototype.setState=function(t){return this.enabled?(this.previousState=this.getState(),"number"==typeof t.x&&(this.x=t.x),"number"==typeof t.y&&(this.y=t.y),"number"==typeof t.angle&&(this.angle=t.angle),"number"==typeof t.ratio&&(this.ratio=t.ratio),this.hasState(this.previousState)||this.emit("updated",this.getState()),this):this},e.prototype.animate=function(t,e,r){var i=this;if(this.enabled){var o=Object.assign({},s.ANIMATE_DEFAULTS,e),n="function"==typeof o.easing?o.easing:h.default[o.easing],a=Date.now(),l=this.getState(),u=function(){var e=(Date.now()-a)/o.duration;if(e>=1)return i.nextFrame=null,i.setState(t),void(i.animationCallback&&(i.animationCallback.call(null),i.animationCallback=void 0));var r=n(e),s={};t.x&&(s.x=l.x+(t.x-l.x)*r),t.y&&(s.y=l.y+(t.y-l.y)*r),t.angle&&(s.angle=l.angle+(t.angle-l.angle)*r),t.ratio&&(s.ratio=l.ratio+(t.ratio-l.ratio)*r),i.setState(s),i.nextFrame=(0,c.requestFrame)(u)};this.nextFrame?((0,c.cancelFrame)(this.nextFrame),this.animationCallback&&this.animationCallback.call(null),this.nextFrame=(0,c.requestFrame)(u)):u(),this.animationCallback=r}},e.prototype.animatedZoom=function(t){if(t){if("number"==typeof t)return this.animate({ratio:this.ratio/t});this.animate({ratio:this.ratio/(t.factor||l)},t)}else this.animate({ratio:this.ratio/l})},e.prototype.animatedUnzoom=function(t){if(t){if("number"==typeof t)return this.animate({ratio:this.ratio*t});this.animate({ratio:this.ratio*(t.factor||l)},t)}else this.animate({ratio:this.ratio*l})},e.prototype.animatedReset=function(t){this.animate({x:.5,y:.5,ratio:1,angle:0},t)},e.prototype.copy=function(){return e.from(this.getState())},e}(a.EventEmitter);e.default=u},291:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)});Object.defineProperty(e,"__esModule",{value:!0}),e.getWheelDelta=e.getTouchCoords=e.getTouchesArray=e.getMouseCoords=e.getPosition=e.getY=e.getX=void 0;var n=r(187);function a(t){if(void 0!==t.offsetX)return t.offsetX;if(void 0!==t.clientX)return t.clientX;throw new Error("Captor: could not extract x from event.")}function s(t){if(void 0!==t.offsetY)return t.offsetY;if(void 0!==t.clientY)return t.clientY;throw new Error("Captor: could not extract y from event.")}function h(t){return{x:a(t),y:s(t)}}function c(t){for(var e=[],r=0,i=Math.min(t.length,2);r<i;r++)e.push(t[r]);return e}e.getX=a,e.getY=s,e.getPosition=h,e.getMouseCoords=function(t){return{x:a(t),y:s(t),clientX:t.clientX,clientY:t.clientY,ctrlKey:t.ctrlKey,metaKey:t.metaKey,altKey:t.altKey,shiftKey:t.shiftKey,preventDefault:t.preventDefault.bind(t),original:t}},e.getTouchesArray=c,e.getTouchCoords=function(t){return{touches:c(t.touches).map(h),ctrlKey:t.ctrlKey,metaKey:t.metaKey,altKey:t.altKey,shiftKey:t.shiftKey,preventDefault:t.preventDefault.bind(t),original:t}},e.getWheelDelta=function(t){if(void 0!==t.deltaY)return-3*t.deltaY/360;if(void 0!==t.detail)return t.detail/-9;throw new Error("Captor: could not extract delta from event.")};var l=function(t){function e(e,r){var i=t.call(this)||this;return i.container=e,i.renderer=r,i}return o(e,t),e}(n.EventEmitter);e.default=l},269:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)}),n=this&&this.__createBinding||(Object.create?function(t,e,r,i){void 0===i&&(i=r),Object.defineProperty(t,i,{enumerable:!0,get:function(){return e[r]}})}:function(t,e,r,i){void 0===i&&(i=r),t[i]=e[r]}),a=this&&this.__setModuleDefault||(Object.create?function(t,e){Object.defineProperty(t,"default",{enumerable:!0,value:e})}:function(t,e){t.default=e}),s=this&&this.__importStar||function(t){if(t&&t.__esModule)return t;var e={};if(null!=t)for(var r in t)"default"!==r&&Object.prototype.hasOwnProperty.call(t,r)&&n(e,t,r);return a(e,t),e};Object.defineProperty(e,"__esModule",{value:!0});var h=s(r(291)),c=function(t){function e(e,r){var i=t.call(this,e,r)||this;return i.enabled=!0,i.draggedEvents=0,i.downStartTime=null,i.lastMouseX=null,i.lastMouseY=null,i.isMouseDown=!1,i.isMoving=!1,i.movingTimeout=null,i.startCameraState=null,i.clicks=0,i.doubleClickTimeout=null,i.currentWheelDirection=0,i.handleClick=i.handleClick.bind(i),i.handleRightClick=i.handleRightClick.bind(i),i.handleDown=i.handleDown.bind(i),i.handleUp=i.handleUp.bind(i),i.handleMove=i.handleMove.bind(i),i.handleWheel=i.handleWheel.bind(i),i.handleOut=i.handleOut.bind(i),e.addEventListener("click",i.handleClick,!1),e.addEventListener("contextmenu",i.handleRightClick,!1),e.addEventListener("mousedown",i.handleDown,!1),e.addEventListener("mousemove",i.handleMove,!1),e.addEventListener("wheel",i.handleWheel,!1),e.addEventListener("mouseout",i.handleOut,!1),document.addEventListener("mouseup",i.handleUp,!1),i}return o(e,t),e.prototype.kill=function(){var t=this.container;t.removeEventListener("click",this.handleClick),t.removeEventListener("contextmenu",this.handleRightClick),t.removeEventListener("mousedown",this.handleDown),t.removeEventListener("mousemove",this.handleMove),t.removeEventListener("wheel",this.handleWheel),t.removeEventListener("mouseout",this.handleOut),document.removeEventListener("mouseup",this.handleUp)},e.prototype.handleClick=function(t){var e=this;if(this.enabled){if(this.clicks++,2===this.clicks)return this.clicks=0,"number"==typeof this.doubleClickTimeout&&(clearTimeout(this.doubleClickTimeout),this.doubleClickTimeout=null),this.handleDoubleClick(t);setTimeout((function(){e.clicks=0,e.doubleClickTimeout=null}),300),this.draggedEvents<3&&this.emit("click",(0,h.getMouseCoords)(t))}},e.prototype.handleRightClick=function(t){this.enabled&&this.emit("rightClick",(0,h.getMouseCoords)(t))},e.prototype.handleDoubleClick=function(t){if(this.enabled){var e=this.renderer.getCamera(),r=e.getState().ratio/2.2;return e.animate(this.renderer.getViewportZoomedState({x:(0,h.getX)(t),y:(0,h.getY)(t)},r),{easing:"quadraticInOut",duration:200}),t.preventDefault?t.preventDefault():t.returnValue=!1,t.stopPropagation(),!1}},e.prototype.handleDown=function(t){this.enabled&&(this.startCameraState=this.renderer.getCamera().getState(),this.lastMouseX=(0,h.getX)(t),this.lastMouseY=(0,h.getY)(t),this.draggedEvents=0,this.downStartTime=Date.now(),t.which,this.isMouseDown=!0,this.emit("mousedown",(0,h.getMouseCoords)(t)))},e.prototype.handleUp=function(t){var e=this;if(this.enabled&&this.isMouseDown){var r=this.renderer.getCamera();this.isMouseDown=!1,"number"==typeof this.movingTimeout&&(clearTimeout(this.movingTimeout),this.movingTimeout=null);var i=(0,h.getX)(t),o=(0,h.getY)(t),n=r.getState(),a=r.getPreviousState()||{x:0,y:0};this.isMoving?r.animate({x:n.x+3*(n.x-a.x),y:n.y+3*(n.y-a.y)},{duration:200,easing:"quadraticOut"}):this.lastMouseX===i&&this.lastMouseY===o||r.setState({x:n.x,y:n.y}),this.isMoving=!1,setTimeout((function(){e.draggedEvents=0,e.renderer.refresh()}),0),this.emit("mouseup",(0,h.getMouseCoords)(t))}},e.prototype.handleMove=function(t){var e=this;if(this.enabled){if(this.emit("mousemove",(0,h.getMouseCoords)(t)),this.isMouseDown){this.isMoving=!0,this.draggedEvents++,"number"==typeof this.movingTimeout&&clearTimeout(this.movingTimeout),this.movingTimeout=window.setTimeout((function(){e.movingTimeout=null,e.isMoving=!1}),100);var r=this.renderer.getCamera(),i=(0,h.getX)(t),o=(0,h.getY)(t),n=this.renderer.viewportToFramedGraph({x:this.lastMouseX,y:this.lastMouseY}),a=this.renderer.viewportToFramedGraph({x:i,y:o}),s=n.x-a.x,c=n.y-a.y,l=r.getState(),u=l.x+s,d=l.y+c;r.setState({x:u,y:d}),this.lastMouseX=i,this.lastMouseY=o}return t.preventDefault?t.preventDefault():t.returnValue=!1,t.stopPropagation(),!1}},e.prototype.handleWheel=function(t){var e=this;if(t.preventDefault?t.preventDefault():t.returnValue=!1,t.stopPropagation(),!this.enabled)return!1;var r=(0,h.getWheelDelta)(t);if(!r)return!1;var i=r>0?1/1.7:1.7,o=this.renderer.getCamera(),n=o.getState().ratio*i,a=r>0?1:-1,s=Date.now();return this.currentWheelDirection===a&&this.lastWheelTriggerTime&&s-this.lastWheelTriggerTime<50||(o.animate(this.renderer.getViewportZoomedState({x:(0,h.getX)(t),y:(0,h.getY)(t)},n),{easing:"quadraticOut",duration:250},(function(){e.currentWheelDirection=0})),this.currentWheelDirection=a,this.lastWheelTriggerTime=s),!1},e.prototype.handleOut=function(){},e}(h.default);e.default=c},508:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)}),n=this&&this.__createBinding||(Object.create?function(t,e,r,i){void 0===i&&(i=r),Object.defineProperty(t,i,{enumerable:!0,get:function(){return e[r]}})}:function(t,e,r,i){void 0===i&&(i=r),t[i]=e[r]}),a=this&&this.__setModuleDefault||(Object.create?function(t,e){Object.defineProperty(t,"default",{enumerable:!0,value:e})}:function(t,e){t.default=e}),s=this&&this.__importStar||function(t){if(t&&t.__esModule)return t;var e={};if(null!=t)for(var r in t)"default"!==r&&Object.prototype.hasOwnProperty.call(t,r)&&n(e,t,r);return a(e,t),e},h=this&&this.__read||function(t,e){var r="function"==typeof Symbol&&t[Symbol.iterator];if(!r)return t;var i,o,n=r.call(t),a=[];try{for(;(void 0===e||e-- >0)&&!(i=n.next()).done;)a.push(i.value)}catch(t){o={error:t}}finally{try{i&&!i.done&&(r=n.return)&&r.call(n)}finally{if(o)throw o.error}}return a};Object.defineProperty(e,"__esModule",{value:!0});var c=s(r(291)),l=function(t){function e(e,r){var i=t.call(this,e,r)||this;return i.enabled=!0,i.isMoving=!1,i.touchMode=0,i.handleStart=i.handleStart.bind(i),i.handleLeave=i.handleLeave.bind(i),i.handleMove=i.handleMove.bind(i),e.addEventListener("touchstart",i.handleStart,!1),e.addEventListener("touchend",i.handleLeave,!1),e.addEventListener("touchcancel",i.handleLeave,!1),e.addEventListener("touchmove",i.handleMove,!1),i}return o(e,t),e.prototype.kill=function(){var t=this.container;t.removeEventListener("touchstart",this.handleStart),t.removeEventListener("touchend",this.handleLeave),t.removeEventListener("touchcancel",this.handleLeave),t.removeEventListener("touchmove",this.handleMove)},e.prototype.getDimensions=function(){return{width:this.container.offsetWidth,height:this.container.offsetHeight}},e.prototype.dispatchRelatedMouseEvent=function(t,e,r,i){var o=r||(0,c.getPosition)(e.touches[0]),n=new MouseEvent(t,{clientX:o.x,clientY:o.y,altKey:e.altKey,ctrlKey:e.ctrlKey});(i||this.container).dispatchEvent(n)},e.prototype.handleStart=function(t){if(this.enabled){t.preventDefault(),1===t.touches.length&&this.dispatchRelatedMouseEvent("mousedown",t);var e=(0,c.getTouchesArray)(t.touches);if(this.isMoving=!0,this.touchMode=e.length,this.startCameraState=this.renderer.getCamera().getState(),this.startTouchesPositions=e.map(c.getPosition),this.lastTouchesPositions=this.startTouchesPositions,2===this.touchMode){var r=h(this.startTouchesPositions,2),i=r[0],o=i.x,n=i.y,a=r[1],s=a.x,l=a.y;this.startTouchesAngle=Math.atan2(l-n,s-o),this.startTouchesDistance=Math.sqrt(Math.pow(s-o,2)+Math.pow(l-n,2))}this.emit("touchdown",(0,c.getTouchCoords)(t))}},e.prototype.handleLeave=function(t){if(this.enabled){switch(t.preventDefault(),0===t.touches.length&&this.lastTouchesPositions&&this.lastTouchesPositions.length&&(this.dispatchRelatedMouseEvent("mouseup",t,this.lastTouchesPositions[0],document),this.dispatchRelatedMouseEvent("click",t,this.lastTouchesPositions[0])),this.movingTimeout&&(this.isMoving=!1,clearTimeout(this.movingTimeout)),this.touchMode){case 2:if(1===t.touches.length){this.handleStart(t),t.preventDefault();break}case 1:if(this.isMoving){var e=this.renderer.getCamera(),r=e.getState(),i=e.getPreviousState()||{x:0,y:0};e.animate({x:r.x+3*(r.x-i.x),y:r.y+3*(r.y-i.y)},{duration:200,easing:"quadraticOut"})}this.isMoving=!1,this.touchMode=0}this.emit("touchup",(0,c.getTouchCoords)(t))}},e.prototype.handleMove=function(t){var e,r=this;if(this.enabled){t.preventDefault(),1===t.touches.length&&this.dispatchRelatedMouseEvent("mousemove",t);var i=this.startCameraState,o=(0,c.getTouchesArray)(t.touches).map(c.getPosition);switch(this.lastTouchesPositions=o,this.isMoving=!0,this.movingTimeout&&clearTimeout(this.movingTimeout),this.movingTimeout=window.setTimeout((function(){r.isMoving=!1}),200),this.touchMode){case 1:var n=this.renderer.viewportToFramedGraph((this.startTouchesPositions||[])[0]),a=n.x,s=n.y,l=this.renderer.viewportToFramedGraph(o[0]),u=l.x,d=l.y;this.renderer.getCamera().setState({x:i.x+a-u,y:i.y+s-d});break;case 2:var f={},p=o[0],g=p.x,v=p.y,m=o[1],y=m.x,_=m.y,b=Math.atan2(_-v,y-g)-this.startTouchesAngle,x=Math.hypot(_-v,y-g)/this.startTouchesDistance;f.ratio=i.ratio/x,f.angle=i.angle+b;var w=this.getDimensions(),L=this.renderer.viewportToFramedGraph((this.startTouchesPositions||[])[0],{cameraState:i}),E=Math.min(w.width,w.height),A=E/w.width,P=E/w.height,T=f.ratio/E;d=v-E/2/P,u=(e=h([(u=g-E/2/A)*Math.cos(-f.angle)-d*Math.sin(-f.angle),d*Math.cos(-f.angle)+u*Math.sin(-f.angle)],2))[0],d=e[1],f.x=L.x-u*T,f.y=L.y+d*T,this.renderer.getCamera().setState(f)}this.emit("touchmove",(0,c.getTouchCoords)(t))}},e}(c.default);e.default=l},730:(t,e)=>{"use strict";Object.defineProperty(e,"__esModule",{value:!0}),e.edgeLabelsToDisplayFromNodes=e.LabelGrid=void 0;var r=function(){function t(t,e){this.key=t,this.size=e}return t.compare=function(t,e){return t.size>e.size?-1:t.size<e.size||t.key>e.key?1:-1},t}(),i=function(){function t(){this.width=0,this.height=0,this.cellSize=0,this.columns=0,this.rows=0,this.cells={}}return t.prototype.resizeAndClear=function(t,e){this.width=t.width,this.height=t.height,this.cellSize=e,this.columns=Math.ceil(t.width/e),this.rows=Math.ceil(t.height/e),this.cells={}},t.prototype.getIndex=function(t){var e=Math.floor(t.x/this.cellSize),r=Math.floor(t.y/this.cellSize);return e*this.columns+r},t.prototype.add=function(t,e,i){var o=new r(t,e),n=this.getIndex(i),a=this.cells[n];a||(a=[],this.cells[n]=a),a.push(o)},t.prototype.organize=function(){for(var t in this.cells)this.cells[t].sort(r.compare)},t.prototype.getLabelsToDisplay=function(t,e){var r=this.cellSize*this.cellSize,i=r/t/t*e/r,o=Math.ceil(i),n=[];for(var a in this.cells)for(var s=this.cells[a],h=0;h<Math.min(o,s.length);h++)n.push(s[h].key);return n},t}();e.LabelGrid=i,e.edgeLabelsToDisplayFromNodes=function(t){var e=t.graph,r=t.hoveredNode,i=t.highlightedNodes,o=t.displayedNodeLabels,n=[];return e.forEachEdge((function(t,e,a,s){(a===r||s===r||i.has(a)||i.has(s)||o.has(a)&&o.has(s))&&n.push(t)})),n}},134:function(t,e,r){"use strict";var i=this&&this.__read||function(t,e){var r="function"==typeof Symbol&&t[Symbol.iterator];if(!r)return t;var i,o,n=r.call(t),a=[];try{for(;(void 0===e||e-- >0)&&!(i=n.next()).done;)a.push(i.value)}catch(t){o={error:t}}finally{try{i&&!i.done&&(r=n.return)&&r.call(n)}finally{if(o)throw o.error}}return a},o=this&&this.__spreadArray||function(t,e,r){if(r||2===arguments.length)for(var i,o=0,n=e.length;o<n;o++)!i&&o in e||(i||(i=Array.prototype.slice.call(e,0,o)),i[o]=e[o]);return t.concat(i||Array.prototype.slice.call(e))},n=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0}),e.rectangleCollidesWithQuad=e.squareCollidesWithQuad=e.getCircumscribedAlignedRectangle=e.isRectangleAligned=void 0;var a=n(r(796)),s=!1;function h(t){return t.x1===t.x2||t.y1===t.y2}function c(t){var e=Math.sqrt(Math.pow(t.x2-t.x1,2)+Math.pow(t.y2-t.y1,2)),r=(t.y1-t.y2)*t.height/e,i=(t.x2-t.x1)*t.height/e,o={x:t.x1,y:t.y1},n={x:t.x2,y:t.y2},a={x:t.x1+r,y:t.y1+i},s={x:t.x2+r,y:t.y2+i},h=Math.min(o.x,n.x,a.x,s.x),c=Math.max(o.x,n.x,a.x,s.x),l=Math.min(o.y,n.y,a.y,s.y);return{x1:h,y1:l,x2:c,y2:l,height:Math.max(o.y,n.y,a.y,s.y)-l}}function l(t,e,r,i,o,n,a){return t<i+n&&t+r>i&&e<o+a&&e+r>o}function u(t,e,r,i,o,n,a,s){return t<o+a&&t+r>o&&e<n+s&&e+i>n}function d(t,e,r,i,o,n){var a=t<r+o/2;return e<i+n/2?a?1:2:a?3:4}e.isRectangleAligned=h,e.getCircumscribedAlignedRectangle=c,e.squareCollidesWithQuad=l,e.rectangleCollidesWithQuad=u;var f=function(){function t(t){var e;void 0===t&&(t={}),this.containers=((e={}).outside=[],e),this.cache=null,this.lastRectangle=null;var r=Math.pow(4,5);this.data=new Float32Array((4*r-1)/3*4),t.boundaries?this.resize(t.boundaries):this.resize({x:0,y:0,width:1,height:1})}return t.prototype.add=function(t,e,r,i){return function(t,e,r,i,o,n,a){for(var h=o-a,c=n-a,u=2*a,d=0,f=0;;){if(d>=5)return r[f]=r[f]||[],void r[f].push(i);var p=4*f+4,g=4*f+8,v=4*f+12,m=4*f+16,y=l(h,c,u,e[p+0],e[p+1],e[p+2],e[p+3]),_=l(h,c,u,e[g+0],e[g+1],e[g+2],e[g+3]),b=l(h,c,u,e[v+0],e[v+1],e[v+2],e[v+3]),x=l(h,c,u,e[m+0],e[m+1],e[m+2],e[m+3]),w=[y,_,b,x].reduce((function(t,e){return e?t+1:t}),0);if(0===w&&0===d)return r.outside.push(i),void(!s&&r.outside.length>=5&&(s=!0,console.warn("sigma/quadtree.insertNode: At least 5 nodes are outside the global quadtree zone. You might have a problem with the normalization function or the custom bounding box.")));if(0===w)throw new Error("sigma/quadtree.insertNode: no collision (level: "+d+", key: "+i+", x: "+o+", y: "+n+", size: "+a+").");if(3===w)throw new Error("sigma/quadtree.insertNode: 3 impossible collisions (level: "+d+", key: "+i+", x: "+o+", y: "+n+", size: "+a+").");if(w>1)return r[f]=r[f]||[],void r[f].push(i);d++,y&&(f=p),_&&(f=g),b&&(f=v),x&&(f=m)}}(0,this.data,this.containers,t,e,r,i),this},t.prototype.resize=function(t){this.clear(),this.data[0]=t.x,this.data[1]=t.y,this.data[2]=t.width,this.data[3]=t.height,function(t,e){for(var r=[0,0];r.length;){var i=r.pop(),o=r.pop(),n=4*o+4,a=4*o+8,s=4*o+12,h=4*o+16,c=e[o+0],l=e[o+1],u=e[o+2]/2,d=e[o+3]/2;e[n+0]=c,e[n+1]=l,e[n+2]=u,e[n+3]=d,e[a+0]=c+u,e[a+1]=l,e[a+2]=u,e[a+3]=d,e[s+0]=c,e[s+1]=l+d,e[s+2]=u,e[s+3]=d,e[h+0]=c+u,e[h+1]=l+d,e[h+2]=u,e[h+3]=d,i<4&&(r.push(h,i+1),r.push(s,i+1),r.push(a,i+1),r.push(n,i+1))}}(0,this.data)},t.prototype.clear=function(){var t;return this.containers=((t={}).outside=[],t),this},t.prototype.point=function(t,e){var r=this.containers.outside,n=0,a=0;do{this.containers[n]&&r.push.apply(r,o([],i(this.containers[n]),!1)),n=4*n+4*d(t,e,this.data[n+0],this.data[n+1],this.data[n+2],this.data[n+3]),a++}while(a<=5);return r},t.prototype.rectangle=function(t,e,r,n,s){var l,d=this.lastRectangle;return d&&t===d.x1&&r===d.x2&&e===d.y1&&n===d.y2&&s===d.height||(this.lastRectangle={x1:t,y1:e,x2:r,y2:n,height:s},h(this.lastRectangle)||(this.lastRectangle=c(this.lastRectangle)),this.cache=function(t,e,r,i,o,n,s){for(var h,c=[0,0],l=[];c.length;){var d=c.pop(),f=c.pop();if((h=r[f])&&(0,a.default)(l,h),!(d>=5)){var p=4*f+4,g=4*f+8,v=4*f+12,m=4*f+16,y=u(i,o,n,s,e[p+0],e[p+1],e[p+2],e[p+3]),_=u(i,o,n,s,e[g+0],e[g+1],e[g+2],e[g+3]),b=u(i,o,n,s,e[v+0],e[v+1],e[v+2],e[v+3]),x=u(i,o,n,s,e[m+0],e[m+1],e[m+2],e[m+3]);y&&c.push(p,d+1),_&&c.push(g,d+1),b&&c.push(v,d+1),x&&c.push(m,d+1)}}return l}(0,this.data,this.containers,t,e,Math.abs(t-r)||Math.abs(e-n),s),(l=this.cache).push.apply(l,o([],i(this.containers.outside),!1))),this.cache},t}();e.default=f},607:function(t,e,r){"use strict";var i=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0}),e.Sigma=e.MouseCaptor=e.QuadTree=e.Camera=void 0;var o=i(r(159));e.Sigma=o.default;var n=i(r(764));e.Camera=n.default;var a=i(r(134));e.QuadTree=a.default;var s=i(r(269));e.MouseCaptor=s.default,e.default=o.default},942:(t,e)=>{"use strict";Object.defineProperty(e,"__esModule",{value:!0}),e.default=function(t,e,r,i,o){var n=o.edgeLabelSize,a=o.edgeLabelFont,s=o.edgeLabelWeight,h=e.label;if(h){t.fillStyle=e.color,t.font=s+" "+n+"px "+a;var c,l=t.measureText(h).width,u=(r.x+i.x)/2,d=(r.y+i.y)/2,f=i.x-r.x,p=i.y-r.y,g=Math.sqrt(f*f+p*p);c=f>0?p>0?Math.acos(f/g):Math.asin(p/g):p>0?Math.acos(f/g)+Math.PI:Math.asin(f/g)+Math.PI/2,t.save(),t.translate(u,d),t.rotate(c),t.fillText(h,-l/2,e.size/2+n),t.restore()}}},61:function(t,e,r){"use strict";var i=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0});var o=i(r(622));e.default=function(t,e,r){var i=r.labelSize,n=r.labelFont,a=r.labelWeight;if(t.font=a+" "+i+"px "+n,t.fillStyle="#FFF",t.shadowOffsetX=0,t.shadowOffsetY=0,t.shadowBlur=8,t.shadowColor="#000","string"==typeof e.label){var s=t.measureText(e.label).width,h=Math.round(s+5),c=Math.round(i+4),l=Math.max(e.size,i/2)+2,u=Math.asin(c/2/l),d=Math.sqrt(Math.abs(Math.pow(l,2)-Math.pow(c/2,2)));t.beginPath(),t.moveTo(e.x+d,e.y+c/2),t.lineTo(e.x+l+h,e.y+c/2),t.lineTo(e.x+l+h,e.y-c/2),t.lineTo(e.x+d,e.y-c/2),t.arc(e.x,e.y,l,u,-u),t.closePath(),t.fill()}else t.beginPath(),t.arc(e.x,e.y,e.size+2,0,2*Math.PI),t.closePath(),t.fill();t.shadowOffsetX=0,t.shadowOffsetY=0,t.shadowBlur=0,(0,o.default)(t,e,r)}},622:(t,e)=>{"use strict";Object.defineProperty(e,"__esModule",{value:!0}),e.default=function(t,e,r){if(e.label){var i=r.labelSize,o=r.labelFont,n=r.labelWeight;t.fillStyle="#000",t.font=n+" "+i+"px "+o,t.fillText(e.label,e.x+e.size+3,e.y+i/3)}}},195:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)});Object.defineProperty(e,"__esModule",{value:!0}),e.createEdgeCompoundProgram=e.AbstractEdgeProgram=void 0;var n=function(t){function e(e,r,i,o,n){return t.call(this,e,r,i,o,n)||this}return o(e,t),e}(r(171).AbstractProgram);e.AbstractEdgeProgram=n,e.createEdgeCompoundProgram=function(t){return function(){function e(e,r){this.programs=t.map((function(t){return new t(e,r)}))}return e.prototype.bufferData=function(){this.programs.forEach((function(t){return t.bufferData()}))},e.prototype.allocate=function(t){this.programs.forEach((function(e){return e.allocate(t)}))},e.prototype.bind=function(){},e.prototype.computeIndices=function(){this.programs.forEach((function(t){return t.computeIndices()}))},e.prototype.render=function(t){this.programs.forEach((function(e){e.bind(),e.bufferData(),e.render(t)}))},e.prototype.process=function(t,e,r,i,o){this.programs.forEach((function(n){return n.process(t,e,r,i,o)}))},e}()}},909:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)});Object.defineProperty(e,"__esModule",{value:!0}),e.createNodeCompoundProgram=e.AbstractNodeProgram=void 0;var n=function(t){function e(e,r,i,o,n){var a=t.call(this,e,r,i,o,n)||this;a.positionLocation=e.getAttribLocation(a.program,"a_position"),a.sizeLocation=e.getAttribLocation(a.program,"a_size"),a.colorLocation=e.getAttribLocation(a.program,"a_color");var s=e.getUniformLocation(a.program,"u_matrix");if(null===s)throw new Error("AbstractNodeProgram: error while getting matrixLocation");a.matrixLocation=s;var h=e.getUniformLocation(a.program,"u_ratio");if(null===h)throw new Error("AbstractNodeProgram: error while getting ratioLocation");a.ratioLocation=h;var c=e.getUniformLocation(a.program,"u_scale");if(null===c)throw new Error("AbstractNodeProgram: error while getting scaleLocation");return a.scaleLocation=c,a}return o(e,t),e.prototype.bind=function(){var t=this.gl;t.enableVertexAttribArray(this.positionLocation),t.enableVertexAttribArray(this.sizeLocation),t.enableVertexAttribArray(this.colorLocation),t.vertexAttribPointer(this.positionLocation,2,t.FLOAT,!1,this.attributes*Float32Array.BYTES_PER_ELEMENT,0),t.vertexAttribPointer(this.sizeLocation,1,t.FLOAT,!1,this.attributes*Float32Array.BYTES_PER_ELEMENT,8),t.vertexAttribPointer(this.colorLocation,4,t.UNSIGNED_BYTE,!0,this.attributes*Float32Array.BYTES_PER_ELEMENT,12)},e}(r(171).AbstractProgram);e.AbstractNodeProgram=n,e.createNodeCompoundProgram=function(t){return function(){function e(e,r){this.programs=t.map((function(t){return new t(e,r)}))}return e.prototype.bufferData=function(){this.programs.forEach((function(t){return t.bufferData()}))},e.prototype.allocate=function(t){this.programs.forEach((function(e){return e.allocate(t)}))},e.prototype.bind=function(){},e.prototype.render=function(t){this.programs.forEach((function(e){return e.render(t)}))},e.prototype.process=function(t,e,r){this.programs.forEach((function(i){return i.process(t,e,r)}))},e}()}},171:(t,e,r)=>{"use strict";Object.defineProperty(e,"__esModule",{value:!0}),e.AbstractProgram=void 0;var i=r(706),o=function(){function t(t,e,r,o,n){this.array=new Float32Array,this.points=o,this.attributes=n,this.gl=t,this.vertexShaderSource=e,this.fragmentShaderSource=r;var a=t.createBuffer();if(null===a)throw new Error("AbstractProgram: error while creating the buffer");this.buffer=a,t.bindBuffer(t.ARRAY_BUFFER,this.buffer),this.vertexShader=(0,i.loadVertexShader)(t,this.vertexShaderSource),this.fragmentShader=(0,i.loadFragmentShader)(t,this.fragmentShaderSource),this.program=(0,i.loadProgram)(t,[this.vertexShader,this.fragmentShader])}return t.prototype.bufferData=function(){var t=this.gl;t.bufferData(t.ARRAY_BUFFER,this.array,t.DYNAMIC_DRAW)},t.prototype.allocate=function(t){this.array=new Float32Array(this.points*this.attributes*t)},t}();e.AbstractProgram=o},569:function(t,e,r){"use strict";var i=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0});var o=r(195),n=i(r(805)),a=i(r(483)),s=(0,o.createEdgeCompoundProgram)([a.default,n.default]);e.default=s},805:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)}),n=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0});var a=r(928),s=n(r(912)),h=n(r(973)),c=r(195),l=10,u=function(t){function e(e){var r=t.call(this,e,s.default,h.default,3,l)||this;r.positionLocation=e.getAttribLocation(r.program,"a_position"),r.colorLocation=e.getAttribLocation(r.program,"a_color"),r.normalLocation=e.getAttribLocation(r.program,"a_normal"),r.thicknessLocation=e.getAttribLocation(r.program,"a_thickness"),r.radiusLocation=e.getAttribLocation(r.program,"a_radius"),r.barycentricLocation=e.getAttribLocation(r.program,"a_barycentric");var i=e.getUniformLocation(r.program,"u_scale");if(null===i)throw new Error("EdgeArrowHeadProgram: error while getting scaleLocation");r.scaleLocation=i;var o=e.getUniformLocation(r.program,"u_matrix");if(null===o)throw new Error("EdgeArrowHeadProgram: error while getting matrixLocation");r.matrixLocation=o;var n=e.getUniformLocation(r.program,"u_cameraRatio");if(null===n)throw new Error("EdgeArrowHeadProgram: error while getting cameraRatioLocation");r.cameraRatioLocation=n;var a=e.getUniformLocation(r.program,"u_viewportRatio");if(null===a)throw new Error("EdgeArrowHeadProgram: error while getting viewportRatioLocation");r.viewportRatioLocation=a;var c=e.getUniformLocation(r.program,"u_thicknessRatio");if(null===c)throw new Error("EdgeArrowHeadProgram: error while getting thicknessRatioLocation");return r.thicknessRatioLocation=c,r.bind(),r}return o(e,t),e.prototype.bind=function(){var t=this.gl;t.enableVertexAttribArray(this.positionLocation),t.enableVertexAttribArray(this.normalLocation),t.enableVertexAttribArray(this.thicknessLocation),t.enableVertexAttribArray(this.radiusLocation),t.enableVertexAttribArray(this.colorLocation),t.enableVertexAttribArray(this.barycentricLocation),t.vertexAttribPointer(this.positionLocation,2,t.FLOAT,!1,l*Float32Array.BYTES_PER_ELEMENT,0),t.vertexAttribPointer(this.normalLocation,2,t.FLOAT,!1,l*Float32Array.BYTES_PER_ELEMENT,8),t.vertexAttribPointer(this.thicknessLocation,1,t.FLOAT,!1,l*Float32Array.BYTES_PER_ELEMENT,16),t.vertexAttribPointer(this.radiusLocation,1,t.FLOAT,!1,l*Float32Array.BYTES_PER_ELEMENT,20),t.vertexAttribPointer(this.colorLocation,4,t.UNSIGNED_BYTE,!0,l*Float32Array.BYTES_PER_ELEMENT,24),t.vertexAttribPointer(this.barycentricLocation,3,t.FLOAT,!1,l*Float32Array.BYTES_PER_ELEMENT,28)},e.prototype.computeIndices=function(){},e.prototype.process=function(t,e,r,i,o){if(i)for(var n=30*o,s=n+30;n<s;n++)this.array[n]=0;else{var h=r.size||1,c=e.size||1,l=t.x,u=t.y,d=e.x,f=e.y,p=(0,a.floatColor)(r.color),g=d-l,v=f-u,m=g*g+v*v,y=0,_=0;m&&(y=-v*(m=1/Math.sqrt(m)),_=g*m);var b=30*o,x=this.array;x[b++]=d,x[b++]=f,x[b++]=-y,x[b++]=-_,x[b++]=h,x[b++]=c,x[b++]=p,x[b++]=1,x[b++]=0,x[b++]=0,x[b++]=d,x[b++]=f,x[b++]=-y,x[b++]=-_,x[b++]=h,x[b++]=c,x[b++]=p,x[b++]=0,x[b++]=1,x[b++]=0,x[b++]=d,x[b++]=f,x[b++]=-y,x[b++]=-_,x[b++]=h,x[b++]=c,x[b++]=p,x[b++]=0,x[b++]=0,x[b]=1}},e.prototype.render=function(t){var e=this.gl,r=this.program;e.useProgram(r),e.uniform1f(this.scaleLocation,t.scalingRatio),e.uniformMatrix3fv(this.matrixLocation,!1,t.matrix),e.uniform1f(this.cameraRatioLocation,t.ratio),e.uniform1f(this.viewportRatioLocation,1/Math.min(t.width,t.height)),e.uniform1f(this.thicknessRatioLocation,1/Math.pow(t.ratio,t.edgesPowRatio)),e.drawArrays(e.TRIANGLES,0,this.array.length/l)},e}(c.AbstractEdgeProgram);e.default=u},483:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)}),n=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0});var a=r(195),s=r(928),h=n(r(620)),c=n(r(498)),l=function(t){function e(e){var r=t.call(this,e,h.default,c.default,4,7)||this,i=e.createBuffer();if(null===i)throw new Error("EdgeClampedProgram: error while getting resolutionLocation");r.indicesBuffer=i,r.positionLocation=e.getAttribLocation(r.program,"a_position"),r.colorLocation=e.getAttribLocation(r.program,"a_color"),r.normalLocation=e.getAttribLocation(r.program,"a_normal"),r.thicknessLocation=e.getAttribLocation(r.program,"a_thickness"),r.radiusLocation=e.getAttribLocation(r.program,"a_radius");var o=e.getUniformLocation(r.program,"u_scale");if(null===o)throw new Error("EdgeClampedProgram: error while getting scaleLocation");r.scaleLocation=o;var n=e.getUniformLocation(r.program,"u_matrix");if(null===n)throw new Error("EdgeClampedProgram: error while getting matrixLocation");r.matrixLocation=n;var a=e.getUniformLocation(r.program,"u_cameraRatio");if(null===a)throw new Error("EdgeClampedProgram: error while getting cameraRatioLocation");r.cameraRatioLocation=a;var l=e.getUniformLocation(r.program,"u_viewportRatio");if(null===l)throw new Error("EdgeClampedProgram: error while getting viewportRatioLocation");r.viewportRatioLocation=l;var u=e.getUniformLocation(r.program,"u_thicknessRatio");if(null===u)throw new Error("EdgeClampedProgram: error while getting thicknessRatioLocation");return r.thicknessRatioLocation=u,r.canUse32BitsIndices=(0,s.canUse32BitsIndices)(e),r.IndicesArray=r.canUse32BitsIndices?Uint32Array:Uint16Array,r.indicesArray=new r.IndicesArray,r.indicesType=r.canUse32BitsIndices?e.UNSIGNED_INT:e.UNSIGNED_SHORT,r.bind(),r}return o(e,t),e.prototype.bind=function(){var t=this.gl;t.bindBuffer(t.ELEMENT_ARRAY_BUFFER,this.indicesBuffer),t.enableVertexAttribArray(this.positionLocation),t.enableVertexAttribArray(this.normalLocation),t.enableVertexAttribArray(this.thicknessLocation),t.enableVertexAttribArray(this.colorLocation),t.enableVertexAttribArray(this.radiusLocation),t.vertexAttribPointer(this.positionLocation,2,t.FLOAT,!1,7*Float32Array.BYTES_PER_ELEMENT,0),t.vertexAttribPointer(this.normalLocation,2,t.FLOAT,!1,7*Float32Array.BYTES_PER_ELEMENT,8),t.vertexAttribPointer(this.thicknessLocation,1,t.FLOAT,!1,7*Float32Array.BYTES_PER_ELEMENT,16),t.vertexAttribPointer(this.colorLocation,4,t.UNSIGNED_BYTE,!0,7*Float32Array.BYTES_PER_ELEMENT,20),t.vertexAttribPointer(this.radiusLocation,1,t.FLOAT,!1,7*Float32Array.BYTES_PER_ELEMENT,24)},e.prototype.process=function(t,e,r,i,o){if(i)for(var n=28*o,a=n+28;n<a;n++)this.array[n]=0;else{var h=r.size||1,c=t.x,l=t.y,u=e.x,d=e.y,f=e.size||1,p=(0,s.floatColor)(r.color),g=u-c,v=d-l,m=g*g+v*v,y=0,_=0;m&&(y=-v*(m=1/Math.sqrt(m)),_=g*m);var b=28*o,x=this.array;x[b++]=c,x[b++]=l,x[b++]=y,x[b++]=_,x[b++]=h,x[b++]=p,x[b++]=0,x[b++]=c,x[b++]=l,x[b++]=-y,x[b++]=-_,x[b++]=h,x[b++]=p,x[b++]=0,x[b++]=u,x[b++]=d,x[b++]=y,x[b++]=_,x[b++]=h,x[b++]=p,x[b++]=f,x[b++]=u,x[b++]=d,x[b++]=-y,x[b++]=-_,x[b++]=h,x[b++]=p,x[b]=-f}},e.prototype.computeIndices=function(){for(var t=this.array.length/7,e=t+t/2,r=new this.IndicesArray(e),i=0,o=0;i<t;i+=4)r[o++]=i,r[o++]=i+1,r[o++]=i+2,r[o++]=i+2,r[o++]=i+1,r[o++]=i+3;this.indicesArray=r},e.prototype.bufferData=function(){t.prototype.bufferData.call(this);var e=this.gl;e.bufferData(e.ELEMENT_ARRAY_BUFFER,this.indicesArray,e.STATIC_DRAW)},e.prototype.render=function(t){var e=this.gl,r=this.program;e.useProgram(r),e.uniform1f(this.scaleLocation,t.scalingRatio),e.uniformMatrix3fv(this.matrixLocation,!1,t.matrix),e.uniform1f(this.cameraRatioLocation,t.ratio),e.uniform1f(this.viewportRatioLocation,1/Math.min(t.width,t.height)),e.uniform1f(this.thicknessRatioLocation,1/Math.pow(t.ratio,t.edgesPowRatio)),e.drawElements(e.TRIANGLES,this.indicesArray.length,this.indicesType,0)},e}(a.AbstractEdgeProgram);e.default=l},753:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)}),n=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0});var a=r(928),s=n(r(223)),h=n(r(498)),c=function(t){function e(e){var r=t.call(this,e,s.default,h.default,4,6)||this,i=e.createBuffer();if(null===i)throw new Error("EdgeProgram: error while getting resolutionLocation");r.indicesBuffer=i,r.positionLocation=e.getAttribLocation(r.program,"a_position"),r.colorLocation=e.getAttribLocation(r.program,"a_color"),r.normalLocation=e.getAttribLocation(r.program,"a_normal"),r.thicknessLocation=e.getAttribLocation(r.program,"a_thickness");var o=e.getUniformLocation(r.program,"u_scale");if(null===o)throw new Error("EdgeProgram: error while getting scaleLocation");r.scaleLocation=o;var n=e.getUniformLocation(r.program,"u_matrix");if(null===n)throw new Error("EdgeProgram: error while getting matrixLocation");r.matrixLocation=n;var c=e.getUniformLocation(r.program,"u_cameraRatio");if(null===c)throw new Error("EdgeProgram: error while getting cameraRatioLocation");r.cameraRatioLocation=c;var l=e.getUniformLocation(r.program,"u_viewportRatio");if(null===l)throw new Error("EdgeProgram: error while getting viewportRatioLocation");r.viewportRatioLocation=l;var u=e.getUniformLocation(r.program,"u_thicknessRatio");if(null===u)throw new Error("EdgeProgram: error while getting thicknessRatioLocation");return r.thicknessRatioLocation=u,r.canUse32BitsIndices=(0,a.canUse32BitsIndices)(e),r.IndicesArray=r.canUse32BitsIndices?Uint32Array:Uint16Array,r.indicesArray=new r.IndicesArray,r.indicesType=r.canUse32BitsIndices?e.UNSIGNED_INT:e.UNSIGNED_SHORT,r.bind(),r}return o(e,t),e.prototype.bind=function(){var t=this.gl;t.bindBuffer(t.ELEMENT_ARRAY_BUFFER,this.indicesBuffer),t.enableVertexAttribArray(this.positionLocation),t.enableVertexAttribArray(this.normalLocation),t.enableVertexAttribArray(this.thicknessLocation),t.enableVertexAttribArray(this.colorLocation),t.vertexAttribPointer(this.positionLocation,2,t.FLOAT,!1,6*Float32Array.BYTES_PER_ELEMENT,0),t.vertexAttribPointer(this.normalLocation,2,t.FLOAT,!1,6*Float32Array.BYTES_PER_ELEMENT,8),t.vertexAttribPointer(this.thicknessLocation,1,t.FLOAT,!1,6*Float32Array.BYTES_PER_ELEMENT,16),t.vertexAttribPointer(this.colorLocation,4,t.UNSIGNED_BYTE,!0,6*Float32Array.BYTES_PER_ELEMENT,20)},e.prototype.computeIndices=function(){for(var t=this.array.length/6,e=t+t/2,r=new this.IndicesArray(e),i=0,o=0;i<t;i+=4)r[o++]=i,r[o++]=i+1,r[o++]=i+2,r[o++]=i+2,r[o++]=i+1,r[o++]=i+3;this.indicesArray=r},e.prototype.bufferData=function(){t.prototype.bufferData.call(this);var e=this.gl;e.bufferData(e.ELEMENT_ARRAY_BUFFER,this.indicesArray,e.STATIC_DRAW)},e.prototype.process=function(t,e,r,i,o){if(i)for(var n=24*o,s=n+24;n<s;n++)this.array[n]=0;else{var h=r.size||1,c=t.x,l=t.y,u=e.x,d=e.y,f=(0,a.floatColor)(r.color),p=u-c,g=d-l,v=p*p+g*g,m=0,y=0;v&&(m=-g*(v=1/Math.sqrt(v)),y=p*v);var _=24*o,b=this.array;b[_++]=c,b[_++]=l,b[_++]=m,b[_++]=y,b[_++]=h,b[_++]=f,b[_++]=c,b[_++]=l,b[_++]=-m,b[_++]=-y,b[_++]=h,b[_++]=f,b[_++]=u,b[_++]=d,b[_++]=m,b[_++]=y,b[_++]=h,b[_++]=f,b[_++]=u,b[_++]=d,b[_++]=-m,b[_++]=-y,b[_++]=h,b[_]=f}},e.prototype.render=function(t){var e=this.gl,r=this.program;e.useProgram(r),e.uniform1f(this.scaleLocation,t.scalingRatio),e.uniformMatrix3fv(this.matrixLocation,!1,t.matrix),e.uniform1f(this.cameraRatioLocation,t.ratio),e.uniform1f(this.viewportRatioLocation,1/Math.min(t.width,t.height)),e.uniform1f(this.thicknessRatioLocation,1/Math.pow(t.ratio,t.edgesPowRatio)),e.drawElements(e.TRIANGLES,this.indicesArray.length,this.indicesType,0)},e}(r(195).AbstractEdgeProgram);e.default=c},582:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)}),n=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0});var a=r(928),s=n(r(106)),h=n(r(262)),c=function(t){function e(e){var r=t.call(this,e,s.default,h.default,1,4)||this;return r.bind(),r}return o(e,t),e.prototype.process=function(t,e,r){var i=this.array,o=1*r*4;if(e)return i[o++]=0,i[o++]=0,i[o++]=0,void(i[o++]=0);var n=(0,a.floatColor)(t.color);i[o++]=t.x,i[o++]=t.y,i[o++]=t.size,i[o]=n},e.prototype.render=function(t){var e=this.gl,r=this.program;e.useProgram(r),e.uniform1f(this.ratioLocation,1/Math.pow(t.ratio,t.nodesPowRatio)),e.uniform1f(this.scaleLocation,t.scalingRatio),e.uniformMatrix3fv(this.matrixLocation,!1,t.matrix),e.drawArrays(e.POINTS,0,this.array.length/4)},e}(r(909).AbstractNodeProgram);e.default=c},706:(t,e)=>{"use strict";function r(t,e,r){var i="VERTEX"===t?e.VERTEX_SHADER:e.FRAGMENT_SHADER,o=e.createShader(i);if(null===o)throw new Error("loadShader: error while creating the shader");if(e.shaderSource(o,r),e.compileShader(o),!e.getShaderParameter(o,e.COMPILE_STATUS)){var n=e.getShaderInfoLog(o);throw e.deleteShader(o),new Error("loadShader: error while compiling the shader:\n"+n+"\n"+r)}return o}Object.defineProperty(e,"__esModule",{value:!0}),e.loadProgram=e.loadFragmentShader=e.loadVertexShader=void 0,e.loadVertexShader=function(t,e){return r("VERTEX",t,e)},e.loadFragmentShader=function(t,e){return r("FRAGMENT",t,e)},e.loadProgram=function(t,e){var r,i,o=t.createProgram();if(null===o)throw new Error("loadProgram: error while creating the program.");for(r=0,i=e.length;r<i;r++)t.attachShader(o,e[r]);if(t.linkProgram(o),!t.getProgramParameter(o,t.LINK_STATUS))throw t.deleteProgram(o),new Error("loadProgram: error while linking the program.");return o}},310:function(t,e,r){"use strict";var i=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0}),e.DEFAULT_SETTINGS=e.validateSettings=void 0;var o=i(r(622)),n=i(r(61)),a=i(r(942)),s=i(r(582)),h=i(r(753)),c=i(r(569));e.validateSettings=function(t){if("number"!=typeof t.labelDensity||t.labelDensity<0)throw new Error("Settings: invalid `labelDensity`. Expecting a positive number.")},e.DEFAULT_SETTINGS={hideEdgesOnMove:!1,hideLabelsOnMove:!1,renderLabels:!0,renderEdgeLabels:!1,defaultNodeColor:"#999",defaultNodeType:"circle",defaultEdgeColor:"#ccc",defaultEdgeType:"line",labelFont:"Arial",labelSize:14,labelWeight:"normal",edgeLabelFont:"Arial",edgeLabelSize:14,edgeLabelWeight:"normal",stagePadding:30,labelDensity:.07,labelGridCellSize:60,labelRenderedSizeThreshold:6,nodeReducer:null,edgeReducer:null,zIndex:!1,labelRenderer:o.default,hoverRenderer:n.default,edgeLabelRenderer:a.default,nodeProgramClasses:{circle:s.default},edgeProgramClasses:{arrow:c.default,line:h.default}}},159:function(t,e,r){"use strict";var i,o=this&&this.__extends||(i=function(t,e){return i=Object.setPrototypeOf||{__proto__:[]}instanceof Array&&function(t,e){t.__proto__=e}||function(t,e){for(var r in e)Object.prototype.hasOwnProperty.call(e,r)&&(t[r]=e[r])},i(t,e)},function(t,e){if("function"!=typeof e&&null!==e)throw new TypeError("Class extends value "+String(e)+" is not a constructor or null");function r(){this.constructor=t}i(t,e),t.prototype=null===e?Object.create(e):(r.prototype=e.prototype,new r)}),n=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0});var a=r(187),s=n(r(886)),h=n(r(764)),c=n(r(269)),l=n(r(134)),u=r(928),d=r(730),f=r(310),p=n(r(508)),g=r(700),v=s.default.nodeExtent,m=(0,u.getPixelRatio)(),y=(0,u.getPixelRatio)();function _(t,e,r){if(!r.hasOwnProperty("x")||!r.hasOwnProperty("y"))throw new Error('Sigma: could not find a valid position (x, y) for node "'+e+'". All your nodes must have a number "x" and "y". Maybe your forgot to apply a layout or your "nodeReducer" is not returning the correct data?');return r.color||(r.color=t.defaultNodeColor),r.label||""===r.label||(r.label=null),void 0!==r.label&&null!==r.label?r.label=""+r.label:r.label=null,r.size||(r.size=2),r.hasOwnProperty("hidden")||(r.hidden=!1),r.hasOwnProperty("highlighted")||(r.highlighted=!1),r.type&&""!==r.type||(r.type=t.defaultNodeType),r.zIndex||(r.zIndex=0),r}function b(t,e,r){return r.color||(r.color=t.defaultEdgeColor),r.label||(r.label=""),r.size||(r.size=.5),r.hasOwnProperty("hidden")||(r.hidden=!1),r.type&&""!==r.type||(r.type=t.defaultEdgeType),r.zIndex||(r.zIndex=0),r}var x=function(t){function e(e,r,i){void 0===i&&(i={});var o=t.call(this)||this;if(o.elements={},o.canvasContexts={},o.webGLContexts={},o.activeListeners={},o.quadtree=new l.default,o.labelGrid=new d.LabelGrid,o.nodeDataCache={},o.edgeDataCache={},o.nodeKeyToIndex={},o.edgeKeyToIndex={},o.nodeExtent={x:[0,1],y:[0,1]},o.matrix=(0,g.identity)(),o.invMatrix=(0,g.identity)(),o.customBBox=null,o.normalizationFunction=(0,u.createNormalizationFunction)({x:[-1/0,1/0],y:[-1/0,1/0]}),o.cameraSizeRatio=1,o.width=0,o.height=0,o.displayedLabels=new Set,o.highlightedNodes=new Set,o.hoveredNode=null,o.renderFrame=null,o.renderHighlightedNodesFrame=null,o.needToProcess=!1,o.needToSoftProcess=!1,o.nodePrograms={},o.hoverNodePrograms={},o.edgePrograms={},o.settings=(0,u.assignDeep)({},f.DEFAULT_SETTINGS,i),(0,f.validateSettings)(o.settings),(0,u.validateGraph)(e),!(r instanceof HTMLElement))throw new Error("Sigma: container should be an html element.");o.graph=e,o.container=r,o.initializeCache(),o.createWebGLContext("edges"),o.createWebGLContext("nodes"),o.createCanvasContext("edgeLabels"),o.createCanvasContext("labels"),o.createCanvasContext("hovers"),o.createWebGLContext("hoverNodes"),o.createCanvasContext("mouse");var n=o.webGLContexts.nodes;for(var a in n.blendFunc(n.ONE,n.ONE_MINUS_SRC_ALPHA),n.enable(n.BLEND),(n=o.webGLContexts.edges).blendFunc(n.ONE,n.ONE_MINUS_SRC_ALPHA),n.enable(n.BLEND),o.settings.nodeProgramClasses){var s=o.settings.nodeProgramClasses[a];o.nodePrograms[a]=new s(o.webGLContexts.nodes,o),o.hoverNodePrograms[a]=new s(o.webGLContexts.hoverNodes,o)}for(var a in o.settings.edgeProgramClasses){var v=o.settings.edgeProgramClasses[a];o.edgePrograms[a]=new v(o.webGLContexts.edges,o)}return o.resize(),o.camera=new h.default,o.bindCameraHandlers(),o.mouseCaptor=new c.default(o.elements.mouse,o),o.touchCaptor=new p.default(o.elements.mouse,o),o.bindEventHandlers(),o.bindGraphHandlers(),o.process(),o.render(),o}return o(e,t),e.prototype.createCanvas=function(t){var e=(0,u.createElement)("canvas",{position:"absolute"},{class:"sigma-"+t});return this.elements[t]=e,this.container.appendChild(e),e},e.prototype.createCanvasContext=function(t){var e=this.createCanvas(t);return this.canvasContexts[t]=e.getContext("2d",{preserveDrawingBuffer:!1,antialias:!1}),this},e.prototype.createWebGLContext=function(t){var e,r=this.createCanvas(t),i={preserveDrawingBuffer:!1,antialias:!1};return(e=r.getContext("webgl2",i))||(e=r.getContext("webgl",i)),e||(e=r.getContext("experimental-webgl",i)),this.webGLContexts[t]=e,this},e.prototype.initializeCache=function(){var t=this,e=this.graph,r=0;e.forEachNode((function(e){t.nodeKeyToIndex[e]=r++})),r=0,e.forEachEdge((function(e){t.edgeKeyToIndex[e]=r++}))},e.prototype.bindCameraHandlers=function(){var t=this;return this.activeListeners.camera=function(){t._scheduleRefresh()},this.camera.on("updated",this.activeListeners.camera),this},e.prototype.bindEventHandlers=function(){var t=this;this.activeListeners.handleResize=function(){t.needToSoftProcess=!0,t._scheduleRefresh()},window.addEventListener("resize",this.activeListeners.handleResize);var e=function(t,e,r,i,o){return t>r-o&&t<r+o&&e>i-o&&e<i+o&&Math.sqrt(Math.pow(t-r,2)+Math.pow(e-i,2))<o},r=function(e,r){var i=t.viewportToFramedGraph({x:e,y:r});return t.quadtree.point(i.x,1-i.y)};this.activeListeners.handleMove=function(i){for(var o=r(i.x,i.y),n=1/0,a=null,s=0,h=o.length;s<h;s++){var c=o[s],l=t.nodeDataCache[c],u=t.framedGraphToViewport(l),d=t.scaleSize(l.size);if(e(i.x,i.y,u.x,u.y,d)){var f=Math.sqrt(Math.pow(i.x-u.x,2)+Math.pow(i.y-u.y,2));f<n&&(n=f,a=c)}}return a&&t.hoveredNode!==a&&!t.nodeDataCache[a].hidden?(t.hoveredNode&&t.emit("leaveNode",{node:t.hoveredNode}),t.hoveredNode=a,t.emit("enterNode",{node:a}),void t.scheduleHighlightedNodesRender()):t.hoveredNode&&(l=t.nodeDataCache[t.hoveredNode],u=t.framedGraphToViewport(l),d=t.scaleSize(l.size),!e(i.x,i.y,u.x,u.y,d))?(c=t.hoveredNode,t.hoveredNode=null,t.emit("leaveNode",{node:c}),t.scheduleHighlightedNodesRender()):void 0};var i=function(i){return function(o){for(var n=r(o.x,o.y),a=0,s=n.length;a<s;a++){var h=n[a],c=t.nodeDataCache[h],l=t.framedGraphToViewport(c),u=t.scaleSize(c.size);if(e(o.x,o.y,l.x,l.y,u))return t.emit(i+"Node",{node:h,captor:o,event:o})}return t.emit(i+"Stage",{event:o})}};return this.activeListeners.handleClick=i("click"),this.activeListeners.handleRightClick=i("rightClick"),this.activeListeners.handleDown=i("down"),this.mouseCaptor.on("mousemove",this.activeListeners.handleMove),this.mouseCaptor.on("click",this.activeListeners.handleClick),this.mouseCaptor.on("rightClick",this.activeListeners.handleRightClick),this.mouseCaptor.on("mousedown",this.activeListeners.handleDown),this},e.prototype.bindGraphHandlers=function(){var t=this,e=this.graph;return this.activeListeners.graphUpdate=function(){t.needToProcess=!0,t._scheduleRefresh()},this.activeListeners.softGraphUpdate=function(){t.needToSoftProcess=!0,t._scheduleRefresh()},this.activeListeners.addNodeGraphUpdate=function(r){t.nodeKeyToIndex[r.key]=e.order-1,t.activeListeners.graphUpdate()},this.activeListeners.addEdgeGraphUpdate=function(r){t.nodeKeyToIndex[r.key]=e.order-1,t.activeListeners.graphUpdate()},e.on("nodeAdded",this.activeListeners.addNodeGraphUpdate),e.on("nodeDropped",this.activeListeners.graphUpdate),e.on("nodeAttributesUpdated",this.activeListeners.softGraphUpdate),e.on("eachNodeAttributesUpdated",this.activeListeners.graphUpdate),e.on("edgeAdded",this.activeListeners.addEdgeGraphUpdate),e.on("edgeDropped",this.activeListeners.graphUpdate),e.on("edgeAttributesUpdated",this.activeListeners.softGraphUpdate),e.on("eachEdgeAttributesUpdated",this.activeListeners.graphUpdate),e.on("edgesCleared",this.activeListeners.graphUpdate),e.on("cleared",this.activeListeners.graphUpdate),this},e.prototype.process=function(t){var e=this;void 0===t&&(t=!1);var r=this.graph,i=this.settings,o=this.getDimensions(),n=[1/0,-1/0],a=[1/0,-1/0];this.quadtree.clear(),this.labelGrid.resizeAndClear(o,i.labelGridCellSize),this.highlightedNodes=new Set,this.nodeExtent=v(r,["x","y"]);var s=new h.default,c=(0,u.matrixFromCamera)(s.getState(),this.getDimensions(),this.getGraphDimensions(),this.getSetting("stagePadding")||0);this.normalizationFunction=(0,u.createNormalizationFunction)(this.customBBox||this.nodeExtent);for(var l={},d=r.nodes(),f=0,p=d.length;f<p;f++){var g=d[f],m=Object.assign({},r.getNodeAttributes(g));i.nodeReducer&&(m=i.nodeReducer(g,m)),l[(x=_(this.settings,g,m)).type]=(l[x.type]||0)+1,this.nodeDataCache[g]=x,this.normalizationFunction.applyTo(x),this.settings.zIndex&&(x.zIndex<n[0]&&(n[0]=x.zIndex),x.zIndex>n[1]&&(n[1]=x.zIndex))}for(var y in l){if(!this.nodePrograms.hasOwnProperty(y))throw new Error('Sigma: could not find a suitable program for node type "'+y+'"!');t||this.nodePrograms[y].allocate(l[y]),l[y]=0}for(this.settings.zIndex&&n[0]!==n[1]&&(d=(0,u.zIndexOrdering)(n,(function(t){return e.nodeDataCache[t].zIndex}),d)),f=0,p=d.length;f<p;f++){g=d[f];var x=this.nodeDataCache[g];this.quadtree.add(g,x.x,1-x.y,x.size/this.width),x.label&&this.labelGrid.add(g,x.size,this.framedGraphToViewport(x,{matrix:c})),this.nodePrograms[x.type].process(x,x.hidden,l[x.type]++),x.highlighted&&!x.hidden&&this.highlightedNodes.add(g),this.nodeKeyToIndex[g]=f}this.labelGrid.organize();var w={},L=r.edges();for(f=0,p=L.length;f<p;f++){var E=L[f];m=Object.assign({},r.getEdgeAttributes(E)),i.edgeReducer&&(m=i.edgeReducer(E,m)),w[(x=b(this.settings,0,m)).type]=(w[x.type]||0)+1,this.edgeDataCache[E]=x,this.settings.zIndex&&(x.zIndex<a[0]&&(a[0]=x.zIndex),x.zIndex>a[1]&&(a[1]=x.zIndex))}for(var y in w){if(!this.edgePrograms.hasOwnProperty(y))throw new Error('Sigma: could not find a suitable program for edge type "'+y+'"!');t||this.edgePrograms[y].allocate(w[y]),w[y]=0}for(this.settings.zIndex&&a[0]!==a[1]&&(L=(0,u.zIndexOrdering)(a,(function(t){return e.edgeDataCache[t].zIndex}),L)),f=0,p=L.length;f<p;f++){E=L[f],x=this.edgeDataCache[E];var A=r.extremities(E),P=this.nodeDataCache[A[0]],T=this.nodeDataCache[A[1]],C=x.hidden||P.hidden||T.hidden;this.edgePrograms[x.type].process(P,T,x,C,w[x.type]++),this.nodeKeyToIndex[E]=f}for(var y in this.edgePrograms){var S=this.edgePrograms[y];t||"function"!=typeof S.computeIndices||S.computeIndices()}return this},e.prototype._refresh=function(){return this.needToProcess?this.process():this.needToSoftProcess&&this.process(!0),this.needToProcess=!1,this.needToSoftProcess=!1,this.render(),this},e.prototype._scheduleRefresh=function(){var t=this;return this.renderFrame||(this.renderFrame=(0,u.requestFrame)((function(){t._refresh(),t.renderFrame=null}))),this},e.prototype.renderLabels=function(){if(!this.settings.renderLabels)return this;var t,e=this.camera.getState();if(e.ratio>=1)t=new Set(this.graph.nodes());else{var r=this.viewRectangle();t=new Set(this.quadtree.rectangle(r.x1,1-r.y1,r.x2,1-r.y2,r.height))}var i=this.labelGrid.getLabelsToDisplay(e.ratio,this.settings.labelDensity);this.displayedLabels=new Set;for(var o=this.canvasContexts.labels,n=0,a=i.length;n<a;n++){var s=i[n],h=this.nodeDataCache[s];if(!h.hidden){var c=this.framedGraphToViewport(h),l=c.x,u=c.y,d=this.scaleSize(h.size);d<this.settings.labelRenderedSizeThreshold||t.has(s)&&(this.displayedLabels.add(s),this.settings.labelRenderer(o,{key:s,label:h.label,color:"#000",size:d,x:l,y:u},this.settings))}}return this},e.prototype.renderEdgeLabels=function(){if(!this.settings.renderEdgeLabels)return this;var t=this.canvasContexts.edgeLabels;t.clearRect(0,0,this.width,this.height);for(var e=(0,d.edgeLabelsToDisplayFromNodes)({graph:this.graph,hoveredNode:this.hoveredNode,displayedNodeLabels:this.displayedLabels,highlightedNodes:this.highlightedNodes}),r=0,i=e.length;r<i;r++){var o=e[r],n=this.graph.extremities(o),a=this.nodeDataCache[n[0]],s=this.nodeDataCache[n[1]],h=this.edgeDataCache[e[r]];if(!(h.hidden||a.hidden||s.hidden)){var c=this.framedGraphToViewport(a),l=c.x,u=c.y,f=this.framedGraphToViewport(s),p=f.x,g=f.y,v=this.scaleSize(h.size);this.settings.edgeLabelRenderer(t,{key:o,label:h.label,color:h.color,size:v},{key:n[0],x:l,y:u},{key:n[1],x:p,y:g},this.settings)}}return this},e.prototype.renderHighlightedNodes=function(){var t=this,e=this.canvasContexts.hovers;e.clearRect(0,0,this.width,this.height);var r=[];this.hoveredNode&&!this.nodeDataCache[this.hoveredNode].hidden&&r.push(this.hoveredNode),this.highlightedNodes.forEach((function(e){e!==t.hoveredNode&&r.push(e)})),r.forEach((function(r){return function(r){var i=t.nodeDataCache[r],o=t.framedGraphToViewport(i),n=o.x,a=o.y,s=t.scaleSize(i.size);t.settings.hoverRenderer(e,{key:r,label:i.label,color:i.color,size:s,x:n,y:a},t.settings)}(r)}));var i={};for(var o in r.forEach((function(e){var r=t.nodeDataCache[e].type;i[r]=(i[r]||0)+1})),this.hoverNodePrograms)this.hoverNodePrograms[o].allocate(i[o]||0),i[o]=0;for(var o in r.forEach((function(e){var r=t.nodeDataCache[e];t.hoverNodePrograms[r.type].process(r,r.hidden,i[r.type]++)})),this.hoverNodePrograms){var n=this.hoverNodePrograms[o];n.bind(),n.bufferData(),n.render({matrix:this.matrix,width:this.width,height:this.height,ratio:this.camera.ratio,nodesPowRatio:.5,scalingRatio:y})}},e.prototype.scheduleHighlightedNodesRender=function(){var t=this;this.renderHighlightedNodesFrame||this.renderFrame||(this.renderHighlightedNodesFrame=(0,u.requestFrame)((function(){t.renderHighlightedNodesFrame=null,t.renderHighlightedNodes(),t.renderEdgeLabels()})))},e.prototype.render=function(){if(this.renderFrame&&((0,u.cancelFrame)(this.renderFrame),this.renderFrame=null,this.needToProcess=!1,this.needToSoftProcess=!1),this.resize(),this.clear(),this.updateCachedValues(),!this.graph.order)return this;var t=this.mouseCaptor,e=this.camera.isAnimated()||t.isMoving||t.draggedEvents||t.currentWheelDirection,r=this.camera.getState(),i=this.getDimensions(),o=this.getGraphDimensions(),n=this.getSetting("stagePadding")||0;for(var a in this.matrix=(0,u.matrixFromCamera)(r,i,o,n),this.invMatrix=(0,u.matrixFromCamera)(r,i,o,n,!0),this.nodePrograms)(s=this.nodePrograms[a]).bind(),s.bufferData(),s.render({matrix:this.matrix,width:this.width,height:this.height,ratio:r.ratio,nodesPowRatio:.5,scalingRatio:y});if(!this.settings.hideEdgesOnMove||!e)for(var a in this.edgePrograms){var s;(s=this.edgePrograms[a]).bind(),s.bufferData(),s.render({matrix:this.matrix,width:this.width,height:this.height,ratio:r.ratio,edgesPowRatio:.5,scalingRatio:y})}return this.settings.hideLabelsOnMove&&e||(this.renderLabels(),this.renderEdgeLabels(),this.renderHighlightedNodes(),this.emit("afterRender")),this},e.prototype.updateCachedValues=function(){var t=this.camera.getState().ratio;this.cameraSizeRatio=Math.pow(t,.5)},e.prototype.getCamera=function(){return this.camera},e.prototype.getGraph=function(){return this.graph},e.prototype.getMouseCaptor=function(){return this.mouseCaptor},e.prototype.getTouchCaptor=function(){return this.touchCaptor},e.prototype.getDimensions=function(){return{width:this.width,height:this.height}},e.prototype.getGraphDimensions=function(){var t=this.customBBox||this.nodeExtent;return{width:t.x[1]-t.x[0]||1,height:t.y[1]-t.y[0]||1}},e.prototype.getNodeDisplayData=function(t){var e=this.nodeDataCache[t];return e?Object.assign({},e):void 0},e.prototype.getEdgeDisplayData=function(t){var e=this.edgeDataCache[t];return e?Object.assign({},e):void 0},e.prototype.getSetting=function(t){return this.settings[t]},e.prototype.setSetting=function(t,e){return this.settings[t]=e,(0,f.validateSettings)(this.settings),this.needToProcess=!0,this._scheduleRefresh(),this},e.prototype.updateSetting=function(t,e){return this.settings[t]=e(this.settings[t]),(0,f.validateSettings)(this.settings),this.needToProcess=!0,this._scheduleRefresh(),this},e.prototype.resize=function(){var t=this.width,e=this.height;if(this.width=this.container.offsetWidth,this.height=this.container.offsetHeight,0===this.width)throw new Error("Sigma: container has no width.");if(0===this.height)throw new Error("Sigma: container has no height.");if(t===this.width&&e===this.height)return this;for(var r in this.elements){var i=this.elements[r];i.style.width=this.width+"px",i.style.height=this.height+"px"}for(var r in this.canvasContexts)this.elements[r].setAttribute("width",this.width*m+"px"),this.elements[r].setAttribute("height",this.height*m+"px"),1!==m&&this.canvasContexts[r].scale(m,m);for(var r in this.webGLContexts)this.elements[r].setAttribute("width",this.width*y+"px"),this.elements[r].setAttribute("height",this.height*y+"px"),this.webGLContexts[r].viewport(0,0,this.width*y,this.height*y);return this},e.prototype.clear=function(){return this.webGLContexts.nodes.clear(this.webGLContexts.nodes.COLOR_BUFFER_BIT),this.webGLContexts.edges.clear(this.webGLContexts.edges.COLOR_BUFFER_BIT),this.canvasContexts.labels.clearRect(0,0,this.width,this.height),this.canvasContexts.hovers.clearRect(0,0,this.width,this.height),this.canvasContexts.edgeLabels.clearRect(0,0,this.width,this.height),this},e.prototype.refresh=function(){return this.needToProcess=!0,this._refresh(),this},e.prototype.scheduleRefresh=function(){return this.needToProcess=!0,this._scheduleRefresh(),this},e.prototype.getViewportZoomedState=function(t,e){var r=this.camera.getState(),i=r.ratio,o=r.angle,n=r.x,a=r.y,s=e/i,h={x:this.width/2,y:this.height/2},c=this.viewportToFramedGraph(t),l=this.viewportToFramedGraph(h);return{angle:o,x:(c.x-l.x)*(1-s)+n,y:(c.y-l.y)*(1-s)+a,ratio:e}},e.prototype.viewRectangle=function(){var t=0*this.width/8,e=0*this.height/8,r=this.viewportToFramedGraph({x:0-t,y:0-e}),i=this.viewportToFramedGraph({x:this.width+t,y:0-e}),o=this.viewportToFramedGraph({x:0,y:this.height+e});return{x1:r.x,y1:r.y,x2:i.x,y2:i.y,height:i.y-o.y}},e.prototype.framedGraphToViewport=function(t,e){void 0===e&&(e={});var r=!!e.cameraState||!!e.viewportDimensions||!!e.graphDimensions,i=e.matrix?e.matrix:r?(0,u.matrixFromCamera)(e.cameraState||this.camera.getState(),e.viewportDimensions||this.getDimensions(),e.graphDimensions||this.getGraphDimensions(),e.padding||this.getSetting("stagePadding")||0):this.matrix,o=[t.x,t.y,1],n=(0,g.multiplyVec)(i,o);return{x:(1+n[0])*this.width/2,y:(1-n[1])*this.height/2}},e.prototype.viewportToFramedGraph=function(t,e){void 0===e&&(e={});var r=!!e.cameraState||!!e.viewportDimensions||!e.graphDimensions,i=e.matrix?e.matrix:r?(0,u.matrixFromCamera)(e.cameraState||this.camera.getState(),e.viewportDimensions||this.getDimensions(),e.graphDimensions||this.getGraphDimensions(),e.padding||this.getSetting("stagePadding")||0,!0):this.invMatrix,o=[t.x/this.width*2-1,1-t.y/this.height*2,1],n=(0,g.multiplyVec)(i,o);return{x:n[0],y:n[1]}},e.prototype.viewportToGraph=function(t){return this.normalizationFunction.inverse(this.viewportToFramedGraph(t))},e.prototype.graphToViewport=function(t){return this.framedGraphToViewport(this.normalizationFunction(t))},e.prototype.getBBox=function(){return v(this.graph,["x","y"])},e.prototype.getCustomBBox=function(){return this.customBBox},e.prototype.setCustomBBox=function(t){return this.customBBox=t,this._scheduleRefresh(),this},e.prototype.kill=function(){var t=this.graph;this.emit("kill"),this.removeAllListeners(),this.camera.removeListener("updated",this.activeListeners.camera),window.removeEventListener("resize",this.activeListeners.handleResize),this.mouseCaptor.kill(),this.touchCaptor.kill(),t.removeListener("nodeAdded",this.activeListeners.addNodeGraphUpdate),t.removeListener("nodeDropped",this.activeListeners.graphUpdate),t.removeListener("nodeAttributesUpdated",this.activeListeners.softGraphUpdate),t.removeListener("eachNodeAttributesUpdated",this.activeListeners.graphUpdate),t.removeListener("edgeAdded",this.activeListeners.addEdgeGraphUpdate),t.removeListener("edgeDropped",this.activeListeners.graphUpdate),t.removeListener("edgeAttributesUpdated",this.activeListeners.softGraphUpdate),t.removeListener("eachEdgeAttributesUpdated",this.activeListeners.graphUpdate),t.removeListener("edgesCleared",this.activeListeners.graphUpdate),t.removeListener("cleared",this.activeListeners.graphUpdate),this.quadtree=new l.default,this.nodeDataCache={},this.edgeDataCache={},this.highlightedNodes.clear(),this.renderFrame&&((0,u.cancelFrame)(this.renderFrame),this.renderFrame=null),this.renderHighlightedNodesFrame&&((0,u.cancelFrame)(this.renderHighlightedNodesFrame),this.renderHighlightedNodesFrame=null);for(var e=this.container;e.firstChild;)e.removeChild(e.firstChild)},e.prototype.scaleSize=function(t){return t/this.cameraSizeRatio},e}(a.EventEmitter);e.default=x},751:function(t,e,r){"use strict";var i=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0}),e.animateNodes=e.ANIMATE_DEFAULTS=void 0;var o=r(928),n=i(r(358));e.ANIMATE_DEFAULTS={easing:"quadraticInOut",duration:150},e.animateNodes=function(t,r,i,a){var s=Object.assign({},e.ANIMATE_DEFAULTS,i),h="function"==typeof s.easing?s.easing:n.default[s.easing],c=Date.now(),l={};for(var u in r){var d=r[u];for(var f in l[u]={},d)l[u][f]=t.getNodeAttribute(u,f)}var p=null,g=function(){var e=(Date.now()-c)/s.duration;if(e>=1){for(var i in r){var n=r[i];for(var u in n)t.setNodeAttribute(i,u,n[u])}"function"==typeof a&&a()}else{for(var i in e=h(e),r){n=r[i];var d=l[i];for(var u in n)t.setNodeAttribute(i,u,n[u]*e+d[u]*(1-e))}p=(0,o.requestFrame)(g)}};return g(),function(){p&&(0,o.cancelFrame)(p)}}},358:(t,e)=>{"use strict";Object.defineProperty(e,"__esModule",{value:!0}),e.cubicInOut=e.cubicOut=e.cubicIn=e.quadraticInOut=e.quadraticOut=e.quadraticIn=e.linear=void 0,e.linear=function(t){return t},e.quadraticIn=function(t){return t*t},e.quadraticOut=function(t){return t*(2-t)},e.quadraticInOut=function(t){return(t*=2)<1?.5*t*t:-.5*(--t*(t-2)-1)},e.cubicIn=function(t){return t*t*t},e.cubicOut=function(t){return--t*t*t+1},e.cubicInOut=function(t){return(t*=2)<1?.5*t*t*t:.5*((t-=2)*t*t+2)};var r={linear:e.linear,quadraticIn:e.quadraticIn,quadraticOut:e.quadraticOut,quadraticInOut:e.quadraticInOut,cubicIn:e.cubicIn,cubicOut:e.cubicOut,cubicInOut:e.cubicInOut};e.default=r},928:function(t,e,r){"use strict";var i=this&&this.__read||function(t,e){var r="function"==typeof Symbol&&t[Symbol.iterator];if(!r)return t;var i,o,n=r.call(t),a=[];try{for(;(void 0===e||e-- >0)&&!(i=n.next()).done;)a.push(i.value)}catch(t){o={error:t}}finally{try{i&&!i.done&&(r=n.return)&&r.call(n)}finally{if(o)throw o.error}}return a},o=this&&this.__importDefault||function(t){return t&&t.__esModule?t:{default:t}};Object.defineProperty(e,"__esModule",{value:!0}),e.validateGraph=e.canUse32BitsIndices=e.extractPixel=e.matrixFromCamera=e.getCorrectionRatio=e.floatColor=e.zIndexOrdering=e.createNormalizationFunction=e.getPixelRatio=e.createElement=e.cancelFrame=e.requestFrame=e.assignDeep=e.isPlainObject=void 0;var n=o(r(186)),a=r(700);function s(t){return"object"==typeof t&&null!==t&&t.constructor===Object}e.isPlainObject=s,e.assignDeep=function t(e){for(var r=[],i=1;i<arguments.length;i++)r[i-1]=arguments[i];e=e||{};for(var o=0,n=r.length;o<n;o++){var a=r[o];if(a)for(var h in a)s(a[h])?e[h]=t(e[h],a[h]):e[h]=a[h]}return e},e.requestFrame="undefined"!=typeof requestAnimationFrame?function(t){return requestAnimationFrame(t)}:function(t){return setTimeout(t,0)},e.cancelFrame="undefined"!=typeof cancelAnimationFrame?function(t){return cancelAnimationFrame(t)}:function(t){return clearTimeout(t)},e.createElement=function(t,e,r){var i=document.createElement(t);if(e)for(var o in e)i.style[o]=e[o];if(r)for(var o in r)i.setAttribute(o,r[o]);return i},e.getPixelRatio=function(){return void 0!==window.devicePixelRatio?window.devicePixelRatio:1},e.createNormalizationFunction=function(t){var e=i(t.x,2),r=e[0],o=e[1],n=i(t.y,2),a=n[0],s=n[1],h=Math.max(o-r,s-a);0===h&&(h=1);var c=(o+r)/2,l=(s+a)/2,u=function(t){return{x:.5+(t.x-c)/h,y:.5+(t.y-l)/h}};return u.applyTo=function(t){t.x=.5+(t.x-c)/h,t.y=.5+(t.y-l)/h},u.inverse=function(t){return{x:c+h*(t.x-.5),y:l+h*(t.y-.5)}},u.ratio=h,u},e.zIndexOrdering=function(t,e,r){return r.sort((function(t,r){var i=e(t)||0,o=e(r)||0;return i<o?-1:i>o?1:0}))};var h={},c=new Int8Array(4),l=new Int32Array(c.buffer,0,1),u=new Float32Array(c.buffer,0,1),d=/^\s*rgba?\s*\(/,f=/^\s*rgba?\s*\(\s*([0-9]*)\s*,\s*([0-9]*)\s*,\s*([0-9]*)(?:\s*,\s*(.*)?)?\)\s*$/;function p(t,e){var r=t.height/t.width,i=e.height/e.width;return r<1&&i>1||r>1&&i<1?1:Math.min(Math.max(i,1/i),Math.max(1/r,r))}e.floatColor=function(t){if(void 0!==h[t])return h[t];var e=0,r=0,i=0,o=1;if("#"===t[0])4===t.length?(e=parseInt(t.charAt(1)+t.charAt(1),16),r=parseInt(t.charAt(2)+t.charAt(2),16),i=parseInt(t.charAt(3)+t.charAt(3),16)):(e=parseInt(t.charAt(1)+t.charAt(2),16),r=parseInt(t.charAt(3)+t.charAt(4),16),i=parseInt(t.charAt(5)+t.charAt(6),16));else if(d.test(t)){var n=t.match(f);n&&(e=+n[1],r=+n[2],i=+n[3],n[4]&&(o=+n[4]))}o=255*o|0,l[0]=4278190079&(o<<24|i<<16|r<<8|e);var a=u[0];return h[t]=a,a},e.getCorrectionRatio=p,e.matrixFromCamera=function(t,e,r,i,o){var n=t.angle,s=t.ratio,h=t.x,c=t.y,l=e.width,u=e.height,d=(0,a.identity)(),f=Math.min(l,u)-2*i,g=p(e,r);return o?((0,a.multiply)(d,(0,a.translate)((0,a.identity)(),h,c)),(0,a.multiply)(d,(0,a.scale)((0,a.identity)(),s)),(0,a.multiply)(d,(0,a.rotate)((0,a.identity)(),n)),(0,a.multiply)(d,(0,a.scale)((0,a.identity)(),l/f/2/g,u/f/2/g))):((0,a.multiply)(d,(0,a.scale)((0,a.identity)(),f/l*2*g,f/u*2*g)),(0,a.multiply)(d,(0,a.rotate)((0,a.identity)(),-n)),(0,a.multiply)(d,(0,a.scale)((0,a.identity)(),1/s)),(0,a.multiply)(d,(0,a.translate)((0,a.identity)(),-h,-c))),d},e.extractPixel=function(t,e,r,i){var o=i||new Uint8Array(4);return t.readPixels(e,r,1,1,t.RGBA,t.UNSIGNED_BYTE,o),o},e.canUse32BitsIndices=function(t){return"undefined"!=typeof WebGL2RenderingContext&&t instanceof WebGL2RenderingContext||!!t.getExtension("OES_element_index_uint")},e.validateGraph=function(t){if(!(0,n.default)(t))throw new Error("Sigma: invalid graph instance.");t.forEachNode((function(t,e){if(!Number.isFinite(e.x)||!Number.isFinite(e.y))throw new Error("Sigma: Coordinates of node "+t+" are invalid. A node must have a numeric 'x' and 'y' attribute.")}))}},700:(t,e)=>{"use strict";Object.defineProperty(e,"__esModule",{value:!0}),e.multiplyVec=e.multiply=e.translate=e.rotate=e.scale=e.identity=void 0,e.identity=function(){return Float32Array.of(1,0,0,0,1,0,0,0,1)},e.scale=function(t,e,r){return t[0]=e,t[4]="number"==typeof r?r:e,t},e.rotate=function(t,e){var r=Math.sin(e),i=Math.cos(e);return t[0]=i,t[1]=r,t[3]=-r,t[4]=i,t},e.translate=function(t,e,r){return t[6]=e,t[7]=r,t},e.multiply=function(t,e){var r=t[0],i=t[1],o=t[2],n=t[3],a=t[4],s=t[5],h=t[6],c=t[7],l=t[8],u=e[0],d=e[1],f=e[2],p=e[3],g=e[4],v=e[5],m=e[6],y=e[7],_=e[8];return t[0]=u*r+d*n+f*h,t[1]=u*i+d*a+f*c,t[2]=u*o+d*s+f*l,t[3]=p*r+g*n+v*h,t[4]=p*i+g*a+v*c,t[5]=p*o+g*s+v*l,t[6]=m*r+y*n+_*h,t[7]=m*i+y*a+_*c,t[8]=m*o+y*s+_*l,t},e.multiplyVec=function(t,e){var r=t[0],i=t[1],o=t[2],n=t[3],a=t[4],s=t[5],h=t[6],c=t[7],l=t[8],u=e[0],d=e[1],f=e[2],p=Array.isArray(e)?[0,0,0]:Float32Array.of(0,0,0);return p[0]=u*r+d*n+f*h,p[1]=u*i+d*a+f*c,p[2]=u*o+d*s+f*l,p}}},e={};function r(i){var o=e[i];if(void 0!==o)return o.exports;var n=e[i]={exports:{}};return t[i].call(n.exports,n,n.exports,r),n.exports}return r.d=(t,e)=>{for(var i in e)r.o(e,i)&&!r.o(t,i)&&Object.defineProperty(t,i,{enumerable:!0,get:e[i]})},r.o=(t,e)=>Object.prototype.hasOwnProperty.call(t,e),r.r=t=>{"undefined"!=typeof Symbol&&Symbol.toStringTag&&Object.defineProperty(t,Symbol.toStringTag,{value:"Module"}),Object.defineProperty(t,"__esModule",{value:!0})},r(607)})()})); \ No newline at end of file