diff --git a/.vscode/settings.json b/.vscode/settings.json new file mode 100644 index 0000000..2971aef --- /dev/null +++ b/.vscode/settings.json @@ -0,0 +1,11 @@ +{ + "editor.insertSpaces": true, + "editor.tabSize": 2, + "editor.rulers" : [100], + "files.encoding": "utf8", + "files.eol": "\n", + "files.insertFinalNewline": true, + "files.trimFinalNewlines": true, + "files.trimTrailingWhitespace": true, + "search.usePCRE2": true +} \ No newline at end of file diff --git a/Main.lean b/Main.lean deleted file mode 100644 index 65e1902..0000000 --- a/Main.lean +++ /dev/null @@ -1,4 +0,0 @@ -import WidgetKit - -def main : IO Unit := - IO.println s!"Welcome to WidgetKit, nice." diff --git a/README.md b/README.md index e7867ba..9421b64 100644 --- a/README.md +++ b/README.md @@ -1,17 +1,64 @@ -# Development kit for Lean 4 widgets +# WidgetKit -Authors: Wojciech Nawrocki, E.W.Ayers +WidgetKit is a library of user interface components for [Lean 4](https://leanprover.github.io/). It +supports building: +- symbolic visualizations of mathematical objects and data structures +- data visualizations +- interfaces for tactics and tactic modes +- custom visual proof modes +- custom goal state displays +- interfaces for entering or editing expressions in a domain-specific manner +- really anything that has to do with the infoview + +WidgetKit essentially supersedes [user widgets](https://leanprover.github.io/lean4/doc/examples/widgets.lean.html). +It is just as general, but more user-friendly. + +Authors: Wojciech Nawrocki, E.W.Ayers with contributions from Tomáš Skřivan # Usage -## Building the demos: +## Viewing the demos +The easiest way to get started is to clone a **release tag** of WidgetKit and run +`lake build :release`, as follows: + +```bash +# You should replace v0.0.1 with the latest version published under Releases +git clone https://github.com/EdAyers/WidgetKit --branch v0.0.1 +cd WidgetKit/ +lake build :release ``` -lake build + +After doing this you will hopefully be able to view the demos in `WidgetKit/Demos/`. If the demo +contains a `#html` command, put your cursors over it to see a widget in the infoview. Top tip: use +the pushpin icon to keep the widget in view. You can then live code your widgets. + +## Using WidgetKit as a dependency + +Put this in your `lakefile.lean`: +```lean +-- You should replace v0.0.1 with the latest version published under Releases +require widgetkit from git "https://github.com/EdAyers/WidgetKit"@"v0.0.1" ``` -Now go to the demo folder in VSCode. -Putting your cursor on any `#widget` or `#html` will show you a widget in the infoview. Top tip: use the pushpin icon to keep the widget in view. You can then live code your widgets. +Note that [developing WidgetKit](#developing-widgetkit) involves building TypeScript code with NPM. +When depending on `WidgetKit` but not writing any custom TypeScript yourself, you likely want to +avoid you or your users having to run NPM. To support this, WidgetKit is configured to use Lake's +[cloud releases](https://github.com/leanprover/lake/#cloud-releases) feature which will +automatically fetch pre-built files *as long as* you require a release tag rather than our `main` +branch. This is why the snippet above does that. + +## Developing WidgetKit + +You must have NPM installed (it is part of Node.js). Run `lake build` to build the TypeScript UI +code as well as all Lean modules. Run `lake build widgetJsAll` to just build the TypeScript. Outputs +of `npm` are not shown by default - use `lake build -v [widgetJsAll]` to see them. + +⚠️ We use the `include_str` term elaborator to splice the JavaScript produced this way into our Lean +files. The JS is stored in `build/js/`. Note however that due to Lake issue [#86](https://github.com/leanprover/lake/issues/86), +rebuilding the `.js` will *not* rebuild the Lean file that includes it. To ensure freshness you may +have to resort to hacks like adding a random comment in the file that uses `include_str` in order to +ensure it gets rebuilt. Alternatively, you can run `lake clean` to delete the build directory. # Features @@ -20,6 +67,8 @@ Putting your cursor on any `#widget` or `#html` will show you a widget in the in JSON-like syntax. Invoke with `json%`, escape with `$( _ )` ```lean +import WidgetKit.Data.Json +open scoped WidgetKit.Json #eval json% { hello : "world", cheese : ["edam", "cheddar", {kind : "spicy", rank : 100.2}], @@ -33,30 +82,27 @@ JSON-like syntax. Invoke with `json%`, escape with `$( _ )` ## JSX-like syntax ```lean -import WidgetKit.HtmlWidget -open Lean.Widget.Jsx +import WidgetKit.Component.HtmlDisplay +open scoped WidgetKit.Jsx -- click on the line below to see it in your infoview! #html You can use HTML in lean! {toString <| 1 + 2>} ``` -We also support all elements that are exposed by the [Recharts library](https://recharts.org/en-US/api), so you can make your own plots. See `src/Demo/Plot.lean` for an example. +See `WidgetKit/Demos/Basic.lean` and `WidgetKit/Demos/Svg.lean` for a more advanced example. -## Animated html +~~We also support all elements that are exposed by the [Recharts library](https://recharts.org/en-US/api), +so you can make your own plots. See `src/Demo/Plot.lean` for an example.~~ (Currently broken.) -As a hidden feature, you can also make animated widgets by passing an array of `Widget.Html` objects to the `staticHtml` widget. This works particularly well with the rechart plotting library, which eases between different plots. -You can see an example of how to do this in `src/Demo/Plot.lean` +## Custom `Expr` displays -# Development +Just like delaborators and unexpanders allow you to customize how expressions are displayed as text, +WidgetKit allows "delaborating" into (potentially interactive) HTML. See +`WidgetKit/Demos/Presentation.lean`. -WidgetKit has the built, bundled JavaScript files checked in to the git repository (under widget/dist/*.js), so you do not need to build these yourself. -If you want to rebuild the javascript, you need to do - -``` -cd widget -npm install -cd .. -lake build widgets -``` +## Animated HTML -This should overwrite the files in `widget/dist`. \ No newline at end of file +~~As a hidden feature, you can also make animated widgets by passing an array of `Widget.Html` +objects to the `staticHtml` widget. This works particularly well with the rechart plotting library, +which eases between different plots. You can see an example of how to do this in +`src/Demo/Plot.lean`.~~ (Currently broken.) diff --git a/WidgetKit.lean b/WidgetKit.lean index cccc5a2..d4f6af5 100644 --- a/WidgetKit.lean +++ b/WidgetKit.lean @@ -1,4 +1,9 @@ -import WidgetKit.Html -import WidgetKit.HtmlWidget -import WidgetKit.Json -import WidgetKit.Svg +import WidgetKit.Component.Basic +import WidgetKit.Component.HtmlDisplay +import WidgetKit.Component.InteractiveSvg +import WidgetKit.Component.PenroseDiagram +import WidgetKit.Data.Html +import WidgetKit.Data.Json +import WidgetKit.Data.Svg +import WidgetKit.Presentation.Expr +import WidgetKit.Presentation.Goal diff --git a/WidgetKit/Compat.lean b/WidgetKit/Compat.lean new file mode 100644 index 0000000..3749b01 --- /dev/null +++ b/WidgetKit/Compat.lean @@ -0,0 +1,186 @@ +import Lean.Attributes +import Lean.Widget.UserWidget + +import WidgetKit.Data.Json + +/-! +A compatibility layer aimed at redefining the user widget API in terms of modules and components. +New features: +- component props have Lean types +- props can be RpcEncodable +- we distinguish between an arbitrary 'widget module' (any ES module) and a 'widget component', + that is a module which can be rendered +- moreover only 'panel widget components' can appear as top-level panels in the infoview + +TODO: If the design works out, it could replace the current Lean core definitions. +-/ + +namespace WidgetKit +open Lean Server Elab + +deriving instance TypeName for Expr + +abbrev LazyEncodable α := StateM RpcObjectStore α + +instance : RpcEncodable (LazyEncodable Json) where + rpcEncode fn := fn + rpcDecode j := return return j + +-- back from exile +structure ExprWithCtx where + ci : Elab.ContextInfo + lctx : LocalContext + expr : Expr + deriving TypeName + +def ExprWithCtx.runMetaM (e : ExprWithCtx) (x : Expr → MetaM α) : IO α := + e.ci.runMetaM e.lctx (x e.expr) + +def ExprWithCtx.save (e : Expr) : MetaM ExprWithCtx := + return { + ci := ← ContextInfo.save + lctx := ← getLCtx + expr := e + } + +open Command in +/-- HACK: around https://leanprover.zulipchat.com/#narrow/stream/341532-lean4-dev/topic/Should.20instance.20names.20inherit.20macro.20scopes.3F -/ +elab "#mkrpcenc" n:ident : command => do + elabCommand <| ← `( + namespace $n + deriving instance Lean.Server.RpcEncodable for $n + end $n + ) + +def widgetDefPostfix : Name := `userWidgetDef + +-- NOTE: Compared to core, this is almost like UserWidgetDefinition but with a different "attitude": +-- the module itself need not be a user widget, i.e. it can also be a support library. It doesn't +-- need a displayable `name`. +structure Module where + /-- An arbitrary JS [module](https://developer.mozilla.org/en-US/docs/Web/JavaScript/Guide/Modules) + intended for use in user widgets. + + To initialize this field from an external JS file, use `include_str "path"/"to"/"file.js"`. + However **beware** that this does not register a dependency with Lake and your module will not + automatically be rebuilt when the `.js` file changes. -/ + javascript : String + +-- This could maybe be a macro but good luck actually writing it. +open Lean Meta Elab Command in +initialize + registerBuiltinAttribute { + name := `widget_module + descr := "Registers a widget module. Its type must extend WidgetKit.Module." + applicationTime := AttributeApplicationTime.afterCompilation + -- The implementation is a hack due to the fact that widgetSource is tied to the storage + -- of user widgets. I think a single widgetModuleRegistry should suffice. TODO fix in core. + add := fun nm _ _ => do + -- The type is not checked; really we just need it to have a `javascript : String` field. + let proc : CommandElabM Unit := do + elabCommand <| ← `(command| + @[widget] + def $(mkIdent <| nm ++ widgetDefPostfix) : Lean.Widget.UserWidgetDefinition where + name := $(quote nm.toString) + javascript := $(mkIdent nm).javascript) + let ctx ← read + let st ← get + -- Cursed manual CommandElabM => CoreM lift punning all fields + let (_, st') ← proc.run { st, ctx with tacticCache? := none } |>.run { st, ctx with } + set { st' with : Core.State } + : AttributeImpl } + +-- "goal widget"/"at cursor widget"/"panel widget"/"info block widget" +structure PanelWidgetInfo where + /-- Name of the `widget_module` to show. -/ + id : Name + props : LazyEncodable Json + -- Compatibility hack. Remove if in core. + infoId : Name + deriving TypeName + +/-- A widget component with a resolved source hash and its props. -/ +structure WidgetInstance where + /-- Name of the `widget_module`. -/ + id : Name + srcHash : UInt64 + props : LazyEncodable Json + -- Compatibility hack. Remove if in core. + infoId : Name + +#mkrpcenc WidgetInstance + +structure PanelWidgetInstance extends WidgetInstance where + range? : Option Lsp.Range + +#mkrpcenc PanelWidgetInstance + +structure GetPanelWidgetsParams where + pos : Lsp.Position + deriving FromJson, ToJson + +structure GetPanelWidgetsResponse where + widgets : Array PanelWidgetInstance + +#mkrpcenc GetPanelWidgetsResponse + +def customInfosAt? (text : FileMap) (t : InfoTree) (hoverPos : String.Pos) : List CustomInfo := + t.deepestNodes fun + | _ctx, i@(Info.ofCustomInfo ci), _cs => do + if let (some pos, some tailPos) := (i.pos?, i.tailPos?) then + let trailSize := i.stx.getTrailingSize + -- show info at EOF even if strictly outside token + trail + let atEOF := tailPos.byteIdx + trailSize == text.source.endPos.byteIdx + guard <| pos ≤ hoverPos ∧ (hoverPos.byteIdx < tailPos.byteIdx + trailSize || atEOF) + return ci + else + failure + | _, _, _ => none + +open RequestM in +@[server_rpc_method] +def getPanelWidgets (args : GetPanelWidgetsParams) : RequestM (RequestTask GetPanelWidgetsResponse) := do + let doc ← readDoc + let filemap := doc.meta.text + let pos := filemap.lspPosToUtf8Pos args.pos + withWaitFindSnap doc (·.endPos >= pos) (notFoundX := return ⟨∅⟩) fun snap => do + let ws := customInfosAt? filemap snap.infoTree pos + let mut widgets := #[] + for w in ws do + let some wi := w.value.get? PanelWidgetInfo + | IO.eprintln "Oops! Unknown ofCustomInfo found. TODO add kinds" + return {widgets := #[]} + let some widgetDef := Widget.userWidgetRegistry.find? snap.env (wi.id ++ widgetDefPostfix) + | throw <| RequestError.invalidParams s!"No registered widget source with id {wi.id}" + widgets := widgets.push { wi with + srcHash := widgetDef.javascriptHash + range? := String.Range.toLspRange filemap <$> Syntax.getRange? w.stx + } + return {widgets} + +@[widget] +def metaWidget : Lean.Widget.UserWidgetDefinition where + -- The header is sometimes briefly visible before compat.tsx loads and hides it + name := "Loading WidgetKit.." + javascript := include_str ".." / "build" / "js" / "compat.js" + +open scoped Json in +/-- Save the data of a panel widget which will be displayed whenever the text cursor is on `stx`. +`id` must be the name of a definition annotated with `@[widget_module]`. + +Note that to be a good citizen which fits in the infoview, a panel widget should be a block element +with some way to collapse it, for example by having `
` be the top-level tag. -/ +def savePanelWidgetInfo [Monad m] [MonadInfoTree m] [MonadNameGenerator m] + (stx : Syntax) (id : Name) (props : LazyEncodable Json) : m Unit := do + let infoId := `panelWidget ++ (← mkFreshId) + pushInfoLeaf <| .ofUserWidgetInfo { + stx + widgetId := ``metaWidget + props := json% { + infoId : $(infoId) + } + } + let wi : PanelWidgetInfo := { id, props, infoId } + pushInfoLeaf <| .ofCustomInfo { stx, value := Dynamic.mk wi } + +end WidgetKit diff --git a/WidgetKit/Component/Basic.lean b/WidgetKit/Component/Basic.lean new file mode 100644 index 0000000..974279c --- /dev/null +++ b/WidgetKit/Component/Basic.lean @@ -0,0 +1,62 @@ +import Lean.Widget.InteractiveCode +import WidgetKit.Compat + +namespace WidgetKit + +/-- A component is a widget module whose `default` export is a +[React component](https://reactjs.org/docs/components-and-props.html). The definition must be +annotated with `@[widget_module]` to be accessible from the infoview. + +## Execution environment + +The JS environment in which components execute provides a fixed set of libraries accessible via +direct `import`, notably `@leanprover/infoview`. All React contexts exported from +`@leanprover/infoview` are usable from components. + +## Lean encoding of props + +`Props` is expected to have a `Lean.Server.RpcEncodable` instance. The JS export should then have +type `function(props: JsProps & { pos : DocumentPosition }): React.ReactNode` where `JsProps` is the +JSON encoding of `Props` and `DocumentPosition` is defined in `@leanprover/infoview`. -/ +structure Component (Props : Type) extends Module + +open Lean + +structure InteractiveCodeProps where + fmt : Widget.CodeWithInfos + +#mkrpcenc InteractiveCodeProps + +@[widget_module] +def InteractiveCode : Component InteractiveCodeProps where + javascript := " + import { InteractiveCode } from '@leanprover/infoview' + import * as React from 'react' + export default function(props) { + return React.createElement(InteractiveCode, props) + }" + +structure InteractiveExprProps where + expr : Server.WithRpcRef ExprWithCtx + +#mkrpcenc InteractiveExprProps + +@[server_rpc_method] +def ppExprTagged : InteractiveExprProps → Server.RequestM (Server.RequestTask Widget.CodeWithInfos) + | ⟨⟨expr⟩⟩ => Server.RequestM.asTask <| expr.runMetaM Widget.ppExprTagged + +/-- Pretty-print and present a `Lean.Expr` as interactive text. + +If you are producing a `Component`in a `MetaM` context where `Widget.ppExprTagged` can be used +directly, it is preferrable to use that and the `InteractiveCode` component instead. This avoids +one extra roundtrip to make the RPC call. -/ +@[widget_module] +def InteractiveExpr : Component InteractiveExprProps where + javascript := include_str ".." / ".." / "build" / "js" / "interactiveExpr.js" + +/-- These are the props passed to a panel widget, i.e. a component which can appear as a top-level +panel in the infoview. For example, a goal state display. See `savePanelWidgetInfo`. -/ +-- TODO: This contains the fields described in `userWidget.tsx` +structure PanelWidgetProps where + +end WidgetKit diff --git a/WidgetKit/Component/HtmlDisplay.lean b/WidgetKit/Component/HtmlDisplay.lean new file mode 100644 index 0000000..768c2e0 --- /dev/null +++ b/WidgetKit/Component/HtmlDisplay.lean @@ -0,0 +1,53 @@ +import Lean.Server.Rpc.Basic + +import WidgetKit.Data.Html + +namespace WidgetKit +open Lean Server + +structure HtmlDisplayProps where + html : EncodableHtml + +#mkrpcenc HtmlDisplayProps + +@[widget_module] +def HtmlDisplay : Component HtmlDisplayProps where + javascript := include_str ".." / ".." / "build" / "js" / "htmlDisplay.js" + +@[widget_module] +def HtmlDisplayPanel : Component HtmlDisplayProps where + javascript := include_str ".." / ".." / "build" / "js" / "htmlDisplayPanel.js" + +open Elab in +unsafe def evalEncodableHtmlUnsafe (stx : Term) : TermElabM EncodableHtml := do + let htmlT := mkConst ``EncodableHtml + Term.evalTerm EncodableHtml htmlT stx + +open Elab in +@[implemented_by evalEncodableHtmlUnsafe] +opaque evalEncodableHtml : Term → TermElabM EncodableHtml + +syntax (name := htmlCmd) "#html " term : command + +open Elab Command Json in +@[command_elab htmlCmd] +def elabHtmlCmd : CommandElab := fun + | stx@`(#html $t:term) => + runTermElabM fun _ => do + let ht ← evalEncodableHtml <| ← `(WidgetKit.EncodableHtml.ofHtml $t) + savePanelWidgetInfo stx ``HtmlDisplayPanel do + return json% { html: $(← rpcEncode ht) } + | stx => throwError "Unexpected syntax {stx}." + +syntax (name := htmlTac) "html! " term : tactic + +open Elab Tactic Json in +@[tactic htmlTac] +def elabHtmlTac : Tactic + | stx@`(tactic| html! $t:term) => do + let ht ← evalEncodableHtml <| ← `(WidgetKit.EncodableHtml.ofHtml $t) + savePanelWidgetInfo stx ``HtmlDisplayPanel do + return json% { html: $(← rpcEncode ht) } + | stx => throwError "Unexpected syntax {stx}." + +end WidgetKit diff --git a/WidgetKit/InteractiveSvg.lean b/WidgetKit/Component/InteractiveSvg.lean similarity index 85% rename from WidgetKit/InteractiveSvg.lean rename to WidgetKit/Component/InteractiveSvg.lean index ad9c2ba..8522435 100644 --- a/WidgetKit/InteractiveSvg.lean +++ b/WidgetKit/Component/InteractiveSvg.lean @@ -1,12 +1,9 @@ -import Lean.Elab +import WidgetKit.Data.Svg -import WidgetKit.Svg +namespace WidgetKit +open Lean -open Lean.Widget.Jsx -open Lean Widget - - -private def Float.toInt (x : Float) : Int := +private def _root_.Float.toInt (x : Float) : Int := if x >= 0 then x.toUInt64.toNat else @@ -27,7 +24,7 @@ structure Action where data : Option Json deriving ToJson, FromJson -/-- The input type `State` is any state the user wants to use and update +/-- The input type `State` is any state the user wants to use and update SvgState in addition automatically handles tracking of time, selection and custom data -/ structure SvgState (State : Type) where @@ -38,6 +35,8 @@ structure SvgState (State : Type) where idToData : List (String × Json) deriving ToJson, FromJson +#mkrpcenc SvgState + structure UpdateParams (State : Type) where elapsed : Float actions : Array Action @@ -46,24 +45,25 @@ structure UpdateParams (State : Type) where deriving ToJson, FromJson structure UpdateResult (State : Type) where - html : Widget.Html + html : EncodableHtml state : SvgState State /-- Approximate number of milliseconds to wait before calling again. -/ callbackTime : Option Float := some 33 - deriving ToJson, FromJson + +#mkrpcenc UpdateResult -- maybe add title, refresh rate, initial time?, custom selection rendering structure InteractiveSvg (State : Type) where init : State frame : Svg.Frame - update (time_ms Δt_ms : Float) (action : Action) - (mouseStart mouseEnd : Option (Svg.Point frame)) + update (time_ms Δt_ms : Float) (action : Action) + (mouseStart mouseEnd : Option (Svg.Point frame)) (selectedId : Option String) (getSelectedData : (α : Type) → [FromJson α] → Option α) : State → State render (time_ms : Float) (mouseStart mouseEnd : Option (Svg.Point frame)) : State → Svg frame -open Server RequestM in -def InteractiveSvg.serverRpcMethod {State : Type} (isvg : InteractiveSvg State) (params : UpdateParams State) +open Server RequestM Jsx in +def InteractiveSvg.serverRpcMethod {State : Type} (isvg : InteractiveSvg State) (params : UpdateParams State) : RequestM (RequestTask (UpdateResult State)) := do -- Ideally, each action should have time and mouse position attached @@ -76,14 +76,14 @@ def InteractiveSvg.serverRpcMethod {State : Type} (isvg : InteractiveSvg State) let mut state := params.state.state let mut selected := params.state.selected - let getData := λ (α : Type) [FromJson α] => do - let id ← selected; + let getData := λ (α : Type) [FromJson α] => do + let id ← selected; let data ← idToData[id] match fromJson? (α:=α) data with | .error _ => none | .ok val => some val - + let mouseStart := params.state.mousePos.map λ (i,j) => (i, j) let mouseEnd := params.mousePos.map λ (x,y) => (x.toInt, y.toInt) @@ -103,9 +103,9 @@ def InteractiveSvg.serverRpcMethod {State : Type} (isvg : InteractiveSvg State) time := time + Δt let mut svg := isvg.render time mouseStart mouseEnd state - - let svgState : SvgState State := - { state := state + + let svgState : SvgState State := + { state := state time := params.elapsed selected := selected mousePos := mouseEnd.map λ p => p.toPixels @@ -119,7 +119,7 @@ def InteractiveSvg.serverRpcMethod {State : Type} (isvg : InteractiveSvg State) return RequestTask.pure { - html := + html := EncodableHtml.ofHtml
{svg.toHtml}
, @@ -128,4 +128,4 @@ def InteractiveSvg.serverRpcMethod {State : Type} (isvg : InteractiveSvg State) } end Svg - +end WidgetKit diff --git a/WidgetKit/Component/PenroseDiagram.lean b/WidgetKit/Component/PenroseDiagram.lean new file mode 100644 index 0000000..3bc0925 --- /dev/null +++ b/WidgetKit/Component/PenroseDiagram.lean @@ -0,0 +1,28 @@ +import WidgetKit.Component.Basic +import WidgetKit.Data.Html + +namespace WidgetKit +open Lean Server + +structure PenroseDiagramProps where + embeds : Array (String × EncodableHtml) + dsl : String + sty : String + sub : String + deriving Inhabited + +#mkrpcenc PenroseDiagramProps + +/-- Displays the given diagram using [Penrose](https://penrose.cs.cmu.edu/). The website contains +explanations of how to write domain (`dsl`), style (`sty`), and substance (`sub`) programs. + +The diagram may also contain embedded HTML trees which are specified in `embeds`. Each embed is HTML +together with the name of an object `x` in the substance program which *must* be assigned +a `x.textBox : Rectangle` field in the style program. This rectangle will be replaced with the HTML +tree. +-/ +@[widget_module] +def PenroseDiagram : Component PenroseDiagramProps where + javascript := include_str ".." / ".." / "build" / "js" / "penroseDisplay.js" + +end WidgetKit diff --git a/WidgetKit/Data/Html.lean b/WidgetKit/Data/Html.lean new file mode 100644 index 0000000..abe3b6b --- /dev/null +++ b/WidgetKit/Data/Html.lean @@ -0,0 +1,155 @@ +/- + Copyright (c) 2021 Wojciech Nawrocki. All rights reserved. + Released under Apache 2.0 license as described in the file LICENSE. + Authors: Wojciech Nawrocki, Sebastian Ullrich + -/ +import Lean.Data.Json.FromToJson +import Lean.Parser +import Lean.PrettyPrinter.Delaborator.Basic +import Lean.Server.Rpc.Basic + +import WidgetKit.Component.Basic + +/-! We define a representation of HTML trees together with a JSX-like DSL for writing them. -/ + +namespace WidgetKit +open Lean Server + +/-- A HTML tree which may contain widget components. -/ +inductive Html where + /-- An `element "tag" attrs children` represents `{...children}`. -/ + | element : String → Array (String × Json) → Array Html → Html + /-- Raw HTML text.-/ + | text : String → Html + /-- A `component Foo props children` represents `{...children}`. -/ + -- TODO: The universe lift is unfortunate. + | component {Props} [RpcEncodable Props] : Component Props → Props → Array Html → Html + +/-- A variant of `Html` which lives in `Type`, at the cost of some type safety. + +Unfortunately we cannot build `RpcEncodable Html` because `Html : Type 1` and `RpcEncodable` is not +universe-polymorphic. We can, however, do it for `EncodableHtml`. -/ +inductive EncodableHtml where + | element : String → Array (String × Json) → Array EncodableHtml → EncodableHtml + | text : String → EncodableHtml + | component : UInt64 → LazyEncodable Json → Array EncodableHtml → EncodableHtml + deriving Inhabited + +#mkrpcenc EncodableHtml + +partial def EncodableHtml.ofHtml : Html → EncodableHtml + | .element t as cs => element t as (cs.map ofHtml) + | .text s => text s + | @Html.component _ _ c ps cs => component (hash c.javascript) (rpcEncode ps) (cs.map ofHtml) + +namespace Jsx +open Parser PrettyPrinter + +declare_syntax_cat jsxElement +declare_syntax_cat jsxChild +declare_syntax_cat jsxAttr +declare_syntax_cat jsxAttrVal + +-- See https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/expected.20parser.20to.20return.20exactly.20one.20syntax.20object +-- def jsxAttrVal : Parser := strLit <|> group ("{" >> termParser >> "}") +scoped syntax str : jsxAttrVal +scoped syntax group("{" term "}") : jsxAttrVal +scoped syntax ident "=" jsxAttrVal : jsxAttr + +-- JSXTextCharacter : SourceCharacter but not one of {, <, > or } +def jsxText : Parser := + withAntiquot (mkAntiquot "jsxText" `jsxText) { + fn := fun c s => + let startPos := s.pos + let s := takeWhile1Fn (not ∘ "{<>}$".contains) "expected JSX text" c s + mkNodeToken `jsxText startPos c s } + +def getJsxText : TSyntax `jsxText → String + | stx => stx.raw[0].getAtomVal + +@[combinator_formatter WidgetKit.Jsx.jsxText] +def jsxText.formatter : Formatter := + Formatter.visitAtom `jsxText +@[combinator_parenthesizer WidgetKit.Jsx.jsxText] +def jsxText.parenthesizer : Parenthesizer := + Parenthesizer.visitToken + +scoped syntax "<" ident jsxAttr* "/>" : jsxElement +scoped syntax "<" ident jsxAttr* ">" jsxChild* "" : jsxElement + +scoped syntax jsxText : jsxChild +-- TODO(WN): expand `{... $t}` as list of children +scoped syntax "{" term "}" : jsxChild +scoped syntax jsxElement : jsxChild + +scoped syntax:max jsxElement : term + +def transformTag (n m : Ident) (ns : Array Ident) (vs : Array (TSyntax `jsxAttrVal)) + (cs : Array (TSyntax `jsxChild)) : MacroM Term := do + if n.getId != m.getId then + Macro.throwErrorAt m s!"expected " + let cs ← cs.mapM fun + | `(jsxChild| $t:jsxText) => `(Html.text $(quote <| getJsxText t)) + | `(jsxChild| { $t }) => return t + | `(jsxChild| $e:jsxElement) => `(term| $e:jsxElement) + | _ => unreachable! + let vs : Array (TSyntax `term) := vs.map fun + | `(jsxAttrVal| $s:str) => s + | `(jsxAttrVal| { $t:term }) => t + | _ => unreachable! + let tag := toString n.getId + -- Uppercase tags are parsed as components + if tag.get? 0 |>.filter (·.isUpper) |>.isSome then + `(Html.component $n { $[$ns:ident := $vs],* } #[ $[$cs],* ]) + -- Lowercase tags are parsed as standard HTML + else + let ns := ns.map (quote <| toString ·.getId) + `(Html.element $(quote tag) #[ $[($ns, $vs)],* ] #[ $[$cs],* ]) + + +/-- Support for writing HTML trees directly, using XML-like angle bracket syntax. It work very +similarly to [JSX](https://reactjs.org/docs/introducing-jsx.html) in JavaScript. The syntax is +enabled using `open scoped WidgetKit.Jsx`. + +Lowercase tags are interpreted as standard HTML whereas uppercase ones are expected to be +`WidgetKit.Component`s. -/ +macro_rules + | `(<$n:ident $[$ns:ident = $vs:jsxAttrVal]* />) => transformTag n n ns vs #[] + | `(<$n:ident $[$ns:ident = $vs:jsxAttrVal]* >$cs*) => transformTag n m ns vs cs + +open Lean Delaborator SubExpr + +@[delab app.WidgetKit.Html.text] +def delabHtmlText : Delab := do + withAppArg delab + +@[delab app.WidgetKit.Html.element] +def delabHtmlElement : Delab := do + let e ← getExpr + -- `Html.element tag attrs children` + let #[tag, _, _] := e.getAppArgs | failure + + let .lit (.strVal s) := tag | failure + let tag := mkIdent s + + let attrs ← withAppFn (withAppArg delab) + let `(term| #[ $[($as:str, $vs)],* ] ) := attrs | failure + let attrs : Array (TSyntax `jsxAttr) ← as.zip vs |>.mapM fun (a, v) => do + let attr := mkIdent a.getString + `(jsxAttr| $attr:ident={ $v }) + + let children ← withAppArg delab + let `(term| #[ $cs,* ]) := children | failure + let cs : Array (TSyntax `jsxChild) ← (@id (TSyntaxArray `term) cs).mapM fun c => do + if let some s := c.raw.isStrLit? then + let txt := mkNode `jsxText #[mkAtom s] + `(jsxChild| $txt:jsxText) + -- hack. + else if c.raw[0].getKind == ``WidgetKit.Jsx.«jsxElement<__>_» then + `(jsxChild| $(⟨c.raw⟩):jsxElement) + else + `(jsxChild| { $c }) + `(term| < $tag $[$attrs]* > $[$cs]* ) + +end Jsx +end WidgetKit diff --git a/WidgetKit/Json.lean b/WidgetKit/Data/Json.lean similarity index 73% rename from WidgetKit/Json.lean rename to WidgetKit/Data/Json.lean index b8a3fc6..09a8729 100644 --- a/WidgetKit/Json.lean +++ b/WidgetKit/Data/Json.lean @@ -4,16 +4,15 @@ Authors: E.W.Ayers -/ import Lean.Data.Json -import Lean.Meta.Basic -import Lean.Elab.Term -import Lean.Elab.Eval /-! -# Json-like syntax for Lean. +# JSON-like syntax for Lean. Now you can write ```lean +open scoped WidgetKit.Json + #eval json% { hello : "world", cheese : ["edam", "cheddar", {kind : "spicy", rank : 100.2}], @@ -23,9 +22,11 @@ Now you can write lookACalc: $(23 + 54 * 2) } ``` - -/ +-/ + +namespace WidgetKit.Json +open Lean -open Lean Parser declare_syntax_cat jso declare_syntax_cat jso_field declare_syntax_cat jso_ident @@ -56,20 +57,20 @@ instance : ToJson Float where let j := if x < 0.0 then -j else j Json.num j -syntax "[" jso,* "]" : jso -syntax "-"? scientific : jso -syntax "-"? num : jso -syntax str : jso -syntax "true" : jso -syntax "false" : jso -syntax "null" : jso -syntax ident : jso_ident -syntax "$(" term ")" : jso_ident -syntax str : jso_ident -syntax jso_ident ": " jso : jso_field -syntax "{" jso_field,* "}" : jso -syntax "$(" term ")" : jso -syntax "json% " jso : term +scoped syntax "[" jso,* "]" : jso +scoped syntax "-"? scientific : jso +scoped syntax "-"? num : jso +scoped syntax str : jso +scoped syntax "true" : jso +scoped syntax "false" : jso +scoped syntax "null" : jso +scoped syntax ident : jso_ident +scoped syntax "$(" term ")" : jso_ident +scoped syntax str : jso_ident +scoped syntax jso_ident ": " jso : jso_field +scoped syntax "{" jso_field,* "}" : jso +scoped syntax "$(" term ")" : jso +scoped syntax "json% " jso : term macro_rules | `(json% $($t)) => `(Lean.toJson $t) @@ -90,12 +91,4 @@ macro_rules | stx => panic! s!"unrecognized ident syntax {stx}" `(Lean.Json.mkObj [$[($ks, json% $vs)],*]) -open Lean Elab Meta Term in -unsafe def Lean.evalJsonUnsafe (stx : Syntax) : TermElabM Json := do - let JsonT := mkConst ``Json - let jsonStx : TSyntax `jso := TSyntax.mk stx - Elab.Term.evalTerm Json JsonT (← `(json% $jsonStx)) - -open Lean Elab in -@[implemented_by evalJsonUnsafe] -opaque Lean.evalJson : Syntax → TermElabM Json +end WidgetKit.Json diff --git a/WidgetKit/Svg.lean b/WidgetKit/Data/Svg.lean similarity index 72% rename from WidgetKit/Svg.lean rename to WidgetKit/Data/Svg.lean index 77089ac..aaa6f26 100644 --- a/WidgetKit/Svg.lean +++ b/WidgetKit/Data/Svg.lean @@ -1,17 +1,16 @@ -import WidgetKit.HtmlWidget -import WidgetKit.Json +import WidgetKit.Data.Json +import WidgetKit.Data.Html -open Lean.Widget.Jsx -open Lean Widget +namespace WidgetKit +open Lean - -private def Float.toInt (x : Float) : Int := +private def _root_.Float.toInt (x : Float) : Int := if x >= 0 then x.toUInt64.toNat else -((-x).toUInt64.toNat) -private def Int.toFloat (i : Int) : Float := +private def _root_.Int.toFloat (i : Int) : Float := if i >= 0 then i.toNat.toFloat else @@ -50,7 +49,7 @@ deriving Inhabited, ToJson, FromJson instance (f) : Coe (Float×Float) (Point f) := ⟨λ (x,y) => .abs x y⟩ instance (f) : Coe (Int×Int) (Point f) := ⟨λ (i,j) => .px i j⟩ -def Point.toPixels {f : Frame} (p : Point f) : Int × Int := +def Point.toPixels {f : Frame} (p : Point f) : Int × Int := match p with | .px x y => (x,y) | .abs x y => @@ -59,12 +58,12 @@ def Point.toPixels {f : Frame} (p : Point f) : Int × Int := let j := ((f.ymax - y) / Δx).floor.toInt (i, j) -def Point.toAbsolute {f : Frame} (p : Point f) : Float × Float := +def Point.toAbsolute {f : Frame} (p : Point f) : Float × Float := match p with | .abs x y => (x,y) | .px i j => let Δx := f.pixelSize - let x := f.xmin + (i.toFloat + 0.5) * Δx + let x := f.xmin + (i.toFloat + 0.5) * Δx let y := f.ymax - (j.toFloat + 0.5) * Δx (x,y) @@ -88,26 +87,26 @@ inductive Shape (f : Frame) where deriving ToJson, FromJson def Shape.toHtmlData {f : Frame} : Shape f → String × Array (String × Json) -| .line src trg => +| .line src trg => let (x1,y1) := src.toPixels let (x2,y2) := trg.toPixels ("line", #[("x1", x1), ("y1", y1), ("x2", x2), ("y2", y2)]) -| .circle center radius => +| .circle center radius => let (cx,cy) := center.toPixels let r := radius.toPixels ("circle", #[("cx", cx), ("cy", cy), ("r", r)]) -| .polyline points => - let pts := points +| .polyline points => + let pts := points |>.map (λ p => let (x,y) := p.toPixels; s!"{x},{y}") |>.foldl (init := "") (λ s p => s ++ " " ++ p) ("polyline", #[("points", pts)]) -| .polygon points => - let pts := points +| .polygon points => + let pts := points |>.map (λ p => let (x,y) := p.toPixels; s!"{x},{y}") |>.foldl (init := "") (λ s p => s ++ " " ++ p) ("polygon", #[("fillRule", "nonzero"), ("points", pts)]) - + structure Element (f : Frame) where shape : Shape f strokeColor := (none : Option Color) @@ -132,21 +131,21 @@ def Element.setData {α : Type} {f} (elem : Element f) (a : α) [ToJson α] := def Element.toHtml {f : Frame} (e : Element f) : Html := Id.run do let mut (tag, args) := e.shape.toHtmlData - if let .some color ← e.strokeColor then + if let .some color := e.strokeColor then args := args.push ("stroke", color.toStringRGB) - if let .some width ← e.strokeWidth then + if let .some width := e.strokeWidth then args := args.push ("strokeWidth", width.toPixels) - if let .some color ← e.fillColor then + if let .some color := e.fillColor then args := args.push ("fill", color.toStringRGB) else args := args.push ("fill", "none") - if let .some id ← e.id then + if let .some id := e.id then args := args.push ("id", id) - if let .some data ← e.data then + if let .some data := e.data then args := args.push ("data", data) return .element tag args #[] @@ -158,69 +157,47 @@ def Element.toHtml {f : Frame} (e : Element f) : Html := Id.run do end Svg -def mkIdToIdx {f} (elements : Array (Svg.Element f)) : HashMap String (Fin elements.size) := +def mkIdToIdx {f} (elements : Array (Svg.Element f)) : HashMap String (Fin elements.size) := let idToIdx := (elements |>.mapIdx (λ idx el => (idx,el))) -- zip with index |>.filterMap (λ (idx,el) => el.id.map (λ id => (id,idx))) -- keep only elements with specified id - |>.toList + |>.toList |> HashMap.ofList idToIdx structure Svg (f : Svg.Frame) where elements : Array (Svg.Element f) - idToIdx := mkIdToIdx elements - -def Svg.toHtml {f : Frame} (svg : Svg f) : Html := - .element "svg" - #[("xmlns", "http://www.w3.org/2000/svg"), - ("version", "1.1"), - ("width", f.width), - ("height", f.height)] + idToIdx := mkIdToIdx elements + +namespace Svg + +def toHtml {f : Frame} (svg : Svg f) : Html := + .element "svg" + #[("xmlns", "http://www.w3.org/2000/svg"), + ("version", "1.1"), + ("width", f.width), + ("height", f.height)] (svg.elements.map λ e => e.toHtml) -def Svg.idToDataList {f} (svg : Svg f) : List (String × Json) := - svg.elements.foldr (init := []) (λ e l => +def idToDataList {f} (svg : Svg f) : List (String × Json) := + svg.elements.foldr (init := []) (λ e l => match e.id, e.data with | some id, some data => (id,data)::l | _, _ => l) -def Svg.idToData {f} (svg : Svg f) : HashMap String Json := +def idToData {f} (svg : Svg f) : HashMap String Json := HashMap.ofList svg.idToDataList - + instance {f} : GetElem (Svg f) Nat (Svg.Element f) (λ svg idx => idx < svg.elements.size) where getElem svg i h := svg.elements[i] instance {f} : GetElem (Svg f) String (Option (Svg.Element f)) (λ _ _ => True) where getElem svg id _ := svg.idToIdx[id].map (λ idx => svg.elements[idx]) -def Svg.getData {f} (svg : Svg f) (id : String) : Option Json := +def getData {f} (svg : Svg f) (id : String) : Option Json := match svg[id] with | none => none | some elem => elem.data - -section Example - - open Svg - - private def frame : Frame where - xmin := -2 - ymin := -2 - xSize := 4 - width := 400 - height := 400 - - private def svg : Svg frame := - { elements := - #[line (0.,0.) (1.,0.) |>.setStroke (1.,0.,0.) (.px 2), - line (1.,0.) (0.,1.) |>.setStroke (0.,1.,0.) (.px 2), - line (0.,1.) (0.,0.) |>.setStroke (0.,0.,1.) (.px 2), - circle (0.,0.) (.abs 0.1) |>.setStroke (0.,0.,0.) (.px 2) |>.setFill (0.,1.,1.) |>.setId "point1", - circle (1.,0.) (.abs 0.1) |>.setStroke (0.,0.,0.) (.px 2) |>.setFill (1.,0.,1.) |>.setId "point2", - circle (0.,1.) (.abs 0.1) |>.setStroke (0.,0.,0.) (.px 2) |>.setFill (1.,1.,0.) |>.setId "point3"] } - - -- #eval toJson svg.toHtml - - #html svg.toHtml - -end Example +end Svg +end WidgetKit diff --git a/WidgetKit/Demos/Basic.lean b/WidgetKit/Demos/Basic.lean index 5908c07..9042488 100644 --- a/WidgetKit/Demos/Basic.lean +++ b/WidgetKit/Demos/Basic.lean @@ -1,14 +1,13 @@ -import WidgetKit.HtmlWidget -open Lean.Widget.Jsx -- ⟵ remember this! +import WidgetKit.Component.HtmlDisplay -def x := You can use HTML in lean! {toString $ 4 + 5}
+open scoped WidgetKit.Jsx -- ⟵ remember this! + +def x := You can use HTML in lean! {.text <| toString <| 4 + 5}
#html x -open scoped Lean.Widget.Jsx in theorem ghjk : True := by html! What, HTML in Lean?! html! And another! html! trivial - diff --git a/WidgetKit/Demos/Graphics2D.lean b/WidgetKit/Demos/Graphics2D.lean index 7c42912..777313d 100644 --- a/WidgetKit/Demos/Graphics2D.lean +++ b/WidgetKit/Demos/Graphics2D.lean @@ -1,9 +1,9 @@ -import WidgetKit.HtmlWidget -import WidgetKit.Json - -open Lean.Widget.Jsx -open Lean Widget +import WidgetKit.Data.Json +import WidgetKit.Component.HtmlDisplay +open Lean +open WidgetKit +open scoped WidgetKit.Jsx namespace Svg @@ -60,26 +60,26 @@ inductive Shape where deriving ToJson, FromJson def Shape.toHtmlData (frame : Frame) : Shape → String × Array (String × Json) -| .line src trg => +| .line src trg => let (x1,y1) := src.toPixels frame let (x2,y2) := trg.toPixels frame ("line", #[("x1", x1), ("y1", y1), ("x2", x2), ("y2", y2)]) -| .circle center radius => +| .circle center radius => let (cx,cy) := center.toPixels frame let r := radius.toPixels frame ("circle", #[("cx", cx), ("cy", cy), ("r", r)]) -| .polyline points => - let pts := points +| .polyline points => + let pts := points |>.map (λ p => let (x,y) := p.toPixels frame; s!"{x},{y}") |>.foldl (init := "") (λ s p => s ++ " " ++ p) ("path", #[("points", pts)]) -| .polygon points => - let pts := points +| .polygon points => + let pts := points |>.map (λ p => let (x,y) := p.toPixels frame; s!"{x},{y}") |>.foldl (init := "") (λ s p => s ++ " " ++ p) ("polygon", #[("fillRule", "nonzero"), ("points", pts)]) - + structure Element where shape : Shape strokeColor := (none : Option Color) @@ -91,13 +91,13 @@ deriving ToJson, FromJson def Element.toHtml (frame : Frame) (e : Element) : Html := Id.run do let mut (tag, args) := e.shape.toHtmlData frame - if let .some color ← e.strokeColor then + if let .some color := e.strokeColor then args := args.push ("stroke", color.toRGB) - if let .some width ← e.strokeWidth then + if let .some width := e.strokeWidth then args := args.push ("strokeWidth", width.toPixels frame) - if let .some color ← e.fillColor then + if let .some color := e.fillColor then args := args.push ("fill", color.toRGB) - if let .some id ← e.id then + if let .some id := e.id then args := args.push ("click", id) return .element tag args #[] @@ -106,11 +106,11 @@ def Element.toHtml (frame : Frame) (e : Element) : Html := Id.run do end Svg -def mkIdToIdx (elements : Array Svg.Element) : HashMap String (Fin elements.size) := +def mkIdToIdx (elements : Array Svg.Element) : HashMap String (Fin elements.size) := let idToIdx := (elements |>.mapIdx (λ idx el => (idx,el))) -- zip with index |>.filterMap (λ (idx,el) => el.id.map (λ id => (id,idx))) -- keep only elements with specified id - |>.toList + |>.toList |> HashMap.ofList idToIdx @@ -118,22 +118,22 @@ def mkIdToIdx (elements : Array Svg.Element) : HashMap String (Fin elements.size structure Svg where elements : Array Svg.Element frame : Svg.Frame - idToIdx := mkIdToIdx elements + idToIdx := mkIdToIdx elements -def Svg.toHtml (svg : Svg) : Html := - .element "svg" - #[("xmlns", "http://www.w3.org/2000/svg"), - ("version", "1.1"), - ("width", svg.frame.width), - ("height", svg.frame.height)] +def Svg.toHtml (svg : Svg) : Html := + .element "svg" + #[("xmlns", "http://www.w3.org/2000/svg"), + ("version", "1.1"), + ("width", svg.frame.width), + ("height", svg.frame.height)] (svg.elements.map λ e => e.toHtml svg.frame) - + instance : GetElem Svg Nat Svg.Element (λ svg idx => idx < svg.elements.size) where getElem svg i h := svg.elements[i] instance : GetElem Svg String (Option Svg.Element) (λ _ _ => True) where - getElem svg id _ := + getElem svg id _ := match svg.idToIdx[id] with | some idx => svg.elements[idx] | none => none @@ -150,8 +150,8 @@ private def frame : Frame where height := 400 -private def svg : Svg := - { elements := +private def svg : Svg := + { elements := #[{ shape := .line ⟨0,0⟩ ⟨1,0⟩, strokeWidth := some (.pixels 2), strokeColor := some ⟨1,0,0⟩, id := some "line1"}, { shape := .line ⟨1,0⟩ ⟨0,1⟩, strokeWidth := some (.pixels 2), strokeColor := some ⟨0,1,0⟩, id := some "line2"}, { shape := .line ⟨0,1⟩ ⟨0,0⟩, strokeWidth := some (.pixels 2), strokeColor := some ⟨0,0,1⟩, id := some "line3"}, @@ -161,6 +161,6 @@ private def svg : Svg := ], frame := frame } -#eval toJson svg.toHtml +#html svg.toHtml end Example diff --git a/WidgetKit/Demos/InteractiveSvg.lean b/WidgetKit/Demos/InteractiveSvg.lean index 9ffcdf4..9f63fb8 100644 --- a/WidgetKit/Demos/InteractiveSvg.lean +++ b/WidgetKit/Demos/InteractiveSvg.lean @@ -1,16 +1,15 @@ -import WidgetKit.InteractiveSvg +import WidgetKit.Component.InteractiveSvg +import WidgetKit.Component.HtmlDisplay -open Lean.Widget.Jsx -open Lean Widget - -open Svg +open Lean +open WidgetKit Svg Jsx abbrev State := Array (Float × Float) def isvg : InteractiveSvg State where init := #[(-0.5, -0.5), (0.5, -0.5), (0.5, 0.5), (-0.5, 0.5)] - frame := + frame := { xmin := -1 ymin := -1 xSize := 2 @@ -18,22 +17,22 @@ def isvg : InteractiveSvg State where height := 400 } update time Δt action mouseStart mouseEnd selected getData state := - match getData Nat, mouseEnd with + match getData Nat, mouseEnd with | some id, some p => state.set! id p.toAbsolute | _, _ => state - render time mouseStart mouseEnd state := - { - elements := - let mousePointer := + render time mouseStart mouseEnd state := + { + elements := + let mousePointer := match mouseStart, mouseEnd with - | some s, some e => - #[ + | some s, some e => + #[ Svg.circle e (.px 5) |>.setFill (1.,1.,1.), Svg.line s e |>.setStroke (1.,1.,1.) (.px 2) ] | _, _ => #[] - let circles := (state.mapIdx fun idx (p : Float × Float) => + let circles := (state.mapIdx fun idx (p : Float × Float) => Svg.circle p (.abs 0.2) |>.setFill (0.7,0.7,0.7) |>.setId s!"circle{idx}" |>.setData idx.1 ) mousePointer.append circles @@ -44,14 +43,13 @@ open Server RequestM in @[server_rpc_method] def updateSvg (params : UpdateParams State) : RequestM (RequestTask (UpdateResult State)) := isvg.serverRpcMethod params -@[widget] -def svgWidget : UserWidgetDefinition where - name := "Interactive SVG" - javascript := include_str ".." / ".." / "widget" / "dist" / "interactiveSvg.js" - +-- TODO: the tsx file is pretty broken +@[widget_module] +def SvgWidget : Component (UpdateResult State) where + javascript := include_str ".." / ".." / "build" / "js" / "interactiveSvg.js" def init : UpdateResult State := { - html :=
Init!!!
, + html := EncodableHtml.ofHtml
Init!!!
, state := { state := isvg.init time := 0 selected := none @@ -59,5 +57,4 @@ def init : UpdateResult State := { idToData := isvg.render 0 none none isvg.init |>.idToDataList} } -#widget svgWidget (toJson init) - +#html diff --git a/WidgetKit/Demos/Plot.lean b/WidgetKit/Demos/Plot.lean index 3d6062f..38663ac 100644 --- a/WidgetKit/Demos/Plot.lean +++ b/WidgetKit/Demos/Plot.lean @@ -1,13 +1,18 @@ -import WidgetKit.HtmlWidget -import WidgetKit.Json -open Lean.Widget.Jsx -open Lean Widget +import WidgetKit.Data.Json +import WidgetKit.Component.HtmlDisplay + +open scoped WidgetKit.Jsx +open scoped WidgetKit.Json + +open Lean WidgetKit def fn (t : Float) (x : Float): Float := 50 * (x - 0.25) * (x - 0.5) * (x - 0.7) + 0.1 * (x * 40 - t * 2 * 3.141).sin +-- TODO: Currently the demo does not work because we are missing Recharts components +#exit -def Plot (fn : Float → Float) (steps := 100) : Widget.Html := +def Plot (fn : Float → Float) (steps := 100) : Html := let jsonData : Json := List.range (steps + 1) |>.toArray |> Array.map (fun (x : Nat) => let x : Float := x.toFloat / steps.toFloat; (x, fn x)) @@ -34,4 +39,3 @@ def mkFrames (fn : Float → Float → Float) (steps := 100) : Array Widget.Html -- put your cursor on the below line to see an animated widget #widget staticHtmlWidget json% {frames : $(toJson (mkFrames fn)), framesPerSecond : 60} - diff --git a/WidgetKit/Demos/Presentation.lean b/WidgetKit/Demos/Presentation.lean new file mode 100644 index 0000000..6b033c4 --- /dev/null +++ b/WidgetKit/Demos/Presentation.lean @@ -0,0 +1,18 @@ +import WidgetKit.Presentation.Goal + +open WidgetKit Jsx + +@[expr_presenter] +def presenter : ExprPresenter where + userName := "With octopodes" + isApplicable _ := return true + present e := + return EncodableHtml.ofHtml + + {.text "🐙 "}{.text " 🐙"} + + +example (h : 2 + 2 = 5) : 2 + 2 = 4 := by + withSelectionDisplay + -- Place cursor here and select subexpressions in the goal with shift-click + rfl diff --git a/WidgetKit/Demos/Svg.lean b/WidgetKit/Demos/Svg.lean new file mode 100644 index 0000000..7c848de --- /dev/null +++ b/WidgetKit/Demos/Svg.lean @@ -0,0 +1,24 @@ +import WidgetKit.Data.Svg +import WidgetKit.Component.HtmlDisplay + +open WidgetKit Svg + +private def frame : Frame where + xmin := -2 + ymin := -2 + xSize := 4 + width := 400 + height := 400 + +private def svg : Svg frame := + { elements := + #[line (0.,0.) (1.,0.) |>.setStroke (1.,0.,0.) (.px 2), + line (1.,0.) (0.,1.) |>.setStroke (0.,1.,0.) (.px 2), + line (0.,1.) (0.,0.) |>.setStroke (0.,0.,1.) (.px 2), + circle (0.,0.) (.abs 0.1) |>.setStroke (0.,0.,0.) (.px 2) |>.setFill (0.,1.,1.) |>.setId "point1", + circle (1.,0.) (.abs 0.1) |>.setStroke (0.,0.,0.) (.px 2) |>.setFill (1.,0.,1.) |>.setId "point2", + circle (0.,1.) (.abs 0.1) |>.setStroke (0.,0.,0.) (.px 2) |>.setFill (1.,1.,0.) |>.setId "point3"] } + +-- #eval toJson svg.toHtml + +#html svg.toHtml diff --git a/WidgetKit/Html.lean b/WidgetKit/Html.lean deleted file mode 100644 index 85ad6b4..0000000 --- a/WidgetKit/Html.lean +++ /dev/null @@ -1,87 +0,0 @@ -/- - Copyright (c) 2021 Wojciech Nawrocki. All rights reserved. - Released under Apache 2.0 license as described in the file LICENSE. - Authors: Wojciech Nawrocki, Sebastian Ullrich - -/ -import Lean.Data.Json.FromToJson -import Lean.Parser - -/-! This module defines: -- a representation of HTML trees -- together with a JSX-like DSL for writing them -- and widget support for visualizing any type as HTML. -/ - -namespace Lean.Widget - -inductive Html where - -- TODO(WN): it's nameless for shorter JSON; re-add names when we have deriving strategies for From/ToJson - -- element (tag : String) (attrs : Array HtmlAttribute) (children : Array Html) - | element : String → Array (String × Json) → Array Html → Html - | text : String → Html - deriving BEq, Inhabited, FromJson, ToJson - -instance : Coe String Html := - ⟨Html.text⟩ - -namespace Jsx -open Parser PrettyPrinter - -declare_syntax_cat jsxElement -declare_syntax_cat jsxChild - -def jsxAttrVal : Parser := strLit <|> group ("{" >> termParser >> "}") -def jsxAttr : Parser := ident >> "=" >> jsxAttrVal - --- JSXTextCharacter : SourceCharacter but not one of {, <, > or } -def jsxText : Parser := - withAntiquot (mkAntiquot "jsxText" `jsxText) { - fn := fun c s => - let startPos := s.pos - let s := takeWhile1Fn (not ∘ "{<>}$".contains) "expected JSX text" c s - mkNodeToken `jsxText startPos c s } - -@[combinator_formatter Lean.Widget.Jsx.jsxText] def jsxText.formatter : Formatter := pure () -@[combinator_parenthesizer Lean.Widget.Jsx.jsxText] def jsxText.parenthesizer : Parenthesizer := pure () - -scoped syntax "<" ident jsxAttr* "/>" : jsxElement -scoped syntax "<" ident jsxAttr* ">" jsxChild* "" : jsxElement - -scoped syntax jsxText : jsxChild -scoped syntax "{" term "}" : jsxChild -scoped syntax jsxElement : jsxChild - -scoped syntax:max jsxElement : term - -macro_rules - | `(<$n $[$ns = $vs]* />) => - let ns := ns.map (quote <| toString ·.getId) - let vs : Array (TSyntax `term) := vs.map fun - | `(jsxAttrVal| $s:str) => s - | `(jsxAttrVal| { $t:term }) => t - | _ => unreachable! - `(Html.element $(quote <| toString n.getId) #[ $[($ns, toJson $vs)],* ] #[]) - | `(<$n $[$ns = $vs]* >$cs*) => - if n.getId == m.getId then do - let ns := ns.map (quote <| toString ·.getId) - let vs : Array (TSyntax `term) := vs.map fun - | `(jsxAttrVal| $s:str) => s - | `(jsxAttrVal| { $t:term }) => t - | _ => unreachable! - let cs ← cs.mapM fun - | `(jsxChild|$t:jsxText) => `(Html.text $(quote t.raw[0].getAtomVal)) - -- TODO(WN): elab as list of children if type is `t Html` where `Foldable t` - | `(jsxChild|{$t}) => return t - | `(jsxChild|$e:jsxElement) => `($e:jsxElement) - | _ => unreachable! - let tag := toString n.getId - `(Html.element $(quote tag) #[ $[($ns, toJson $vs)],* ] #[ $[$cs],* ]) - else Macro.throwError ("expected ") - -end Jsx - -/-- A type which implements `ToHtmlFormat` will be visualized -as the resulting HTML in editors which support it. -/ -class ToHtmlFormat (α : Type u) where - formatHtml : α → Html - -end Lean.Widget \ No newline at end of file diff --git a/WidgetKit/HtmlWidget.lean b/WidgetKit/HtmlWidget.lean deleted file mode 100644 index 29cd31a..0000000 --- a/WidgetKit/HtmlWidget.lean +++ /dev/null @@ -1,44 +0,0 @@ -/- Author: Wojciech Nawrocki -/ -import Lean.Widget.UserWidget -import Lean.Elab.Eval -import WidgetKit.Html - -open Lean Widget in -@[widget] -def staticHtmlWidget : UserWidgetDefinition where - name := "HTML Display" - javascript := include_str ".." / "widget" / "dist" / "staticHtml.js" - -open Lean Elab Widget in -unsafe def evalHtmlUnsafe (stx : Syntax) : TermElabM Html := do - let htmlT := mkConst ``Html - Term.evalTerm Html htmlT stx - -open Lean Elab Widget in -@[implemented_by evalHtmlUnsafe] -opaque evalHtml : Syntax → TermElabM Html - -syntax (name := htmlCmd) "#html " term : command - -open Lean Meta Elab Command in -@[command_elab htmlCmd] -def elabHtmlCmd : CommandElab := fun - | stx@`(#html $t:term) => - runTermElabM fun _ => do - let id := `staticHtmlWidget - let ht ← evalHtml t - let props := Json.mkObj [("html", toJson ht)] - Lean.Widget.saveWidgetInfo id props stx - | stx => throwError "Unexpected syntax {stx}." - -syntax (name := htmlTac) "html! " term : tactic - -open Lean Elab Tactic in -@[tactic htmlTac] -def elabHtmlTac : Tactic - | stx@`(tactic| html! $t:term) => do - let id := `staticHtmlWidget - let ht ← evalHtml t - let props := Json.mkObj [("html", toJson ht)] - Lean.Widget.saveWidgetInfo id props stx - | stx => throwError "Unexpected syntax {stx}." diff --git a/WidgetKit/Presentation/Expr.lean b/WidgetKit/Presentation/Expr.lean new file mode 100644 index 0000000..c409b34 --- /dev/null +++ b/WidgetKit/Presentation/Expr.lean @@ -0,0 +1,106 @@ +import WidgetKit.Data.Html + +namespace WidgetKit +open Lean Server + +/-- An `Expr` presenter is similar to a delaborator but outputs HTML trees instead of syntax, and +the output HTML can contain elements which interact with the original `Expr` in some way. We call +interactive outputs with a reference to the original input *presentations*. -/ +structure ExprPresenter where + /-- A user-friendly name for this presenter. For example, "LaTeX". -/ + userName : String + /- TODO: there is a general problem of writing env extensions which store an extendable list of + functions to run on `Expr`s, but not all of which are applicable (actually, most are not) to any + single `Expr`. Invoking them in sequence is O(n); we should better use sth like `DiscrTree`. + Registering new entries would need to extend the DiscrTree, perhaps like + `registerSelf : DiscrTree ExprPresenter → DiscrTree ExprPresenter`. + Dispatching on just one constant like e.g. delaborators (`app.MyType.myCtr`) is not sufficient + because one entry may apply to multiple expressions of a given form which could be represented + as a schematic with mvars, say `@ofNat ? 0 ?`. -/ + /-- Should quickly determine if the `Expr` is within this presenter's domain of applicability. + For example it could check for a constant like the `` `name `` in ``@[delab `name]``. -/ + isApplicable : Expr → MetaM Bool + /-- *Must* return `some _` or throw when `isApplicable` is `true`. -/ + present : Expr → MetaM (Option EncodableHtml) + +initialize exprPresenters : TagAttribute ← + registerTagAttribute `expr_presenter + "Register an Expr presenter. It must have the type `WidgetKit.ExprPresenter`." + (validate := fun nm => do + let const ← getConstInfo nm + if !const.type.isConstOf ``ExprPresenter then + throwError m!"type mismatch, expected {mkConst ``ExprPresenter} but got {const.type}" + return ()) + +private unsafe def evalExprPresenterUnsafe (env : Environment) (opts : Options) + (constName : Name) : Except String ExprPresenter := + env.evalConstCheck ExprPresenter opts ``ExprPresenter constName + +@[implemented_by evalExprPresenterUnsafe] +opaque evalExprPresenter (env : Environment) (opts : Options) (constName : Name) : + Except String ExprPresenter + +structure ApplicableExprPresentersParams where + expr : WithRpcRef ExprWithCtx + +#mkrpcenc ApplicableExprPresentersParams + +structure ExprPresenterId where + name : Name + userName : String + deriving FromJson, ToJson + +structure ApplicableExprPresenters where + presenters : Array ExprPresenterId + deriving FromJson, ToJson + +@[server_rpc_method] +def applicableExprPresenters : ApplicableExprPresentersParams → + RequestM (RequestTask ApplicableExprPresenters) + | ⟨⟨expr⟩⟩ => RequestM.asTask do + let mut presenters : Array ExprPresenterId := #[] + let ci := expr.ci + for nm in exprPresenters.ext.getState expr.ci.env do + match evalExprPresenter ci.env ci.options nm with + | .ok p => + if ← expr.runMetaM p.isApplicable then + presenters := presenters.push ⟨nm, p.userName⟩ + | .error e => + throw <| RequestError.internalError s!"Failed to evaluate Expr presenter '{nm}': {e}" + return { presenters } + +structure GetExprPresentationParams extends ApplicableExprPresentersParams where + /-- Name of the presenter to use. -/ + name : Name + +#mkrpcenc GetExprPresentationParams + +@[server_rpc_method] +def getExprPresentation : GetExprPresentationParams → + RequestM (RequestTask EncodableHtml) + | { expr := ⟨expr⟩, name } => RequestM.asTask do + let ci := expr.ci + if !exprPresenters.hasTag ci.env name then + throw <| RequestError.invalidParams s!"The constant '{name}' is not an Expr presenter." + match evalExprPresenter ci.env ci.options name with + | .ok p => + let some ret ← expr.runMetaM p.present + | throw <| RequestError.internalError <| + s!"Got none from {name}.present e, expected some _ because {name}.isApplicable e " ++ + s!"returned true, where e := {expr.expr}" + return ret + | .error e => + throw <| RequestError.internalError s!"Failed to evaluate Expr presenter '{name}': {e}" + +structure ExprPresentationProps where + expr : WithRpcRef ExprWithCtx + +#mkrpcenc ExprPresentationProps + +/-- This component shows a selection of all known and applicable `WidgetKit.ExprPresenter`s which +are used to render the expression when selected. By default `WidgetKit.InteractiveExpr` is shown. -/ +@[widget_module] +def ExprPresentation : Component ExprPresentationProps where + javascript := include_str ".." / ".." / "build" / "js" / "exprPresentation.js" + +end WidgetKit diff --git a/WidgetKit/Presentation/Goal.lean b/WidgetKit/Presentation/Goal.lean new file mode 100644 index 0000000..4406261 --- /dev/null +++ b/WidgetKit/Presentation/Goal.lean @@ -0,0 +1,66 @@ +import Lean.Elab.Tactic +import Lean.Meta.ExprLens + +import WidgetKit.Component.Basic +import WidgetKit.Presentation.Expr -- Needed for RPC calls in PresentSelectionPanel + +namespace WidgetKit +open Lean Server + +structure GoalsLocationsToExprsParams where + locations : Array (WithRpcRef Elab.ContextInfo × SubExpr.GoalsLocation) + +#mkrpcenc GoalsLocationsToExprsParams + +structure GoalsLocationsToExprsResponse where + exprs : Array (WithRpcRef ExprWithCtx) + +#mkrpcenc GoalsLocationsToExprsResponse + +/-- Compute expressions corresponding to the given `GoalsLocation`s. -/ +@[server_rpc_method] +def goalsLocationsToExprs (args : GoalsLocationsToExprsParams) : + RequestM (RequestTask GoalsLocationsToExprsResponse) := + RequestM.asTask do + let mut exprs := #[] + for ⟨⟨ci⟩, loc⟩ in args.locations do + exprs := exprs.push ⟨← ci.runMetaM {} <| go loc.mvarId loc.loc⟩ + return { exprs } +where + go (mvarId : MVarId) : SubExpr.GoalLocation → MetaM ExprWithCtx + | .hyp fv => + mvarId.withContext <| + ExprWithCtx.save (mkFVar fv) + | .hypType fv pos => mvarId.withContext do + let tp ← Meta.inferType (mkFVar fv) + Meta.viewSubexpr (visit := fun _ => ExprWithCtx.save) pos tp + | .hypValue fv pos => mvarId.withContext do + let some val ← fv.getValue? + | throwError "fvar {mkFVar fv} is not a let-binding" + Meta.viewSubexpr (visit := fun _ => ExprWithCtx.save) pos val + | .target pos => mvarId.withContext do + let tp ← Meta.inferType (mkMVar mvarId) + Meta.viewSubexpr (visit := fun _ => ExprWithCtx.save) pos tp + +@[widget_module] +def PresentSelectionPanel : Component PanelWidgetProps where + javascript := include_str ".." / ".." / "build" / "js" / "presentSelection.js" + +/-- Displays any expressions selected in the goal state using registered `WidgetKit.ExprPresenter`s. +Expressions can be selected using shift-click -/ +syntax (name := withSelectionDisplayTacStx) "withSelectionDisplay " tacticSeq : tactic + +open Elab Tactic in +@[tactic withSelectionDisplayTacStx] +def withSelectionDisplay : Tactic + | stx@`(tactic| withSelectionDisplay $seq) => do + savePanelWidgetInfo stx ``PresentSelectionPanel (pure .null) + evalTacticSeq seq + | _ => throwUnsupportedSyntax + +-- TODO: replace the panel with an extensible one which invokes an arbitrary +-- `@[goals_presenter]` where `GoalsPresenter ≈ PanelWidgetProps → MetaM EncodableHtml` +-- with access to the entire goal state in order to display it in a "global" way. +-- In short, implement custom tactic state displays. + +end WidgetKit diff --git a/lakefile.lean b/lakefile.lean index c93c655..5627bf1 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -1,18 +1,11 @@ import Lake open Lake DSL System -package widgetKit { - -- add package configuration options here +package widgetkit { + preferReleaseBuild := true } -lean_lib WidgetKit { - -- add library configuration options here -} - -@[default_target] -lean_exe widgetKit { - root := `Main -} +lean_lib WidgetKit {} /-! Widget build -/ @@ -33,7 +26,7 @@ target widgetPackageLock : FilePath := do def widgetTsxTarget (pkg : Package) (tsxName : String) (deps : Array (BuildJob FilePath)) [Fact (pkg.name = _package.name)] : IndexBuildM (BuildJob FilePath) := do - let jsFile := widgetDir / "dist" / s!"{tsxName}.js" + let jsFile := pkg.buildDir / "js" / s!"{tsxName}.js" let deps := deps ++ #[ ← inputFile <| widgetDir / "src" / s!"{tsxName}.tsx", ← inputFile <| widgetDir / "rollup.config.js", @@ -47,10 +40,21 @@ def widgetTsxTarget (pkg : Package) (tsxName : String) (deps : Array (BuildJob F cwd := some widgetDir } -target widgets (pkg : Package) : Array FilePath := do - let dependencies : Array FilePath := #[] -- ← add extra deps here - let dependencies ← dependencies.mapM fun f => inputFile (widgetDir / "src" / f) - let items := #["staticHtml", "interactiveSvg"] - let xs ← items.mapM (fun item => widgetTsxTarget pkg item dependencies) - BuildJob.collectArray xs +target widgetJsAll (pkg : Package) : Array FilePath := do + let fs ← (widgetDir / "src").readDir + let tsxs : Array FilePath := fs.filterMap fun f => + let p := f.path; if let some "tsx" := p.extension then some p else none + -- Conservatively, every .js build depends on all the .tsx source files. + let deps ← liftM <| tsxs.mapM inputFile + let jobs ← tsxs.mapM fun tsx => widgetTsxTarget pkg tsx.fileStem.get! deps + BuildJob.collectArray jobs +@[default_target] +target all (pkg : Package) : Unit := do + let some lib := pkg.findLeanLib? ``WidgetKit | + error "cannot find lean_lib target" + let job₁ ← fetch (pkg.target ``widgetJsAll) + let _ ← job₁.await + let job₂ ← lib.recBuildLean + let _ ← job₂.await + return .nil diff --git a/lean-toolchain b/lean-toolchain index 5364bcb..c76d115 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2022-11-15 \ No newline at end of file +leanprover/lean4:nightly-2023-01-29 diff --git a/widget/dist/staticHtml.js b/widget/dist/staticHtml.js deleted file mode 100644 index c90bcdb..0000000 --- a/widget/dist/staticHtml.js +++ /dev/null @@ -1,41405 +0,0 @@ -const global = window; - -import * as React from 'react'; -import React__default, { isValidElement, PureComponent, cloneElement, createRef, forwardRef, useState, useRef, useImperativeHandle, useEffect, Component, Children, createElement } from 'react'; -import require$$2, { findDOMNode } from 'react-dom'; -import { InteractiveCode } from '@leanprover/infoview'; - -var commonjsGlobal$1 = typeof globalThis !== 'undefined' ? globalThis : "object" !== 'undefined' ? window : typeof global !== 'undefined' ? global : typeof self !== 'undefined' ? self : {}; - -function getDefaultExportFromCjs (x) { - return x && x.__esModule && Object.prototype.hasOwnProperty.call(x, 'default') ? x['default'] : x; -} - -function getAugmentedNamespace(n) { - var f = n.default; - if (typeof f == "function") { - var a = function () { - return f.apply(this, arguments); - }; - a.prototype = f.prototype; - } else a = {}; - Object.defineProperty(a, '__esModule', {value: true}); - Object.keys(n).forEach(function (k) { - var d = Object.getOwnPropertyDescriptor(n, k); - Object.defineProperty(a, k, d.get ? d : { - enumerable: true, - get: function () { - return n[k]; - } - }); - }); - return a; -} - -var classnames = {exports: {}}; - -/*! - Copyright (c) 2018 Jed Watson. - Licensed under the MIT License (MIT), see - http://jedwatson.github.io/classnames -*/ - -(function (module) { - /* global define */ - - (function () { - - var hasOwn = {}.hasOwnProperty; - - function classNames() { - var classes = []; - - for (var i = 0; i < arguments.length; i++) { - var arg = arguments[i]; - if (!arg) continue; - - var argType = typeof arg; - - if (argType === 'string' || argType === 'number') { - classes.push(arg); - } else if (Array.isArray(arg)) { - if (arg.length) { - var inner = classNames.apply(null, arg); - if (inner) { - classes.push(inner); - } - } - } else if (argType === 'object') { - if (arg.toString !== Object.prototype.toString && !arg.toString.toString().includes('[native code]')) { - classes.push(arg.toString()); - continue; - } - - for (var key in arg) { - if (hasOwn.call(arg, key) && arg[key]) { - classes.push(key); - } - } - } - } - - return classes.join(' '); - } - - if (module.exports) { - classNames.default = classNames; - module.exports = classNames; - } else { - window.classNames = classNames; - } - }()); -} (classnames)); - -var classNames = classnames.exports; - -/** - * Checks if `value` is the - * [language type](http://www.ecma-international.org/ecma-262/7.0/#sec-ecmascript-language-types) - * of `Object`. (e.g. arrays, functions, objects, regexes, `new Number(0)`, and `new String('')`) - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is an object, else `false`. - * @example - * - * _.isObject({}); - * // => true - * - * _.isObject([1, 2, 3]); - * // => true - * - * _.isObject(_.noop); - * // => true - * - * _.isObject(null); - * // => false - */ - -function isObject$b(value) { - var type = typeof value; - return value != null && (type == 'object' || type == 'function'); -} - -var isObject_1$1 = isObject$b; - -function _typeof$z(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$z = function _typeof(obj) { return typeof obj; }; } else { _typeof$z = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$z(obj); } -var SVGContainerPropKeys = ['viewBox', 'children']; -var SVGElementPropKeys = ['aria-activedescendant', 'aria-atomic', 'aria-autocomplete', 'aria-busy', 'aria-checked', 'aria-colcount', 'aria-colindex', 'aria-colspan', 'aria-controls', 'aria-current', 'aria-describedby', 'aria-details', 'aria-disabled', 'aria-errormessage', 'aria-expanded', 'aria-flowto', 'aria-haspopup', 'aria-hidden', 'aria-invalid', 'aria-keyshortcuts', 'aria-label', 'aria-labelledby', 'aria-level', 'aria-live', 'aria-modal', 'aria-multiline', 'aria-multiselectable', 'aria-orientation', 'aria-owns', 'aria-placeholder', 'aria-posinset', 'aria-pressed', 'aria-readonly', 'aria-relevant', 'aria-required', 'aria-roledescription', 'aria-rowcount', 'aria-rowindex', 'aria-rowspan', 'aria-selected', 'aria-setsize', 'aria-sort', 'aria-valuemax', 'aria-valuemin', 'aria-valuenow', 'aria-valuetext', 'className', 'color', 'height', 'id', 'lang', 'max', 'media', 'method', 'min', 'name', 'style', 'target', 'type', 'width', 'role', 'tabIndex', 'accentHeight', 'accumulate', 'additive', 'alignmentBaseline', 'allowReorder', 'alphabetic', 'amplitude', 'arabicForm', 'ascent', 'attributeName', 'attributeType', 'autoReverse', 'azimuth', 'baseFrequency', 'baselineShift', 'baseProfile', 'bbox', 'begin', 'bias', 'by', 'calcMode', 'capHeight', 'clip', 'clipPath', 'clipPathUnits', 'clipRule', 'colorInterpolation', 'colorInterpolationFilters', 'colorProfile', 'colorRendering', 'contentScriptType', 'contentStyleType', 'cursor', 'cx', 'cy', 'd', 'decelerate', 'descent', 'diffuseConstant', 'direction', 'display', 'divisor', 'dominantBaseline', 'dur', 'dx', 'dy', 'edgeMode', 'elevation', 'enableBackground', 'end', 'exponent', 'externalResourcesRequired', 'fill', 'fillOpacity', 'fillRule', 'filter', 'filterRes', 'filterUnits', 'floodColor', 'floodOpacity', 'focusable', 'fontFamily', 'fontSize', 'fontSizeAdjust', 'fontStretch', 'fontStyle', 'fontVariant', 'fontWeight', 'format', 'from', 'fx', 'fy', 'g1', 'g2', 'glyphName', 'glyphOrientationHorizontal', 'glyphOrientationVertical', 'glyphRef', 'gradientTransform', 'gradientUnits', 'hanging', 'horizAdvX', 'horizOriginX', 'href', 'ideographic', 'imageRendering', 'in2', 'in', 'intercept', 'k1', 'k2', 'k3', 'k4', 'k', 'kernelMatrix', 'kernelUnitLength', 'kerning', 'keyPoints', 'keySplines', 'keyTimes', 'lengthAdjust', 'letterSpacing', 'lightingColor', 'limitingConeAngle', 'local', 'markerEnd', 'markerHeight', 'markerMid', 'markerStart', 'markerUnits', 'markerWidth', 'mask', 'maskContentUnits', 'maskUnits', 'mathematical', 'mode', 'numOctaves', 'offset', 'opacity', 'operator', 'order', 'orient', 'orientation', 'origin', 'overflow', 'overlinePosition', 'overlineThickness', 'paintOrder', 'panose1', 'pathLength', 'patternContentUnits', 'patternTransform', 'patternUnits', 'pointerEvents', 'points', 'pointsAtX', 'pointsAtY', 'pointsAtZ', 'preserveAlpha', 'preserveAspectRatio', 'primitiveUnits', 'r', 'radius', 'refX', 'refY', 'renderingIntent', 'repeatCount', 'repeatDur', 'requiredExtensions', 'requiredFeatures', 'restart', 'result', 'rotate', 'rx', 'ry', 'seed', 'shapeRendering', 'slope', 'spacing', 'specularConstant', 'specularExponent', 'speed', 'spreadMethod', 'startOffset', 'stdDeviation', 'stemh', 'stemv', 'stitchTiles', 'stopColor', 'stopOpacity', 'strikethroughPosition', 'strikethroughThickness', 'string', 'stroke', 'strokeDasharray', 'strokeDashoffset', 'strokeLinecap', 'strokeLinejoin', 'strokeMiterlimit', 'strokeOpacity', 'strokeWidth', 'surfaceScale', 'systemLanguage', 'tableValues', 'targetX', 'targetY', 'textAnchor', 'textDecoration', 'textLength', 'textRendering', 'to', 'transform', 'u1', 'u2', 'underlinePosition', 'underlineThickness', 'unicode', 'unicodeBidi', 'unicodeRange', 'unitsPerEm', 'vAlphabetic', 'values', 'vectorEffect', 'version', 'vertAdvY', 'vertOriginX', 'vertOriginY', 'vHanging', 'vIdeographic', 'viewTarget', 'visibility', 'vMathematical', 'widths', 'wordSpacing', 'writingMode', 'x1', 'x2', 'x', 'xChannelSelector', 'xHeight', 'xlinkActuate', 'xlinkArcrole', 'xlinkHref', 'xlinkRole', 'xlinkShow', 'xlinkTitle', 'xlinkType', 'xmlBase', 'xmlLang', 'xmlns', 'xmlnsXlink', 'xmlSpace', 'y1', 'y2', 'y', 'yChannelSelector', 'z', 'zoomAndPan', 'ref', 'key', 'angle']; -var EventKeys = ['dangerouslySetInnerHTML', 'onCopy', 'onCopyCapture', 'onCut', 'onCutCapture', 'onPaste', 'onPasteCapture', 'onCompositionEnd', 'onCompositionEndCapture', 'onCompositionStart', 'onCompositionStartCapture', 'onCompositionUpdate', 'onCompositionUpdateCapture', 'onFocus', 'onFocusCapture', 'onBlur', 'onBlurCapture', 'onChange', 'onChangeCapture', 'onBeforeInput', 'onBeforeInputCapture', 'onInput', 'onInputCapture', 'onReset', 'onResetCapture', 'onSubmit', 'onSubmitCapture', 'onInvalid', 'onInvalidCapture', 'onLoad', 'onLoadCapture', 'onError', 'onErrorCapture', 'onKeyDown', 'onKeyDownCapture', 'onKeyPress', 'onKeyPressCapture', 'onKeyUp', 'onKeyUpCapture', 'onAbort', 'onAbortCapture', 'onCanPlay', 'onCanPlayCapture', 'onCanPlayThrough', 'onCanPlayThroughCapture', 'onDurationChange', 'onDurationChangeCapture', 'onEmptied', 'onEmptiedCapture', 'onEncrypted', 'onEncryptedCapture', 'onEnded', 'onEndedCapture', 'onLoadedData', 'onLoadedDataCapture', 'onLoadedMetadata', 'onLoadedMetadataCapture', 'onLoadStart', 'onLoadStartCapture', 'onPause', 'onPauseCapture', 'onPlay', 'onPlayCapture', 'onPlaying', 'onPlayingCapture', 'onProgress', 'onProgressCapture', 'onRateChange', 'onRateChangeCapture', 'onSeeked', 'onSeekedCapture', 'onSeeking', 'onSeekingCapture', 'onStalled', 'onStalledCapture', 'onSuspend', 'onSuspendCapture', 'onTimeUpdate', 'onTimeUpdateCapture', 'onVolumeChange', 'onVolumeChangeCapture', 'onWaiting', 'onWaitingCapture', 'onAuxClick', 'onAuxClickCapture', 'onClick', 'onClickCapture', 'onContextMenu', 'onContextMenuCapture', 'onDoubleClick', 'onDoubleClickCapture', 'onDrag', 'onDragCapture', 'onDragEnd', 'onDragEndCapture', 'onDragEnter', 'onDragEnterCapture', 'onDragExit', 'onDragExitCapture', 'onDragLeave', 'onDragLeaveCapture', 'onDragOver', 'onDragOverCapture', 'onDragStart', 'onDragStartCapture', 'onDrop', 'onDropCapture', 'onMouseDown', 'onMouseDownCapture', 'onMouseEnter', 'onMouseLeave', 'onMouseMove', 'onMouseMoveCapture', 'onMouseOut', 'onMouseOutCapture', 'onMouseOver', 'onMouseOverCapture', 'onMouseUp', 'onMouseUpCapture', 'onSelect', 'onSelectCapture', 'onTouchCancel', 'onTouchCancelCapture', 'onTouchEnd', 'onTouchEndCapture', 'onTouchMove', 'onTouchMoveCapture', 'onTouchStart', 'onTouchStartCapture', 'onPointerDown', 'onPointerDownCapture', 'onPointerMove', 'onPointerMoveCapture', 'onPointerUp', 'onPointerUpCapture', 'onPointerCancel', 'onPointerCancelCapture', 'onPointerEnter', 'onPointerEnterCapture', 'onPointerLeave', 'onPointerLeaveCapture', 'onPointerOver', 'onPointerOverCapture', 'onPointerOut', 'onPointerOutCapture', 'onGotPointerCapture', 'onGotPointerCaptureCapture', 'onLostPointerCapture', 'onLostPointerCaptureCapture', 'onScroll', 'onScrollCapture', 'onWheel', 'onWheelCapture', 'onAnimationStart', 'onAnimationStartCapture', 'onAnimationEnd', 'onAnimationEndCapture', 'onAnimationIteration', 'onAnimationIterationCapture', 'onTransitionEnd', 'onTransitionEndCapture']; // Animation Types => TODO: Should be moved when react-smooth is typescriptified. - -var filterProps = function filterProps(props, includeEvents, isSvg) { - if (!props || typeof props === 'function' || typeof props === 'boolean') { - return null; - } - - var inputProps = props; - - if ( /*#__PURE__*/isValidElement(props)) { - inputProps = props.props; - } - - if (!isObject_1$1(inputProps)) { - return null; - } - - var out = {}; - Object.keys(inputProps).forEach(function (key) { - // viewBox only exist in - if (SVGElementPropKeys.includes(key) || isSvg && SVGContainerPropKeys.includes(key) || includeEvents && EventKeys.includes(key)) { - out[key] = inputProps[key]; - } - }); - return out; -}; -var adaptEventHandlers = function adaptEventHandlers(props, newHandler) { - if (!props || typeof props === 'function' || typeof props === 'boolean') { - return null; - } - - var inputProps = props; - - if ( /*#__PURE__*/isValidElement(props)) { - inputProps = props.props; - } - - if (!isObject_1$1(inputProps)) { - return null; - } - - var out = {}; - Object.keys(inputProps).forEach(function (key) { - if (EventKeys.includes(key)) { - out[key] = newHandler || function (e) { - return inputProps[key](inputProps, e); - }; - } - }); - return out; -}; - -var getEventHandlerOfChild = function getEventHandlerOfChild(originalHandler, data, index) { - return function (e) { - originalHandler(data, index, e); - return null; - }; -}; - -var adaptEventsOfChild = function adaptEventsOfChild(props, data, index) { - if (!isObject_1$1(props) || _typeof$z(props) !== 'object') { - return null; - } - - var out = null; - Object.keys(props).forEach(function (key) { - var item = props[key]; - - if (EventKeys.includes(key) && typeof item === 'function') { - if (!out) out = {}; - out[key] = getEventHandlerOfChild(item, data, index); - } - }); - return out; -}; - -function _extends$z() { _extends$z = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$z.apply(this, arguments); } - -function _objectWithoutProperties$k(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$l(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$l(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } -function Surface(props) { - var children = props.children, - width = props.width, - height = props.height, - viewBox = props.viewBox, - className = props.className, - style = props.style, - others = _objectWithoutProperties$k(props, ["children", "width", "height", "viewBox", "className", "style"]); - - var svgView = viewBox || { - width: width, - height: height, - x: 0, - y: 0 - }; - var layerClass = classNames('recharts-surface', className); - return /*#__PURE__*/React__default.createElement("svg", _extends$z({}, filterProps(others, true, true), { - className: layerClass, - width: width, - height: height, - style: style, - viewBox: "".concat(svgView.x, " ").concat(svgView.y, " ").concat(svgView.width, " ").concat(svgView.height), - version: "1.1" - }), /*#__PURE__*/React__default.createElement("title", null, props.title), /*#__PURE__*/React__default.createElement("desc", null, props.desc), children); -} - -function _extends$y() { _extends$y = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$y.apply(this, arguments); } - -function _objectWithoutProperties$j(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$k(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$k(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } -var Layer = /*#__PURE__*/React__default.forwardRef(function (props, ref) { - var children = props.children, - className = props.className, - others = _objectWithoutProperties$j(props, ["children", "className"]); - - var layerClass = classNames('recharts-layer', className); - return /*#__PURE__*/React__default.createElement("g", _extends$y({ - className: layerClass - }, filterProps(others, true), { - ref: ref - }), children); -}); - -/** Detect free variable `global` from Node.js. */ - -var freeGlobal$3 = typeof commonjsGlobal$1 == 'object' && commonjsGlobal$1 && commonjsGlobal$1.Object === Object && commonjsGlobal$1; - -var _freeGlobal$1 = freeGlobal$3; - -var freeGlobal$2 = _freeGlobal$1; - -/** Detect free variable `self`. */ -var freeSelf$1 = typeof self == 'object' && self && self.Object === Object && self; - -/** Used as a reference to the global object. */ -var root$c = freeGlobal$2 || freeSelf$1 || Function('return this')(); - -var _root$1 = root$c; - -var root$b = _root$1; - -/** Built-in value references. */ -var Symbol$9 = root$b.Symbol; - -var _Symbol$1 = Symbol$9; - -var Symbol$8 = _Symbol$1; - -/** Used for built-in method references. */ -var objectProto$g = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$d = objectProto$g.hasOwnProperty; - -/** - * Used to resolve the - * [`toStringTag`](http://ecma-international.org/ecma-262/7.0/#sec-object.prototype.tostring) - * of values. - */ -var nativeObjectToString$3 = objectProto$g.toString; - -/** Built-in value references. */ -var symToStringTag$3 = Symbol$8 ? Symbol$8.toStringTag : undefined; - -/** - * A specialized version of `baseGetTag` which ignores `Symbol.toStringTag` values. - * - * @private - * @param {*} value The value to query. - * @returns {string} Returns the raw `toStringTag`. - */ -function getRawTag$3(value) { - var isOwn = hasOwnProperty$d.call(value, symToStringTag$3), - tag = value[symToStringTag$3]; - - try { - value[symToStringTag$3] = undefined; - var unmasked = true; - } catch (e) {} - - var result = nativeObjectToString$3.call(value); - if (unmasked) { - if (isOwn) { - value[symToStringTag$3] = tag; - } else { - delete value[symToStringTag$3]; - } - } - return result; -} - -var _getRawTag$1 = getRawTag$3; - -/** Used for built-in method references. */ - -var objectProto$f = Object.prototype; - -/** - * Used to resolve the - * [`toStringTag`](http://ecma-international.org/ecma-262/7.0/#sec-object.prototype.tostring) - * of values. - */ -var nativeObjectToString$2 = objectProto$f.toString; - -/** - * Converts `value` to a string using `Object.prototype.toString`. - * - * @private - * @param {*} value The value to convert. - * @returns {string} Returns the converted string. - */ -function objectToString$3(value) { - return nativeObjectToString$2.call(value); -} - -var _objectToString$1 = objectToString$3; - -var Symbol$7 = _Symbol$1, - getRawTag$2 = _getRawTag$1, - objectToString$2 = _objectToString$1; - -/** `Object#toString` result references. */ -var nullTag$1 = '[object Null]', - undefinedTag$1 = '[object Undefined]'; - -/** Built-in value references. */ -var symToStringTag$2 = Symbol$7 ? Symbol$7.toStringTag : undefined; - -/** - * The base implementation of `getTag` without fallbacks for buggy environments. - * - * @private - * @param {*} value The value to query. - * @returns {string} Returns the `toStringTag`. - */ -function baseGetTag$a(value) { - if (value == null) { - return value === undefined ? undefinedTag$1 : nullTag$1; - } - return (symToStringTag$2 && symToStringTag$2 in Object(value)) - ? getRawTag$2(value) - : objectToString$2(value); -} - -var _baseGetTag$1 = baseGetTag$a; - -var baseGetTag$9 = _baseGetTag$1, - isObject$a = isObject_1$1; - -/** `Object#toString` result references. */ -var asyncTag = '[object AsyncFunction]', - funcTag$2 = '[object Function]', - genTag$1 = '[object GeneratorFunction]', - proxyTag = '[object Proxy]'; - -/** - * Checks if `value` is classified as a `Function` object. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a function, else `false`. - * @example - * - * _.isFunction(_); - * // => true - * - * _.isFunction(/abc/); - * // => false - */ -function isFunction$3(value) { - if (!isObject$a(value)) { - return false; - } - // The use of `Object#toString` avoids issues with the `typeof` operator - // in Safari 9 which returns 'object' for typed arrays and other constructors. - var tag = baseGetTag$9(value); - return tag == funcTag$2 || tag == genTag$1 || tag == asyncTag || tag == proxyTag; -} - -var isFunction_1 = isFunction$3; - -/** - * Removes all key-value entries from the list cache. - * - * @private - * @name clear - * @memberOf ListCache - */ - -function listCacheClear$1() { - this.__data__ = []; - this.size = 0; -} - -var _listCacheClear = listCacheClear$1; - -/** - * Performs a - * [`SameValueZero`](http://ecma-international.org/ecma-262/7.0/#sec-samevaluezero) - * comparison between two values to determine if they are equivalent. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to compare. - * @param {*} other The other value to compare. - * @returns {boolean} Returns `true` if the values are equivalent, else `false`. - * @example - * - * var object = { 'a': 1 }; - * var other = { 'a': 1 }; - * - * _.eq(object, object); - * // => true - * - * _.eq(object, other); - * // => false - * - * _.eq('a', 'a'); - * // => true - * - * _.eq('a', Object('a')); - * // => false - * - * _.eq(NaN, NaN); - * // => true - */ - -function eq$4(value, other) { - return value === other || (value !== value && other !== other); -} - -var eq_1 = eq$4; - -var eq$3 = eq_1; - -/** - * Gets the index at which the `key` is found in `array` of key-value pairs. - * - * @private - * @param {Array} array The array to inspect. - * @param {*} key The key to search for. - * @returns {number} Returns the index of the matched value, else `-1`. - */ -function assocIndexOf$4(array, key) { - var length = array.length; - while (length--) { - if (eq$3(array[length][0], key)) { - return length; - } - } - return -1; -} - -var _assocIndexOf = assocIndexOf$4; - -var assocIndexOf$3 = _assocIndexOf; - -/** Used for built-in method references. */ -var arrayProto = Array.prototype; - -/** Built-in value references. */ -var splice = arrayProto.splice; - -/** - * Removes `key` and its value from the list cache. - * - * @private - * @name delete - * @memberOf ListCache - * @param {string} key The key of the value to remove. - * @returns {boolean} Returns `true` if the entry was removed, else `false`. - */ -function listCacheDelete$1(key) { - var data = this.__data__, - index = assocIndexOf$3(data, key); - - if (index < 0) { - return false; - } - var lastIndex = data.length - 1; - if (index == lastIndex) { - data.pop(); - } else { - splice.call(data, index, 1); - } - --this.size; - return true; -} - -var _listCacheDelete = listCacheDelete$1; - -var assocIndexOf$2 = _assocIndexOf; - -/** - * Gets the list cache value for `key`. - * - * @private - * @name get - * @memberOf ListCache - * @param {string} key The key of the value to get. - * @returns {*} Returns the entry value. - */ -function listCacheGet$1(key) { - var data = this.__data__, - index = assocIndexOf$2(data, key); - - return index < 0 ? undefined : data[index][1]; -} - -var _listCacheGet = listCacheGet$1; - -var assocIndexOf$1 = _assocIndexOf; - -/** - * Checks if a list cache value for `key` exists. - * - * @private - * @name has - * @memberOf ListCache - * @param {string} key The key of the entry to check. - * @returns {boolean} Returns `true` if an entry for `key` exists, else `false`. - */ -function listCacheHas$1(key) { - return assocIndexOf$1(this.__data__, key) > -1; -} - -var _listCacheHas = listCacheHas$1; - -var assocIndexOf = _assocIndexOf; - -/** - * Sets the list cache `key` to `value`. - * - * @private - * @name set - * @memberOf ListCache - * @param {string} key The key of the value to set. - * @param {*} value The value to set. - * @returns {Object} Returns the list cache instance. - */ -function listCacheSet$1(key, value) { - var data = this.__data__, - index = assocIndexOf(data, key); - - if (index < 0) { - ++this.size; - data.push([key, value]); - } else { - data[index][1] = value; - } - return this; -} - -var _listCacheSet = listCacheSet$1; - -var listCacheClear = _listCacheClear, - listCacheDelete = _listCacheDelete, - listCacheGet = _listCacheGet, - listCacheHas = _listCacheHas, - listCacheSet = _listCacheSet; - -/** - * Creates an list cache object. - * - * @private - * @constructor - * @param {Array} [entries] The key-value pairs to cache. - */ -function ListCache$4(entries) { - var index = -1, - length = entries == null ? 0 : entries.length; - - this.clear(); - while (++index < length) { - var entry = entries[index]; - this.set(entry[0], entry[1]); - } -} - -// Add methods to `ListCache`. -ListCache$4.prototype.clear = listCacheClear; -ListCache$4.prototype['delete'] = listCacheDelete; -ListCache$4.prototype.get = listCacheGet; -ListCache$4.prototype.has = listCacheHas; -ListCache$4.prototype.set = listCacheSet; - -var _ListCache = ListCache$4; - -var ListCache$3 = _ListCache; - -/** - * Removes all key-value entries from the stack. - * - * @private - * @name clear - * @memberOf Stack - */ -function stackClear$1() { - this.__data__ = new ListCache$3; - this.size = 0; -} - -var _stackClear = stackClear$1; - -/** - * Removes `key` and its value from the stack. - * - * @private - * @name delete - * @memberOf Stack - * @param {string} key The key of the value to remove. - * @returns {boolean} Returns `true` if the entry was removed, else `false`. - */ - -function stackDelete$1(key) { - var data = this.__data__, - result = data['delete'](key); - - this.size = data.size; - return result; -} - -var _stackDelete = stackDelete$1; - -/** - * Gets the stack value for `key`. - * - * @private - * @name get - * @memberOf Stack - * @param {string} key The key of the value to get. - * @returns {*} Returns the entry value. - */ - -function stackGet$1(key) { - return this.__data__.get(key); -} - -var _stackGet = stackGet$1; - -/** - * Checks if a stack value for `key` exists. - * - * @private - * @name has - * @memberOf Stack - * @param {string} key The key of the entry to check. - * @returns {boolean} Returns `true` if an entry for `key` exists, else `false`. - */ - -function stackHas$1(key) { - return this.__data__.has(key); -} - -var _stackHas = stackHas$1; - -var root$a = _root$1; - -/** Used to detect overreaching core-js shims. */ -var coreJsData$1 = root$a['__core-js_shared__']; - -var _coreJsData = coreJsData$1; - -var coreJsData = _coreJsData; - -/** Used to detect methods masquerading as native. */ -var maskSrcKey = (function() { - var uid = /[^.]+$/.exec(coreJsData && coreJsData.keys && coreJsData.keys.IE_PROTO || ''); - return uid ? ('Symbol(src)_1.' + uid) : ''; -}()); - -/** - * Checks if `func` has its source masked. - * - * @private - * @param {Function} func The function to check. - * @returns {boolean} Returns `true` if `func` is masked, else `false`. - */ -function isMasked$1(func) { - return !!maskSrcKey && (maskSrcKey in func); -} - -var _isMasked = isMasked$1; - -/** Used for built-in method references. */ - -var funcProto$2 = Function.prototype; - -/** Used to resolve the decompiled source of functions. */ -var funcToString$2 = funcProto$2.toString; - -/** - * Converts `func` to its source code. - * - * @private - * @param {Function} func The function to convert. - * @returns {string} Returns the source code. - */ -function toSource$2(func) { - if (func != null) { - try { - return funcToString$2.call(func); - } catch (e) {} - try { - return (func + ''); - } catch (e) {} - } - return ''; -} - -var _toSource = toSource$2; - -var isFunction$2 = isFunction_1, - isMasked = _isMasked, - isObject$9 = isObject_1$1, - toSource$1 = _toSource; - -/** - * Used to match `RegExp` - * [syntax characters](http://ecma-international.org/ecma-262/7.0/#sec-patterns). - */ -var reRegExpChar = /[\\^$.*+?()[\]{}|]/g; - -/** Used to detect host constructors (Safari). */ -var reIsHostCtor = /^\[object .+?Constructor\]$/; - -/** Used for built-in method references. */ -var funcProto$1 = Function.prototype, - objectProto$e = Object.prototype; - -/** Used to resolve the decompiled source of functions. */ -var funcToString$1 = funcProto$1.toString; - -/** Used to check objects for own properties. */ -var hasOwnProperty$c = objectProto$e.hasOwnProperty; - -/** Used to detect if a method is native. */ -var reIsNative = RegExp('^' + - funcToString$1.call(hasOwnProperty$c).replace(reRegExpChar, '\\$&') - .replace(/hasOwnProperty|(function).*?(?=\\\()| for .+?(?=\\\])/g, '$1.*?') + '$' -); - -/** - * The base implementation of `_.isNative` without bad shim checks. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a native function, - * else `false`. - */ -function baseIsNative$1(value) { - if (!isObject$9(value) || isMasked(value)) { - return false; - } - var pattern = isFunction$2(value) ? reIsNative : reIsHostCtor; - return pattern.test(toSource$1(value)); -} - -var _baseIsNative = baseIsNative$1; - -/** - * Gets the value at `key` of `object`. - * - * @private - * @param {Object} [object] The object to query. - * @param {string} key The key of the property to get. - * @returns {*} Returns the property value. - */ - -function getValue$2(object, key) { - return object == null ? undefined : object[key]; -} - -var _getValue = getValue$2; - -var baseIsNative = _baseIsNative, - getValue$1 = _getValue; - -/** - * Gets the native function at `key` of `object`. - * - * @private - * @param {Object} object The object to query. - * @param {string} key The key of the method to get. - * @returns {*} Returns the function if it's native, else `undefined`. - */ -function getNative$7(object, key) { - var value = getValue$1(object, key); - return baseIsNative(value) ? value : undefined; -} - -var _getNative = getNative$7; - -var getNative$6 = _getNative, - root$9 = _root$1; - -/* Built-in method references that are verified to be native. */ -var Map$4 = getNative$6(root$9, 'Map'); - -var _Map = Map$4; - -var getNative$5 = _getNative; - -/* Built-in method references that are verified to be native. */ -var nativeCreate$4 = getNative$5(Object, 'create'); - -var _nativeCreate = nativeCreate$4; - -var nativeCreate$3 = _nativeCreate; - -/** - * Removes all key-value entries from the hash. - * - * @private - * @name clear - * @memberOf Hash - */ -function hashClear$1() { - this.__data__ = nativeCreate$3 ? nativeCreate$3(null) : {}; - this.size = 0; -} - -var _hashClear = hashClear$1; - -/** - * Removes `key` and its value from the hash. - * - * @private - * @name delete - * @memberOf Hash - * @param {Object} hash The hash to modify. - * @param {string} key The key of the value to remove. - * @returns {boolean} Returns `true` if the entry was removed, else `false`. - */ - -function hashDelete$1(key) { - var result = this.has(key) && delete this.__data__[key]; - this.size -= result ? 1 : 0; - return result; -} - -var _hashDelete = hashDelete$1; - -var nativeCreate$2 = _nativeCreate; - -/** Used to stand-in for `undefined` hash values. */ -var HASH_UNDEFINED$2 = '__lodash_hash_undefined__'; - -/** Used for built-in method references. */ -var objectProto$d = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$b = objectProto$d.hasOwnProperty; - -/** - * Gets the hash value for `key`. - * - * @private - * @name get - * @memberOf Hash - * @param {string} key The key of the value to get. - * @returns {*} Returns the entry value. - */ -function hashGet$1(key) { - var data = this.__data__; - if (nativeCreate$2) { - var result = data[key]; - return result === HASH_UNDEFINED$2 ? undefined : result; - } - return hasOwnProperty$b.call(data, key) ? data[key] : undefined; -} - -var _hashGet = hashGet$1; - -var nativeCreate$1 = _nativeCreate; - -/** Used for built-in method references. */ -var objectProto$c = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$a = objectProto$c.hasOwnProperty; - -/** - * Checks if a hash value for `key` exists. - * - * @private - * @name has - * @memberOf Hash - * @param {string} key The key of the entry to check. - * @returns {boolean} Returns `true` if an entry for `key` exists, else `false`. - */ -function hashHas$1(key) { - var data = this.__data__; - return nativeCreate$1 ? (data[key] !== undefined) : hasOwnProperty$a.call(data, key); -} - -var _hashHas = hashHas$1; - -var nativeCreate = _nativeCreate; - -/** Used to stand-in for `undefined` hash values. */ -var HASH_UNDEFINED$1 = '__lodash_hash_undefined__'; - -/** - * Sets the hash `key` to `value`. - * - * @private - * @name set - * @memberOf Hash - * @param {string} key The key of the value to set. - * @param {*} value The value to set. - * @returns {Object} Returns the hash instance. - */ -function hashSet$1(key, value) { - var data = this.__data__; - this.size += this.has(key) ? 0 : 1; - data[key] = (nativeCreate && value === undefined) ? HASH_UNDEFINED$1 : value; - return this; -} - -var _hashSet = hashSet$1; - -var hashClear = _hashClear, - hashDelete = _hashDelete, - hashGet = _hashGet, - hashHas = _hashHas, - hashSet = _hashSet; - -/** - * Creates a hash object. - * - * @private - * @constructor - * @param {Array} [entries] The key-value pairs to cache. - */ -function Hash$1(entries) { - var index = -1, - length = entries == null ? 0 : entries.length; - - this.clear(); - while (++index < length) { - var entry = entries[index]; - this.set(entry[0], entry[1]); - } -} - -// Add methods to `Hash`. -Hash$1.prototype.clear = hashClear; -Hash$1.prototype['delete'] = hashDelete; -Hash$1.prototype.get = hashGet; -Hash$1.prototype.has = hashHas; -Hash$1.prototype.set = hashSet; - -var _Hash = Hash$1; - -var Hash = _Hash, - ListCache$2 = _ListCache, - Map$3 = _Map; - -/** - * Removes all key-value entries from the map. - * - * @private - * @name clear - * @memberOf MapCache - */ -function mapCacheClear$1() { - this.size = 0; - this.__data__ = { - 'hash': new Hash, - 'map': new (Map$3 || ListCache$2), - 'string': new Hash - }; -} - -var _mapCacheClear = mapCacheClear$1; - -/** - * Checks if `value` is suitable for use as unique object key. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is suitable, else `false`. - */ - -function isKeyable$1(value) { - var type = typeof value; - return (type == 'string' || type == 'number' || type == 'symbol' || type == 'boolean') - ? (value !== '__proto__') - : (value === null); -} - -var _isKeyable = isKeyable$1; - -var isKeyable = _isKeyable; - -/** - * Gets the data for `map`. - * - * @private - * @param {Object} map The map to query. - * @param {string} key The reference key. - * @returns {*} Returns the map data. - */ -function getMapData$4(map, key) { - var data = map.__data__; - return isKeyable(key) - ? data[typeof key == 'string' ? 'string' : 'hash'] - : data.map; -} - -var _getMapData = getMapData$4; - -var getMapData$3 = _getMapData; - -/** - * Removes `key` and its value from the map. - * - * @private - * @name delete - * @memberOf MapCache - * @param {string} key The key of the value to remove. - * @returns {boolean} Returns `true` if the entry was removed, else `false`. - */ -function mapCacheDelete$1(key) { - var result = getMapData$3(this, key)['delete'](key); - this.size -= result ? 1 : 0; - return result; -} - -var _mapCacheDelete = mapCacheDelete$1; - -var getMapData$2 = _getMapData; - -/** - * Gets the map value for `key`. - * - * @private - * @name get - * @memberOf MapCache - * @param {string} key The key of the value to get. - * @returns {*} Returns the entry value. - */ -function mapCacheGet$1(key) { - return getMapData$2(this, key).get(key); -} - -var _mapCacheGet = mapCacheGet$1; - -var getMapData$1 = _getMapData; - -/** - * Checks if a map value for `key` exists. - * - * @private - * @name has - * @memberOf MapCache - * @param {string} key The key of the entry to check. - * @returns {boolean} Returns `true` if an entry for `key` exists, else `false`. - */ -function mapCacheHas$1(key) { - return getMapData$1(this, key).has(key); -} - -var _mapCacheHas = mapCacheHas$1; - -var getMapData = _getMapData; - -/** - * Sets the map `key` to `value`. - * - * @private - * @name set - * @memberOf MapCache - * @param {string} key The key of the value to set. - * @param {*} value The value to set. - * @returns {Object} Returns the map cache instance. - */ -function mapCacheSet$1(key, value) { - var data = getMapData(this, key), - size = data.size; - - data.set(key, value); - this.size += data.size == size ? 0 : 1; - return this; -} - -var _mapCacheSet = mapCacheSet$1; - -var mapCacheClear = _mapCacheClear, - mapCacheDelete = _mapCacheDelete, - mapCacheGet = _mapCacheGet, - mapCacheHas = _mapCacheHas, - mapCacheSet = _mapCacheSet; - -/** - * Creates a map cache object to store key-value pairs. - * - * @private - * @constructor - * @param {Array} [entries] The key-value pairs to cache. - */ -function MapCache$3(entries) { - var index = -1, - length = entries == null ? 0 : entries.length; - - this.clear(); - while (++index < length) { - var entry = entries[index]; - this.set(entry[0], entry[1]); - } -} - -// Add methods to `MapCache`. -MapCache$3.prototype.clear = mapCacheClear; -MapCache$3.prototype['delete'] = mapCacheDelete; -MapCache$3.prototype.get = mapCacheGet; -MapCache$3.prototype.has = mapCacheHas; -MapCache$3.prototype.set = mapCacheSet; - -var _MapCache = MapCache$3; - -var ListCache$1 = _ListCache, - Map$2 = _Map, - MapCache$2 = _MapCache; - -/** Used as the size to enable large array optimizations. */ -var LARGE_ARRAY_SIZE$1 = 200; - -/** - * Sets the stack `key` to `value`. - * - * @private - * @name set - * @memberOf Stack - * @param {string} key The key of the value to set. - * @param {*} value The value to set. - * @returns {Object} Returns the stack cache instance. - */ -function stackSet$1(key, value) { - var data = this.__data__; - if (data instanceof ListCache$1) { - var pairs = data.__data__; - if (!Map$2 || (pairs.length < LARGE_ARRAY_SIZE$1 - 1)) { - pairs.push([key, value]); - this.size = ++data.size; - return this; - } - data = this.__data__ = new MapCache$2(pairs); - } - data.set(key, value); - this.size = data.size; - return this; -} - -var _stackSet = stackSet$1; - -var ListCache = _ListCache, - stackClear = _stackClear, - stackDelete = _stackDelete, - stackGet = _stackGet, - stackHas = _stackHas, - stackSet = _stackSet; - -/** - * Creates a stack cache object to store key-value pairs. - * - * @private - * @constructor - * @param {Array} [entries] The key-value pairs to cache. - */ -function Stack$3(entries) { - var data = this.__data__ = new ListCache(entries); - this.size = data.size; -} - -// Add methods to `Stack`. -Stack$3.prototype.clear = stackClear; -Stack$3.prototype['delete'] = stackDelete; -Stack$3.prototype.get = stackGet; -Stack$3.prototype.has = stackHas; -Stack$3.prototype.set = stackSet; - -var _Stack = Stack$3; - -/** Used to stand-in for `undefined` hash values. */ - -var HASH_UNDEFINED = '__lodash_hash_undefined__'; - -/** - * Adds `value` to the array cache. - * - * @private - * @name add - * @memberOf SetCache - * @alias push - * @param {*} value The value to cache. - * @returns {Object} Returns the cache instance. - */ -function setCacheAdd$1(value) { - this.__data__.set(value, HASH_UNDEFINED); - return this; -} - -var _setCacheAdd = setCacheAdd$1; - -/** - * Checks if `value` is in the array cache. - * - * @private - * @name has - * @memberOf SetCache - * @param {*} value The value to search for. - * @returns {number} Returns `true` if `value` is found, else `false`. - */ - -function setCacheHas$1(value) { - return this.__data__.has(value); -} - -var _setCacheHas = setCacheHas$1; - -var MapCache$1 = _MapCache, - setCacheAdd = _setCacheAdd, - setCacheHas = _setCacheHas; - -/** - * - * Creates an array cache object to store unique values. - * - * @private - * @constructor - * @param {Array} [values] The values to cache. - */ -function SetCache$2(values) { - var index = -1, - length = values == null ? 0 : values.length; - - this.__data__ = new MapCache$1; - while (++index < length) { - this.add(values[index]); - } -} - -// Add methods to `SetCache`. -SetCache$2.prototype.add = SetCache$2.prototype.push = setCacheAdd; -SetCache$2.prototype.has = setCacheHas; - -var _SetCache = SetCache$2; - -/** - * A specialized version of `_.some` for arrays without support for iteratee - * shorthands. - * - * @private - * @param {Array} [array] The array to iterate over. - * @param {Function} predicate The function invoked per iteration. - * @returns {boolean} Returns `true` if any element passes the predicate check, - * else `false`. - */ - -function arraySome$2(array, predicate) { - var index = -1, - length = array == null ? 0 : array.length; - - while (++index < length) { - if (predicate(array[index], index, array)) { - return true; - } - } - return false; -} - -var _arraySome = arraySome$2; - -/** - * Checks if a `cache` value for `key` exists. - * - * @private - * @param {Object} cache The cache to query. - * @param {string} key The key of the entry to check. - * @returns {boolean} Returns `true` if an entry for `key` exists, else `false`. - */ - -function cacheHas$2(cache, key) { - return cache.has(key); -} - -var _cacheHas = cacheHas$2; - -var SetCache$1 = _SetCache, - arraySome$1 = _arraySome, - cacheHas$1 = _cacheHas; - -/** Used to compose bitmasks for value comparisons. */ -var COMPARE_PARTIAL_FLAG$5 = 1, - COMPARE_UNORDERED_FLAG$3 = 2; - -/** - * A specialized version of `baseIsEqualDeep` for arrays with support for - * partial deep comparisons. - * - * @private - * @param {Array} array The array to compare. - * @param {Array} other The other array to compare. - * @param {number} bitmask The bitmask flags. See `baseIsEqual` for more details. - * @param {Function} customizer The function to customize comparisons. - * @param {Function} equalFunc The function to determine equivalents of values. - * @param {Object} stack Tracks traversed `array` and `other` objects. - * @returns {boolean} Returns `true` if the arrays are equivalent, else `false`. - */ -function equalArrays$2(array, other, bitmask, customizer, equalFunc, stack) { - var isPartial = bitmask & COMPARE_PARTIAL_FLAG$5, - arrLength = array.length, - othLength = other.length; - - if (arrLength != othLength && !(isPartial && othLength > arrLength)) { - return false; - } - // Check that cyclic values are equal. - var arrStacked = stack.get(array); - var othStacked = stack.get(other); - if (arrStacked && othStacked) { - return arrStacked == other && othStacked == array; - } - var index = -1, - result = true, - seen = (bitmask & COMPARE_UNORDERED_FLAG$3) ? new SetCache$1 : undefined; - - stack.set(array, other); - stack.set(other, array); - - // Ignore non-index properties. - while (++index < arrLength) { - var arrValue = array[index], - othValue = other[index]; - - if (customizer) { - var compared = isPartial - ? customizer(othValue, arrValue, index, other, array, stack) - : customizer(arrValue, othValue, index, array, other, stack); - } - if (compared !== undefined) { - if (compared) { - continue; - } - result = false; - break; - } - // Recursively compare arrays (susceptible to call stack limits). - if (seen) { - if (!arraySome$1(other, function(othValue, othIndex) { - if (!cacheHas$1(seen, othIndex) && - (arrValue === othValue || equalFunc(arrValue, othValue, bitmask, customizer, stack))) { - return seen.push(othIndex); - } - })) { - result = false; - break; - } - } else if (!( - arrValue === othValue || - equalFunc(arrValue, othValue, bitmask, customizer, stack) - )) { - result = false; - break; - } - } - stack['delete'](array); - stack['delete'](other); - return result; -} - -var _equalArrays = equalArrays$2; - -var root$8 = _root$1; - -/** Built-in value references. */ -var Uint8Array$2 = root$8.Uint8Array; - -var _Uint8Array = Uint8Array$2; - -/** - * Converts `map` to its key-value pairs. - * - * @private - * @param {Object} map The map to convert. - * @returns {Array} Returns the key-value pairs. - */ - -function mapToArray$1(map) { - var index = -1, - result = Array(map.size); - - map.forEach(function(value, key) { - result[++index] = [key, value]; - }); - return result; -} - -var _mapToArray = mapToArray$1; - -/** - * Converts `set` to an array of its values. - * - * @private - * @param {Object} set The set to convert. - * @returns {Array} Returns the values. - */ - -function setToArray$3(set) { - var index = -1, - result = Array(set.size); - - set.forEach(function(value) { - result[++index] = value; - }); - return result; -} - -var _setToArray = setToArray$3; - -var Symbol$6 = _Symbol$1, - Uint8Array$1 = _Uint8Array, - eq$2 = eq_1, - equalArrays$1 = _equalArrays, - mapToArray = _mapToArray, - setToArray$2 = _setToArray; - -/** Used to compose bitmasks for value comparisons. */ -var COMPARE_PARTIAL_FLAG$4 = 1, - COMPARE_UNORDERED_FLAG$2 = 2; - -/** `Object#toString` result references. */ -var boolTag$4 = '[object Boolean]', - dateTag$3 = '[object Date]', - errorTag$2 = '[object Error]', - mapTag$5 = '[object Map]', - numberTag$4 = '[object Number]', - regexpTag$3 = '[object RegExp]', - setTag$5 = '[object Set]', - stringTag$4 = '[object String]', - symbolTag$4 = '[object Symbol]'; - -var arrayBufferTag$3 = '[object ArrayBuffer]', - dataViewTag$4 = '[object DataView]'; - -/** Used to convert symbols to primitives and strings. */ -var symbolProto$2 = Symbol$6 ? Symbol$6.prototype : undefined, - symbolValueOf$1 = symbolProto$2 ? symbolProto$2.valueOf : undefined; - -/** - * A specialized version of `baseIsEqualDeep` for comparing objects of - * the same `toStringTag`. - * - * **Note:** This function only supports comparing values with tags of - * `Boolean`, `Date`, `Error`, `Number`, `RegExp`, or `String`. - * - * @private - * @param {Object} object The object to compare. - * @param {Object} other The other object to compare. - * @param {string} tag The `toStringTag` of the objects to compare. - * @param {number} bitmask The bitmask flags. See `baseIsEqual` for more details. - * @param {Function} customizer The function to customize comparisons. - * @param {Function} equalFunc The function to determine equivalents of values. - * @param {Object} stack Tracks traversed `object` and `other` objects. - * @returns {boolean} Returns `true` if the objects are equivalent, else `false`. - */ -function equalByTag$1(object, other, tag, bitmask, customizer, equalFunc, stack) { - switch (tag) { - case dataViewTag$4: - if ((object.byteLength != other.byteLength) || - (object.byteOffset != other.byteOffset)) { - return false; - } - object = object.buffer; - other = other.buffer; - - case arrayBufferTag$3: - if ((object.byteLength != other.byteLength) || - !equalFunc(new Uint8Array$1(object), new Uint8Array$1(other))) { - return false; - } - return true; - - case boolTag$4: - case dateTag$3: - case numberTag$4: - // Coerce booleans to `1` or `0` and dates to milliseconds. - // Invalid dates are coerced to `NaN`. - return eq$2(+object, +other); - - case errorTag$2: - return object.name == other.name && object.message == other.message; - - case regexpTag$3: - case stringTag$4: - // Coerce regexes to strings and treat strings, primitives and objects, - // as equal. See http://www.ecma-international.org/ecma-262/7.0/#sec-regexp.prototype.tostring - // for more details. - return object == (other + ''); - - case mapTag$5: - var convert = mapToArray; - - case setTag$5: - var isPartial = bitmask & COMPARE_PARTIAL_FLAG$4; - convert || (convert = setToArray$2); - - if (object.size != other.size && !isPartial) { - return false; - } - // Assume cyclic values are equal. - var stacked = stack.get(object); - if (stacked) { - return stacked == other; - } - bitmask |= COMPARE_UNORDERED_FLAG$2; - - // Recursively compare objects (susceptible to call stack limits). - stack.set(object, other); - var result = equalArrays$1(convert(object), convert(other), bitmask, customizer, equalFunc, stack); - stack['delete'](object); - return result; - - case symbolTag$4: - if (symbolValueOf$1) { - return symbolValueOf$1.call(object) == symbolValueOf$1.call(other); - } - } - return false; -} - -var _equalByTag = equalByTag$1; - -/** - * Appends the elements of `values` to `array`. - * - * @private - * @param {Array} array The array to modify. - * @param {Array} values The values to append. - * @returns {Array} Returns `array`. - */ - -function arrayPush$3(array, values) { - var index = -1, - length = values.length, - offset = array.length; - - while (++index < length) { - array[offset + index] = values[index]; - } - return array; -} - -var _arrayPush = arrayPush$3; - -/** - * Checks if `value` is classified as an `Array` object. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is an array, else `false`. - * @example - * - * _.isArray([1, 2, 3]); - * // => true - * - * _.isArray(document.body.children); - * // => false - * - * _.isArray('abc'); - * // => false - * - * _.isArray(_.noop); - * // => false - */ - -var isArray$f = Array.isArray; - -var isArray_1 = isArray$f; - -var arrayPush$2 = _arrayPush, - isArray$e = isArray_1; - -/** - * The base implementation of `getAllKeys` and `getAllKeysIn` which uses - * `keysFunc` and `symbolsFunc` to get the enumerable property names and - * symbols of `object`. - * - * @private - * @param {Object} object The object to query. - * @param {Function} keysFunc The function to get the keys of `object`. - * @param {Function} symbolsFunc The function to get the symbols of `object`. - * @returns {Array} Returns the array of property names and symbols. - */ -function baseGetAllKeys$2(object, keysFunc, symbolsFunc) { - var result = keysFunc(object); - return isArray$e(object) ? result : arrayPush$2(result, symbolsFunc(object)); -} - -var _baseGetAllKeys = baseGetAllKeys$2; - -/** - * A specialized version of `_.filter` for arrays without support for - * iteratee shorthands. - * - * @private - * @param {Array} [array] The array to iterate over. - * @param {Function} predicate The function invoked per iteration. - * @returns {Array} Returns the new filtered array. - */ - -function arrayFilter$1(array, predicate) { - var index = -1, - length = array == null ? 0 : array.length, - resIndex = 0, - result = []; - - while (++index < length) { - var value = array[index]; - if (predicate(value, index, array)) { - result[resIndex++] = value; - } - } - return result; -} - -var _arrayFilter = arrayFilter$1; - -/** - * This method returns a new empty array. - * - * @static - * @memberOf _ - * @since 4.13.0 - * @category Util - * @returns {Array} Returns the new empty array. - * @example - * - * var arrays = _.times(2, _.stubArray); - * - * console.log(arrays); - * // => [[], []] - * - * console.log(arrays[0] === arrays[1]); - * // => false - */ - -function stubArray$2() { - return []; -} - -var stubArray_1 = stubArray$2; - -var arrayFilter = _arrayFilter, - stubArray$1 = stubArray_1; - -/** Used for built-in method references. */ -var objectProto$b = Object.prototype; - -/** Built-in value references. */ -var propertyIsEnumerable$1 = objectProto$b.propertyIsEnumerable; - -/* Built-in method references for those with the same name as other `lodash` methods. */ -var nativeGetSymbols$1 = Object.getOwnPropertySymbols; - -/** - * Creates an array of the own enumerable symbols of `object`. - * - * @private - * @param {Object} object The object to query. - * @returns {Array} Returns the array of symbols. - */ -var getSymbols$3 = !nativeGetSymbols$1 ? stubArray$1 : function(object) { - if (object == null) { - return []; - } - object = Object(object); - return arrayFilter(nativeGetSymbols$1(object), function(symbol) { - return propertyIsEnumerable$1.call(object, symbol); - }); -}; - -var _getSymbols = getSymbols$3; - -/** - * The base implementation of `_.times` without support for iteratee shorthands - * or max array length checks. - * - * @private - * @param {number} n The number of times to invoke `iteratee`. - * @param {Function} iteratee The function invoked per iteration. - * @returns {Array} Returns the array of results. - */ - -function baseTimes$1(n, iteratee) { - var index = -1, - result = Array(n); - - while (++index < n) { - result[index] = iteratee(index); - } - return result; -} - -var _baseTimes = baseTimes$1; - -/** - * Checks if `value` is object-like. A value is object-like if it's not `null` - * and has a `typeof` result of "object". - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is object-like, else `false`. - * @example - * - * _.isObjectLike({}); - * // => true - * - * _.isObjectLike([1, 2, 3]); - * // => true - * - * _.isObjectLike(_.noop); - * // => false - * - * _.isObjectLike(null); - * // => false - */ - -function isObjectLike$c(value) { - return value != null && typeof value == 'object'; -} - -var isObjectLike_1$1 = isObjectLike$c; - -var baseGetTag$8 = _baseGetTag$1, - isObjectLike$b = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var argsTag$3 = '[object Arguments]'; - -/** - * The base implementation of `_.isArguments`. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is an `arguments` object, - */ -function baseIsArguments$1(value) { - return isObjectLike$b(value) && baseGetTag$8(value) == argsTag$3; -} - -var _baseIsArguments = baseIsArguments$1; - -var baseIsArguments = _baseIsArguments, - isObjectLike$a = isObjectLike_1$1; - -/** Used for built-in method references. */ -var objectProto$a = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$9 = objectProto$a.hasOwnProperty; - -/** Built-in value references. */ -var propertyIsEnumerable = objectProto$a.propertyIsEnumerable; - -/** - * Checks if `value` is likely an `arguments` object. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is an `arguments` object, - * else `false`. - * @example - * - * _.isArguments(function() { return arguments; }()); - * // => true - * - * _.isArguments([1, 2, 3]); - * // => false - */ -var isArguments$3 = baseIsArguments(function() { return arguments; }()) ? baseIsArguments : function(value) { - return isObjectLike$a(value) && hasOwnProperty$9.call(value, 'callee') && - !propertyIsEnumerable.call(value, 'callee'); -}; - -var isArguments_1 = isArguments$3; - -var isBuffer$3 = {exports: {}}; - -/** - * This method returns `false`. - * - * @static - * @memberOf _ - * @since 4.13.0 - * @category Util - * @returns {boolean} Returns `false`. - * @example - * - * _.times(2, _.stubFalse); - * // => [false, false] - */ - -function stubFalse() { - return false; -} - -var stubFalse_1 = stubFalse; - -(function (module, exports) { - var root = _root$1, - stubFalse = stubFalse_1; - - /** Detect free variable `exports`. */ - var freeExports = exports && !exports.nodeType && exports; - - /** Detect free variable `module`. */ - var freeModule = freeExports && 'object' == 'object' && module && !module.nodeType && module; - - /** Detect the popular CommonJS extension `module.exports`. */ - var moduleExports = freeModule && freeModule.exports === freeExports; - - /** Built-in value references. */ - var Buffer = moduleExports ? root.Buffer : undefined; - - /* Built-in method references for those with the same name as other `lodash` methods. */ - var nativeIsBuffer = Buffer ? Buffer.isBuffer : undefined; - - /** - * Checks if `value` is a buffer. - * - * @static - * @memberOf _ - * @since 4.3.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a buffer, else `false`. - * @example - * - * _.isBuffer(new Buffer(2)); - * // => true - * - * _.isBuffer(new Uint8Array(2)); - * // => false - */ - var isBuffer = nativeIsBuffer || stubFalse; - - module.exports = isBuffer; -} (isBuffer$3, isBuffer$3.exports)); - -/** Used as references for various `Number` constants. */ - -var MAX_SAFE_INTEGER$1 = 9007199254740991; - -/** Used to detect unsigned integer values. */ -var reIsUint = /^(?:0|[1-9]\d*)$/; - -/** - * Checks if `value` is a valid array-like index. - * - * @private - * @param {*} value The value to check. - * @param {number} [length=MAX_SAFE_INTEGER] The upper bounds of a valid index. - * @returns {boolean} Returns `true` if `value` is a valid index, else `false`. - */ -function isIndex$3(value, length) { - var type = typeof value; - length = length == null ? MAX_SAFE_INTEGER$1 : length; - - return !!length && - (type == 'number' || - (type != 'symbol' && reIsUint.test(value))) && - (value > -1 && value % 1 == 0 && value < length); -} - -var _isIndex = isIndex$3; - -/** Used as references for various `Number` constants. */ - -var MAX_SAFE_INTEGER = 9007199254740991; - -/** - * Checks if `value` is a valid array-like length. - * - * **Note:** This method is loosely based on - * [`ToLength`](http://ecma-international.org/ecma-262/7.0/#sec-tolength). - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a valid length, else `false`. - * @example - * - * _.isLength(3); - * // => true - * - * _.isLength(Number.MIN_VALUE); - * // => false - * - * _.isLength(Infinity); - * // => false - * - * _.isLength('3'); - * // => false - */ -function isLength$3(value) { - return typeof value == 'number' && - value > -1 && value % 1 == 0 && value <= MAX_SAFE_INTEGER; -} - -var isLength_1 = isLength$3; - -var baseGetTag$7 = _baseGetTag$1, - isLength$2 = isLength_1, - isObjectLike$9 = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var argsTag$2 = '[object Arguments]', - arrayTag$2 = '[object Array]', - boolTag$3 = '[object Boolean]', - dateTag$2 = '[object Date]', - errorTag$1 = '[object Error]', - funcTag$1 = '[object Function]', - mapTag$4 = '[object Map]', - numberTag$3 = '[object Number]', - objectTag$4 = '[object Object]', - regexpTag$2 = '[object RegExp]', - setTag$4 = '[object Set]', - stringTag$3 = '[object String]', - weakMapTag$2 = '[object WeakMap]'; - -var arrayBufferTag$2 = '[object ArrayBuffer]', - dataViewTag$3 = '[object DataView]', - float32Tag$2 = '[object Float32Array]', - float64Tag$2 = '[object Float64Array]', - int8Tag$2 = '[object Int8Array]', - int16Tag$2 = '[object Int16Array]', - int32Tag$2 = '[object Int32Array]', - uint8Tag$2 = '[object Uint8Array]', - uint8ClampedTag$2 = '[object Uint8ClampedArray]', - uint16Tag$2 = '[object Uint16Array]', - uint32Tag$2 = '[object Uint32Array]'; - -/** Used to identify `toStringTag` values of typed arrays. */ -var typedArrayTags = {}; -typedArrayTags[float32Tag$2] = typedArrayTags[float64Tag$2] = -typedArrayTags[int8Tag$2] = typedArrayTags[int16Tag$2] = -typedArrayTags[int32Tag$2] = typedArrayTags[uint8Tag$2] = -typedArrayTags[uint8ClampedTag$2] = typedArrayTags[uint16Tag$2] = -typedArrayTags[uint32Tag$2] = true; -typedArrayTags[argsTag$2] = typedArrayTags[arrayTag$2] = -typedArrayTags[arrayBufferTag$2] = typedArrayTags[boolTag$3] = -typedArrayTags[dataViewTag$3] = typedArrayTags[dateTag$2] = -typedArrayTags[errorTag$1] = typedArrayTags[funcTag$1] = -typedArrayTags[mapTag$4] = typedArrayTags[numberTag$3] = -typedArrayTags[objectTag$4] = typedArrayTags[regexpTag$2] = -typedArrayTags[setTag$4] = typedArrayTags[stringTag$3] = -typedArrayTags[weakMapTag$2] = false; - -/** - * The base implementation of `_.isTypedArray` without Node.js optimizations. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a typed array, else `false`. - */ -function baseIsTypedArray$1(value) { - return isObjectLike$9(value) && - isLength$2(value.length) && !!typedArrayTags[baseGetTag$7(value)]; -} - -var _baseIsTypedArray = baseIsTypedArray$1; - -/** - * The base implementation of `_.unary` without support for storing metadata. - * - * @private - * @param {Function} func The function to cap arguments for. - * @returns {Function} Returns the new capped function. - */ - -function baseUnary$4(func) { - return function(value) { - return func(value); - }; -} - -var _baseUnary = baseUnary$4; - -var _nodeUtil = {exports: {}}; - -(function (module, exports) { - var freeGlobal = _freeGlobal$1; - - /** Detect free variable `exports`. */ - var freeExports = exports && !exports.nodeType && exports; - - /** Detect free variable `module`. */ - var freeModule = freeExports && 'object' == 'object' && module && !module.nodeType && module; - - /** Detect the popular CommonJS extension `module.exports`. */ - var moduleExports = freeModule && freeModule.exports === freeExports; - - /** Detect free variable `process` from Node.js. */ - var freeProcess = moduleExports && freeGlobal.process; - - /** Used to access faster Node.js helpers. */ - var nodeUtil = (function() { - try { - // Use `util.types` for Node.js 10+. - var types = freeModule && freeModule.require && freeModule.require('util').types; - - if (types) { - return types; - } - - // Legacy `process.binding('util')` for Node.js < 10. - return freeProcess && freeProcess.binding && freeProcess.binding('util'); - } catch (e) {} - }()); - - module.exports = nodeUtil; -} (_nodeUtil, _nodeUtil.exports)); - -var baseIsTypedArray = _baseIsTypedArray, - baseUnary$3 = _baseUnary, - nodeUtil$2 = _nodeUtil.exports; - -/* Node.js helper references. */ -var nodeIsTypedArray = nodeUtil$2 && nodeUtil$2.isTypedArray; - -/** - * Checks if `value` is classified as a typed array. - * - * @static - * @memberOf _ - * @since 3.0.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a typed array, else `false`. - * @example - * - * _.isTypedArray(new Uint8Array); - * // => true - * - * _.isTypedArray([]); - * // => false - */ -var isTypedArray$2 = nodeIsTypedArray ? baseUnary$3(nodeIsTypedArray) : baseIsTypedArray; - -var isTypedArray_1 = isTypedArray$2; - -var baseTimes = _baseTimes, - isArguments$2 = isArguments_1, - isArray$d = isArray_1, - isBuffer$2 = isBuffer$3.exports, - isIndex$2 = _isIndex, - isTypedArray$1 = isTypedArray_1; - -/** Used for built-in method references. */ -var objectProto$9 = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$8 = objectProto$9.hasOwnProperty; - -/** - * Creates an array of the enumerable property names of the array-like `value`. - * - * @private - * @param {*} value The value to query. - * @param {boolean} inherited Specify returning inherited property names. - * @returns {Array} Returns the array of property names. - */ -function arrayLikeKeys$2(value, inherited) { - var isArr = isArray$d(value), - isArg = !isArr && isArguments$2(value), - isBuff = !isArr && !isArg && isBuffer$2(value), - isType = !isArr && !isArg && !isBuff && isTypedArray$1(value), - skipIndexes = isArr || isArg || isBuff || isType, - result = skipIndexes ? baseTimes(value.length, String) : [], - length = result.length; - - for (var key in value) { - if ((inherited || hasOwnProperty$8.call(value, key)) && - !(skipIndexes && ( - // Safari 9 has enumerable `arguments.length` in strict mode. - key == 'length' || - // Node.js 0.10 has enumerable non-index properties on buffers. - (isBuff && (key == 'offset' || key == 'parent')) || - // PhantomJS 2 has enumerable non-index properties on typed arrays. - (isType && (key == 'buffer' || key == 'byteLength' || key == 'byteOffset')) || - // Skip index properties. - isIndex$2(key, length) - ))) { - result.push(key); - } - } - return result; -} - -var _arrayLikeKeys = arrayLikeKeys$2; - -/** Used for built-in method references. */ - -var objectProto$8 = Object.prototype; - -/** - * Checks if `value` is likely a prototype object. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a prototype, else `false`. - */ -function isPrototype$3(value) { - var Ctor = value && value.constructor, - proto = (typeof Ctor == 'function' && Ctor.prototype) || objectProto$8; - - return value === proto; -} - -var _isPrototype = isPrototype$3; - -/** - * Creates a unary function that invokes `func` with its argument transformed. - * - * @private - * @param {Function} func The function to wrap. - * @param {Function} transform The argument transform. - * @returns {Function} Returns the new function. - */ - -function overArg$2(func, transform) { - return function(arg) { - return func(transform(arg)); - }; -} - -var _overArg = overArg$2; - -var overArg$1 = _overArg; - -/* Built-in method references for those with the same name as other `lodash` methods. */ -var nativeKeys$1 = overArg$1(Object.keys, Object); - -var _nativeKeys = nativeKeys$1; - -var isPrototype$2 = _isPrototype, - nativeKeys = _nativeKeys; - -/** Used for built-in method references. */ -var objectProto$7 = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$7 = objectProto$7.hasOwnProperty; - -/** - * The base implementation of `_.keys` which doesn't treat sparse arrays as dense. - * - * @private - * @param {Object} object The object to query. - * @returns {Array} Returns the array of property names. - */ -function baseKeys$1(object) { - if (!isPrototype$2(object)) { - return nativeKeys(object); - } - var result = []; - for (var key in Object(object)) { - if (hasOwnProperty$7.call(object, key) && key != 'constructor') { - result.push(key); - } - } - return result; -} - -var _baseKeys = baseKeys$1; - -var isFunction$1 = isFunction_1, - isLength$1 = isLength_1; - -/** - * Checks if `value` is array-like. A value is considered array-like if it's - * not a function and has a `value.length` that's an integer greater than or - * equal to `0` and less than or equal to `Number.MAX_SAFE_INTEGER`. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is array-like, else `false`. - * @example - * - * _.isArrayLike([1, 2, 3]); - * // => true - * - * _.isArrayLike(document.body.children); - * // => true - * - * _.isArrayLike('abc'); - * // => true - * - * _.isArrayLike(_.noop); - * // => false - */ -function isArrayLike$6(value) { - return value != null && isLength$1(value.length) && !isFunction$1(value); -} - -var isArrayLike_1 = isArrayLike$6; - -var arrayLikeKeys$1 = _arrayLikeKeys, - baseKeys = _baseKeys, - isArrayLike$5 = isArrayLike_1; - -/** - * Creates an array of the own enumerable property names of `object`. - * - * **Note:** Non-object values are coerced to objects. See the - * [ES spec](http://ecma-international.org/ecma-262/7.0/#sec-object.keys) - * for more details. - * - * @static - * @since 0.1.0 - * @memberOf _ - * @category Object - * @param {Object} object The object to query. - * @returns {Array} Returns the array of property names. - * @example - * - * function Foo() { - * this.a = 1; - * this.b = 2; - * } - * - * Foo.prototype.c = 3; - * - * _.keys(new Foo); - * // => ['a', 'b'] (iteration order is not guaranteed) - * - * _.keys('hi'); - * // => ['0', '1'] - */ -function keys$6(object) { - return isArrayLike$5(object) ? arrayLikeKeys$1(object) : baseKeys(object); -} - -var keys_1 = keys$6; - -var baseGetAllKeys$1 = _baseGetAllKeys, - getSymbols$2 = _getSymbols, - keys$5 = keys_1; - -/** - * Creates an array of own enumerable property names and symbols of `object`. - * - * @private - * @param {Object} object The object to query. - * @returns {Array} Returns the array of property names and symbols. - */ -function getAllKeys$2(object) { - return baseGetAllKeys$1(object, keys$5, getSymbols$2); -} - -var _getAllKeys = getAllKeys$2; - -var getAllKeys$1 = _getAllKeys; - -/** Used to compose bitmasks for value comparisons. */ -var COMPARE_PARTIAL_FLAG$3 = 1; - -/** Used for built-in method references. */ -var objectProto$6 = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$6 = objectProto$6.hasOwnProperty; - -/** - * A specialized version of `baseIsEqualDeep` for objects with support for - * partial deep comparisons. - * - * @private - * @param {Object} object The object to compare. - * @param {Object} other The other object to compare. - * @param {number} bitmask The bitmask flags. See `baseIsEqual` for more details. - * @param {Function} customizer The function to customize comparisons. - * @param {Function} equalFunc The function to determine equivalents of values. - * @param {Object} stack Tracks traversed `object` and `other` objects. - * @returns {boolean} Returns `true` if the objects are equivalent, else `false`. - */ -function equalObjects$1(object, other, bitmask, customizer, equalFunc, stack) { - var isPartial = bitmask & COMPARE_PARTIAL_FLAG$3, - objProps = getAllKeys$1(object), - objLength = objProps.length, - othProps = getAllKeys$1(other), - othLength = othProps.length; - - if (objLength != othLength && !isPartial) { - return false; - } - var index = objLength; - while (index--) { - var key = objProps[index]; - if (!(isPartial ? key in other : hasOwnProperty$6.call(other, key))) { - return false; - } - } - // Check that cyclic values are equal. - var objStacked = stack.get(object); - var othStacked = stack.get(other); - if (objStacked && othStacked) { - return objStacked == other && othStacked == object; - } - var result = true; - stack.set(object, other); - stack.set(other, object); - - var skipCtor = isPartial; - while (++index < objLength) { - key = objProps[index]; - var objValue = object[key], - othValue = other[key]; - - if (customizer) { - var compared = isPartial - ? customizer(othValue, objValue, key, other, object, stack) - : customizer(objValue, othValue, key, object, other, stack); - } - // Recursively compare objects (susceptible to call stack limits). - if (!(compared === undefined - ? (objValue === othValue || equalFunc(objValue, othValue, bitmask, customizer, stack)) - : compared - )) { - result = false; - break; - } - skipCtor || (skipCtor = key == 'constructor'); - } - if (result && !skipCtor) { - var objCtor = object.constructor, - othCtor = other.constructor; - - // Non `Object` object instances with different constructors are not equal. - if (objCtor != othCtor && - ('constructor' in object && 'constructor' in other) && - !(typeof objCtor == 'function' && objCtor instanceof objCtor && - typeof othCtor == 'function' && othCtor instanceof othCtor)) { - result = false; - } - } - stack['delete'](object); - stack['delete'](other); - return result; -} - -var _equalObjects = equalObjects$1; - -var getNative$4 = _getNative, - root$7 = _root$1; - -/* Built-in method references that are verified to be native. */ -var DataView$2 = getNative$4(root$7, 'DataView'); - -var _DataView = DataView$2; - -var getNative$3 = _getNative, - root$6 = _root$1; - -/* Built-in method references that are verified to be native. */ -var Promise$2 = getNative$3(root$6, 'Promise'); - -var _Promise = Promise$2; - -var getNative$2 = _getNative, - root$5 = _root$1; - -/* Built-in method references that are verified to be native. */ -var Set$3 = getNative$2(root$5, 'Set'); - -var _Set = Set$3; - -var getNative$1 = _getNative, - root$4 = _root$1; - -/* Built-in method references that are verified to be native. */ -var WeakMap$2 = getNative$1(root$4, 'WeakMap'); - -var _WeakMap = WeakMap$2; - -var DataView$1 = _DataView, - Map$1 = _Map, - Promise$1 = _Promise, - Set$2 = _Set, - WeakMap$1 = _WeakMap, - baseGetTag$6 = _baseGetTag$1, - toSource = _toSource; - -/** `Object#toString` result references. */ -var mapTag$3 = '[object Map]', - objectTag$3 = '[object Object]', - promiseTag = '[object Promise]', - setTag$3 = '[object Set]', - weakMapTag$1 = '[object WeakMap]'; - -var dataViewTag$2 = '[object DataView]'; - -/** Used to detect maps, sets, and weakmaps. */ -var dataViewCtorString = toSource(DataView$1), - mapCtorString = toSource(Map$1), - promiseCtorString = toSource(Promise$1), - setCtorString = toSource(Set$2), - weakMapCtorString = toSource(WeakMap$1); - -/** - * Gets the `toStringTag` of `value`. - * - * @private - * @param {*} value The value to query. - * @returns {string} Returns the `toStringTag`. - */ -var getTag$4 = baseGetTag$6; - -// Fallback for data views, maps, sets, and weak maps in IE 11 and promises in Node.js < 6. -if ((DataView$1 && getTag$4(new DataView$1(new ArrayBuffer(1))) != dataViewTag$2) || - (Map$1 && getTag$4(new Map$1) != mapTag$3) || - (Promise$1 && getTag$4(Promise$1.resolve()) != promiseTag) || - (Set$2 && getTag$4(new Set$2) != setTag$3) || - (WeakMap$1 && getTag$4(new WeakMap$1) != weakMapTag$1)) { - getTag$4 = function(value) { - var result = baseGetTag$6(value), - Ctor = result == objectTag$3 ? value.constructor : undefined, - ctorString = Ctor ? toSource(Ctor) : ''; - - if (ctorString) { - switch (ctorString) { - case dataViewCtorString: return dataViewTag$2; - case mapCtorString: return mapTag$3; - case promiseCtorString: return promiseTag; - case setCtorString: return setTag$3; - case weakMapCtorString: return weakMapTag$1; - } - } - return result; - }; -} - -var _getTag = getTag$4; - -var Stack$2 = _Stack, - equalArrays = _equalArrays, - equalByTag = _equalByTag, - equalObjects = _equalObjects, - getTag$3 = _getTag, - isArray$c = isArray_1, - isBuffer$1 = isBuffer$3.exports, - isTypedArray = isTypedArray_1; - -/** Used to compose bitmasks for value comparisons. */ -var COMPARE_PARTIAL_FLAG$2 = 1; - -/** `Object#toString` result references. */ -var argsTag$1 = '[object Arguments]', - arrayTag$1 = '[object Array]', - objectTag$2 = '[object Object]'; - -/** Used for built-in method references. */ -var objectProto$5 = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$5 = objectProto$5.hasOwnProperty; - -/** - * A specialized version of `baseIsEqual` for arrays and objects which performs - * deep comparisons and tracks traversed objects enabling objects with circular - * references to be compared. - * - * @private - * @param {Object} object The object to compare. - * @param {Object} other The other object to compare. - * @param {number} bitmask The bitmask flags. See `baseIsEqual` for more details. - * @param {Function} customizer The function to customize comparisons. - * @param {Function} equalFunc The function to determine equivalents of values. - * @param {Object} [stack] Tracks traversed `object` and `other` objects. - * @returns {boolean} Returns `true` if the objects are equivalent, else `false`. - */ -function baseIsEqualDeep$1(object, other, bitmask, customizer, equalFunc, stack) { - var objIsArr = isArray$c(object), - othIsArr = isArray$c(other), - objTag = objIsArr ? arrayTag$1 : getTag$3(object), - othTag = othIsArr ? arrayTag$1 : getTag$3(other); - - objTag = objTag == argsTag$1 ? objectTag$2 : objTag; - othTag = othTag == argsTag$1 ? objectTag$2 : othTag; - - var objIsObj = objTag == objectTag$2, - othIsObj = othTag == objectTag$2, - isSameTag = objTag == othTag; - - if (isSameTag && isBuffer$1(object)) { - if (!isBuffer$1(other)) { - return false; - } - objIsArr = true; - objIsObj = false; - } - if (isSameTag && !objIsObj) { - stack || (stack = new Stack$2); - return (objIsArr || isTypedArray(object)) - ? equalArrays(object, other, bitmask, customizer, equalFunc, stack) - : equalByTag(object, other, objTag, bitmask, customizer, equalFunc, stack); - } - if (!(bitmask & COMPARE_PARTIAL_FLAG$2)) { - var objIsWrapped = objIsObj && hasOwnProperty$5.call(object, '__wrapped__'), - othIsWrapped = othIsObj && hasOwnProperty$5.call(other, '__wrapped__'); - - if (objIsWrapped || othIsWrapped) { - var objUnwrapped = objIsWrapped ? object.value() : object, - othUnwrapped = othIsWrapped ? other.value() : other; - - stack || (stack = new Stack$2); - return equalFunc(objUnwrapped, othUnwrapped, bitmask, customizer, stack); - } - } - if (!isSameTag) { - return false; - } - stack || (stack = new Stack$2); - return equalObjects(object, other, bitmask, customizer, equalFunc, stack); -} - -var _baseIsEqualDeep = baseIsEqualDeep$1; - -var baseIsEqualDeep = _baseIsEqualDeep, - isObjectLike$8 = isObjectLike_1$1; - -/** - * The base implementation of `_.isEqual` which supports partial comparisons - * and tracks traversed objects. - * - * @private - * @param {*} value The value to compare. - * @param {*} other The other value to compare. - * @param {boolean} bitmask The bitmask flags. - * 1 - Unordered comparison - * 2 - Partial comparison - * @param {Function} [customizer] The function to customize comparisons. - * @param {Object} [stack] Tracks traversed `value` and `other` objects. - * @returns {boolean} Returns `true` if the values are equivalent, else `false`. - */ -function baseIsEqual$3(value, other, bitmask, customizer, stack) { - if (value === other) { - return true; - } - if (value == null || other == null || (!isObjectLike$8(value) && !isObjectLike$8(other))) { - return value !== value && other !== other; - } - return baseIsEqualDeep(value, other, bitmask, customizer, baseIsEqual$3, stack); -} - -var _baseIsEqual = baseIsEqual$3; - -var Stack$1 = _Stack, - baseIsEqual$2 = _baseIsEqual; - -/** Used to compose bitmasks for value comparisons. */ -var COMPARE_PARTIAL_FLAG$1 = 1, - COMPARE_UNORDERED_FLAG$1 = 2; - -/** - * The base implementation of `_.isMatch` without support for iteratee shorthands. - * - * @private - * @param {Object} object The object to inspect. - * @param {Object} source The object of property values to match. - * @param {Array} matchData The property names, values, and compare flags to match. - * @param {Function} [customizer] The function to customize comparisons. - * @returns {boolean} Returns `true` if `object` is a match, else `false`. - */ -function baseIsMatch$1(object, source, matchData, customizer) { - var index = matchData.length, - length = index, - noCustomizer = !customizer; - - if (object == null) { - return !length; - } - object = Object(object); - while (index--) { - var data = matchData[index]; - if ((noCustomizer && data[2]) - ? data[1] !== object[data[0]] - : !(data[0] in object) - ) { - return false; - } - } - while (++index < length) { - data = matchData[index]; - var key = data[0], - objValue = object[key], - srcValue = data[1]; - - if (noCustomizer && data[2]) { - if (objValue === undefined && !(key in object)) { - return false; - } - } else { - var stack = new Stack$1; - if (customizer) { - var result = customizer(objValue, srcValue, key, object, source, stack); - } - if (!(result === undefined - ? baseIsEqual$2(srcValue, objValue, COMPARE_PARTIAL_FLAG$1 | COMPARE_UNORDERED_FLAG$1, customizer, stack) - : result - )) { - return false; - } - } - } - return true; -} - -var _baseIsMatch = baseIsMatch$1; - -var isObject$8 = isObject_1$1; - -/** - * Checks if `value` is suitable for strict equality comparisons, i.e. `===`. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` if suitable for strict - * equality comparisons, else `false`. - */ -function isStrictComparable$2(value) { - return value === value && !isObject$8(value); -} - -var _isStrictComparable = isStrictComparable$2; - -var isStrictComparable$1 = _isStrictComparable, - keys$4 = keys_1; - -/** - * Gets the property names, values, and compare flags of `object`. - * - * @private - * @param {Object} object The object to query. - * @returns {Array} Returns the match data of `object`. - */ -function getMatchData$1(object) { - var result = keys$4(object), - length = result.length; - - while (length--) { - var key = result[length], - value = object[key]; - - result[length] = [key, value, isStrictComparable$1(value)]; - } - return result; -} - -var _getMatchData = getMatchData$1; - -/** - * A specialized version of `matchesProperty` for source values suitable - * for strict equality comparisons, i.e. `===`. - * - * @private - * @param {string} key The key of the property to get. - * @param {*} srcValue The value to match. - * @returns {Function} Returns the new spec function. - */ - -function matchesStrictComparable$2(key, srcValue) { - return function(object) { - if (object == null) { - return false; - } - return object[key] === srcValue && - (srcValue !== undefined || (key in Object(object))); - }; -} - -var _matchesStrictComparable = matchesStrictComparable$2; - -var baseIsMatch = _baseIsMatch, - getMatchData = _getMatchData, - matchesStrictComparable$1 = _matchesStrictComparable; - -/** - * The base implementation of `_.matches` which doesn't clone `source`. - * - * @private - * @param {Object} source The object of property values to match. - * @returns {Function} Returns the new spec function. - */ -function baseMatches$1(source) { - var matchData = getMatchData(source); - if (matchData.length == 1 && matchData[0][2]) { - return matchesStrictComparable$1(matchData[0][0], matchData[0][1]); - } - return function(object) { - return object === source || baseIsMatch(object, source, matchData); - }; -} - -var _baseMatches = baseMatches$1; - -var baseGetTag$5 = _baseGetTag$1, - isObjectLike$7 = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var symbolTag$3 = '[object Symbol]'; - -/** - * Checks if `value` is classified as a `Symbol` primitive or object. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a symbol, else `false`. - * @example - * - * _.isSymbol(Symbol.iterator); - * // => true - * - * _.isSymbol('abc'); - * // => false - */ -function isSymbol$8(value) { - return typeof value == 'symbol' || - (isObjectLike$7(value) && baseGetTag$5(value) == symbolTag$3); -} - -var isSymbol_1$1 = isSymbol$8; - -var isArray$b = isArray_1, - isSymbol$7 = isSymbol_1$1; - -/** Used to match property names within property paths. */ -var reIsDeepProp = /\.|\[(?:[^[\]]*|(["'])(?:(?!\1)[^\\]|\\.)*?\1)\]/, - reIsPlainProp = /^\w*$/; - -/** - * Checks if `value` is a property name and not a property path. - * - * @private - * @param {*} value The value to check. - * @param {Object} [object] The object to query keys on. - * @returns {boolean} Returns `true` if `value` is a property name, else `false`. - */ -function isKey$3(value, object) { - if (isArray$b(value)) { - return false; - } - var type = typeof value; - if (type == 'number' || type == 'symbol' || type == 'boolean' || - value == null || isSymbol$7(value)) { - return true; - } - return reIsPlainProp.test(value) || !reIsDeepProp.test(value) || - (object != null && value in Object(object)); -} - -var _isKey = isKey$3; - -var MapCache = _MapCache; - -/** Error message constants. */ -var FUNC_ERROR_TEXT$4 = 'Expected a function'; - -/** - * Creates a function that memoizes the result of `func`. If `resolver` is - * provided, it determines the cache key for storing the result based on the - * arguments provided to the memoized function. By default, the first argument - * provided to the memoized function is used as the map cache key. The `func` - * is invoked with the `this` binding of the memoized function. - * - * **Note:** The cache is exposed as the `cache` property on the memoized - * function. Its creation may be customized by replacing the `_.memoize.Cache` - * constructor with one whose instances implement the - * [`Map`](http://ecma-international.org/ecma-262/7.0/#sec-properties-of-the-map-prototype-object) - * method interface of `clear`, `delete`, `get`, `has`, and `set`. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Function - * @param {Function} func The function to have its output memoized. - * @param {Function} [resolver] The function to resolve the cache key. - * @returns {Function} Returns the new memoized function. - * @example - * - * var object = { 'a': 1, 'b': 2 }; - * var other = { 'c': 3, 'd': 4 }; - * - * var values = _.memoize(_.values); - * values(object); - * // => [1, 2] - * - * values(other); - * // => [3, 4] - * - * object.a = 2; - * values(object); - * // => [1, 2] - * - * // Modify the result cache. - * values.cache.set(object, ['a', 'b']); - * values(object); - * // => ['a', 'b'] - * - * // Replace `_.memoize.Cache`. - * _.memoize.Cache = WeakMap; - */ -function memoize$2(func, resolver) { - if (typeof func != 'function' || (resolver != null && typeof resolver != 'function')) { - throw new TypeError(FUNC_ERROR_TEXT$4); - } - var memoized = function() { - var args = arguments, - key = resolver ? resolver.apply(this, args) : args[0], - cache = memoized.cache; - - if (cache.has(key)) { - return cache.get(key); - } - var result = func.apply(this, args); - memoized.cache = cache.set(key, result) || cache; - return result; - }; - memoized.cache = new (memoize$2.Cache || MapCache); - return memoized; -} - -// Expose `MapCache`. -memoize$2.Cache = MapCache; - -var memoize_1 = memoize$2; - -var memoize$1 = memoize_1; - -/** Used as the maximum memoize cache size. */ -var MAX_MEMOIZE_SIZE = 500; - -/** - * A specialized version of `_.memoize` which clears the memoized function's - * cache when it exceeds `MAX_MEMOIZE_SIZE`. - * - * @private - * @param {Function} func The function to have its output memoized. - * @returns {Function} Returns the new memoized function. - */ -function memoizeCapped$1(func) { - var result = memoize$1(func, function(key) { - if (cache.size === MAX_MEMOIZE_SIZE) { - cache.clear(); - } - return key; - }); - - var cache = result.cache; - return result; -} - -var _memoizeCapped = memoizeCapped$1; - -var memoizeCapped = _memoizeCapped; - -/** Used to match property names within property paths. */ -var rePropName = /[^.[\]]+|\[(?:(-?\d+(?:\.\d+)?)|(["'])((?:(?!\2)[^\\]|\\.)*?)\2)\]|(?=(?:\.|\[\])(?:\.|\[\]|$))/g; - -/** Used to match backslashes in property paths. */ -var reEscapeChar = /\\(\\)?/g; - -/** - * Converts `string` to a property path array. - * - * @private - * @param {string} string The string to convert. - * @returns {Array} Returns the property path array. - */ -var stringToPath$1 = memoizeCapped(function(string) { - var result = []; - if (string.charCodeAt(0) === 46 /* . */) { - result.push(''); - } - string.replace(rePropName, function(match, number, quote, subString) { - result.push(quote ? subString.replace(reEscapeChar, '$1') : (number || match)); - }); - return result; -}); - -var _stringToPath = stringToPath$1; - -/** - * A specialized version of `_.map` for arrays without support for iteratee - * shorthands. - * - * @private - * @param {Array} [array] The array to iterate over. - * @param {Function} iteratee The function invoked per iteration. - * @returns {Array} Returns the new mapped array. - */ - -function arrayMap$4(array, iteratee) { - var index = -1, - length = array == null ? 0 : array.length, - result = Array(length); - - while (++index < length) { - result[index] = iteratee(array[index], index, array); - } - return result; -} - -var _arrayMap = arrayMap$4; - -var Symbol$5 = _Symbol$1, - arrayMap$3 = _arrayMap, - isArray$a = isArray_1, - isSymbol$6 = isSymbol_1$1; - -/** Used as references for various `Number` constants. */ -var INFINITY$3 = 1 / 0; - -/** Used to convert symbols to primitives and strings. */ -var symbolProto$1 = Symbol$5 ? Symbol$5.prototype : undefined, - symbolToString = symbolProto$1 ? symbolProto$1.toString : undefined; - -/** - * The base implementation of `_.toString` which doesn't convert nullish - * values to empty strings. - * - * @private - * @param {*} value The value to process. - * @returns {string} Returns the string. - */ -function baseToString$1(value) { - // Exit early for strings to avoid a performance hit in some environments. - if (typeof value == 'string') { - return value; - } - if (isArray$a(value)) { - // Recursively convert values (susceptible to call stack limits). - return arrayMap$3(value, baseToString$1) + ''; - } - if (isSymbol$6(value)) { - return symbolToString ? symbolToString.call(value) : ''; - } - var result = (value + ''); - return (result == '0' && (1 / value) == -INFINITY$3) ? '-0' : result; -} - -var _baseToString = baseToString$1; - -var baseToString = _baseToString; - -/** - * Converts `value` to a string. An empty string is returned for `null` - * and `undefined` values. The sign of `-0` is preserved. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to convert. - * @returns {string} Returns the converted string. - * @example - * - * _.toString(null); - * // => '' - * - * _.toString(-0); - * // => '-0' - * - * _.toString([1, 2, 3]); - * // => '1,2,3' - */ -function toString$2(value) { - return value == null ? '' : baseToString(value); -} - -var toString_1 = toString$2; - -var isArray$9 = isArray_1, - isKey$2 = _isKey, - stringToPath = _stringToPath, - toString$1 = toString_1; - -/** - * Casts `value` to a path array if it's not one. - * - * @private - * @param {*} value The value to inspect. - * @param {Object} [object] The object to query keys on. - * @returns {Array} Returns the cast property path array. - */ -function castPath$4(value, object) { - if (isArray$9(value)) { - return value; - } - return isKey$2(value, object) ? [value] : stringToPath(toString$1(value)); -} - -var _castPath = castPath$4; - -var isSymbol$5 = isSymbol_1$1; - -/** Used as references for various `Number` constants. */ -var INFINITY$2 = 1 / 0; - -/** - * Converts `value` to a string key if it's not a string or symbol. - * - * @private - * @param {*} value The value to inspect. - * @returns {string|symbol} Returns the key. - */ -function toKey$5(value) { - if (typeof value == 'string' || isSymbol$5(value)) { - return value; - } - var result = (value + ''); - return (result == '0' && (1 / value) == -INFINITY$2) ? '-0' : result; -} - -var _toKey = toKey$5; - -var castPath$3 = _castPath, - toKey$4 = _toKey; - -/** - * The base implementation of `_.get` without support for default values. - * - * @private - * @param {Object} object The object to query. - * @param {Array|string} path The path of the property to get. - * @returns {*} Returns the resolved value. - */ -function baseGet$4(object, path) { - path = castPath$3(path, object); - - var index = 0, - length = path.length; - - while (object != null && index < length) { - object = object[toKey$4(path[index++])]; - } - return (index && index == length) ? object : undefined; -} - -var _baseGet = baseGet$4; - -var baseGet$3 = _baseGet; - -/** - * Gets the value at `path` of `object`. If the resolved value is - * `undefined`, the `defaultValue` is returned in its place. - * - * @static - * @memberOf _ - * @since 3.7.0 - * @category Object - * @param {Object} object The object to query. - * @param {Array|string} path The path of the property to get. - * @param {*} [defaultValue] The value returned for `undefined` resolved values. - * @returns {*} Returns the resolved value. - * @example - * - * var object = { 'a': [{ 'b': { 'c': 3 } }] }; - * - * _.get(object, 'a[0].b.c'); - * // => 3 - * - * _.get(object, ['a', '0', 'b', 'c']); - * // => 3 - * - * _.get(object, 'a.b.c', 'default'); - * // => 'default' - */ -function get$1(object, path, defaultValue) { - var result = object == null ? undefined : baseGet$3(object, path); - return result === undefined ? defaultValue : result; -} - -var get_1 = get$1; - -/** - * The base implementation of `_.hasIn` without support for deep paths. - * - * @private - * @param {Object} [object] The object to query. - * @param {Array|string} key The key to check. - * @returns {boolean} Returns `true` if `key` exists, else `false`. - */ - -function baseHasIn$1(object, key) { - return object != null && key in Object(object); -} - -var _baseHasIn = baseHasIn$1; - -var castPath$2 = _castPath, - isArguments$1 = isArguments_1, - isArray$8 = isArray_1, - isIndex$1 = _isIndex, - isLength = isLength_1, - toKey$3 = _toKey; - -/** - * Checks if `path` exists on `object`. - * - * @private - * @param {Object} object The object to query. - * @param {Array|string} path The path to check. - * @param {Function} hasFunc The function to check properties. - * @returns {boolean} Returns `true` if `path` exists, else `false`. - */ -function hasPath$1(object, path, hasFunc) { - path = castPath$2(path, object); - - var index = -1, - length = path.length, - result = false; - - while (++index < length) { - var key = toKey$3(path[index]); - if (!(result = object != null && hasFunc(object, key))) { - break; - } - object = object[key]; - } - if (result || ++index != length) { - return result; - } - length = object == null ? 0 : object.length; - return !!length && isLength(length) && isIndex$1(key, length) && - (isArray$8(object) || isArguments$1(object)); -} - -var _hasPath = hasPath$1; - -var baseHasIn = _baseHasIn, - hasPath = _hasPath; - -/** - * Checks if `path` is a direct or inherited property of `object`. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Object - * @param {Object} object The object to query. - * @param {Array|string} path The path to check. - * @returns {boolean} Returns `true` if `path` exists, else `false`. - * @example - * - * var object = _.create({ 'a': _.create({ 'b': 2 }) }); - * - * _.hasIn(object, 'a'); - * // => true - * - * _.hasIn(object, 'a.b'); - * // => true - * - * _.hasIn(object, ['a', 'b']); - * // => true - * - * _.hasIn(object, 'b'); - * // => false - */ -function hasIn$1(object, path) { - return object != null && hasPath(object, path, baseHasIn); -} - -var hasIn_1 = hasIn$1; - -var baseIsEqual$1 = _baseIsEqual, - get = get_1, - hasIn = hasIn_1, - isKey$1 = _isKey, - isStrictComparable = _isStrictComparable, - matchesStrictComparable = _matchesStrictComparable, - toKey$2 = _toKey; - -/** Used to compose bitmasks for value comparisons. */ -var COMPARE_PARTIAL_FLAG = 1, - COMPARE_UNORDERED_FLAG = 2; - -/** - * The base implementation of `_.matchesProperty` which doesn't clone `srcValue`. - * - * @private - * @param {string} path The path of the property to get. - * @param {*} srcValue The value to match. - * @returns {Function} Returns the new spec function. - */ -function baseMatchesProperty$1(path, srcValue) { - if (isKey$1(path) && isStrictComparable(srcValue)) { - return matchesStrictComparable(toKey$2(path), srcValue); - } - return function(object) { - var objValue = get(object, path); - return (objValue === undefined && objValue === srcValue) - ? hasIn(object, path) - : baseIsEqual$1(srcValue, objValue, COMPARE_PARTIAL_FLAG | COMPARE_UNORDERED_FLAG); - }; -} - -var _baseMatchesProperty = baseMatchesProperty$1; - -/** - * This method returns the first argument it receives. - * - * @static - * @since 0.1.0 - * @memberOf _ - * @category Util - * @param {*} value Any value. - * @returns {*} Returns `value`. - * @example - * - * var object = { 'a': 1 }; - * - * console.log(_.identity(object) === object); - * // => true - */ - -function identity$a(value) { - return value; -} - -var identity_1 = identity$a; - -/** - * The base implementation of `_.property` without support for deep paths. - * - * @private - * @param {string} key The key of the property to get. - * @returns {Function} Returns the new accessor function. - */ - -function baseProperty$1(key) { - return function(object) { - return object == null ? undefined : object[key]; - }; -} - -var _baseProperty = baseProperty$1; - -var baseGet$2 = _baseGet; - -/** - * A specialized version of `baseProperty` which supports deep paths. - * - * @private - * @param {Array|string} path The path of the property to get. - * @returns {Function} Returns the new accessor function. - */ -function basePropertyDeep$1(path) { - return function(object) { - return baseGet$2(object, path); - }; -} - -var _basePropertyDeep = basePropertyDeep$1; - -var baseProperty = _baseProperty, - basePropertyDeep = _basePropertyDeep, - isKey = _isKey, - toKey$1 = _toKey; - -/** - * Creates a function that returns the value at `path` of a given object. - * - * @static - * @memberOf _ - * @since 2.4.0 - * @category Util - * @param {Array|string} path The path of the property to get. - * @returns {Function} Returns the new accessor function. - * @example - * - * var objects = [ - * { 'a': { 'b': 2 } }, - * { 'a': { 'b': 1 } } - * ]; - * - * _.map(objects, _.property('a.b')); - * // => [2, 1] - * - * _.map(_.sortBy(objects, _.property(['a', 'b'])), 'a.b'); - * // => [1, 2] - */ -function property$1(path) { - return isKey(path) ? baseProperty(toKey$1(path)) : basePropertyDeep(path); -} - -var property_1 = property$1; - -var baseMatches = _baseMatches, - baseMatchesProperty = _baseMatchesProperty, - identity$9 = identity_1, - isArray$7 = isArray_1, - property = property_1; - -/** - * The base implementation of `_.iteratee`. - * - * @private - * @param {*} [value=_.identity] The value to convert to an iteratee. - * @returns {Function} Returns the iteratee. - */ -function baseIteratee$b(value) { - // Don't store the `typeof` result in a variable to avoid a JIT bug in Safari 9. - // See https://bugs.webkit.org/show_bug.cgi?id=156034 for more details. - if (typeof value == 'function') { - return value; - } - if (value == null) { - return identity$9; - } - if (typeof value == 'object') { - return isArray$7(value) - ? baseMatchesProperty(value[0], value[1]) - : baseMatches(value); - } - return property(value); -} - -var _baseIteratee = baseIteratee$b; - -/** - * The base implementation of `_.findIndex` and `_.findLastIndex` without - * support for iteratee shorthands. - * - * @private - * @param {Array} array The array to inspect. - * @param {Function} predicate The function invoked per iteration. - * @param {number} fromIndex The index to search from. - * @param {boolean} [fromRight] Specify iterating from right to left. - * @returns {number} Returns the index of the matched value, else `-1`. - */ - -function baseFindIndex$2(array, predicate, fromIndex, fromRight) { - var length = array.length, - index = fromIndex + (fromRight ? 1 : -1); - - while ((fromRight ? index-- : ++index < length)) { - if (predicate(array[index], index, array)) { - return index; - } - } - return -1; -} - -var _baseFindIndex = baseFindIndex$2; - -/** - * The base implementation of `_.isNaN` without support for number objects. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is `NaN`, else `false`. - */ - -function baseIsNaN$1(value) { - return value !== value; -} - -var _baseIsNaN = baseIsNaN$1; - -/** - * A specialized version of `_.indexOf` which performs strict equality - * comparisons of values, i.e. `===`. - * - * @private - * @param {Array} array The array to inspect. - * @param {*} value The value to search for. - * @param {number} fromIndex The index to search from. - * @returns {number} Returns the index of the matched value, else `-1`. - */ - -function strictIndexOf$1(array, value, fromIndex) { - var index = fromIndex - 1, - length = array.length; - - while (++index < length) { - if (array[index] === value) { - return index; - } - } - return -1; -} - -var _strictIndexOf = strictIndexOf$1; - -var baseFindIndex$1 = _baseFindIndex, - baseIsNaN = _baseIsNaN, - strictIndexOf = _strictIndexOf; - -/** - * The base implementation of `_.indexOf` without `fromIndex` bounds checks. - * - * @private - * @param {Array} array The array to inspect. - * @param {*} value The value to search for. - * @param {number} fromIndex The index to search from. - * @returns {number} Returns the index of the matched value, else `-1`. - */ -function baseIndexOf$1(array, value, fromIndex) { - return value === value - ? strictIndexOf(array, value, fromIndex) - : baseFindIndex$1(array, baseIsNaN, fromIndex); -} - -var _baseIndexOf = baseIndexOf$1; - -var baseIndexOf = _baseIndexOf; - -/** - * A specialized version of `_.includes` for arrays without support for - * specifying an index to search from. - * - * @private - * @param {Array} [array] The array to inspect. - * @param {*} target The value to search for. - * @returns {boolean} Returns `true` if `target` is found, else `false`. - */ -function arrayIncludes$1(array, value) { - var length = array == null ? 0 : array.length; - return !!length && baseIndexOf(array, value, 0) > -1; -} - -var _arrayIncludes = arrayIncludes$1; - -/** - * This function is like `arrayIncludes` except that it accepts a comparator. - * - * @private - * @param {Array} [array] The array to inspect. - * @param {*} target The value to search for. - * @param {Function} comparator The comparator invoked per element. - * @returns {boolean} Returns `true` if `target` is found, else `false`. - */ - -function arrayIncludesWith$1(array, value, comparator) { - var index = -1, - length = array == null ? 0 : array.length; - - while (++index < length) { - if (comparator(value, array[index])) { - return true; - } - } - return false; -} - -var _arrayIncludesWith = arrayIncludesWith$1; - -/** - * This method returns `undefined`. - * - * @static - * @memberOf _ - * @since 2.3.0 - * @category Util - * @example - * - * _.times(2, _.noop); - * // => [undefined, undefined] - */ - -function noop$3() { - // No operation performed. -} - -var noop_1 = noop$3; - -var Set$1 = _Set, - noop$2 = noop_1, - setToArray$1 = _setToArray; - -/** Used as references for various `Number` constants. */ -var INFINITY$1 = 1 / 0; - -/** - * Creates a set object of `values`. - * - * @private - * @param {Array} values The values to add to the set. - * @returns {Object} Returns the new set. - */ -var createSet$1 = !(Set$1 && (1 / setToArray$1(new Set$1([,-0]))[1]) == INFINITY$1) ? noop$2 : function(values) { - return new Set$1(values); -}; - -var _createSet = createSet$1; - -var SetCache = _SetCache, - arrayIncludes = _arrayIncludes, - arrayIncludesWith = _arrayIncludesWith, - cacheHas = _cacheHas, - createSet = _createSet, - setToArray = _setToArray; - -/** Used as the size to enable large array optimizations. */ -var LARGE_ARRAY_SIZE = 200; - -/** - * The base implementation of `_.uniqBy` without support for iteratee shorthands. - * - * @private - * @param {Array} array The array to inspect. - * @param {Function} [iteratee] The iteratee invoked per element. - * @param {Function} [comparator] The comparator invoked per element. - * @returns {Array} Returns the new duplicate free array. - */ -function baseUniq$1(array, iteratee, comparator) { - var index = -1, - includes = arrayIncludes, - length = array.length, - isCommon = true, - result = [], - seen = result; - - if (comparator) { - isCommon = false; - includes = arrayIncludesWith; - } - else if (length >= LARGE_ARRAY_SIZE) { - var set = iteratee ? null : createSet(array); - if (set) { - return setToArray(set); - } - isCommon = false; - includes = cacheHas; - seen = new SetCache; - } - else { - seen = iteratee ? [] : result; - } - outer: - while (++index < length) { - var value = array[index], - computed = iteratee ? iteratee(value) : value; - - value = (comparator || value !== 0) ? value : 0; - if (isCommon && computed === computed) { - var seenIndex = seen.length; - while (seenIndex--) { - if (seen[seenIndex] === computed) { - continue outer; - } - } - if (iteratee) { - seen.push(computed); - } - result.push(value); - } - else if (!includes(seen, computed, comparator)) { - if (seen !== result) { - seen.push(computed); - } - result.push(value); - } - } - return result; -} - -var _baseUniq = baseUniq$1; - -var baseIteratee$a = _baseIteratee, - baseUniq = _baseUniq; - -/** - * This method is like `_.uniq` except that it accepts `iteratee` which is - * invoked for each element in `array` to generate the criterion by which - * uniqueness is computed. The order of result values is determined by the - * order they occur in the array. The iteratee is invoked with one argument: - * (value). - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Array - * @param {Array} array The array to inspect. - * @param {Function} [iteratee=_.identity] The iteratee invoked per element. - * @returns {Array} Returns the new duplicate free array. - * @example - * - * _.uniqBy([2.1, 1.2, 2.3], Math.floor); - * // => [2.1, 1.2] - * - * // The `_.property` iteratee shorthand. - * _.uniqBy([{ 'x': 1 }, { 'x': 2 }, { 'x': 1 }], 'x'); - * // => [{ 'x': 1 }, { 'x': 2 }] - */ -function uniqBy(array, iteratee) { - return (array && array.length) ? baseUniq(array, baseIteratee$a(iteratee)) : []; -} - -var uniqBy_1 = uniqBy; - -/** - * The base implementation of `_.slice` without an iteratee call guard. - * - * @private - * @param {Array} array The array to slice. - * @param {number} [start=0] The start position. - * @param {number} [end=array.length] The end position. - * @returns {Array} Returns the slice of `array`. - */ - -function baseSlice$2(array, start, end) { - var index = -1, - length = array.length; - - if (start < 0) { - start = -start > length ? 0 : (length + start); - } - end = end > length ? length : end; - if (end < 0) { - end += length; - } - length = start > end ? 0 : ((end - start) >>> 0); - start >>>= 0; - - var result = Array(length); - while (++index < length) { - result[index] = array[index + start]; - } - return result; -} - -var _baseSlice = baseSlice$2; - -var baseSlice$1 = _baseSlice; - -/** - * Casts `array` to a slice if it's needed. - * - * @private - * @param {Array} array The array to inspect. - * @param {number} start The start position. - * @param {number} [end=array.length] The end position. - * @returns {Array} Returns the cast slice. - */ -function castSlice$1(array, start, end) { - var length = array.length; - end = end === undefined ? length : end; - return (!start && end >= length) ? array : baseSlice$1(array, start, end); -} - -var _castSlice = castSlice$1; - -/** Used to compose unicode character classes. */ - -var rsAstralRange$1 = '\\ud800-\\udfff', - rsComboMarksRange$1 = '\\u0300-\\u036f', - reComboHalfMarksRange$1 = '\\ufe20-\\ufe2f', - rsComboSymbolsRange$1 = '\\u20d0-\\u20ff', - rsComboRange$1 = rsComboMarksRange$1 + reComboHalfMarksRange$1 + rsComboSymbolsRange$1, - rsVarRange$1 = '\\ufe0e\\ufe0f'; - -/** Used to compose unicode capture groups. */ -var rsZWJ$1 = '\\u200d'; - -/** Used to detect strings with [zero-width joiners or code points from the astral planes](http://eev.ee/blog/2015/09/12/dark-corners-of-unicode/). */ -var reHasUnicode = RegExp('[' + rsZWJ$1 + rsAstralRange$1 + rsComboRange$1 + rsVarRange$1 + ']'); - -/** - * Checks if `string` contains Unicode symbols. - * - * @private - * @param {string} string The string to inspect. - * @returns {boolean} Returns `true` if a symbol is found, else `false`. - */ -function hasUnicode$2(string) { - return reHasUnicode.test(string); -} - -var _hasUnicode = hasUnicode$2; - -/** - * Converts an ASCII `string` to an array. - * - * @private - * @param {string} string The string to convert. - * @returns {Array} Returns the converted array. - */ - -function asciiToArray$1(string) { - return string.split(''); -} - -var _asciiToArray = asciiToArray$1; - -/** Used to compose unicode character classes. */ - -var rsAstralRange = '\\ud800-\\udfff', - rsComboMarksRange = '\\u0300-\\u036f', - reComboHalfMarksRange = '\\ufe20-\\ufe2f', - rsComboSymbolsRange = '\\u20d0-\\u20ff', - rsComboRange = rsComboMarksRange + reComboHalfMarksRange + rsComboSymbolsRange, - rsVarRange = '\\ufe0e\\ufe0f'; - -/** Used to compose unicode capture groups. */ -var rsAstral = '[' + rsAstralRange + ']', - rsCombo = '[' + rsComboRange + ']', - rsFitz = '\\ud83c[\\udffb-\\udfff]', - rsModifier = '(?:' + rsCombo + '|' + rsFitz + ')', - rsNonAstral = '[^' + rsAstralRange + ']', - rsRegional = '(?:\\ud83c[\\udde6-\\uddff]){2}', - rsSurrPair = '[\\ud800-\\udbff][\\udc00-\\udfff]', - rsZWJ = '\\u200d'; - -/** Used to compose unicode regexes. */ -var reOptMod = rsModifier + '?', - rsOptVar = '[' + rsVarRange + ']?', - rsOptJoin = '(?:' + rsZWJ + '(?:' + [rsNonAstral, rsRegional, rsSurrPair].join('|') + ')' + rsOptVar + reOptMod + ')*', - rsSeq = rsOptVar + reOptMod + rsOptJoin, - rsSymbol = '(?:' + [rsNonAstral + rsCombo + '?', rsCombo, rsRegional, rsSurrPair, rsAstral].join('|') + ')'; - -/** Used to match [string symbols](https://mathiasbynens.be/notes/javascript-unicode). */ -var reUnicode = RegExp(rsFitz + '(?=' + rsFitz + ')|' + rsSymbol + rsSeq, 'g'); - -/** - * Converts a Unicode `string` to an array. - * - * @private - * @param {string} string The string to convert. - * @returns {Array} Returns the converted array. - */ -function unicodeToArray$1(string) { - return string.match(reUnicode) || []; -} - -var _unicodeToArray = unicodeToArray$1; - -var asciiToArray = _asciiToArray, - hasUnicode$1 = _hasUnicode, - unicodeToArray = _unicodeToArray; - -/** - * Converts `string` to an array. - * - * @private - * @param {string} string The string to convert. - * @returns {Array} Returns the converted array. - */ -function stringToArray$1(string) { - return hasUnicode$1(string) - ? unicodeToArray(string) - : asciiToArray(string); -} - -var _stringToArray = stringToArray$1; - -var castSlice = _castSlice, - hasUnicode = _hasUnicode, - stringToArray = _stringToArray, - toString = toString_1; - -/** - * Creates a function like `_.lowerFirst`. - * - * @private - * @param {string} methodName The name of the `String` case method to use. - * @returns {Function} Returns the new case function. - */ -function createCaseFirst$1(methodName) { - return function(string) { - string = toString(string); - - var strSymbols = hasUnicode(string) - ? stringToArray(string) - : undefined; - - var chr = strSymbols - ? strSymbols[0] - : string.charAt(0); - - var trailing = strSymbols - ? castSlice(strSymbols, 1).join('') - : string.slice(1); - - return chr[methodName]() + trailing; - }; -} - -var _createCaseFirst = createCaseFirst$1; - -var createCaseFirst = _createCaseFirst; - -/** - * Converts the first character of `string` to upper case. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category String - * @param {string} [string=''] The string to convert. - * @returns {string} Returns the converted string. - * @example - * - * _.upperFirst('fred'); - * // => 'Fred' - * - * _.upperFirst('FRED'); - * // => 'FRED' - */ -var upperFirst = createCaseFirst('toUpperCase'); - -var upperFirst_1 = upperFirst; - -const pi$1 = Math.PI, - tau$1 = 2 * pi$1, - epsilon$1 = 1e-6, - tauEpsilon = tau$1 - epsilon$1; - -function Path() { - this._x0 = this._y0 = // start of current subpath - this._x1 = this._y1 = null; // end of current subpath - this._ = ""; -} - -function path() { - return new Path; -} - -Path.prototype = path.prototype = { - constructor: Path, - moveTo: function(x, y) { - this._ += "M" + (this._x0 = this._x1 = +x) + "," + (this._y0 = this._y1 = +y); - }, - closePath: function() { - if (this._x1 !== null) { - this._x1 = this._x0, this._y1 = this._y0; - this._ += "Z"; - } - }, - lineTo: function(x, y) { - this._ += "L" + (this._x1 = +x) + "," + (this._y1 = +y); - }, - quadraticCurveTo: function(x1, y1, x, y) { - this._ += "Q" + (+x1) + "," + (+y1) + "," + (this._x1 = +x) + "," + (this._y1 = +y); - }, - bezierCurveTo: function(x1, y1, x2, y2, x, y) { - this._ += "C" + (+x1) + "," + (+y1) + "," + (+x2) + "," + (+y2) + "," + (this._x1 = +x) + "," + (this._y1 = +y); - }, - arcTo: function(x1, y1, x2, y2, r) { - x1 = +x1, y1 = +y1, x2 = +x2, y2 = +y2, r = +r; - var x0 = this._x1, - y0 = this._y1, - x21 = x2 - x1, - y21 = y2 - y1, - x01 = x0 - x1, - y01 = y0 - y1, - l01_2 = x01 * x01 + y01 * y01; - - // Is the radius negative? Error. - if (r < 0) throw new Error("negative radius: " + r); - - // Is this path empty? Move to (x1,y1). - if (this._x1 === null) { - this._ += "M" + (this._x1 = x1) + "," + (this._y1 = y1); - } - - // Or, is (x1,y1) coincident with (x0,y0)? Do nothing. - else if (!(l01_2 > epsilon$1)); - - // Or, are (x0,y0), (x1,y1) and (x2,y2) collinear? - // Equivalently, is (x1,y1) coincident with (x2,y2)? - // Or, is the radius zero? Line to (x1,y1). - else if (!(Math.abs(y01 * x21 - y21 * x01) > epsilon$1) || !r) { - this._ += "L" + (this._x1 = x1) + "," + (this._y1 = y1); - } - - // Otherwise, draw an arc! - else { - var x20 = x2 - x0, - y20 = y2 - y0, - l21_2 = x21 * x21 + y21 * y21, - l20_2 = x20 * x20 + y20 * y20, - l21 = Math.sqrt(l21_2), - l01 = Math.sqrt(l01_2), - l = r * Math.tan((pi$1 - Math.acos((l21_2 + l01_2 - l20_2) / (2 * l21 * l01))) / 2), - t01 = l / l01, - t21 = l / l21; - - // If the start tangent is not coincident with (x0,y0), line to. - if (Math.abs(t01 - 1) > epsilon$1) { - this._ += "L" + (x1 + t01 * x01) + "," + (y1 + t01 * y01); - } - - this._ += "A" + r + "," + r + ",0,0," + (+(y01 * x20 > x01 * y20)) + "," + (this._x1 = x1 + t21 * x21) + "," + (this._y1 = y1 + t21 * y21); - } - }, - arc: function(x, y, r, a0, a1, ccw) { - x = +x, y = +y, r = +r, ccw = !!ccw; - var dx = r * Math.cos(a0), - dy = r * Math.sin(a0), - x0 = x + dx, - y0 = y + dy, - cw = 1 ^ ccw, - da = ccw ? a0 - a1 : a1 - a0; - - // Is the radius negative? Error. - if (r < 0) throw new Error("negative radius: " + r); - - // Is this path empty? Move to (x0,y0). - if (this._x1 === null) { - this._ += "M" + x0 + "," + y0; - } - - // Or, is (x0,y0) not coincident with the previous point? Line to (x0,y0). - else if (Math.abs(this._x1 - x0) > epsilon$1 || Math.abs(this._y1 - y0) > epsilon$1) { - this._ += "L" + x0 + "," + y0; - } - - // Is this arc empty? We’re done. - if (!r) return; - - // Does the angle go the wrong way? Flip the direction. - if (da < 0) da = da % tau$1 + tau$1; - - // Is this a complete circle? Draw two arcs to complete the circle. - if (da > tauEpsilon) { - this._ += "A" + r + "," + r + ",0,1," + cw + "," + (x - dx) + "," + (y - dy) + "A" + r + "," + r + ",0,1," + cw + "," + (this._x1 = x0) + "," + (this._y1 = y0); - } - - // Is this arc non-empty? Draw an arc! - else if (da > epsilon$1) { - this._ += "A" + r + "," + r + ",0," + (+(da >= pi$1)) + "," + cw + "," + (this._x1 = x + r * Math.cos(a1)) + "," + (this._y1 = y + r * Math.sin(a1)); - } - }, - rect: function(x, y, w, h) { - this._ += "M" + (this._x0 = this._x1 = +x) + "," + (this._y0 = this._y1 = +y) + "h" + (+w) + "v" + (+h) + "h" + (-w) + "Z"; - }, - toString: function() { - return this._; - } -}; - -function constant$3(x) { - return function constant() { - return x; - }; -} - -var epsilon = 1e-12; -var pi = Math.PI; -var tau = 2 * pi; - -function array(x) { - return typeof x === "object" && "length" in x - ? x // Array, TypedArray, NodeList, array-like - : Array.from(x); // Map, Set, iterable, string, or anything else -} - -function Linear(context) { - this._context = context; -} - -Linear.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._point = 0; - }, - lineEnd: function() { - if (this._line || (this._line !== 0 && this._point === 1)) this._context.closePath(); - this._line = 1 - this._line; - }, - point: function(x, y) { - x = +x, y = +y; - switch (this._point) { - case 0: this._point = 1; this._line ? this._context.lineTo(x, y) : this._context.moveTo(x, y); break; - case 1: this._point = 2; // proceed - default: this._context.lineTo(x, y); break; - } - } -}; - -function curveLinear(context) { - return new Linear(context); -} - -function x$1(p) { - return p[0]; -} - -function y$1(p) { - return p[1]; -} - -function shapeLine(x, y) { - var defined = constant$3(true), - context = null, - curve = curveLinear, - output = null; - - x = typeof x === "function" ? x : (x === undefined) ? x$1 : constant$3(x); - y = typeof y === "function" ? y : (y === undefined) ? y$1 : constant$3(y); - - function line(data) { - var i, - n = (data = array(data)).length, - d, - defined0 = false, - buffer; - - if (context == null) output = curve(buffer = path()); - - for (i = 0; i <= n; ++i) { - if (!(i < n && defined(d = data[i], i, data)) === defined0) { - if (defined0 = !defined0) output.lineStart(); - else output.lineEnd(); - } - if (defined0) output.point(+x(d, i, data), +y(d, i, data)); - } - - if (buffer) return output = null, buffer + "" || null; - } - - line.x = function(_) { - return arguments.length ? (x = typeof _ === "function" ? _ : constant$3(+_), line) : x; - }; - - line.y = function(_) { - return arguments.length ? (y = typeof _ === "function" ? _ : constant$3(+_), line) : y; - }; - - line.defined = function(_) { - return arguments.length ? (defined = typeof _ === "function" ? _ : constant$3(!!_), line) : defined; - }; - - line.curve = function(_) { - return arguments.length ? (curve = _, context != null && (output = curve(context)), line) : curve; - }; - - line.context = function(_) { - return arguments.length ? (_ == null ? context = output = null : output = curve(context = _), line) : context; - }; - - return line; -} - -function shapeArea(x0, y0, y1) { - var x1 = null, - defined = constant$3(true), - context = null, - curve = curveLinear, - output = null; - - x0 = typeof x0 === "function" ? x0 : (x0 === undefined) ? x$1 : constant$3(+x0); - y0 = typeof y0 === "function" ? y0 : (y0 === undefined) ? constant$3(0) : constant$3(+y0); - y1 = typeof y1 === "function" ? y1 : (y1 === undefined) ? y$1 : constant$3(+y1); - - function area(data) { - var i, - j, - k, - n = (data = array(data)).length, - d, - defined0 = false, - buffer, - x0z = new Array(n), - y0z = new Array(n); - - if (context == null) output = curve(buffer = path()); - - for (i = 0; i <= n; ++i) { - if (!(i < n && defined(d = data[i], i, data)) === defined0) { - if (defined0 = !defined0) { - j = i; - output.areaStart(); - output.lineStart(); - } else { - output.lineEnd(); - output.lineStart(); - for (k = i - 1; k >= j; --k) { - output.point(x0z[k], y0z[k]); - } - output.lineEnd(); - output.areaEnd(); - } - } - if (defined0) { - x0z[i] = +x0(d, i, data), y0z[i] = +y0(d, i, data); - output.point(x1 ? +x1(d, i, data) : x0z[i], y1 ? +y1(d, i, data) : y0z[i]); - } - } - - if (buffer) return output = null, buffer + "" || null; - } - - function arealine() { - return shapeLine().defined(defined).curve(curve).context(context); - } - - area.x = function(_) { - return arguments.length ? (x0 = typeof _ === "function" ? _ : constant$3(+_), x1 = null, area) : x0; - }; - - area.x0 = function(_) { - return arguments.length ? (x0 = typeof _ === "function" ? _ : constant$3(+_), area) : x0; - }; - - area.x1 = function(_) { - return arguments.length ? (x1 = _ == null ? null : typeof _ === "function" ? _ : constant$3(+_), area) : x1; - }; - - area.y = function(_) { - return arguments.length ? (y0 = typeof _ === "function" ? _ : constant$3(+_), y1 = null, area) : y0; - }; - - area.y0 = function(_) { - return arguments.length ? (y0 = typeof _ === "function" ? _ : constant$3(+_), area) : y0; - }; - - area.y1 = function(_) { - return arguments.length ? (y1 = _ == null ? null : typeof _ === "function" ? _ : constant$3(+_), area) : y1; - }; - - area.lineX0 = - area.lineY0 = function() { - return arealine().x(x0).y(y0); - }; - - area.lineY1 = function() { - return arealine().x(x0).y(y1); - }; - - area.lineX1 = function() { - return arealine().x(x1).y(y0); - }; - - area.defined = function(_) { - return arguments.length ? (defined = typeof _ === "function" ? _ : constant$3(!!_), area) : defined; - }; - - area.curve = function(_) { - return arguments.length ? (curve = _, context != null && (output = curve(context)), area) : curve; - }; - - area.context = function(_) { - return arguments.length ? (_ == null ? context = output = null : output = curve(context = _), area) : context; - }; - - return area; -} - -var symbolCircle = { - draw: function(context, size) { - var r = Math.sqrt(size / pi); - context.moveTo(r, 0); - context.arc(0, 0, r, 0, tau); - } -}; - -var symbolCross = { - draw: function(context, size) { - var r = Math.sqrt(size / 5) / 2; - context.moveTo(-3 * r, -r); - context.lineTo(-r, -r); - context.lineTo(-r, -3 * r); - context.lineTo(r, -3 * r); - context.lineTo(r, -r); - context.lineTo(3 * r, -r); - context.lineTo(3 * r, r); - context.lineTo(r, r); - context.lineTo(r, 3 * r); - context.lineTo(-r, 3 * r); - context.lineTo(-r, r); - context.lineTo(-3 * r, r); - context.closePath(); - } -}; - -var tan30 = Math.sqrt(1 / 3), - tan30_2 = tan30 * 2; - -var symbolDiamond = { - draw: function(context, size) { - var y = Math.sqrt(size / tan30_2), - x = y * tan30; - context.moveTo(0, -y); - context.lineTo(x, 0); - context.lineTo(0, y); - context.lineTo(-x, 0); - context.closePath(); - } -}; - -var ka = 0.89081309152928522810, - kr = Math.sin(pi / 10) / Math.sin(7 * pi / 10), - kx = Math.sin(tau / 10) * kr, - ky = -Math.cos(tau / 10) * kr; - -var symbolStar = { - draw: function(context, size) { - var r = Math.sqrt(size * ka), - x = kx * r, - y = ky * r; - context.moveTo(0, -r); - context.lineTo(x, y); - for (var i = 1; i < 5; ++i) { - var a = tau * i / 5, - c = Math.cos(a), - s = Math.sin(a); - context.lineTo(s * r, -c * r); - context.lineTo(c * x - s * y, s * x + c * y); - } - context.closePath(); - } -}; - -var symbolSquare = { - draw: function(context, size) { - var w = Math.sqrt(size), - x = -w / 2; - context.rect(x, x, w, w); - } -}; - -var sqrt3 = Math.sqrt(3); - -var symbolTriangle = { - draw: function(context, size) { - var y = -Math.sqrt(size / (sqrt3 * 3)); - context.moveTo(0, y * 2); - context.lineTo(-sqrt3 * y, -y); - context.lineTo(sqrt3 * y, -y); - context.closePath(); - } -}; - -var c$1 = -0.5, - s = Math.sqrt(3) / 2, - k$1 = 1 / Math.sqrt(12), - a = (k$1 / 2 + 1) * 3; - -var symbolWye = { - draw: function(context, size) { - var r = Math.sqrt(size / a), - x0 = r / 2, - y0 = r * k$1, - x1 = x0, - y1 = r * k$1 + r, - x2 = -x1, - y2 = y1; - context.moveTo(x0, y0); - context.lineTo(x1, y1); - context.lineTo(x2, y2); - context.lineTo(c$1 * x0 - s * y0, s * x0 + c$1 * y0); - context.lineTo(c$1 * x1 - s * y1, s * x1 + c$1 * y1); - context.lineTo(c$1 * x2 - s * y2, s * x2 + c$1 * y2); - context.lineTo(c$1 * x0 + s * y0, c$1 * y0 - s * x0); - context.lineTo(c$1 * x1 + s * y1, c$1 * y1 - s * x1); - context.lineTo(c$1 * x2 + s * y2, c$1 * y2 - s * x2); - context.closePath(); - } -}; - -function shapeSymbol(type, size) { - var context = null; - type = typeof type === "function" ? type : constant$3(type || symbolCircle); - size = typeof size === "function" ? size : constant$3(size === undefined ? 64 : +size); - - function symbol() { - var buffer; - if (!context) context = buffer = path(); - type.apply(this, arguments).draw(context, +size.apply(this, arguments)); - if (buffer) return context = null, buffer + "" || null; - } - - symbol.type = function(_) { - return arguments.length ? (type = typeof _ === "function" ? _ : constant$3(_), symbol) : type; - }; - - symbol.size = function(_) { - return arguments.length ? (size = typeof _ === "function" ? _ : constant$3(+_), symbol) : size; - }; - - symbol.context = function(_) { - return arguments.length ? (context = _ == null ? null : _, symbol) : context; - }; - - return symbol; -} - -function noop$1() {} - -function point$4(that, x, y) { - that._context.bezierCurveTo( - (2 * that._x0 + that._x1) / 3, - (2 * that._y0 + that._y1) / 3, - (that._x0 + 2 * that._x1) / 3, - (that._y0 + 2 * that._y1) / 3, - (that._x0 + 4 * that._x1 + x) / 6, - (that._y0 + 4 * that._y1 + y) / 6 - ); -} - -function Basis(context) { - this._context = context; -} - -Basis.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x0 = this._x1 = - this._y0 = this._y1 = NaN; - this._point = 0; - }, - lineEnd: function() { - switch (this._point) { - case 3: point$4(this, this._x1, this._y1); // proceed - case 2: this._context.lineTo(this._x1, this._y1); break; - } - if (this._line || (this._line !== 0 && this._point === 1)) this._context.closePath(); - this._line = 1 - this._line; - }, - point: function(x, y) { - x = +x, y = +y; - switch (this._point) { - case 0: this._point = 1; this._line ? this._context.lineTo(x, y) : this._context.moveTo(x, y); break; - case 1: this._point = 2; break; - case 2: this._point = 3; this._context.lineTo((5 * this._x0 + this._x1) / 6, (5 * this._y0 + this._y1) / 6); // proceed - default: point$4(this, x, y); break; - } - this._x0 = this._x1, this._x1 = x; - this._y0 = this._y1, this._y1 = y; - } -}; - -function curveBasis(context) { - return new Basis(context); -} - -function BasisClosed(context) { - this._context = context; -} - -BasisClosed.prototype = { - areaStart: noop$1, - areaEnd: noop$1, - lineStart: function() { - this._x0 = this._x1 = this._x2 = this._x3 = this._x4 = - this._y0 = this._y1 = this._y2 = this._y3 = this._y4 = NaN; - this._point = 0; - }, - lineEnd: function() { - switch (this._point) { - case 1: { - this._context.moveTo(this._x2, this._y2); - this._context.closePath(); - break; - } - case 2: { - this._context.moveTo((this._x2 + 2 * this._x3) / 3, (this._y2 + 2 * this._y3) / 3); - this._context.lineTo((this._x3 + 2 * this._x2) / 3, (this._y3 + 2 * this._y2) / 3); - this._context.closePath(); - break; - } - case 3: { - this.point(this._x2, this._y2); - this.point(this._x3, this._y3); - this.point(this._x4, this._y4); - break; - } - } - }, - point: function(x, y) { - x = +x, y = +y; - switch (this._point) { - case 0: this._point = 1; this._x2 = x, this._y2 = y; break; - case 1: this._point = 2; this._x3 = x, this._y3 = y; break; - case 2: this._point = 3; this._x4 = x, this._y4 = y; this._context.moveTo((this._x0 + 4 * this._x1 + x) / 6, (this._y0 + 4 * this._y1 + y) / 6); break; - default: point$4(this, x, y); break; - } - this._x0 = this._x1, this._x1 = x; - this._y0 = this._y1, this._y1 = y; - } -}; - -function curveBasisClosed(context) { - return new BasisClosed(context); -} - -function BasisOpen(context) { - this._context = context; -} - -BasisOpen.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x0 = this._x1 = - this._y0 = this._y1 = NaN; - this._point = 0; - }, - lineEnd: function() { - if (this._line || (this._line !== 0 && this._point === 3)) this._context.closePath(); - this._line = 1 - this._line; - }, - point: function(x, y) { - x = +x, y = +y; - switch (this._point) { - case 0: this._point = 1; break; - case 1: this._point = 2; break; - case 2: this._point = 3; var x0 = (this._x0 + 4 * this._x1 + x) / 6, y0 = (this._y0 + 4 * this._y1 + y) / 6; this._line ? this._context.lineTo(x0, y0) : this._context.moveTo(x0, y0); break; - case 3: this._point = 4; // proceed - default: point$4(this, x, y); break; - } - this._x0 = this._x1, this._x1 = x; - this._y0 = this._y1, this._y1 = y; - } -}; - -function curveBasisOpen(context) { - return new BasisOpen(context); -} - -function Bundle(context, beta) { - this._basis = new Basis(context); - this._beta = beta; -} - -Bundle.prototype = { - lineStart: function() { - this._x = []; - this._y = []; - this._basis.lineStart(); - }, - lineEnd: function() { - var x = this._x, - y = this._y, - j = x.length - 1; - - if (j > 0) { - var x0 = x[0], - y0 = y[0], - dx = x[j] - x0, - dy = y[j] - y0, - i = -1, - t; - - while (++i <= j) { - t = i / j; - this._basis.point( - this._beta * x[i] + (1 - this._beta) * (x0 + t * dx), - this._beta * y[i] + (1 - this._beta) * (y0 + t * dy) - ); - } - } - - this._x = this._y = null; - this._basis.lineEnd(); - }, - point: function(x, y) { - this._x.push(+x); - this._y.push(+y); - } -}; - -((function custom(beta) { - - function bundle(context) { - return beta === 1 ? new Basis(context) : new Bundle(context, beta); - } - - bundle.beta = function(beta) { - return custom(+beta); - }; - - return bundle; -}))(0.85); - -function point$3(that, x, y) { - that._context.bezierCurveTo( - that._x1 + that._k * (that._x2 - that._x0), - that._y1 + that._k * (that._y2 - that._y0), - that._x2 + that._k * (that._x1 - x), - that._y2 + that._k * (that._y1 - y), - that._x2, - that._y2 - ); -} - -function Cardinal(context, tension) { - this._context = context; - this._k = (1 - tension) / 6; -} - -Cardinal.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x0 = this._x1 = this._x2 = - this._y0 = this._y1 = this._y2 = NaN; - this._point = 0; - }, - lineEnd: function() { - switch (this._point) { - case 2: this._context.lineTo(this._x2, this._y2); break; - case 3: point$3(this, this._x1, this._y1); break; - } - if (this._line || (this._line !== 0 && this._point === 1)) this._context.closePath(); - this._line = 1 - this._line; - }, - point: function(x, y) { - x = +x, y = +y; - switch (this._point) { - case 0: this._point = 1; this._line ? this._context.lineTo(x, y) : this._context.moveTo(x, y); break; - case 1: this._point = 2; this._x1 = x, this._y1 = y; break; - case 2: this._point = 3; // proceed - default: point$3(this, x, y); break; - } - this._x0 = this._x1, this._x1 = this._x2, this._x2 = x; - this._y0 = this._y1, this._y1 = this._y2, this._y2 = y; - } -}; - -((function custom(tension) { - - function cardinal(context) { - return new Cardinal(context, tension); - } - - cardinal.tension = function(tension) { - return custom(+tension); - }; - - return cardinal; -}))(0); - -function CardinalClosed(context, tension) { - this._context = context; - this._k = (1 - tension) / 6; -} - -CardinalClosed.prototype = { - areaStart: noop$1, - areaEnd: noop$1, - lineStart: function() { - this._x0 = this._x1 = this._x2 = this._x3 = this._x4 = this._x5 = - this._y0 = this._y1 = this._y2 = this._y3 = this._y4 = this._y5 = NaN; - this._point = 0; - }, - lineEnd: function() { - switch (this._point) { - case 1: { - this._context.moveTo(this._x3, this._y3); - this._context.closePath(); - break; - } - case 2: { - this._context.lineTo(this._x3, this._y3); - this._context.closePath(); - break; - } - case 3: { - this.point(this._x3, this._y3); - this.point(this._x4, this._y4); - this.point(this._x5, this._y5); - break; - } - } - }, - point: function(x, y) { - x = +x, y = +y; - switch (this._point) { - case 0: this._point = 1; this._x3 = x, this._y3 = y; break; - case 1: this._point = 2; this._context.moveTo(this._x4 = x, this._y4 = y); break; - case 2: this._point = 3; this._x5 = x, this._y5 = y; break; - default: point$3(this, x, y); break; - } - this._x0 = this._x1, this._x1 = this._x2, this._x2 = x; - this._y0 = this._y1, this._y1 = this._y2, this._y2 = y; - } -}; - -((function custom(tension) { - - function cardinal(context) { - return new CardinalClosed(context, tension); - } - - cardinal.tension = function(tension) { - return custom(+tension); - }; - - return cardinal; -}))(0); - -function CardinalOpen(context, tension) { - this._context = context; - this._k = (1 - tension) / 6; -} - -CardinalOpen.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x0 = this._x1 = this._x2 = - this._y0 = this._y1 = this._y2 = NaN; - this._point = 0; - }, - lineEnd: function() { - if (this._line || (this._line !== 0 && this._point === 3)) this._context.closePath(); - this._line = 1 - this._line; - }, - point: function(x, y) { - x = +x, y = +y; - switch (this._point) { - case 0: this._point = 1; break; - case 1: this._point = 2; break; - case 2: this._point = 3; this._line ? this._context.lineTo(this._x2, this._y2) : this._context.moveTo(this._x2, this._y2); break; - case 3: this._point = 4; // proceed - default: point$3(this, x, y); break; - } - this._x0 = this._x1, this._x1 = this._x2, this._x2 = x; - this._y0 = this._y1, this._y1 = this._y2, this._y2 = y; - } -}; - -((function custom(tension) { - - function cardinal(context) { - return new CardinalOpen(context, tension); - } - - cardinal.tension = function(tension) { - return custom(+tension); - }; - - return cardinal; -}))(0); - -function point$2(that, x, y) { - var x1 = that._x1, - y1 = that._y1, - x2 = that._x2, - y2 = that._y2; - - if (that._l01_a > epsilon) { - var a = 2 * that._l01_2a + 3 * that._l01_a * that._l12_a + that._l12_2a, - n = 3 * that._l01_a * (that._l01_a + that._l12_a); - x1 = (x1 * a - that._x0 * that._l12_2a + that._x2 * that._l01_2a) / n; - y1 = (y1 * a - that._y0 * that._l12_2a + that._y2 * that._l01_2a) / n; - } - - if (that._l23_a > epsilon) { - var b = 2 * that._l23_2a + 3 * that._l23_a * that._l12_a + that._l12_2a, - m = 3 * that._l23_a * (that._l23_a + that._l12_a); - x2 = (x2 * b + that._x1 * that._l23_2a - x * that._l12_2a) / m; - y2 = (y2 * b + that._y1 * that._l23_2a - y * that._l12_2a) / m; - } - - that._context.bezierCurveTo(x1, y1, x2, y2, that._x2, that._y2); -} - -function CatmullRom(context, alpha) { - this._context = context; - this._alpha = alpha; -} - -CatmullRom.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x0 = this._x1 = this._x2 = - this._y0 = this._y1 = this._y2 = NaN; - this._l01_a = this._l12_a = this._l23_a = - this._l01_2a = this._l12_2a = this._l23_2a = - this._point = 0; - }, - lineEnd: function() { - switch (this._point) { - case 2: this._context.lineTo(this._x2, this._y2); break; - case 3: this.point(this._x2, this._y2); break; - } - if (this._line || (this._line !== 0 && this._point === 1)) this._context.closePath(); - this._line = 1 - this._line; - }, - point: function(x, y) { - x = +x, y = +y; - - if (this._point) { - var x23 = this._x2 - x, - y23 = this._y2 - y; - this._l23_a = Math.sqrt(this._l23_2a = Math.pow(x23 * x23 + y23 * y23, this._alpha)); - } - - switch (this._point) { - case 0: this._point = 1; this._line ? this._context.lineTo(x, y) : this._context.moveTo(x, y); break; - case 1: this._point = 2; break; - case 2: this._point = 3; // proceed - default: point$2(this, x, y); break; - } - - this._l01_a = this._l12_a, this._l12_a = this._l23_a; - this._l01_2a = this._l12_2a, this._l12_2a = this._l23_2a; - this._x0 = this._x1, this._x1 = this._x2, this._x2 = x; - this._y0 = this._y1, this._y1 = this._y2, this._y2 = y; - } -}; - -((function custom(alpha) { - - function catmullRom(context) { - return alpha ? new CatmullRom(context, alpha) : new Cardinal(context, 0); - } - - catmullRom.alpha = function(alpha) { - return custom(+alpha); - }; - - return catmullRom; -}))(0.5); - -function CatmullRomClosed(context, alpha) { - this._context = context; - this._alpha = alpha; -} - -CatmullRomClosed.prototype = { - areaStart: noop$1, - areaEnd: noop$1, - lineStart: function() { - this._x0 = this._x1 = this._x2 = this._x3 = this._x4 = this._x5 = - this._y0 = this._y1 = this._y2 = this._y3 = this._y4 = this._y5 = NaN; - this._l01_a = this._l12_a = this._l23_a = - this._l01_2a = this._l12_2a = this._l23_2a = - this._point = 0; - }, - lineEnd: function() { - switch (this._point) { - case 1: { - this._context.moveTo(this._x3, this._y3); - this._context.closePath(); - break; - } - case 2: { - this._context.lineTo(this._x3, this._y3); - this._context.closePath(); - break; - } - case 3: { - this.point(this._x3, this._y3); - this.point(this._x4, this._y4); - this.point(this._x5, this._y5); - break; - } - } - }, - point: function(x, y) { - x = +x, y = +y; - - if (this._point) { - var x23 = this._x2 - x, - y23 = this._y2 - y; - this._l23_a = Math.sqrt(this._l23_2a = Math.pow(x23 * x23 + y23 * y23, this._alpha)); - } - - switch (this._point) { - case 0: this._point = 1; this._x3 = x, this._y3 = y; break; - case 1: this._point = 2; this._context.moveTo(this._x4 = x, this._y4 = y); break; - case 2: this._point = 3; this._x5 = x, this._y5 = y; break; - default: point$2(this, x, y); break; - } - - this._l01_a = this._l12_a, this._l12_a = this._l23_a; - this._l01_2a = this._l12_2a, this._l12_2a = this._l23_2a; - this._x0 = this._x1, this._x1 = this._x2, this._x2 = x; - this._y0 = this._y1, this._y1 = this._y2, this._y2 = y; - } -}; - -((function custom(alpha) { - - function catmullRom(context) { - return alpha ? new CatmullRomClosed(context, alpha) : new CardinalClosed(context, 0); - } - - catmullRom.alpha = function(alpha) { - return custom(+alpha); - }; - - return catmullRom; -}))(0.5); - -function CatmullRomOpen(context, alpha) { - this._context = context; - this._alpha = alpha; -} - -CatmullRomOpen.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x0 = this._x1 = this._x2 = - this._y0 = this._y1 = this._y2 = NaN; - this._l01_a = this._l12_a = this._l23_a = - this._l01_2a = this._l12_2a = this._l23_2a = - this._point = 0; - }, - lineEnd: function() { - if (this._line || (this._line !== 0 && this._point === 3)) this._context.closePath(); - this._line = 1 - this._line; - }, - point: function(x, y) { - x = +x, y = +y; - - if (this._point) { - var x23 = this._x2 - x, - y23 = this._y2 - y; - this._l23_a = Math.sqrt(this._l23_2a = Math.pow(x23 * x23 + y23 * y23, this._alpha)); - } - - switch (this._point) { - case 0: this._point = 1; break; - case 1: this._point = 2; break; - case 2: this._point = 3; this._line ? this._context.lineTo(this._x2, this._y2) : this._context.moveTo(this._x2, this._y2); break; - case 3: this._point = 4; // proceed - default: point$2(this, x, y); break; - } - - this._l01_a = this._l12_a, this._l12_a = this._l23_a; - this._l01_2a = this._l12_2a, this._l12_2a = this._l23_2a; - this._x0 = this._x1, this._x1 = this._x2, this._x2 = x; - this._y0 = this._y1, this._y1 = this._y2, this._y2 = y; - } -}; - -((function custom(alpha) { - - function catmullRom(context) { - return alpha ? new CatmullRomOpen(context, alpha) : new CardinalOpen(context, 0); - } - - catmullRom.alpha = function(alpha) { - return custom(+alpha); - }; - - return catmullRom; -}))(0.5); - -function LinearClosed(context) { - this._context = context; -} - -LinearClosed.prototype = { - areaStart: noop$1, - areaEnd: noop$1, - lineStart: function() { - this._point = 0; - }, - lineEnd: function() { - if (this._point) this._context.closePath(); - }, - point: function(x, y) { - x = +x, y = +y; - if (this._point) this._context.lineTo(x, y); - else this._point = 1, this._context.moveTo(x, y); - } -}; - -function curveLinearClosed(context) { - return new LinearClosed(context); -} - -function sign(x) { - return x < 0 ? -1 : 1; -} - -// Calculate the slopes of the tangents (Hermite-type interpolation) based on -// the following paper: Steffen, M. 1990. A Simple Method for Monotonic -// Interpolation in One Dimension. Astronomy and Astrophysics, Vol. 239, NO. -// NOV(II), P. 443, 1990. -function slope3(that, x2, y2) { - var h0 = that._x1 - that._x0, - h1 = x2 - that._x1, - s0 = (that._y1 - that._y0) / (h0 || h1 < 0 && -0), - s1 = (y2 - that._y1) / (h1 || h0 < 0 && -0), - p = (s0 * h1 + s1 * h0) / (h0 + h1); - return (sign(s0) + sign(s1)) * Math.min(Math.abs(s0), Math.abs(s1), 0.5 * Math.abs(p)) || 0; -} - -// Calculate a one-sided slope. -function slope2(that, t) { - var h = that._x1 - that._x0; - return h ? (3 * (that._y1 - that._y0) / h - t) / 2 : t; -} - -// According to https://en.wikipedia.org/wiki/Cubic_Hermite_spline#Representations -// "you can express cubic Hermite interpolation in terms of cubic Bézier curves -// with respect to the four values p0, p0 + m0 / 3, p1 - m1 / 3, p1". -function point$1(that, t0, t1) { - var x0 = that._x0, - y0 = that._y0, - x1 = that._x1, - y1 = that._y1, - dx = (x1 - x0) / 3; - that._context.bezierCurveTo(x0 + dx, y0 + dx * t0, x1 - dx, y1 - dx * t1, x1, y1); -} - -function MonotoneX(context) { - this._context = context; -} - -MonotoneX.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x0 = this._x1 = - this._y0 = this._y1 = - this._t0 = NaN; - this._point = 0; - }, - lineEnd: function() { - switch (this._point) { - case 2: this._context.lineTo(this._x1, this._y1); break; - case 3: point$1(this, this._t0, slope2(this, this._t0)); break; - } - if (this._line || (this._line !== 0 && this._point === 1)) this._context.closePath(); - this._line = 1 - this._line; - }, - point: function(x, y) { - var t1 = NaN; - - x = +x, y = +y; - if (x === this._x1 && y === this._y1) return; // Ignore coincident points. - switch (this._point) { - case 0: this._point = 1; this._line ? this._context.lineTo(x, y) : this._context.moveTo(x, y); break; - case 1: this._point = 2; break; - case 2: this._point = 3; point$1(this, slope2(this, t1 = slope3(this, x, y)), t1); break; - default: point$1(this, this._t0, t1 = slope3(this, x, y)); break; - } - - this._x0 = this._x1, this._x1 = x; - this._y0 = this._y1, this._y1 = y; - this._t0 = t1; - } -}; - -function MonotoneY(context) { - this._context = new ReflectContext(context); -} - -(MonotoneY.prototype = Object.create(MonotoneX.prototype)).point = function(x, y) { - MonotoneX.prototype.point.call(this, y, x); -}; - -function ReflectContext(context) { - this._context = context; -} - -ReflectContext.prototype = { - moveTo: function(x, y) { this._context.moveTo(y, x); }, - closePath: function() { this._context.closePath(); }, - lineTo: function(x, y) { this._context.lineTo(y, x); }, - bezierCurveTo: function(x1, y1, x2, y2, x, y) { this._context.bezierCurveTo(y1, x1, y2, x2, y, x); } -}; - -function monotoneX(context) { - return new MonotoneX(context); -} - -function monotoneY(context) { - return new MonotoneY(context); -} - -function Natural(context) { - this._context = context; -} - -Natural.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x = []; - this._y = []; - }, - lineEnd: function() { - var x = this._x, - y = this._y, - n = x.length; - - if (n) { - this._line ? this._context.lineTo(x[0], y[0]) : this._context.moveTo(x[0], y[0]); - if (n === 2) { - this._context.lineTo(x[1], y[1]); - } else { - var px = controlPoints(x), - py = controlPoints(y); - for (var i0 = 0, i1 = 1; i1 < n; ++i0, ++i1) { - this._context.bezierCurveTo(px[0][i0], py[0][i0], px[1][i0], py[1][i0], x[i1], y[i1]); - } - } - } - - if (this._line || (this._line !== 0 && n === 1)) this._context.closePath(); - this._line = 1 - this._line; - this._x = this._y = null; - }, - point: function(x, y) { - this._x.push(+x); - this._y.push(+y); - } -}; - -// See https://www.particleincell.com/2012/bezier-splines/ for derivation. -function controlPoints(x) { - var i, - n = x.length - 1, - m, - a = new Array(n), - b = new Array(n), - r = new Array(n); - a[0] = 0, b[0] = 2, r[0] = x[0] + 2 * x[1]; - for (i = 1; i < n - 1; ++i) a[i] = 1, b[i] = 4, r[i] = 4 * x[i] + 2 * x[i + 1]; - a[n - 1] = 2, b[n - 1] = 7, r[n - 1] = 8 * x[n - 1] + x[n]; - for (i = 1; i < n; ++i) m = a[i] / b[i - 1], b[i] -= m, r[i] -= m * r[i - 1]; - a[n - 1] = r[n - 1] / b[n - 1]; - for (i = n - 2; i >= 0; --i) a[i] = (r[i] - a[i + 1]) / b[i]; - b[n - 1] = (x[n] + a[n - 1]) / 2; - for (i = 0; i < n - 1; ++i) b[i] = 2 * x[i + 1] - a[i + 1]; - return [a, b]; -} - -function curveNatural(context) { - return new Natural(context); -} - -function Step(context, t) { - this._context = context; - this._t = t; -} - -Step.prototype = { - areaStart: function() { - this._line = 0; - }, - areaEnd: function() { - this._line = NaN; - }, - lineStart: function() { - this._x = this._y = NaN; - this._point = 0; - }, - lineEnd: function() { - if (0 < this._t && this._t < 1 && this._point === 2) this._context.lineTo(this._x, this._y); - if (this._line || (this._line !== 0 && this._point === 1)) this._context.closePath(); - if (this._line >= 0) this._t = 1 - this._t, this._line = 1 - this._line; - }, - point: function(x, y) { - x = +x, y = +y; - switch (this._point) { - case 0: this._point = 1; this._line ? this._context.lineTo(x, y) : this._context.moveTo(x, y); break; - case 1: this._point = 2; // proceed - default: { - if (this._t <= 0) { - this._context.lineTo(this._x, y); - this._context.lineTo(x, y); - } else { - var x1 = this._x * (1 - this._t) + x * this._t; - this._context.lineTo(x1, this._y); - this._context.lineTo(x1, y); - } - break; - } - } - this._x = x, this._y = y; - } -}; - -function curveStep(context) { - return new Step(context, 0.5); -} - -function stepBefore(context) { - return new Step(context, 0); -} - -function stepAfter(context) { - return new Step(context, 1); -} - -function stackOffsetNone(series, order) { - if (!((n = series.length) > 1)) return; - for (var i = 1, j, s0, s1 = series[order[0]], n, m = s1.length; i < n; ++i) { - s0 = s1, s1 = series[order[i]]; - for (j = 0; j < m; ++j) { - s1[j][1] += s1[j][0] = isNaN(s0[j][1]) ? s0[j][0] : s0[j][1]; - } - } -} - -function stackOrderNone(series) { - var n = series.length, o = new Array(n); - while (--n >= 0) o[n] = n; - return o; -} - -function stackValue(d, key) { - return d[key]; -} - -function stackSeries(key) { - const series = []; - series.key = key; - return series; -} - -function shapeStack() { - var keys = constant$3([]), - order = stackOrderNone, - offset = stackOffsetNone, - value = stackValue; - - function stack(data) { - var sz = Array.from(keys.apply(this, arguments), stackSeries), - i, n = sz.length, j = -1, - oz; - - for (const d of data) { - for (i = 0, ++j; i < n; ++i) { - (sz[i][j] = [0, +value(d, sz[i].key, j, data)]).data = d; - } - } - - for (i = 0, oz = array(order(sz)); i < n; ++i) { - sz[oz[i]].index = i; - } - - offset(sz, oz); - return sz; - } - - stack.keys = function(_) { - return arguments.length ? (keys = typeof _ === "function" ? _ : constant$3(Array.from(_)), stack) : keys; - }; - - stack.value = function(_) { - return arguments.length ? (value = typeof _ === "function" ? _ : constant$3(+_), stack) : value; - }; - - stack.order = function(_) { - return arguments.length ? (order = _ == null ? stackOrderNone : typeof _ === "function" ? _ : constant$3(Array.from(_)), stack) : order; - }; - - stack.offset = function(_) { - return arguments.length ? (offset = _ == null ? stackOffsetNone : _, stack) : offset; - }; - - return stack; -} - -function stackOffsetExpand(series, order) { - if (!((n = series.length) > 0)) return; - for (var i, n, j = 0, m = series[0].length, y; j < m; ++j) { - for (y = i = 0; i < n; ++i) y += series[i][j][1] || 0; - if (y) for (i = 0; i < n; ++i) series[i][j][1] /= y; - } - stackOffsetNone(series, order); -} - -function stackOffsetSilhouette(series, order) { - if (!((n = series.length) > 0)) return; - for (var j = 0, s0 = series[order[0]], n, m = s0.length; j < m; ++j) { - for (var i = 0, y = 0; i < n; ++i) y += series[i][j][1] || 0; - s0[j][1] += s0[j][0] = -y / 2; - } - stackOffsetNone(series, order); -} - -function stackOffsetWiggle(series, order) { - if (!((n = series.length) > 0) || !((m = (s0 = series[order[0]]).length) > 0)) return; - for (var y = 0, j = 1, s0, m, n; j < m; ++j) { - for (var i = 0, s1 = 0, s2 = 0; i < n; ++i) { - var si = series[order[i]], - sij0 = si[j][1] || 0, - sij1 = si[j - 1][1] || 0, - s3 = (sij0 - sij1) / 2; - for (var k = 0; k < i; ++k) { - var sk = series[order[k]], - skj0 = sk[j][1] || 0, - skj1 = sk[j - 1][1] || 0; - s3 += skj0 - skj1; - } - s1 += sij0, s2 += s3 * sij0; - } - s0[j - 1][1] += s0[j - 1][0] = y; - if (s1) y -= s2 / s1; - } - s0[j - 1][1] += s0[j - 1][0] = y; - stackOffsetNone(series, order); -} - -function _typeof$y(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$y = function _typeof(obj) { return typeof obj; }; } else { _typeof$y = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$y(obj); } - -function _extends$x() { _extends$x = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$x.apply(this, arguments); } - -function _classCallCheck$w(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$w(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$w(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$w(Constructor.prototype, protoProps); if (staticProps) _defineProperties$w(Constructor, staticProps); return Constructor; } - -function _inherits$v(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$v(subClass, superClass); } - -function _setPrototypeOf$v(o, p) { _setPrototypeOf$v = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$v(o, p); } - -function _createSuper$v(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$v(); return function _createSuperInternal() { var Super = _getPrototypeOf$v(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$v(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$v(this, result); }; } - -function _possibleConstructorReturn$v(self, call) { if (call && (_typeof$y(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$v(self); } - -function _assertThisInitialized$v(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$v() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$v(o) { _getPrototypeOf$v = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$v(o); } -var symbolFactories = { - symbolCircle: symbolCircle, - symbolCross: symbolCross, - symbolDiamond: symbolDiamond, - symbolSquare: symbolSquare, - symbolStar: symbolStar, - symbolTriangle: symbolTriangle, - symbolWye: symbolWye -}; -var RADIAN$2 = Math.PI / 180; - -var getSymbolFactory = function getSymbolFactory(type) { - var name = "symbol".concat(upperFirst_1(type)); - return symbolFactories[name] || symbolCircle; -}; - -var calculateAreaSize = function calculateAreaSize(size, sizeType, type) { - if (sizeType === 'area') { - return size; - } - - switch (type) { - case 'cross': - return 5 * size * size / 9; - - case 'diamond': - return 0.5 * size * size / Math.sqrt(3); - - case 'square': - return size * size; - - case 'star': - { - var angle = 18 * RADIAN$2; - return 1.25 * size * size * (Math.tan(angle) - Math.tan(angle * 2) * Math.pow(Math.tan(angle), 2)); - } - - case 'triangle': - return Math.sqrt(3) * size * size / 4; - - case 'wye': - return (21 - 10 * Math.sqrt(3)) * size * size / 8; - - default: - return Math.PI * size * size / 4; - } -}; - -var Symbols = /*#__PURE__*/function (_PureComponent) { - _inherits$v(Symbols, _PureComponent); - - var _super = _createSuper$v(Symbols); - - function Symbols() { - _classCallCheck$w(this, Symbols); - - return _super.apply(this, arguments); - } - - _createClass$w(Symbols, [{ - key: "getPath", - value: - /** - * Calculate the path of curve - * @return {String} path - */ - function getPath() { - var _this$props = this.props, - size = _this$props.size, - sizeType = _this$props.sizeType, - type = _this$props.type; - var symbolFactory = getSymbolFactory(type); - var symbol = shapeSymbol().type(symbolFactory).size(calculateAreaSize(size, sizeType, type)); - return symbol(); - } - }, { - key: "render", - value: function render() { - var _this$props2 = this.props, - className = _this$props2.className, - cx = _this$props2.cx, - cy = _this$props2.cy, - size = _this$props2.size; - - if (cx === +cx && cy === +cy && size === +size) { - return /*#__PURE__*/React__default.createElement("path", _extends$x({}, filterProps(this.props, true), { - className: classNames('recharts-symbols', className), - transform: "translate(".concat(cx, ", ").concat(cy, ")"), - d: this.getPath() - })); - } - - return null; - } - }]); - - return Symbols; -}(PureComponent); -Symbols.defaultProps = { - type: 'circle', - size: 64, - sizeType: 'area' -}; - -Symbols.registerSymbol = function (key, factory) { - symbolFactories["symbol".concat(upperFirst_1(key))] = factory; -}; - -function _typeof$x(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$x = function _typeof(obj) { return typeof obj; }; } else { _typeof$x = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$x(obj); } - -function _extends$w() { _extends$w = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$w.apply(this, arguments); } - -function ownKeys$A(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$A(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$A(Object(source), true).forEach(function (key) { _defineProperty$B(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$A(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$B(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$v(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$v(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$v(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$v(Constructor.prototype, protoProps); if (staticProps) _defineProperties$v(Constructor, staticProps); return Constructor; } - -function _inherits$u(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$u(subClass, superClass); } - -function _setPrototypeOf$u(o, p) { _setPrototypeOf$u = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$u(o, p); } - -function _createSuper$u(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$u(); return function _createSuperInternal() { var Super = _getPrototypeOf$u(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$u(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$u(this, result); }; } - -function _possibleConstructorReturn$u(self, call) { if (call && (_typeof$x(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$u(self); } - -function _assertThisInitialized$u(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$u() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$u(o) { _getPrototypeOf$u = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$u(o); } -var SIZE = 32; -var DefaultLegendContent = /*#__PURE__*/function (_PureComponent) { - _inherits$u(DefaultLegendContent, _PureComponent); - - var _super = _createSuper$u(DefaultLegendContent); - - function DefaultLegendContent() { - _classCallCheck$v(this, DefaultLegendContent); - - return _super.apply(this, arguments); - } - - _createClass$v(DefaultLegendContent, [{ - key: "renderIcon", - value: - /** - * Render the path of icon - * @param {Object} data Data of each legend item - * @return {String} Path element - */ - function renderIcon(data) { - var inactiveColor = this.props.inactiveColor; - var halfSize = SIZE / 2; - var sixthSize = SIZE / 6; - var thirdSize = SIZE / 3; - var color = data.inactive ? inactiveColor : data.color; - - if (data.type === 'plainline') { - return /*#__PURE__*/React__default.createElement("line", { - strokeWidth: 4, - fill: "none", - stroke: color, - strokeDasharray: data.payload.strokeDasharray, - x1: 0, - y1: halfSize, - x2: SIZE, - y2: halfSize, - className: "recharts-legend-icon" - }); - } - - if (data.type === 'line') { - return /*#__PURE__*/React__default.createElement("path", { - strokeWidth: 4, - fill: "none", - stroke: color, - d: "M0,".concat(halfSize, "h").concat(thirdSize, "\n A").concat(sixthSize, ",").concat(sixthSize, ",0,1,1,").concat(2 * thirdSize, ",").concat(halfSize, "\n H").concat(SIZE, "M").concat(2 * thirdSize, ",").concat(halfSize, "\n A").concat(sixthSize, ",").concat(sixthSize, ",0,1,1,").concat(thirdSize, ",").concat(halfSize), - className: "recharts-legend-icon" - }); - } - - if (data.type === 'rect') { - return /*#__PURE__*/React__default.createElement("path", { - stroke: "none", - fill: color, - d: "M0,".concat(SIZE / 8, "h").concat(SIZE, "v").concat(SIZE * 3 / 4, "h").concat(-SIZE, "z"), - className: "recharts-legend-icon" - }); - } - - if ( /*#__PURE__*/React__default.isValidElement(data.legendIcon)) { - var iconProps = _objectSpread$A({}, data); - - delete iconProps.legendIcon; - return /*#__PURE__*/React__default.cloneElement(data.legendIcon, iconProps); - } - - return /*#__PURE__*/React__default.createElement(Symbols, { - fill: color, - cx: halfSize, - cy: halfSize, - size: SIZE, - sizeType: "diameter", - type: data.type - }); - } - /** - * Draw items of legend - * @return {ReactElement} Items - */ - - }, { - key: "renderItems", - value: function renderItems() { - var _this = this; - - var _this$props = this.props, - payload = _this$props.payload, - iconSize = _this$props.iconSize, - layout = _this$props.layout, - formatter = _this$props.formatter, - inactiveColor = _this$props.inactiveColor; - var viewBox = { - x: 0, - y: 0, - width: SIZE, - height: SIZE - }; - var itemStyle = { - display: layout === 'horizontal' ? 'inline-block' : 'block', - marginRight: 10 - }; - var svgStyle = { - display: 'inline-block', - verticalAlign: 'middle', - marginRight: 4 - }; - return payload.map(function (entry, i) { - var _classNames; - - var finalFormatter = entry.formatter || formatter; - var className = classNames((_classNames = { - 'recharts-legend-item': true - }, _defineProperty$B(_classNames, "legend-item-".concat(i), true), _defineProperty$B(_classNames, "inactive", entry.inactive), _classNames)); - - if (entry.type === 'none') { - return null; - } - - var color = entry.inactive ? inactiveColor : entry.color; - return /*#__PURE__*/React__default.createElement("li", _extends$w({ - className: className, - style: itemStyle, - key: "legend-item-".concat(i) // eslint-disable-line react/no-array-index-key - - }, adaptEventsOfChild(_this.props, entry, i)), /*#__PURE__*/React__default.createElement(Surface, { - width: iconSize, - height: iconSize, - viewBox: viewBox, - style: svgStyle - }, _this.renderIcon(entry)), /*#__PURE__*/React__default.createElement("span", { - className: "recharts-legend-item-text", - style: { - color: color - } - }, finalFormatter ? finalFormatter(entry.value, entry, i) : entry.value)); - }); - } - }, { - key: "render", - value: function render() { - var _this$props2 = this.props, - payload = _this$props2.payload, - layout = _this$props2.layout, - align = _this$props2.align; - - if (!payload || !payload.length) { - return null; - } - - var finalStyle = { - padding: 0, - margin: 0, - textAlign: layout === 'horizontal' ? align : 'left' - }; - return /*#__PURE__*/React__default.createElement("ul", { - className: "recharts-default-legend", - style: finalStyle - }, this.renderItems()); - } - }]); - - return DefaultLegendContent; -}(PureComponent); -DefaultLegendContent.displayName = 'Legend'; -DefaultLegendContent.defaultProps = { - iconSize: 14, - layout: 'horizontal', - align: 'center', - verticalAlign: 'middle', - inactiveColor: '#ccc' -}; - -var baseGetTag$4 = _baseGetTag$1, - isObjectLike$6 = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var numberTag$2 = '[object Number]'; - -/** - * Checks if `value` is classified as a `Number` primitive or object. - * - * **Note:** To exclude `Infinity`, `-Infinity`, and `NaN`, which are - * classified as numbers, use the `_.isFinite` method. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a number, else `false`. - * @example - * - * _.isNumber(3); - * // => true - * - * _.isNumber(Number.MIN_VALUE); - * // => true - * - * _.isNumber(Infinity); - * // => true - * - * _.isNumber('3'); - * // => false - */ -function isNumber$2(value) { - return typeof value == 'number' || - (isObjectLike$6(value) && baseGetTag$4(value) == numberTag$2); -} - -var isNumber_1 = isNumber$2; - -var isNumber$1 = isNumber_1; - -/** - * Checks if `value` is `NaN`. - * - * **Note:** This method is based on - * [`Number.isNaN`](https://mdn.io/Number/isNaN) and is not the same as - * global [`isNaN`](https://mdn.io/isNaN) which returns `true` for - * `undefined` and other non-number values. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is `NaN`, else `false`. - * @example - * - * _.isNaN(NaN); - * // => true - * - * _.isNaN(new Number(NaN)); - * // => true - * - * isNaN(undefined); - * // => true - * - * _.isNaN(undefined); - * // => false - */ -function isNaN$1(value) { - // An `NaN` primitive is the only value that is not equal to itself. - // Perform the `toStringTag` check first to avoid errors with some - // ActiveX objects in IE. - return isNumber$1(value) && value != +value; -} - -var _isNaN = isNaN$1; - -var baseGetTag$3 = _baseGetTag$1, - isArray$6 = isArray_1, - isObjectLike$5 = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var stringTag$2 = '[object String]'; - -/** - * Checks if `value` is classified as a `String` primitive or object. - * - * @static - * @since 0.1.0 - * @memberOf _ - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a string, else `false`. - * @example - * - * _.isString('abc'); - * // => true - * - * _.isString(1); - * // => false - */ -function isString(value) { - return typeof value == 'string' || - (!isArray$6(value) && isObjectLike$5(value) && baseGetTag$3(value) == stringTag$2); -} - -var isString_1 = isString; - -var mathSign = function mathSign(value) { - if (value === 0) { - return 0; - } - - if (value > 0) { - return 1; - } - - return -1; -}; -var isPercent = function isPercent(value) { - return isString_1(value) && value.indexOf('%') === value.length - 1; -}; -var isNumber = function isNumber(value) { - return isNumber_1(value) && !_isNaN(value); -}; -var isNumOrStr = function isNumOrStr(value) { - return isNumber(value) || isString_1(value); -}; -var idCounter = 0; -var uniqueId = function uniqueId(prefix) { - var id = ++idCounter; - return "".concat(prefix || '').concat(id); -}; -/** - * Get percent value of a total value - * @param {Number|String} percent A percent - * @param {Number} totalValue Total value - * @param {NUmber} defaultValue The value returned when percent is undefined or invalid - * @param {Boolean} validate If set to be true, the result will be validated - * @return {Number} value - */ - -var getPercentValue = function getPercentValue(percent, totalValue) { - var defaultValue = arguments.length > 2 && arguments[2] !== undefined ? arguments[2] : 0; - var validate = arguments.length > 3 && arguments[3] !== undefined ? arguments[3] : false; - - if (!isNumber(percent) && !isString_1(percent)) { - return defaultValue; - } - - var value; - - if (isPercent(percent)) { - var index = percent.indexOf('%'); - value = totalValue * parseFloat(percent.slice(0, index)) / 100; - } else { - value = +percent; - } - - if (_isNaN(value)) { - value = defaultValue; - } - - if (validate && value > totalValue) { - value = totalValue; - } - - return value; -}; -var getAnyElementOfObject = function getAnyElementOfObject(obj) { - if (!obj) { - return null; - } - - var keys = Object.keys(obj); - - if (keys && keys.length) { - return obj[keys[0]]; - } - - return null; -}; -var hasDuplicate = function hasDuplicate(ary) { - if (!isArray_1(ary)) { - return false; - } - - var len = ary.length; - var cache = {}; - - for (var i = 0; i < len; i++) { - if (!cache[ary[i]]) { - cache[ary[i]] = true; - } else { - return true; - } - } - - return false; -}; -var interpolateNumber$2 = function interpolateNumber(numberA, numberB) { - if (isNumber(numberA) && isNumber(numberB)) { - return function (t) { - return numberA + t * (numberB - numberA); - }; - } - - return function () { - return numberB; - }; -}; -function findEntryInArray(ary, specifiedKey, specifiedValue) { - if (!ary || !ary.length) { - return null; - } - - return ary.find(function (entry) { - return entry && (typeof specifiedKey === 'function' ? specifiedKey(entry) : get_1(entry, specifiedKey)) === specifiedValue; - }); -} -/** - * The least square linear regression - * @param {Array} data The array of points - * @returns {Object} The domain of x, and the parameter of linear function - */ - -var getLinearRegression = function getLinearRegression(data) { - if (!data || !data.length) { - return null; - } - - var len = data.length; - var xsum = 0; - var ysum = 0; - var xysum = 0; - var xxsum = 0; - var xmin = Infinity; - var xmax = -Infinity; - var xcurrent = 0; - var ycurrent = 0; - - for (var i = 0; i < len; i++) { - xcurrent = data[i].cx || 0; - ycurrent = data[i].cy || 0; - xsum += xcurrent; - ysum += ycurrent; - xysum += xcurrent * ycurrent; - xxsum += xcurrent * xcurrent; - xmin = Math.min(xmin, xcurrent); - xmax = Math.max(xmax, xcurrent); - } - - var a = len * xxsum !== xsum * xsum ? (len * xysum - xsum * ysum) / (len * xxsum - xsum * xsum) : 0; - return { - xmin: xmin, - xmax: xmax, - a: a, - b: (ysum - a * xsum) / len - }; -}; - -function _typeof$w(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$w = function _typeof(obj) { return typeof obj; }; } else { _typeof$w = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$w(obj); } - -function ownKeys$z(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$z(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$z(Object(source), true).forEach(function (key) { _defineProperty$A(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$z(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$A(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$u(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$u(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$u(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$u(Constructor.prototype, protoProps); if (staticProps) _defineProperties$u(Constructor, staticProps); return Constructor; } - -function _inherits$t(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$t(subClass, superClass); } - -function _setPrototypeOf$t(o, p) { _setPrototypeOf$t = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$t(o, p); } - -function _createSuper$t(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$t(); return function _createSuperInternal() { var Super = _getPrototypeOf$t(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$t(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$t(this, result); }; } - -function _possibleConstructorReturn$t(self, call) { if (call && (_typeof$w(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$t(self); } - -function _assertThisInitialized$t(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$t() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$t(o) { _getPrototypeOf$t = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$t(o); } - -function _objectWithoutProperties$i(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$j(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$j(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function defaultUniqBy$1(entry) { - return entry.value; -} - -function getUniqPayload$1(option, payload) { - if (option === true) { - return uniqBy_1(payload, defaultUniqBy$1); - } - - if (isFunction_1(option)) { - return uniqBy_1(payload, option); - } - - return payload; -} - -function renderContent$1(content, props) { - if ( /*#__PURE__*/React__default.isValidElement(content)) { - return /*#__PURE__*/React__default.cloneElement(content, props); - } - - if (isFunction_1(content)) { - return /*#__PURE__*/React__default.createElement(content, props); - } - - props.ref; - var otherProps = _objectWithoutProperties$i(props, ["ref"]); - - return /*#__PURE__*/React__default.createElement(DefaultLegendContent, otherProps); -} - -var EPS$2 = 1; -var Legend = /*#__PURE__*/function (_PureComponent) { - _inherits$t(Legend, _PureComponent); - - var _super = _createSuper$t(Legend); - - function Legend() { - var _this; - - _classCallCheck$u(this, Legend); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.wrapperNode = void 0; - _this.state = { - boxWidth: -1, - boxHeight: -1 - }; - return _this; - } - - _createClass$u(Legend, [{ - key: "componentDidMount", - value: function componentDidMount() { - this.updateBBox(); - } - }, { - key: "componentDidUpdate", - value: function componentDidUpdate() { - this.updateBBox(); - } - }, { - key: "getBBox", - value: function getBBox() { - if (this.wrapperNode && this.wrapperNode.getBoundingClientRect) { - return this.wrapperNode.getBoundingClientRect(); - } - - return null; - } - }, { - key: "getBBoxSnapshot", - value: function getBBoxSnapshot() { - var _this$state = this.state, - boxWidth = _this$state.boxWidth, - boxHeight = _this$state.boxHeight; - - if (boxWidth >= 0 && boxHeight >= 0) { - return { - width: boxWidth, - height: boxHeight - }; - } - - return null; - } - }, { - key: "getDefaultPosition", - value: function getDefaultPosition(style) { - var _this$props = this.props, - layout = _this$props.layout, - align = _this$props.align, - verticalAlign = _this$props.verticalAlign, - margin = _this$props.margin, - chartWidth = _this$props.chartWidth, - chartHeight = _this$props.chartHeight; - var hPos, vPos; - - if (!style || (style.left === undefined || style.left === null) && (style.right === undefined || style.right === null)) { - if (align === 'center' && layout === 'vertical') { - var _box = this.getBBoxSnapshot() || { - width: 0 - }; - - hPos = { - left: ((chartWidth || 0) - _box.width) / 2 - }; - } else { - hPos = align === 'right' ? { - right: margin && margin.right || 0 - } : { - left: margin && margin.left || 0 - }; - } - } - - if (!style || (style.top === undefined || style.top === null) && (style.bottom === undefined || style.bottom === null)) { - if (verticalAlign === 'middle') { - var _box2 = this.getBBoxSnapshot() || { - height: 0 - }; - - vPos = { - top: ((chartHeight || 0) - _box2.height) / 2 - }; - } else { - vPos = verticalAlign === 'bottom' ? { - bottom: margin && margin.bottom || 0 - } : { - top: margin && margin.top || 0 - }; - } - } - - return _objectSpread$z(_objectSpread$z({}, hPos), vPos); - } - }, { - key: "updateBBox", - value: function updateBBox() { - var _this$state2 = this.state, - boxWidth = _this$state2.boxWidth, - boxHeight = _this$state2.boxHeight; - var onBBoxUpdate = this.props.onBBoxUpdate; - - if (this.wrapperNode && this.wrapperNode.getBoundingClientRect) { - var _box3 = this.wrapperNode.getBoundingClientRect(); - - if (Math.abs(_box3.width - boxWidth) > EPS$2 || Math.abs(_box3.height - boxHeight) > EPS$2) { - this.setState({ - boxWidth: _box3.width, - boxHeight: _box3.height - }, function () { - if (onBBoxUpdate) { - onBBoxUpdate(_box3); - } - }); - } - } else if (boxWidth !== -1 || boxHeight !== -1) { - this.setState({ - boxWidth: -1, - boxHeight: -1 - }, function () { - if (onBBoxUpdate) { - onBBoxUpdate(null); - } - }); - } - } - }, { - key: "render", - value: function render() { - var _this2 = this; - - var _this$props2 = this.props, - content = _this$props2.content, - width = _this$props2.width, - height = _this$props2.height, - wrapperStyle = _this$props2.wrapperStyle, - payloadUniqBy = _this$props2.payloadUniqBy, - payload = _this$props2.payload; - - var outerStyle = _objectSpread$z(_objectSpread$z({ - position: 'absolute', - width: width || 'auto', - height: height || 'auto' - }, this.getDefaultPosition(wrapperStyle)), wrapperStyle); - - return /*#__PURE__*/React__default.createElement("div", { - className: "recharts-legend-wrapper", - style: outerStyle, - ref: function ref(node) { - _this2.wrapperNode = node; - } - }, renderContent$1(content, _objectSpread$z(_objectSpread$z({}, this.props), {}, { - payload: getUniqPayload$1(payloadUniqBy, payload) - }))); - } - }], [{ - key: "getWithHeight", - value: function getWithHeight(item, chartWidth) { - var layout = item.props.layout; - - if (layout === 'vertical' && isNumber(item.props.height)) { - return { - height: item.props.height - }; - } - - if (layout === 'horizontal') { - return { - width: item.props.width || chartWidth - }; - } - - return null; - } - }]); - - return Legend; -}(PureComponent); -Legend.displayName = 'Legend'; -Legend.defaultProps = { - iconSize: 14, - layout: 'horizontal', - align: 'center', - verticalAlign: 'bottom' -}; - -/** - * Checks if `value` is `null` or `undefined`. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is nullish, else `false`. - * @example - * - * _.isNil(null); - * // => true - * - * _.isNil(void 0); - * // => true - * - * _.isNil(NaN); - * // => false - */ - -function isNil(value) { - return value == null; -} - -var isNil_1 = isNil; - -var lib$2 = {}; - -var Animate$2 = {}; - -var propTypes = {exports: {}}; - -/** - * Copyright (c) 2013-present, Facebook, Inc. - * - * This source code is licensed under the MIT license found in the - * LICENSE file in the root directory of this source tree. - */ - -var ReactPropTypesSecret$1 = 'SECRET_DO_NOT_PASS_THIS_OR_YOU_WILL_BE_FIRED'; - -var ReactPropTypesSecret_1 = ReactPropTypesSecret$1; - -/** - * Copyright (c) 2013-present, Facebook, Inc. - * - * This source code is licensed under the MIT license found in the - * LICENSE file in the root directory of this source tree. - */ - -var ReactPropTypesSecret = ReactPropTypesSecret_1; - -function emptyFunction() {} -function emptyFunctionWithReset() {} -emptyFunctionWithReset.resetWarningCache = emptyFunction; - -var factoryWithThrowingShims = function() { - function shim(props, propName, componentName, location, propFullName, secret) { - if (secret === ReactPropTypesSecret) { - // It is still safe when called from React. - return; - } - var err = new Error( - 'Calling PropTypes validators directly is not supported by the `prop-types` package. ' + - 'Use PropTypes.checkPropTypes() to call them. ' + - 'Read more at http://fb.me/use-check-prop-types' - ); - err.name = 'Invariant Violation'; - throw err; - } shim.isRequired = shim; - function getShim() { - return shim; - } // Important! - // Keep this list in sync with production version in `./factoryWithTypeCheckers.js`. - var ReactPropTypes = { - array: shim, - bigint: shim, - bool: shim, - func: shim, - number: shim, - object: shim, - string: shim, - symbol: shim, - - any: shim, - arrayOf: getShim, - element: shim, - elementType: shim, - instanceOf: getShim, - node: shim, - objectOf: getShim, - oneOf: getShim, - oneOfType: getShim, - shape: getShim, - exact: getShim, - - checkPropTypes: emptyFunctionWithReset, - resetWarningCache: emptyFunction - }; - - ReactPropTypes.PropTypes = ReactPropTypes; - - return ReactPropTypes; -}; - -/** - * Copyright (c) 2013-present, Facebook, Inc. - * - * This source code is licensed under the MIT license found in the - * LICENSE file in the root directory of this source tree. - */ - -{ - // By explicitly using `prop-types` you are opting into new production behavior. - // http://fb.me/prop-types-in-prod - propTypes.exports = factoryWithThrowingShims(); -} - -var fastEquals = {exports: {}}; - -(function (module, exports) { - (function (global, factory) { - factory(exports) ; - })(commonjsGlobal$1, (function (exports) { - var HAS_WEAKSET_SUPPORT = typeof WeakSet === 'function'; - var keys = Object.keys; - /** - * are the values passed strictly equal or both NaN - * - * @param a the value to compare against - * @param b the value to test - * @returns are the values equal by the SameValueZero principle - */ - function sameValueZeroEqual(a, b) { - return a === b || (a !== a && b !== b); - } - /** - * is the value a plain object - * - * @param value the value to test - * @returns is the value a plain object - */ - function isPlainObject(value) { - return value.constructor === Object || value.constructor == null; - } - /** - * is the value promise-like (meaning it is thenable) - * - * @param value the value to test - * @returns is the value promise-like - */ - function isPromiseLike(value) { - return !!value && typeof value.then === 'function'; - } - /** - * is the value passed a react element - * - * @param value the value to test - * @returns is the value a react element - */ - function isReactElement(value) { - return !!(value && value.$$typeof); - } - /** - * in cases where WeakSet is not supported, creates a new custom - * object that mimics the necessary API aspects for cache purposes - * - * @returns the new cache object - */ - function getNewCacheFallback() { - var values = []; - return { - add: function (value) { - values.push(value); - }, - has: function (value) { - return values.indexOf(value) !== -1; - }, - }; - } - /** - * get a new cache object to prevent circular references - * - * @returns the new cache object - */ - var getNewCache = (function (canUseWeakMap) { - if (canUseWeakMap) { - return function _getNewCache() { - return new WeakSet(); - }; - } - return getNewCacheFallback; - })(HAS_WEAKSET_SUPPORT); - /** - * create a custom isEqual handler specific to circular objects - * - * @param [isEqual] the isEqual comparator to use instead of isDeepEqual - * @returns the method to create the `isEqual` function - */ - function createCircularEqualCreator(isEqual) { - return function createCircularEqual(comparator) { - var _comparator = isEqual || comparator; - return function circularEqual(a, b, cache) { - if (cache === void 0) { cache = getNewCache(); } - var isCacheableA = !!a && typeof a === 'object'; - var isCacheableB = !!b && typeof b === 'object'; - if (isCacheableA || isCacheableB) { - var hasA = isCacheableA && cache.has(a); - var hasB = isCacheableB && cache.has(b); - if (hasA || hasB) { - return hasA && hasB; - } - if (isCacheableA) { - cache.add(a); - } - if (isCacheableB) { - cache.add(b); - } - } - return _comparator(a, b, cache); - }; - }; - } - /** - * are the arrays equal in value - * - * @param a the array to test - * @param b the array to test against - * @param isEqual the comparator to determine equality - * @param meta the meta object to pass through - * @returns are the arrays equal - */ - function areArraysEqual(a, b, isEqual, meta) { - var index = a.length; - if (b.length !== index) { - return false; - } - while (index-- > 0) { - if (!isEqual(a[index], b[index], meta)) { - return false; - } - } - return true; - } - /** - * are the maps equal in value - * - * @param a the map to test - * @param b the map to test against - * @param isEqual the comparator to determine equality - * @param meta the meta map to pass through - * @returns are the maps equal - */ - function areMapsEqual(a, b, isEqual, meta) { - var isValueEqual = a.size === b.size; - if (isValueEqual && a.size) { - var matchedIndices_1 = {}; - a.forEach(function (aValue, aKey) { - if (isValueEqual) { - var hasMatch_1 = false; - var matchIndex_1 = 0; - b.forEach(function (bValue, bKey) { - if (!hasMatch_1 && !matchedIndices_1[matchIndex_1]) { - hasMatch_1 = - isEqual(aKey, bKey, meta) && isEqual(aValue, bValue, meta); - if (hasMatch_1) { - matchedIndices_1[matchIndex_1] = true; - } - } - matchIndex_1++; - }); - isValueEqual = hasMatch_1; - } - }); - } - return isValueEqual; - } - var OWNER = '_owner'; - var hasOwnProperty = Function.prototype.bind.call(Function.prototype.call, Object.prototype.hasOwnProperty); - /** - * are the objects equal in value - * - * @param a the object to test - * @param b the object to test against - * @param isEqual the comparator to determine equality - * @param meta the meta object to pass through - * @returns are the objects equal - */ - function areObjectsEqual(a, b, isEqual, meta) { - var keysA = keys(a); - var index = keysA.length; - if (keys(b).length !== index) { - return false; - } - if (index) { - var key = void 0; - while (index-- > 0) { - key = keysA[index]; - if (key === OWNER) { - var reactElementA = isReactElement(a); - var reactElementB = isReactElement(b); - if ((reactElementA || reactElementB) && - reactElementA !== reactElementB) { - return false; - } - } - if (!hasOwnProperty(b, key) || !isEqual(a[key], b[key], meta)) { - return false; - } - } - } - return true; - } - /** - * are the regExps equal in value - * - * @param a the regExp to test - * @param b the regExp to test agains - * @returns are the regExps equal - */ - function areRegExpsEqual(a, b) { - return (a.source === b.source && - a.global === b.global && - a.ignoreCase === b.ignoreCase && - a.multiline === b.multiline && - a.unicode === b.unicode && - a.sticky === b.sticky && - a.lastIndex === b.lastIndex); - } - /** - * are the sets equal in value - * - * @param a the set to test - * @param b the set to test against - * @param isEqual the comparator to determine equality - * @param meta the meta set to pass through - * @returns are the sets equal - */ - function areSetsEqual(a, b, isEqual, meta) { - var isValueEqual = a.size === b.size; - if (isValueEqual && a.size) { - var matchedIndices_2 = {}; - a.forEach(function (aValue) { - if (isValueEqual) { - var hasMatch_2 = false; - var matchIndex_2 = 0; - b.forEach(function (bValue) { - if (!hasMatch_2 && !matchedIndices_2[matchIndex_2]) { - hasMatch_2 = isEqual(aValue, bValue, meta); - if (hasMatch_2) { - matchedIndices_2[matchIndex_2] = true; - } - } - matchIndex_2++; - }); - isValueEqual = hasMatch_2; - } - }); - } - return isValueEqual; - } - - var HAS_MAP_SUPPORT = typeof Map === 'function'; - var HAS_SET_SUPPORT = typeof Set === 'function'; - function createComparator(createIsEqual) { - var isEqual = - /* eslint-disable no-use-before-define */ - typeof createIsEqual === 'function' - ? createIsEqual(comparator) - : comparator; - /* eslint-enable */ - /** - * compare the value of the two objects and return true if they are equivalent in values - * - * @param a the value to test against - * @param b the value to test - * @param [meta] an optional meta object that is passed through to all equality test calls - * @returns are a and b equivalent in value - */ - function comparator(a, b, meta) { - if (a === b) { - return true; - } - if (a && b && typeof a === 'object' && typeof b === 'object') { - if (isPlainObject(a) && isPlainObject(b)) { - return areObjectsEqual(a, b, isEqual, meta); - } - var aShape = Array.isArray(a); - var bShape = Array.isArray(b); - if (aShape || bShape) { - return aShape === bShape && areArraysEqual(a, b, isEqual, meta); - } - aShape = a instanceof Date; - bShape = b instanceof Date; - if (aShape || bShape) { - return (aShape === bShape && sameValueZeroEqual(a.getTime(), b.getTime())); - } - aShape = a instanceof RegExp; - bShape = b instanceof RegExp; - if (aShape || bShape) { - return aShape === bShape && areRegExpsEqual(a, b); - } - if (isPromiseLike(a) || isPromiseLike(b)) { - return a === b; - } - if (HAS_MAP_SUPPORT) { - aShape = a instanceof Map; - bShape = b instanceof Map; - if (aShape || bShape) { - return aShape === bShape && areMapsEqual(a, b, isEqual, meta); - } - } - if (HAS_SET_SUPPORT) { - aShape = a instanceof Set; - bShape = b instanceof Set; - if (aShape || bShape) { - return aShape === bShape && areSetsEqual(a, b, isEqual, meta); - } - } - return areObjectsEqual(a, b, isEqual, meta); - } - return a !== a && b !== b; - } - return comparator; - } - - var deepEqual = createComparator(); - var shallowEqual = createComparator(function () { return sameValueZeroEqual; }); - var circularDeepEqual = createComparator(createCircularEqualCreator()); - var circularShallowEqual = createComparator(createCircularEqualCreator(sameValueZeroEqual)); - - exports.circularDeepEqual = circularDeepEqual; - exports.circularShallowEqual = circularShallowEqual; - exports.createCustomEqual = createComparator; - exports.deepEqual = deepEqual; - exports.sameValueZeroEqual = sameValueZeroEqual; - exports.shallowEqual = shallowEqual; - - Object.defineProperty(exports, '__esModule', { value: true }); - - })); - -} (fastEquals, fastEquals.exports)); - -var AnimateManager = {}; - -var setRafTimeout$1 = {}; - -Object.defineProperty(setRafTimeout$1, "__esModule", { - value: true -}); -setRafTimeout$1.default = setRafTimeout; - -function setRafTimeout(callback) { - var timeout = arguments.length > 1 && arguments[1] !== undefined ? arguments[1] : 0; - var currTime = -1; - - var shouldUpdate = function shouldUpdate(now) { - if (currTime < 0) { - currTime = now; - } - - if (now - currTime > timeout) { - callback(now); - currTime = -1; - } else { - requestAnimationFrame(shouldUpdate); - } - }; - - requestAnimationFrame(shouldUpdate); -} - -Object.defineProperty(AnimateManager, "__esModule", { - value: true -}); -AnimateManager.default = createAnimateManager; - -var _setRafTimeout = _interopRequireDefault$9(setRafTimeout$1); - -function _interopRequireDefault$9(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -function _typeof$v(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$v = function _typeof(obj) { return typeof obj; }; } else { _typeof$v = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$v(obj); } - -function _toArray(arr) { return _arrayWithHoles$b(arr) || _iterableToArray$c(arr) || _unsupportedIterableToArray$j(arr) || _nonIterableRest$b(); } - -function _nonIterableRest$b() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$j(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$j(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$j(o, minLen); } - -function _arrayLikeToArray$j(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArray$c(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithHoles$b(arr) { if (Array.isArray(arr)) return arr; } - -function createAnimateManager() { - var currStyle = {}; - - var handleChange = function handleChange() { - return null; - }; - - var shouldStop = false; - - var setStyle = function setStyle(_style) { - if (shouldStop) { - return; - } - - if (Array.isArray(_style)) { - if (!_style.length) { - return; - } - - var styles = _style; - - var _styles = _toArray(styles), - curr = _styles[0], - restStyles = _styles.slice(1); - - if (typeof curr === 'number') { - (0, _setRafTimeout.default)(setStyle.bind(null, restStyles), curr); - return; - } - - setStyle(curr); - (0, _setRafTimeout.default)(setStyle.bind(null, restStyles)); - return; - } - - if (_typeof$v(_style) === 'object') { - currStyle = _style; - handleChange(currStyle); - } - - if (typeof _style === 'function') { - _style(); - } - }; - - return { - stop: function stop() { - shouldStop = true; - }, - start: function start(style) { - shouldStop = false; - setStyle(style); - }, - subscribe: function subscribe(_handleChange) { - handleChange = _handleChange; - return function () { - handleChange = function handleChange() { - return null; - }; - }; - } - }; -} - -var easing = {}; - -var util = {}; - -Object.defineProperty(util, "__esModule", { - value: true -}); -util.warn = util.getTransitionVal = util.compose = util.translateStyle = util.mapObject = util.debugf = util.debug = util.log = util.generatePrefixStyle = util.getDashCase = util.identity = util.getIntersectionKeys = void 0; - -function ownKeys$y(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$y(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$y(Object(source), true).forEach(function (key) { _defineProperty$z(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$y(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$z(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -/* eslint no-console: 0 */ -var PREFIX_LIST$1 = ['Webkit', 'Moz', 'O', 'ms']; -var IN_LINE_PREFIX_LIST = ['-webkit-', '-moz-', '-o-', '-ms-']; -var IN_COMPATIBLE_PROPERTY = ['transform', 'transformOrigin', 'transition']; - -var getIntersectionKeys = function getIntersectionKeys(preObj, nextObj) { - return [Object.keys(preObj), Object.keys(nextObj)].reduce(function (a, b) { - return a.filter(function (c) { - return b.includes(c); - }); - }); -}; - -util.getIntersectionKeys = getIntersectionKeys; - -var identity$8 = function identity(param) { - return param; -}; -/* - * @description: convert camel case to dash case - * string => string - */ - - -util.identity = identity$8; - -var getDashCase = function getDashCase(name) { - return name.replace(/([A-Z])/g, function (v) { - return "-".concat(v.toLowerCase()); - }); -}; -/* - * @description: add compatible style prefix - * (string, string) => object - */ - - -util.getDashCase = getDashCase; - -var generatePrefixStyle$1 = function generatePrefixStyle(name, value) { - if (IN_COMPATIBLE_PROPERTY.indexOf(name) === -1) { - return _defineProperty$z({}, name, value); - } - - var isTransition = name === 'transition'; - var camelName = name.replace(/(\w)/, function (v) { - return v.toUpperCase(); - }); - var styleVal = value; - return PREFIX_LIST$1.reduce(function (result, property, i) { - if (isTransition) { - styleVal = value.replace(/(transform|transform-origin)/gim, "".concat(IN_LINE_PREFIX_LIST[i], "$1")); - } - - return _objectSpread$y(_objectSpread$y({}, result), {}, _defineProperty$z({}, property + camelName, styleVal)); - }, {}); -}; - -util.generatePrefixStyle = generatePrefixStyle$1; - -var log$1 = function log() { - var _console; - - (_console = console).log.apply(_console, arguments); -}; -/* - * @description: log the value of a varible - * string => any => any - */ - - -util.log = log$1; - -var debug = function debug(name) { - return function (item) { - log$1(name, item); - return item; - }; -}; -/* - * @description: log name, args, return value of a function - * function => function - */ - - -util.debug = debug; - -var debugf = function debugf(tag, f) { - return function () { - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - var res = f.apply(void 0, args); - var name = tag || f.name || 'anonymous function'; - var argNames = "(".concat(args.map(JSON.stringify).join(', '), ")"); - log$1("".concat(name, ": ").concat(argNames, " => ").concat(JSON.stringify(res))); - return res; - }; -}; -/* - * @description: map object on every element in this object. - * (function, object) => object - */ - - -util.debugf = debugf; - -var mapObject = function mapObject(fn, obj) { - return Object.keys(obj).reduce(function (res, key) { - return _objectSpread$y(_objectSpread$y({}, res), {}, _defineProperty$z({}, key, fn(key, obj[key]))); - }, {}); -}; -/* - * @description: add compatible prefix to style - * object => object - */ - - -util.mapObject = mapObject; - -var translateStyle = function translateStyle(style) { - return Object.keys(style).reduce(function (res, key) { - return _objectSpread$y(_objectSpread$y({}, res), generatePrefixStyle$1(key, res[key])); - }, style); -}; - -util.translateStyle = translateStyle; - -var compose$1 = function compose() { - for (var _len2 = arguments.length, args = new Array(_len2), _key2 = 0; _key2 < _len2; _key2++) { - args[_key2] = arguments[_key2]; - } - - if (!args.length) { - return identity$8; - } - - var fns = args.reverse(); // first function can receive multiply arguments - - var firstFn = fns[0]; - var tailsFn = fns.slice(1); - return function () { - return tailsFn.reduce(function (res, fn) { - return fn(res); - }, firstFn.apply(void 0, arguments)); - }; -}; - -util.compose = compose$1; - -var getTransitionVal = function getTransitionVal(props, duration, easing) { - return props.map(function (prop) { - return "".concat(getDashCase(prop), " ").concat(duration, "ms ").concat(easing); - }).join(','); -}; - -util.getTransitionVal = getTransitionVal; - -var warn$1 = function warn(condition, format, a, b, c, d, e, f) { -}; - -util.warn = warn$1; - -Object.defineProperty(easing, "__esModule", { - value: true -}); -easing.configEasing = easing.configSpring = easing.configBezier = void 0; - -var _util$2 = util; - -function _slicedToArray$a(arr, i) { return _arrayWithHoles$a(arr) || _iterableToArrayLimit$a(arr, i) || _unsupportedIterableToArray$i(arr, i) || _nonIterableRest$a(); } - -function _nonIterableRest$a() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _iterableToArrayLimit$a(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$a(arr) { if (Array.isArray(arr)) return arr; } - -function _toConsumableArray$b(arr) { return _arrayWithoutHoles$b(arr) || _iterableToArray$b(arr) || _unsupportedIterableToArray$i(arr) || _nonIterableSpread$b(); } - -function _nonIterableSpread$b() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$i(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$i(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$i(o, minLen); } - -function _iterableToArray$b(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$b(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$i(arr); } - -function _arrayLikeToArray$i(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -var ACCURACY = 1e-4; - -var cubicBezierFactor = function cubicBezierFactor(c1, c2) { - return [0, 3 * c1, 3 * c2 - 6 * c1, 3 * c1 - 3 * c2 + 1]; -}; - -var multyTime = function multyTime(params, t) { - return params.map(function (param, i) { - return param * Math.pow(t, i); - }).reduce(function (pre, curr) { - return pre + curr; - }); -}; - -var cubicBezier = function cubicBezier(c1, c2) { - return function (t) { - var params = cubicBezierFactor(c1, c2); - return multyTime(params, t); - }; -}; - -var derivativeCubicBezier = function derivativeCubicBezier(c1, c2) { - return function (t) { - var params = cubicBezierFactor(c1, c2); - var newParams = [].concat(_toConsumableArray$b(params.map(function (param, i) { - return param * i; - }).slice(1)), [0]); - return multyTime(newParams, t); - }; -}; // calculate cubic-bezier using Newton's method - - -var configBezier = function configBezier() { - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - var x1 = args[0], - y1 = args[1], - x2 = args[2], - y2 = args[3]; - - if (args.length === 1) { - switch (args[0]) { - case 'linear': - x1 = 0.0; - y1 = 0.0; - x2 = 1.0; - y2 = 1.0; - break; - - case 'ease': - x1 = 0.25; - y1 = 0.1; - x2 = 0.25; - y2 = 1.0; - break; - - case 'ease-in': - x1 = 0.42; - y1 = 0.0; - x2 = 1.0; - y2 = 1.0; - break; - - case 'ease-out': - x1 = 0.42; - y1 = 0.0; - x2 = 0.58; - y2 = 1.0; - break; - - case 'ease-in-out': - x1 = 0.0; - y1 = 0.0; - x2 = 0.58; - y2 = 1.0; - break; - - default: - { - var easing = args[0].split('('); - - if (easing[0] === 'cubic-bezier' && easing[1].split(')')[0].split(',').length === 4) { - var _easing$1$split$0$spl = easing[1].split(')')[0].split(',').map(function (x) { - return parseFloat(x); - }); - - var _easing$1$split$0$spl2 = _slicedToArray$a(_easing$1$split$0$spl, 4); - - x1 = _easing$1$split$0$spl2[0]; - y1 = _easing$1$split$0$spl2[1]; - x2 = _easing$1$split$0$spl2[2]; - y2 = _easing$1$split$0$spl2[3]; - } else { - (0, _util$2.warn)(false, '[configBezier]: arguments should be one of ' + 'oneOf \'linear\', \'ease\', \'ease-in\', \'ease-out\', ' + '\'ease-in-out\',\'cubic-bezier(x1,y1,x2,y2)\', instead received %s', args); - } - } - } - } - - (0, _util$2.warn)([x1, x2, y1, y2].every(function (num) { - return typeof num === 'number' && num >= 0 && num <= 1; - }), '[configBezier]: arguments should be x1, y1, x2, y2 of [0, 1] instead received %s', args); - var curveX = cubicBezier(x1, x2); - var curveY = cubicBezier(y1, y2); - var derCurveX = derivativeCubicBezier(x1, x2); - - var rangeValue = function rangeValue(value) { - if (value > 1) { - return 1; - } else if (value < 0) { - return 0; - } - - return value; - }; - - var bezier = function bezier(_t) { - var t = _t > 1 ? 1 : _t; - var x = t; - - for (var i = 0; i < 8; ++i) { - var evalT = curveX(x) - t; - var derVal = derCurveX(x); - - if (Math.abs(evalT - t) < ACCURACY || derVal < ACCURACY) { - return curveY(x); - } - - x = rangeValue(x - evalT / derVal); - } - - return curveY(x); - }; - - bezier.isStepper = false; - return bezier; -}; - -easing.configBezier = configBezier; - -var configSpring = function configSpring() { - var config = arguments.length > 0 && arguments[0] !== undefined ? arguments[0] : {}; - var _config$stiff = config.stiff, - stiff = _config$stiff === void 0 ? 100 : _config$stiff, - _config$damping = config.damping, - damping = _config$damping === void 0 ? 8 : _config$damping, - _config$dt = config.dt, - dt = _config$dt === void 0 ? 17 : _config$dt; - - var stepper = function stepper(currX, destX, currV) { - var FSpring = -(currX - destX) * stiff; - var FDamping = currV * damping; - var newV = currV + (FSpring - FDamping) * dt / 1000; - var newX = currV * dt / 1000 + currX; - - if (Math.abs(newX - destX) < ACCURACY && Math.abs(newV) < ACCURACY) { - return [destX, 0]; - } - - return [newX, newV]; - }; - - stepper.isStepper = true; - stepper.dt = dt; - return stepper; -}; - -easing.configSpring = configSpring; - -var configEasing = function configEasing() { - for (var _len2 = arguments.length, args = new Array(_len2), _key2 = 0; _key2 < _len2; _key2++) { - args[_key2] = arguments[_key2]; - } - - var easing = args[0]; - - if (typeof easing === 'string') { - switch (easing) { - case 'ease': - case 'ease-in-out': - case 'ease-out': - case 'ease-in': - case 'linear': - return configBezier(easing); - - case 'spring': - return configSpring(); - - default: - if (easing.split('(')[0] === 'cubic-bezier') { - return configBezier(easing); - } - - (0, _util$2.warn)(false, '[configEasing]: first argument should be one of \'ease\', \'ease-in\', ' + '\'ease-out\', \'ease-in-out\',\'cubic-bezier(x1,y1,x2,y2)\', \'linear\' and \'spring\', instead received %s', args); - } - } - - if (typeof easing === 'function') { - return easing; - } - - (0, _util$2.warn)(false, '[configEasing]: first argument type should be function or ' + 'string, instead received %s', args); - return null; -}; - -easing.configEasing = configEasing; - -var configUpdate = {}; - -Object.defineProperty(configUpdate, "__esModule", { - value: true -}); -configUpdate.default = void 0; - -var _util$1 = util; - -function _toConsumableArray$a(arr) { return _arrayWithoutHoles$a(arr) || _iterableToArray$a(arr) || _unsupportedIterableToArray$h(arr) || _nonIterableSpread$a(); } - -function _nonIterableSpread$a() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _iterableToArray$a(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$a(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$h(arr); } - -function ownKeys$x(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$x(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$x(Object(source), true).forEach(function (key) { _defineProperty$y(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$x(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$y(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _slicedToArray$9(arr, i) { return _arrayWithHoles$9(arr) || _iterableToArrayLimit$9(arr, i) || _unsupportedIterableToArray$h(arr, i) || _nonIterableRest$9(); } - -function _nonIterableRest$9() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$h(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$h(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$h(o, minLen); } - -function _arrayLikeToArray$h(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit$9(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$9(arr) { if (Array.isArray(arr)) return arr; } - -var alpha = function alpha(begin, end, k) { - return begin + (end - begin) * k; -}; - -var needContinue = function needContinue(_ref) { - var from = _ref.from, - to = _ref.to; - return from !== to; -}; -/* - * @description: cal new from value and velocity in each stepper - * @return: { [styleProperty]: { from, to, velocity } } - */ - - -var calStepperVals = function calStepperVals(easing, preVals, steps) { - var nextStepVals = (0, _util$1.mapObject)(function (key, val) { - if (needContinue(val)) { - var _easing = easing(val.from, val.to, val.velocity), - _easing2 = _slicedToArray$9(_easing, 2), - newX = _easing2[0], - newV = _easing2[1]; - - return _objectSpread$x(_objectSpread$x({}, val), {}, { - from: newX, - velocity: newV - }); - } - - return val; - }, preVals); - - if (steps < 1) { - return (0, _util$1.mapObject)(function (key, val) { - if (needContinue(val)) { - return _objectSpread$x(_objectSpread$x({}, val), {}, { - velocity: alpha(val.velocity, nextStepVals[key].velocity, steps), - from: alpha(val.from, nextStepVals[key].from, steps) - }); - } - - return val; - }, preVals); - } - - return calStepperVals(easing, nextStepVals, steps - 1); -}; // configure update function - - -var _default$5 = function _default(from, to, easing, duration, render) { - var interKeys = (0, _util$1.getIntersectionKeys)(from, to); - var timingStyle = interKeys.reduce(function (res, key) { - return _objectSpread$x(_objectSpread$x({}, res), {}, _defineProperty$y({}, key, [from[key], to[key]])); - }, {}); - var stepperStyle = interKeys.reduce(function (res, key) { - return _objectSpread$x(_objectSpread$x({}, res), {}, _defineProperty$y({}, key, { - from: from[key], - velocity: 0, - to: to[key] - })); - }, {}); - var cafId = -1; - var preTime; - var beginTime; - - var update = function update() { - return null; - }; - - var getCurrStyle = function getCurrStyle() { - return (0, _util$1.mapObject)(function (key, val) { - return val.from; - }, stepperStyle); - }; - - var shouldStopAnimation = function shouldStopAnimation() { - return !Object.values(stepperStyle).filter(needContinue).length; - }; // stepper timing function like spring - - - var stepperUpdate = function stepperUpdate(now) { - if (!preTime) { - preTime = now; - } - - var deltaTime = now - preTime; - var steps = deltaTime / easing.dt; - stepperStyle = calStepperVals(easing, stepperStyle, steps); // get union set and add compatible prefix - - render(_objectSpread$x(_objectSpread$x(_objectSpread$x({}, from), to), getCurrStyle())); - preTime = now; - - if (!shouldStopAnimation()) { - cafId = requestAnimationFrame(update); - } - }; // t => val timing function like cubic-bezier - - - var timingUpdate = function timingUpdate(now) { - if (!beginTime) { - beginTime = now; - } - - var t = (now - beginTime) / duration; - var currStyle = (0, _util$1.mapObject)(function (key, val) { - return alpha.apply(void 0, _toConsumableArray$a(val).concat([easing(t)])); - }, timingStyle); // get union set and add compatible prefix - - render(_objectSpread$x(_objectSpread$x(_objectSpread$x({}, from), to), currStyle)); - - if (t < 1) { - cafId = requestAnimationFrame(update); - } else { - var finalStyle = (0, _util$1.mapObject)(function (key, val) { - return alpha.apply(void 0, _toConsumableArray$a(val).concat([easing(1)])); - }, timingStyle); - render(_objectSpread$x(_objectSpread$x(_objectSpread$x({}, from), to), finalStyle)); - } - }; - - update = easing.isStepper ? stepperUpdate : timingUpdate; // return start animation method - - return function () { - requestAnimationFrame(update); // return stop animation method - - return function () { - cancelAnimationFrame(cafId); - }; - }; -}; - -configUpdate.default = _default$5; - -function _typeof$u(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$u = function _typeof(obj) { return typeof obj; }; } else { _typeof$u = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$u(obj); } - -Object.defineProperty(Animate$2, "__esModule", { - value: true -}); -Animate$2.default = void 0; - -var _react$4 = _interopRequireWildcard$3(React__default); - -var _propTypes$2 = _interopRequireDefault$8(propTypes.exports); - -var _fastEquals = fastEquals.exports; - -var _AnimateManager = _interopRequireDefault$8(AnimateManager); - -var _easing = easing; - -var _configUpdate = _interopRequireDefault$8(configUpdate); - -var _util = util; - -function _interopRequireDefault$8(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -function _getRequireWildcardCache$2() { if (typeof WeakMap !== "function") return null; var cache = new WeakMap(); _getRequireWildcardCache$2 = function _getRequireWildcardCache() { return cache; }; return cache; } - -function _interopRequireWildcard$3(obj) { if (obj && obj.__esModule) { return obj; } if (obj === null || _typeof$u(obj) !== "object" && typeof obj !== "function") { return { default: obj }; } var cache = _getRequireWildcardCache$2(); if (cache && cache.has(obj)) { return cache.get(obj); } var newObj = {}; var hasPropertyDescriptor = Object.defineProperty && Object.getOwnPropertyDescriptor; for (var key in obj) { if (Object.prototype.hasOwnProperty.call(obj, key)) { var desc = hasPropertyDescriptor ? Object.getOwnPropertyDescriptor(obj, key) : null; if (desc && (desc.get || desc.set)) { Object.defineProperty(newObj, key, desc); } else { newObj[key] = obj[key]; } } } newObj.default = obj; if (cache) { cache.set(obj, newObj); } return newObj; } - -function _objectWithoutProperties$h(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$i(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$i(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _toConsumableArray$9(arr) { return _arrayWithoutHoles$9(arr) || _iterableToArray$9(arr) || _unsupportedIterableToArray$g(arr) || _nonIterableSpread$9(); } - -function _nonIterableSpread$9() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$g(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$g(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$g(o, minLen); } - -function _iterableToArray$9(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$9(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$g(arr); } - -function _arrayLikeToArray$g(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function ownKeys$w(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$w(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$w(Object(source), true).forEach(function (key) { _defineProperty$x(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$w(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$x(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$t(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$t(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$t(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$t(Constructor.prototype, protoProps); if (staticProps) _defineProperties$t(Constructor, staticProps); return Constructor; } - -function _inherits$s(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$s(subClass, superClass); } - -function _setPrototypeOf$s(o, p) { _setPrototypeOf$s = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$s(o, p); } - -function _createSuper$s(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$s(); return function _createSuperInternal() { var Super = _getPrototypeOf$s(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$s(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$s(this, result); }; } - -function _possibleConstructorReturn$s(self, call) { if (call && (_typeof$u(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$s(self); } - -function _assertThisInitialized$s(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$s() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$s(o) { _getPrototypeOf$s = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$s(o); } - -var Animate$1 = /*#__PURE__*/function (_PureComponent) { - _inherits$s(Animate, _PureComponent); - - var _super = _createSuper$s(Animate); - - function Animate(props, context) { - var _this; - - _classCallCheck$t(this, Animate); - - _this = _super.call(this, props, context); - var _this$props = _this.props, - isActive = _this$props.isActive, - attributeName = _this$props.attributeName, - from = _this$props.from, - to = _this$props.to, - steps = _this$props.steps, - children = _this$props.children; - _this.handleStyleChange = _this.handleStyleChange.bind(_assertThisInitialized$s(_this)); - _this.changeStyle = _this.changeStyle.bind(_assertThisInitialized$s(_this)); - - if (!isActive) { - _this.state = { - style: {} - }; // if children is a function and animation is not active, set style to 'to' - - if (typeof children === 'function') { - _this.state = { - style: to - }; - } - - return _possibleConstructorReturn$s(_this); - } - - if (steps && steps.length) { - _this.state = { - style: steps[0].style - }; - } else if (from) { - if (typeof children === 'function') { - _this.state = { - style: from - }; - return _possibleConstructorReturn$s(_this); - } - - _this.state = { - style: attributeName ? _defineProperty$x({}, attributeName, from) : from - }; - } else { - _this.state = { - style: {} - }; - } - - return _this; - } - - _createClass$t(Animate, [{ - key: "componentDidMount", - value: function componentDidMount() { - var _this$props2 = this.props, - isActive = _this$props2.isActive, - canBegin = _this$props2.canBegin; - this.mounted = true; - - if (!isActive || !canBegin) { - return; - } - - this.runAnimation(this.props); - } - }, { - key: "componentDidUpdate", - value: function componentDidUpdate(prevProps) { - var _this$props3 = this.props, - isActive = _this$props3.isActive, - canBegin = _this$props3.canBegin, - attributeName = _this$props3.attributeName, - shouldReAnimate = _this$props3.shouldReAnimate; - - if (!canBegin) { - return; - } - - if (!isActive) { - var newState = { - style: attributeName ? _defineProperty$x({}, attributeName, this.props.to) : this.props.to - }; - - if (this.state && this.state.style) { - if (attributeName && this.state.style[attributeName] !== this.props.to || !attributeName && this.state.style !== this.props.to) { - // eslint-disable-next-line react/no-did-update-set-state - this.setState(newState); - } - } - - return; - } - - if ((0, _fastEquals.deepEqual)(prevProps.to, this.props.to) && prevProps.canBegin && prevProps.isActive) { - return; - } - - var isTriggered = !prevProps.canBegin || !prevProps.isActive; - - if (this.manager) { - this.manager.stop(); - } - - if (this.stopJSAnimation) { - this.stopJSAnimation(); - } - - var from = isTriggered || shouldReAnimate ? this.props.from : prevProps.to; - - if (this.state && this.state.style) { - var _newState = { - style: attributeName ? _defineProperty$x({}, attributeName, from) : from - }; - - if (attributeName && this.state.style[attributeName] !== from || !attributeName && this.state.style !== from) { - // eslint-disable-next-line react/no-did-update-set-state - this.setState(_newState); - } - } - - this.runAnimation(_objectSpread$w(_objectSpread$w({}, this.props), {}, { - from: from, - begin: 0 - })); - } - }, { - key: "componentWillUnmount", - value: function componentWillUnmount() { - this.mounted = false; - - if (this.unSubscribe) { - this.unSubscribe(); - } - - if (this.manager) { - this.manager.stop(); - this.manager = null; - } - - if (this.stopJSAnimation) { - this.stopJSAnimation(); - } - } - }, { - key: "runJSAnimation", - value: function runJSAnimation(props) { - var _this2 = this; - - var from = props.from, - to = props.to, - duration = props.duration, - easing = props.easing, - begin = props.begin, - onAnimationEnd = props.onAnimationEnd, - onAnimationStart = props.onAnimationStart; - var startAnimation = (0, _configUpdate.default)(from, to, (0, _easing.configEasing)(easing), duration, this.changeStyle); - - var finalStartAnimation = function finalStartAnimation() { - _this2.stopJSAnimation = startAnimation(); - }; - - this.manager.start([onAnimationStart, begin, finalStartAnimation, duration, onAnimationEnd]); - } - }, { - key: "runStepAnimation", - value: function runStepAnimation(props) { - var _this3 = this; - - var steps = props.steps, - begin = props.begin, - onAnimationStart = props.onAnimationStart; - var _steps$ = steps[0], - initialStyle = _steps$.style, - _steps$$duration = _steps$.duration, - initialTime = _steps$$duration === void 0 ? 0 : _steps$$duration; - - var addStyle = function addStyle(sequence, nextItem, index) { - if (index === 0) { - return sequence; - } - - var duration = nextItem.duration, - _nextItem$easing = nextItem.easing, - easing = _nextItem$easing === void 0 ? 'ease' : _nextItem$easing, - style = nextItem.style, - nextProperties = nextItem.properties, - onAnimationEnd = nextItem.onAnimationEnd; - var preItem = index > 0 ? steps[index - 1] : nextItem; - var properties = nextProperties || Object.keys(style); - - if (typeof easing === 'function' || easing === 'spring') { - return [].concat(_toConsumableArray$9(sequence), [_this3.runJSAnimation.bind(_this3, { - from: preItem.style, - to: style, - duration: duration, - easing: easing - }), duration]); - } - - var transition = (0, _util.getTransitionVal)(properties, duration, easing); - - var newStyle = _objectSpread$w(_objectSpread$w(_objectSpread$w({}, preItem.style), style), {}, { - transition: transition - }); - - return [].concat(_toConsumableArray$9(sequence), [newStyle, duration, onAnimationEnd]).filter(_util.identity); - }; - - return this.manager.start([onAnimationStart].concat(_toConsumableArray$9(steps.reduce(addStyle, [initialStyle, Math.max(initialTime, begin)])), [props.onAnimationEnd])); - } - }, { - key: "runAnimation", - value: function runAnimation(props) { - if (!this.manager) { - this.manager = (0, _AnimateManager.default)(); - } - - var begin = props.begin, - duration = props.duration, - attributeName = props.attributeName, - propsTo = props.to, - easing = props.easing, - onAnimationStart = props.onAnimationStart, - onAnimationEnd = props.onAnimationEnd, - steps = props.steps, - children = props.children; - var manager = this.manager; - this.unSubscribe = manager.subscribe(this.handleStyleChange); - - if (typeof easing === 'function' || typeof children === 'function' || easing === 'spring') { - this.runJSAnimation(props); - return; - } - - if (steps.length > 1) { - this.runStepAnimation(props); - return; - } - - var to = attributeName ? _defineProperty$x({}, attributeName, propsTo) : propsTo; - var transition = (0, _util.getTransitionVal)(Object.keys(to), duration, easing); - manager.start([onAnimationStart, begin, _objectSpread$w(_objectSpread$w({}, to), {}, { - transition: transition - }), duration, onAnimationEnd]); - } - }, { - key: "handleStyleChange", - value: function handleStyleChange(style) { - this.changeStyle(style); - } - }, { - key: "changeStyle", - value: function changeStyle(style) { - if (this.mounted) { - this.setState({ - style: style - }); - } - } - }, { - key: "render", - value: function render() { - var _this$props4 = this.props, - children = _this$props4.children; - _this$props4.begin; - _this$props4.duration; - _this$props4.attributeName; - _this$props4.easing; - var isActive = _this$props4.isActive; - _this$props4.steps; - _this$props4.from; - _this$props4.to; - _this$props4.canBegin; - _this$props4.onAnimationEnd; - _this$props4.shouldReAnimate; - _this$props4.onAnimationReStart; - var others = _objectWithoutProperties$h(_this$props4, ["children", "begin", "duration", "attributeName", "easing", "isActive", "steps", "from", "to", "canBegin", "onAnimationEnd", "shouldReAnimate", "onAnimationReStart"]); - - var count = _react$4.Children.count(children); - - var stateStyle = (0, _util.translateStyle)(this.state.style); - - if (typeof children === 'function') { - return children(stateStyle); - } - - if (!isActive || count === 0) { - return children; - } - - var cloneContainer = function cloneContainer(container) { - var _container$props = container.props, - _container$props$styl = _container$props.style, - style = _container$props$styl === void 0 ? {} : _container$props$styl, - className = _container$props.className; - var res = /*#__PURE__*/(0, _react$4.cloneElement)(container, _objectSpread$w(_objectSpread$w({}, others), {}, { - style: _objectSpread$w(_objectSpread$w({}, style), stateStyle), - className: className - })); - return res; - }; - - if (count === 1) { - return cloneContainer(_react$4.Children.only(children)); - } - - return /*#__PURE__*/_react$4.default.createElement("div", null, _react$4.Children.map(children, function (child) { - return cloneContainer(child); - })); - } - }]); - - return Animate; -}(_react$4.PureComponent); - -Animate$1.displayName = 'Animate'; -Animate$1.propTypes = { - from: _propTypes$2.default.oneOfType([_propTypes$2.default.object, _propTypes$2.default.string]), - to: _propTypes$2.default.oneOfType([_propTypes$2.default.object, _propTypes$2.default.string]), - attributeName: _propTypes$2.default.string, - // animation duration - duration: _propTypes$2.default.number, - begin: _propTypes$2.default.number, - easing: _propTypes$2.default.oneOfType([_propTypes$2.default.string, _propTypes$2.default.func]), - steps: _propTypes$2.default.arrayOf(_propTypes$2.default.shape({ - duration: _propTypes$2.default.number.isRequired, - style: _propTypes$2.default.object.isRequired, - easing: _propTypes$2.default.oneOfType([_propTypes$2.default.oneOf(['ease', 'ease-in', 'ease-out', 'ease-in-out', 'linear']), _propTypes$2.default.func]), - // transition css properties(dash case), optional - properties: _propTypes$2.default.arrayOf('string'), - onAnimationEnd: _propTypes$2.default.func - })), - children: _propTypes$2.default.oneOfType([_propTypes$2.default.node, _propTypes$2.default.func]), - isActive: _propTypes$2.default.bool, - canBegin: _propTypes$2.default.bool, - onAnimationEnd: _propTypes$2.default.func, - // decide if it should reanimate with initial from style when props change - shouldReAnimate: _propTypes$2.default.bool, - onAnimationStart: _propTypes$2.default.func, - onAnimationReStart: _propTypes$2.default.func -}; -Animate$1.defaultProps = { - begin: 0, - duration: 1000, - from: '', - to: '', - attributeName: '', - easing: 'ease', - isActive: true, - canBegin: true, - steps: [], - onAnimationEnd: function onAnimationEnd() {}, - onAnimationStart: function onAnimationStart() {} -}; -var _default$4 = Animate$1; -Animate$2.default = _default$4; - -var AnimateGroup$1 = {}; - -var CSSTransition = {exports: {}}; - -var addClass = {exports: {}}; - -var interopRequireDefault = {exports: {}}; - -(function (module) { - function _interopRequireDefault(obj) { - return obj && obj.__esModule ? obj : { - "default": obj - }; - } - module.exports = _interopRequireDefault, module.exports.__esModule = true, module.exports["default"] = module.exports; -} (interopRequireDefault)); - -var hasClass = {exports: {}}; - -var hasRequiredHasClass; - -function requireHasClass () { - if (hasRequiredHasClass) return hasClass.exports; - hasRequiredHasClass = 1; - (function (module, exports) { - - exports.__esModule = true; - exports.default = hasClass; - - function hasClass(element, className) { - if (element.classList) return !!className && element.classList.contains(className);else return (" " + (element.className.baseVal || element.className) + " ").indexOf(" " + className + " ") !== -1; - } - - module.exports = exports["default"]; -} (hasClass, hasClass.exports)); - return hasClass.exports; -} - -(function (module, exports) { - - var _interopRequireDefault = interopRequireDefault.exports; - - exports.__esModule = true; - exports.default = addClass; - - var _hasClass = _interopRequireDefault(requireHasClass()); - - function addClass(element, className) { - if (element.classList) element.classList.add(className);else if (!(0, _hasClass.default)(element, className)) if (typeof element.className === 'string') element.className = element.className + ' ' + className;else element.setAttribute('class', (element.className && element.className.baseVal || '') + ' ' + className); - } - - module.exports = exports["default"]; -} (addClass, addClass.exports)); - -function replaceClassName(origClass, classToRemove) { - return origClass.replace(new RegExp('(^|\\s)' + classToRemove + '(?:\\s|$)', 'g'), '$1').replace(/\s+/g, ' ').replace(/^\s*|\s*$/g, ''); -} - -var removeClass = function removeClass(element, className) { - if (element.classList) element.classList.remove(className);else if (typeof element.className === 'string') element.className = replaceClassName(element.className, className);else element.setAttribute('class', replaceClassName(element.className && element.className.baseVal || '', className)); -}; - -var Transition$1 = {}; - -var reactLifecyclesCompat_cjs = {}; - -Object.defineProperty(reactLifecyclesCompat_cjs, '__esModule', { value: true }); - -/** - * Copyright (c) 2013-present, Facebook, Inc. - * - * This source code is licensed under the MIT license found in the - * LICENSE file in the root directory of this source tree. - */ - -function componentWillMount() { - // Call this.constructor.gDSFP to support sub-classes. - var state = this.constructor.getDerivedStateFromProps(this.props, this.state); - if (state !== null && state !== undefined) { - this.setState(state); - } -} - -function componentWillReceiveProps(nextProps) { - // Call this.constructor.gDSFP to support sub-classes. - // Use the setState() updater to ensure state isn't stale in certain edge cases. - function updater(prevState) { - var state = this.constructor.getDerivedStateFromProps(nextProps, prevState); - return state !== null && state !== undefined ? state : null; - } - // Binding "this" is important for shallow renderer support. - this.setState(updater.bind(this)); -} - -function componentWillUpdate(nextProps, nextState) { - try { - var prevProps = this.props; - var prevState = this.state; - this.props = nextProps; - this.state = nextState; - this.__reactInternalSnapshotFlag = true; - this.__reactInternalSnapshot = this.getSnapshotBeforeUpdate( - prevProps, - prevState - ); - } finally { - this.props = prevProps; - this.state = prevState; - } -} - -// React may warn about cWM/cWRP/cWU methods being deprecated. -// Add a flag to suppress these warnings for this special case. -componentWillMount.__suppressDeprecationWarning = true; -componentWillReceiveProps.__suppressDeprecationWarning = true; -componentWillUpdate.__suppressDeprecationWarning = true; - -function polyfill(Component) { - var prototype = Component.prototype; - - if (!prototype || !prototype.isReactComponent) { - throw new Error('Can only polyfill class components'); - } - - if ( - typeof Component.getDerivedStateFromProps !== 'function' && - typeof prototype.getSnapshotBeforeUpdate !== 'function' - ) { - return Component; - } - - // If new component APIs are defined, "unsafe" lifecycles won't be called. - // Error if any of these lifecycles are present, - // Because they would work differently between older and newer (16.3+) versions of React. - var foundWillMountName = null; - var foundWillReceivePropsName = null; - var foundWillUpdateName = null; - if (typeof prototype.componentWillMount === 'function') { - foundWillMountName = 'componentWillMount'; - } else if (typeof prototype.UNSAFE_componentWillMount === 'function') { - foundWillMountName = 'UNSAFE_componentWillMount'; - } - if (typeof prototype.componentWillReceiveProps === 'function') { - foundWillReceivePropsName = 'componentWillReceiveProps'; - } else if (typeof prototype.UNSAFE_componentWillReceiveProps === 'function') { - foundWillReceivePropsName = 'UNSAFE_componentWillReceiveProps'; - } - if (typeof prototype.componentWillUpdate === 'function') { - foundWillUpdateName = 'componentWillUpdate'; - } else if (typeof prototype.UNSAFE_componentWillUpdate === 'function') { - foundWillUpdateName = 'UNSAFE_componentWillUpdate'; - } - if ( - foundWillMountName !== null || - foundWillReceivePropsName !== null || - foundWillUpdateName !== null - ) { - var componentName = Component.displayName || Component.name; - var newApiName = - typeof Component.getDerivedStateFromProps === 'function' - ? 'getDerivedStateFromProps()' - : 'getSnapshotBeforeUpdate()'; - - throw Error( - 'Unsafe legacy lifecycles will not be called for components using new component APIs.\n\n' + - componentName + - ' uses ' + - newApiName + - ' but also contains the following legacy lifecycles:' + - (foundWillMountName !== null ? '\n ' + foundWillMountName : '') + - (foundWillReceivePropsName !== null - ? '\n ' + foundWillReceivePropsName - : '') + - (foundWillUpdateName !== null ? '\n ' + foundWillUpdateName : '') + - '\n\nThe above lifecycles should be removed. Learn more about this warning here:\n' + - 'https://fb.me/react-async-component-lifecycle-hooks' - ); - } - - // React <= 16.2 does not support static getDerivedStateFromProps. - // As a workaround, use cWM and cWRP to invoke the new static lifecycle. - // Newer versions of React will ignore these lifecycles if gDSFP exists. - if (typeof Component.getDerivedStateFromProps === 'function') { - prototype.componentWillMount = componentWillMount; - prototype.componentWillReceiveProps = componentWillReceiveProps; - } - - // React <= 16.2 does not support getSnapshotBeforeUpdate. - // As a workaround, use cWU to invoke the new lifecycle. - // Newer versions of React will ignore that lifecycle if gSBU exists. - if (typeof prototype.getSnapshotBeforeUpdate === 'function') { - if (typeof prototype.componentDidUpdate !== 'function') { - throw new Error( - 'Cannot polyfill getSnapshotBeforeUpdate() for components that do not define componentDidUpdate() on the prototype' - ); - } - - prototype.componentWillUpdate = componentWillUpdate; - - var componentDidUpdate = prototype.componentDidUpdate; - - prototype.componentDidUpdate = function componentDidUpdatePolyfill( - prevProps, - prevState, - maybeSnapshot - ) { - // 16.3+ will not execute our will-update method; - // It will pass a snapshot value to did-update though. - // Older versions will require our polyfilled will-update value. - // We need to handle both cases, but can't just check for the presence of "maybeSnapshot", - // Because for <= 15.x versions this might be a "prevContext" object. - // We also can't just check "__reactInternalSnapshot", - // Because get-snapshot might return a falsy value. - // So check for the explicit __reactInternalSnapshotFlag flag to determine behavior. - var snapshot = this.__reactInternalSnapshotFlag - ? this.__reactInternalSnapshot - : maybeSnapshot; - - componentDidUpdate.call(this, prevProps, prevState, snapshot); - }; - } - - return Component; -} - -reactLifecyclesCompat_cjs.polyfill = polyfill; - -_interopRequireDefault$7(propTypes.exports); - -function _interopRequireDefault$7(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -Transition$1.__esModule = true; -Transition$1.default = Transition$1.EXITING = Transition$1.ENTERED = Transition$1.ENTERING = Transition$1.EXITED = Transition$1.UNMOUNTED = void 0; - -var PropTypes = _interopRequireWildcard$2(propTypes.exports); - -var _react$3 = _interopRequireDefault$6(React__default); - -var _reactDom = _interopRequireDefault$6(require$$2); - -var _reactLifecyclesCompat = reactLifecyclesCompat_cjs; - -function _interopRequireDefault$6(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -function _interopRequireWildcard$2(obj) { if (obj && obj.__esModule) { return obj; } else { var newObj = {}; if (obj != null) { for (var key in obj) { if (Object.prototype.hasOwnProperty.call(obj, key)) { var desc = Object.defineProperty && Object.getOwnPropertyDescriptor ? Object.getOwnPropertyDescriptor(obj, key) : {}; if (desc.get || desc.set) { Object.defineProperty(newObj, key, desc); } else { newObj[key] = obj[key]; } } } } newObj.default = obj; return newObj; } } - -function _objectWithoutPropertiesLoose$h(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _inheritsLoose(subClass, superClass) { subClass.prototype = Object.create(superClass.prototype); subClass.prototype.constructor = subClass; subClass.__proto__ = superClass; } - -var UNMOUNTED = 'unmounted'; -Transition$1.UNMOUNTED = UNMOUNTED; -var EXITED = 'exited'; -Transition$1.EXITED = EXITED; -var ENTERING = 'entering'; -Transition$1.ENTERING = ENTERING; -var ENTERED = 'entered'; -Transition$1.ENTERED = ENTERED; -var EXITING = 'exiting'; -/** - * The Transition component lets you describe a transition from one component - * state to another _over time_ with a simple declarative API. Most commonly - * it's used to animate the mounting and unmounting of a component, but can also - * be used to describe in-place transition states as well. - * - * --- - * - * **Note**: `Transition` is a platform-agnostic base component. If you're using - * transitions in CSS, you'll probably want to use - * [`CSSTransition`](https://reactcommunity.org/react-transition-group/css-transition) - * instead. It inherits all the features of `Transition`, but contains - * additional features necessary to play nice with CSS transitions (hence the - * name of the component). - * - * --- - * - * By default the `Transition` component does not alter the behavior of the - * component it renders, it only tracks "enter" and "exit" states for the - * components. It's up to you to give meaning and effect to those states. For - * example we can add styles to a component when it enters or exits: - * - * ```jsx - * import { Transition } from 'react-transition-group'; - * - * const duration = 300; - * - * const defaultStyle = { - * transition: `opacity ${duration}ms ease-in-out`, - * opacity: 0, - * } - * - * const transitionStyles = { - * entering: { opacity: 0 }, - * entered: { opacity: 1 }, - * }; - * - * const Fade = ({ in: inProp }) => ( - * - * {state => ( - *
- * I'm a fade Transition! - *
- * )} - *
- * ); - * ``` - * - * There are 4 main states a Transition can be in: - * - `'entering'` - * - `'entered'` - * - `'exiting'` - * - `'exited'` - * - * Transition state is toggled via the `in` prop. When `true` the component - * begins the "Enter" stage. During this stage, the component will shift from - * its current transition state, to `'entering'` for the duration of the - * transition and then to the `'entered'` stage once it's complete. Let's take - * the following example (we'll use the - * [useState](https://reactjs.org/docs/hooks-reference.html#usestate) hook): - * - * ```jsx - * function App() { - * const [inProp, setInProp] = useState(false); - * return ( - *
- * - * {state => ( - * // ... - * )} - * - * - *
- * ); - * } - * ``` - * - * When the button is clicked the component will shift to the `'entering'` state - * and stay there for 500ms (the value of `timeout`) before it finally switches - * to `'entered'`. - * - * When `in` is `false` the same thing happens except the state moves from - * `'exiting'` to `'exited'`. - */ - -Transition$1.EXITING = EXITING; - -var Transition = -/*#__PURE__*/ -function (_React$Component) { - _inheritsLoose(Transition, _React$Component); - - function Transition(props, context) { - var _this; - - _this = _React$Component.call(this, props, context) || this; - var parentGroup = context.transitionGroup; // In the context of a TransitionGroup all enters are really appears - - var appear = parentGroup && !parentGroup.isMounting ? props.enter : props.appear; - var initialStatus; - _this.appearStatus = null; - - if (props.in) { - if (appear) { - initialStatus = EXITED; - _this.appearStatus = ENTERING; - } else { - initialStatus = ENTERED; - } - } else { - if (props.unmountOnExit || props.mountOnEnter) { - initialStatus = UNMOUNTED; - } else { - initialStatus = EXITED; - } - } - - _this.state = { - status: initialStatus - }; - _this.nextCallback = null; - return _this; - } - - var _proto = Transition.prototype; - - _proto.getChildContext = function getChildContext() { - return { - transitionGroup: null // allows for nested Transitions - - }; - }; - - Transition.getDerivedStateFromProps = function getDerivedStateFromProps(_ref, prevState) { - var nextIn = _ref.in; - - if (nextIn && prevState.status === UNMOUNTED) { - return { - status: EXITED - }; - } - - return null; - }; // getSnapshotBeforeUpdate(prevProps) { - // let nextStatus = null - // if (prevProps !== this.props) { - // const { status } = this.state - // if (this.props.in) { - // if (status !== ENTERING && status !== ENTERED) { - // nextStatus = ENTERING - // } - // } else { - // if (status === ENTERING || status === ENTERED) { - // nextStatus = EXITING - // } - // } - // } - // return { nextStatus } - // } - - - _proto.componentDidMount = function componentDidMount() { - this.updateStatus(true, this.appearStatus); - }; - - _proto.componentDidUpdate = function componentDidUpdate(prevProps) { - var nextStatus = null; - - if (prevProps !== this.props) { - var status = this.state.status; - - if (this.props.in) { - if (status !== ENTERING && status !== ENTERED) { - nextStatus = ENTERING; - } - } else { - if (status === ENTERING || status === ENTERED) { - nextStatus = EXITING; - } - } - } - - this.updateStatus(false, nextStatus); - }; - - _proto.componentWillUnmount = function componentWillUnmount() { - this.cancelNextCallback(); - }; - - _proto.getTimeouts = function getTimeouts() { - var timeout = this.props.timeout; - var exit, enter, appear; - exit = enter = appear = timeout; - - if (timeout != null && typeof timeout !== 'number') { - exit = timeout.exit; - enter = timeout.enter; // TODO: remove fallback for next major - - appear = timeout.appear !== undefined ? timeout.appear : enter; - } - - return { - exit: exit, - enter: enter, - appear: appear - }; - }; - - _proto.updateStatus = function updateStatus(mounting, nextStatus) { - if (mounting === void 0) { - mounting = false; - } - - if (nextStatus !== null) { - // nextStatus will always be ENTERING or EXITING. - this.cancelNextCallback(); - - var node = _reactDom.default.findDOMNode(this); - - if (nextStatus === ENTERING) { - this.performEnter(node, mounting); - } else { - this.performExit(node); - } - } else if (this.props.unmountOnExit && this.state.status === EXITED) { - this.setState({ - status: UNMOUNTED - }); - } - }; - - _proto.performEnter = function performEnter(node, mounting) { - var _this2 = this; - - var enter = this.props.enter; - var appearing = this.context.transitionGroup ? this.context.transitionGroup.isMounting : mounting; - var timeouts = this.getTimeouts(); - var enterTimeout = appearing ? timeouts.appear : timeouts.enter; // no enter animation skip right to ENTERED - // if we are mounting and running this it means appear _must_ be set - - if (!mounting && !enter) { - this.safeSetState({ - status: ENTERED - }, function () { - _this2.props.onEntered(node); - }); - return; - } - - this.props.onEnter(node, appearing); - this.safeSetState({ - status: ENTERING - }, function () { - _this2.props.onEntering(node, appearing); - - _this2.onTransitionEnd(node, enterTimeout, function () { - _this2.safeSetState({ - status: ENTERED - }, function () { - _this2.props.onEntered(node, appearing); - }); - }); - }); - }; - - _proto.performExit = function performExit(node) { - var _this3 = this; - - var exit = this.props.exit; - var timeouts = this.getTimeouts(); // no exit animation skip right to EXITED - - if (!exit) { - this.safeSetState({ - status: EXITED - }, function () { - _this3.props.onExited(node); - }); - return; - } - - this.props.onExit(node); - this.safeSetState({ - status: EXITING - }, function () { - _this3.props.onExiting(node); - - _this3.onTransitionEnd(node, timeouts.exit, function () { - _this3.safeSetState({ - status: EXITED - }, function () { - _this3.props.onExited(node); - }); - }); - }); - }; - - _proto.cancelNextCallback = function cancelNextCallback() { - if (this.nextCallback !== null) { - this.nextCallback.cancel(); - this.nextCallback = null; - } - }; - - _proto.safeSetState = function safeSetState(nextState, callback) { - // This shouldn't be necessary, but there are weird race conditions with - // setState callbacks and unmounting in testing, so always make sure that - // we can cancel any pending setState callbacks after we unmount. - callback = this.setNextCallback(callback); - this.setState(nextState, callback); - }; - - _proto.setNextCallback = function setNextCallback(callback) { - var _this4 = this; - - var active = true; - - this.nextCallback = function (event) { - if (active) { - active = false; - _this4.nextCallback = null; - callback(event); - } - }; - - this.nextCallback.cancel = function () { - active = false; - }; - - return this.nextCallback; - }; - - _proto.onTransitionEnd = function onTransitionEnd(node, timeout, handler) { - this.setNextCallback(handler); - var doesNotHaveTimeoutOrListener = timeout == null && !this.props.addEndListener; - - if (!node || doesNotHaveTimeoutOrListener) { - setTimeout(this.nextCallback, 0); - return; - } - - if (this.props.addEndListener) { - this.props.addEndListener(node, this.nextCallback); - } - - if (timeout != null) { - setTimeout(this.nextCallback, timeout); - } - }; - - _proto.render = function render() { - var status = this.state.status; - - if (status === UNMOUNTED) { - return null; - } - - var _this$props = this.props, - children = _this$props.children, - childProps = _objectWithoutPropertiesLoose$h(_this$props, ["children"]); // filter props for Transtition - - - delete childProps.in; - delete childProps.mountOnEnter; - delete childProps.unmountOnExit; - delete childProps.appear; - delete childProps.enter; - delete childProps.exit; - delete childProps.timeout; - delete childProps.addEndListener; - delete childProps.onEnter; - delete childProps.onEntering; - delete childProps.onEntered; - delete childProps.onExit; - delete childProps.onExiting; - delete childProps.onExited; - - if (typeof children === 'function') { - return children(status, childProps); - } - - var child = _react$3.default.Children.only(children); - - return _react$3.default.cloneElement(child, childProps); - }; - - return Transition; -}(_react$3.default.Component); - -Transition.contextTypes = { - transitionGroup: PropTypes.object -}; -Transition.childContextTypes = { - transitionGroup: function transitionGroup() {} -}; -Transition.propTypes = {}; - -function noop() {} - -Transition.defaultProps = { - in: false, - mountOnEnter: false, - unmountOnExit: false, - appear: false, - enter: true, - exit: true, - onEnter: noop, - onEntering: noop, - onEntered: noop, - onExit: noop, - onExiting: noop, - onExited: noop -}; -Transition.UNMOUNTED = 0; -Transition.EXITED = 1; -Transition.ENTERING = 2; -Transition.ENTERED = 3; -Transition.EXITING = 4; - -var _default$3 = (0, _reactLifecyclesCompat.polyfill)(Transition); - -Transition$1.default = _default$3; - -(function (module, exports) { - - exports.__esModule = true; - exports.default = void 0; - - _interopRequireWildcard(propTypes.exports); - - var _addClass = _interopRequireDefault(addClass.exports); - - var _removeClass = _interopRequireDefault(removeClass); - - var _react = _interopRequireDefault(React__default); - - var _Transition = _interopRequireDefault(Transition$1); - - function _interopRequireDefault(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - - function _interopRequireWildcard(obj) { if (obj && obj.__esModule) { return obj; } else { var newObj = {}; if (obj != null) { for (var key in obj) { if (Object.prototype.hasOwnProperty.call(obj, key)) { var desc = Object.defineProperty && Object.getOwnPropertyDescriptor ? Object.getOwnPropertyDescriptor(obj, key) : {}; if (desc.get || desc.set) { Object.defineProperty(newObj, key, desc); } else { newObj[key] = obj[key]; } } } } newObj.default = obj; return newObj; } } - - function _extends() { _extends = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends.apply(this, arguments); } - - function _inheritsLoose(subClass, superClass) { subClass.prototype = Object.create(superClass.prototype); subClass.prototype.constructor = subClass; subClass.__proto__ = superClass; } - - var addClass$1 = function addClass(node, classes) { - return node && classes && classes.split(' ').forEach(function (c) { - return (0, _addClass.default)(node, c); - }); - }; - - var removeClass$1 = function removeClass(node, classes) { - return node && classes && classes.split(' ').forEach(function (c) { - return (0, _removeClass.default)(node, c); - }); - }; - /** - * A transition component inspired by the excellent - * [ng-animate](http://www.nganimate.org/) library, you should use it if you're - * using CSS transitions or animations. It's built upon the - * [`Transition`](https://reactcommunity.org/react-transition-group/transition) - * component, so it inherits all of its props. - * - * `CSSTransition` applies a pair of class names during the `appear`, `enter`, - * and `exit` states of the transition. The first class is applied and then a - * second `*-active` class in order to activate the CSSS transition. After the - * transition, matching `*-done` class names are applied to persist the - * transition state. - * - * ```jsx - * function App() { - * const [inProp, setInProp] = useState(false); - * return ( - *
- * - *
- * {"I'll receive my-node-* classes"} - *
- *
- * - *
- * ); - * } - * ``` - * - * When the `in` prop is set to `true`, the child component will first receive - * the class `example-enter`, then the `example-enter-active` will be added in - * the next tick. `CSSTransition` [forces a - * reflow](https://github.com/reactjs/react-transition-group/blob/5007303e729a74be66a21c3e2205e4916821524b/src/CSSTransition.js#L208-L215) - * between before adding the `example-enter-active`. This is an important trick - * because it allows us to transition between `example-enter` and - * `example-enter-active` even though they were added immediately one after - * another. Most notably, this is what makes it possible for us to animate - * _appearance_. - * - * ```css - * .my-node-enter { - * opacity: 0; - * } - * .my-node-enter-active { - * opacity: 1; - * transition: opacity 200ms; - * } - * .my-node-exit { - * opacity: 1; - * } - * .my-node-exit-active { - * opacity: 0; - * transition: opacity: 200ms; - * } - * ``` - * - * `*-active` classes represent which styles you want to animate **to**. - */ - - - var CSSTransition = - /*#__PURE__*/ - function (_React$Component) { - _inheritsLoose(CSSTransition, _React$Component); - - function CSSTransition() { - var _this; - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _React$Component.call.apply(_React$Component, [this].concat(args)) || this; - - _this.onEnter = function (node, appearing) { - var _this$getClassNames = _this.getClassNames(appearing ? 'appear' : 'enter'), - className = _this$getClassNames.className; - - _this.removeClasses(node, 'exit'); - - addClass$1(node, className); - - if (_this.props.onEnter) { - _this.props.onEnter(node, appearing); - } - }; - - _this.onEntering = function (node, appearing) { - var _this$getClassNames2 = _this.getClassNames(appearing ? 'appear' : 'enter'), - activeClassName = _this$getClassNames2.activeClassName; - - _this.reflowAndAddClass(node, activeClassName); - - if (_this.props.onEntering) { - _this.props.onEntering(node, appearing); - } - }; - - _this.onEntered = function (node, appearing) { - var appearClassName = _this.getClassNames('appear').doneClassName; - - var enterClassName = _this.getClassNames('enter').doneClassName; - - var doneClassName = appearing ? appearClassName + " " + enterClassName : enterClassName; - - _this.removeClasses(node, appearing ? 'appear' : 'enter'); - - addClass$1(node, doneClassName); - - if (_this.props.onEntered) { - _this.props.onEntered(node, appearing); - } - }; - - _this.onExit = function (node) { - var _this$getClassNames3 = _this.getClassNames('exit'), - className = _this$getClassNames3.className; - - _this.removeClasses(node, 'appear'); - - _this.removeClasses(node, 'enter'); - - addClass$1(node, className); - - if (_this.props.onExit) { - _this.props.onExit(node); - } - }; - - _this.onExiting = function (node) { - var _this$getClassNames4 = _this.getClassNames('exit'), - activeClassName = _this$getClassNames4.activeClassName; - - _this.reflowAndAddClass(node, activeClassName); - - if (_this.props.onExiting) { - _this.props.onExiting(node); - } - }; - - _this.onExited = function (node) { - var _this$getClassNames5 = _this.getClassNames('exit'), - doneClassName = _this$getClassNames5.doneClassName; - - _this.removeClasses(node, 'exit'); - - addClass$1(node, doneClassName); - - if (_this.props.onExited) { - _this.props.onExited(node); - } - }; - - _this.getClassNames = function (type) { - var classNames = _this.props.classNames; - var isStringClassNames = typeof classNames === 'string'; - var prefix = isStringClassNames && classNames ? classNames + '-' : ''; - var className = isStringClassNames ? prefix + type : classNames[type]; - var activeClassName = isStringClassNames ? className + '-active' : classNames[type + 'Active']; - var doneClassName = isStringClassNames ? className + '-done' : classNames[type + 'Done']; - return { - className: className, - activeClassName: activeClassName, - doneClassName: doneClassName - }; - }; - - return _this; - } - - var _proto = CSSTransition.prototype; - - _proto.removeClasses = function removeClasses(node, type) { - var _this$getClassNames6 = this.getClassNames(type), - className = _this$getClassNames6.className, - activeClassName = _this$getClassNames6.activeClassName, - doneClassName = _this$getClassNames6.doneClassName; - - className && removeClass$1(node, className); - activeClassName && removeClass$1(node, activeClassName); - doneClassName && removeClass$1(node, doneClassName); - }; - - _proto.reflowAndAddClass = function reflowAndAddClass(node, className) { - // This is for to force a repaint, - // which is necessary in order to transition styles when adding a class name. - if (className) { - /* eslint-disable no-unused-expressions */ - node && node.scrollTop; - /* eslint-enable no-unused-expressions */ - - addClass$1(node, className); - } - }; - - _proto.render = function render() { - var props = _extends({}, this.props); - - delete props.classNames; - return _react.default.createElement(_Transition.default, _extends({}, props, { - onEnter: this.onEnter, - onEntered: this.onEntered, - onEntering: this.onEntering, - onExit: this.onExit, - onExiting: this.onExiting, - onExited: this.onExited - })); - }; - - return CSSTransition; - }(_react.default.Component); - - CSSTransition.defaultProps = { - classNames: '' - }; - CSSTransition.propTypes = {}; - var _default = CSSTransition; - exports.default = _default; - module.exports = exports["default"]; -} (CSSTransition, CSSTransition.exports)); - -var ReplaceTransition = {exports: {}}; - -var TransitionGroup = {exports: {}}; - -var ChildMapping = {}; - -ChildMapping.__esModule = true; -ChildMapping.getChildMapping = getChildMapping; -ChildMapping.mergeChildMappings = mergeChildMappings; -ChildMapping.getInitialChildMapping = getInitialChildMapping; -ChildMapping.getNextChildMapping = getNextChildMapping; - -var _react$2 = React__default; - -/** - * Given `this.props.children`, return an object mapping key to child. - * - * @param {*} children `this.props.children` - * @return {object} Mapping of key to child - */ -function getChildMapping(children, mapFn) { - var mapper = function mapper(child) { - return mapFn && (0, _react$2.isValidElement)(child) ? mapFn(child) : child; - }; - - var result = Object.create(null); - if (children) _react$2.Children.map(children, function (c) { - return c; - }).forEach(function (child) { - // run the map function here instead so that the key is the computed one - result[child.key] = mapper(child); - }); - return result; -} -/** - * When you're adding or removing children some may be added or removed in the - * same render pass. We want to show *both* since we want to simultaneously - * animate elements in and out. This function takes a previous set of keys - * and a new set of keys and merges them with its best guess of the correct - * ordering. In the future we may expose some of the utilities in - * ReactMultiChild to make this easy, but for now React itself does not - * directly have this concept of the union of prevChildren and nextChildren - * so we implement it here. - * - * @param {object} prev prev children as returned from - * `ReactTransitionChildMapping.getChildMapping()`. - * @param {object} next next children as returned from - * `ReactTransitionChildMapping.getChildMapping()`. - * @return {object} a key set that contains all keys in `prev` and all keys - * in `next` in a reasonable order. - */ - - -function mergeChildMappings(prev, next) { - prev = prev || {}; - next = next || {}; - - function getValueForKey(key) { - return key in next ? next[key] : prev[key]; - } // For each key of `next`, the list of keys to insert before that key in - // the combined list - - - var nextKeysPending = Object.create(null); - var pendingKeys = []; - - for (var prevKey in prev) { - if (prevKey in next) { - if (pendingKeys.length) { - nextKeysPending[prevKey] = pendingKeys; - pendingKeys = []; - } - } else { - pendingKeys.push(prevKey); - } - } - - var i; - var childMapping = {}; - - for (var nextKey in next) { - if (nextKeysPending[nextKey]) { - for (i = 0; i < nextKeysPending[nextKey].length; i++) { - var pendingNextKey = nextKeysPending[nextKey][i]; - childMapping[nextKeysPending[nextKey][i]] = getValueForKey(pendingNextKey); - } - } - - childMapping[nextKey] = getValueForKey(nextKey); - } // Finally, add the keys which didn't appear before any key in `next` - - - for (i = 0; i < pendingKeys.length; i++) { - childMapping[pendingKeys[i]] = getValueForKey(pendingKeys[i]); - } - - return childMapping; -} - -function getProp(child, prop, props) { - return props[prop] != null ? props[prop] : child.props[prop]; -} - -function getInitialChildMapping(props, onExited) { - return getChildMapping(props.children, function (child) { - return (0, _react$2.cloneElement)(child, { - onExited: onExited.bind(null, child), - in: true, - appear: getProp(child, 'appear', props), - enter: getProp(child, 'enter', props), - exit: getProp(child, 'exit', props) - }); - }); -} - -function getNextChildMapping(nextProps, prevChildMapping, onExited) { - var nextChildMapping = getChildMapping(nextProps.children); - var children = mergeChildMappings(prevChildMapping, nextChildMapping); - Object.keys(children).forEach(function (key) { - var child = children[key]; - if (!(0, _react$2.isValidElement)(child)) return; - var hasPrev = key in prevChildMapping; - var hasNext = key in nextChildMapping; - var prevChild = prevChildMapping[key]; - var isLeaving = (0, _react$2.isValidElement)(prevChild) && !prevChild.props.in; // item is new (entering) - - if (hasNext && (!hasPrev || isLeaving)) { - // console.log('entering', key) - children[key] = (0, _react$2.cloneElement)(child, { - onExited: onExited.bind(null, child), - in: true, - exit: getProp(child, 'exit', nextProps), - enter: getProp(child, 'enter', nextProps) - }); - } else if (!hasNext && hasPrev && !isLeaving) { - // item is old (exiting) - // console.log('leaving', key) - children[key] = (0, _react$2.cloneElement)(child, { - in: false - }); - } else if (hasNext && hasPrev && (0, _react$2.isValidElement)(prevChild)) { - // item hasn't changed transition states - // copy over the last transition props; - // console.log('unchanged', key) - children[key] = (0, _react$2.cloneElement)(child, { - onExited: onExited.bind(null, child), - in: prevChild.props.in, - exit: getProp(child, 'exit', nextProps), - enter: getProp(child, 'enter', nextProps) - }); - } - }); - return children; -} - -(function (module, exports) { - - exports.__esModule = true; - exports.default = void 0; - - var _propTypes = _interopRequireDefault(propTypes.exports); - - var _react = _interopRequireDefault(React__default); - - var _reactLifecyclesCompat = reactLifecyclesCompat_cjs; - - var _ChildMapping = ChildMapping; - - function _interopRequireDefault(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - - function _objectWithoutPropertiesLoose(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - - function _extends() { _extends = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends.apply(this, arguments); } - - function _inheritsLoose(subClass, superClass) { subClass.prototype = Object.create(superClass.prototype); subClass.prototype.constructor = subClass; subClass.__proto__ = superClass; } - - function _assertThisInitialized(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - - var values = Object.values || function (obj) { - return Object.keys(obj).map(function (k) { - return obj[k]; - }); - }; - - var defaultProps = { - component: 'div', - childFactory: function childFactory(child) { - return child; - } - /** - * The `` component manages a set of transition components - * (`` and ``) in a list. Like with the transition - * components, `` is a state machine for managing the mounting - * and unmounting of components over time. - * - * Consider the example below. As items are removed or added to the TodoList the - * `in` prop is toggled automatically by the ``. - * - * Note that `` does not define any animation behavior! - * Exactly _how_ a list item animates is up to the individual transition - * component. This means you can mix and match animations across different list - * items. - */ - - }; - - var TransitionGroup = - /*#__PURE__*/ - function (_React$Component) { - _inheritsLoose(TransitionGroup, _React$Component); - - function TransitionGroup(props, context) { - var _this; - - _this = _React$Component.call(this, props, context) || this; - - var handleExited = _this.handleExited.bind(_assertThisInitialized(_assertThisInitialized(_this))); // Initial children should all be entering, dependent on appear - - - _this.state = { - handleExited: handleExited, - firstRender: true - }; - return _this; - } - - var _proto = TransitionGroup.prototype; - - _proto.getChildContext = function getChildContext() { - return { - transitionGroup: { - isMounting: !this.appeared - } - }; - }; - - _proto.componentDidMount = function componentDidMount() { - this.appeared = true; - this.mounted = true; - }; - - _proto.componentWillUnmount = function componentWillUnmount() { - this.mounted = false; - }; - - TransitionGroup.getDerivedStateFromProps = function getDerivedStateFromProps(nextProps, _ref) { - var prevChildMapping = _ref.children, - handleExited = _ref.handleExited, - firstRender = _ref.firstRender; - return { - children: firstRender ? (0, _ChildMapping.getInitialChildMapping)(nextProps, handleExited) : (0, _ChildMapping.getNextChildMapping)(nextProps, prevChildMapping, handleExited), - firstRender: false - }; - }; - - _proto.handleExited = function handleExited(child, node) { - var currentChildMapping = (0, _ChildMapping.getChildMapping)(this.props.children); - if (child.key in currentChildMapping) return; - - if (child.props.onExited) { - child.props.onExited(node); - } - - if (this.mounted) { - this.setState(function (state) { - var children = _extends({}, state.children); - - delete children[child.key]; - return { - children: children - }; - }); - } - }; - - _proto.render = function render() { - var _this$props = this.props, - Component = _this$props.component, - childFactory = _this$props.childFactory, - props = _objectWithoutPropertiesLoose(_this$props, ["component", "childFactory"]); - - var children = values(this.state.children).map(childFactory); - delete props.appear; - delete props.enter; - delete props.exit; - - if (Component === null) { - return children; - } - - return _react.default.createElement(Component, props, children); - }; - - return TransitionGroup; - }(_react.default.Component); - - TransitionGroup.childContextTypes = { - transitionGroup: _propTypes.default.object.isRequired - }; - TransitionGroup.propTypes = {}; - TransitionGroup.defaultProps = defaultProps; - - var _default = (0, _reactLifecyclesCompat.polyfill)(TransitionGroup); - - exports.default = _default; - module.exports = exports["default"]; -} (TransitionGroup, TransitionGroup.exports)); - -(function (module, exports) { - - exports.__esModule = true; - exports.default = void 0; - - _interopRequireDefault(propTypes.exports); - - var _react = _interopRequireDefault(React__default); - - var _reactDom = require$$2; - - var _TransitionGroup = _interopRequireDefault(TransitionGroup.exports); - - function _interopRequireDefault(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - - function _objectWithoutPropertiesLoose(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - - function _inheritsLoose(subClass, superClass) { subClass.prototype = Object.create(superClass.prototype); subClass.prototype.constructor = subClass; subClass.__proto__ = superClass; } - - /** - * The `` component is a specialized `Transition` component - * that animates between two children. - * - * ```jsx - * - *
I appear first
- *
I replace the above
- *
- * ``` - */ - var ReplaceTransition = - /*#__PURE__*/ - function (_React$Component) { - _inheritsLoose(ReplaceTransition, _React$Component); - - function ReplaceTransition() { - var _this; - - for (var _len = arguments.length, _args = new Array(_len), _key = 0; _key < _len; _key++) { - _args[_key] = arguments[_key]; - } - - _this = _React$Component.call.apply(_React$Component, [this].concat(_args)) || this; - - _this.handleEnter = function () { - for (var _len2 = arguments.length, args = new Array(_len2), _key2 = 0; _key2 < _len2; _key2++) { - args[_key2] = arguments[_key2]; - } - - return _this.handleLifecycle('onEnter', 0, args); - }; - - _this.handleEntering = function () { - for (var _len3 = arguments.length, args = new Array(_len3), _key3 = 0; _key3 < _len3; _key3++) { - args[_key3] = arguments[_key3]; - } - - return _this.handleLifecycle('onEntering', 0, args); - }; - - _this.handleEntered = function () { - for (var _len4 = arguments.length, args = new Array(_len4), _key4 = 0; _key4 < _len4; _key4++) { - args[_key4] = arguments[_key4]; - } - - return _this.handleLifecycle('onEntered', 0, args); - }; - - _this.handleExit = function () { - for (var _len5 = arguments.length, args = new Array(_len5), _key5 = 0; _key5 < _len5; _key5++) { - args[_key5] = arguments[_key5]; - } - - return _this.handleLifecycle('onExit', 1, args); - }; - - _this.handleExiting = function () { - for (var _len6 = arguments.length, args = new Array(_len6), _key6 = 0; _key6 < _len6; _key6++) { - args[_key6] = arguments[_key6]; - } - - return _this.handleLifecycle('onExiting', 1, args); - }; - - _this.handleExited = function () { - for (var _len7 = arguments.length, args = new Array(_len7), _key7 = 0; _key7 < _len7; _key7++) { - args[_key7] = arguments[_key7]; - } - - return _this.handleLifecycle('onExited', 1, args); - }; - - return _this; - } - - var _proto = ReplaceTransition.prototype; - - _proto.handleLifecycle = function handleLifecycle(handler, idx, originalArgs) { - var _child$props; - - var children = this.props.children; - - var child = _react.default.Children.toArray(children)[idx]; - - if (child.props[handler]) (_child$props = child.props)[handler].apply(_child$props, originalArgs); - if (this.props[handler]) this.props[handler]((0, _reactDom.findDOMNode)(this)); - }; - - _proto.render = function render() { - var _this$props = this.props, - children = _this$props.children, - inProp = _this$props.in, - props = _objectWithoutPropertiesLoose(_this$props, ["children", "in"]); - - var _React$Children$toArr = _react.default.Children.toArray(children), - first = _React$Children$toArr[0], - second = _React$Children$toArr[1]; - - delete props.onEnter; - delete props.onEntering; - delete props.onEntered; - delete props.onExit; - delete props.onExiting; - delete props.onExited; - return _react.default.createElement(_TransitionGroup.default, props, inProp ? _react.default.cloneElement(first, { - key: 'first', - onEnter: this.handleEnter, - onEntering: this.handleEntering, - onEntered: this.handleEntered - }) : _react.default.cloneElement(second, { - key: 'second', - onEnter: this.handleExit, - onEntering: this.handleExiting, - onEntered: this.handleExited - })); - }; - - return ReplaceTransition; - }(_react.default.Component); - - ReplaceTransition.propTypes = {}; - var _default = ReplaceTransition; - exports.default = _default; - module.exports = exports["default"]; -} (ReplaceTransition, ReplaceTransition.exports)); - -var _CSSTransition = _interopRequireDefault$5(CSSTransition.exports); - -var _ReplaceTransition = _interopRequireDefault$5(ReplaceTransition.exports); - -var _TransitionGroup = _interopRequireDefault$5(TransitionGroup.exports); - -var _Transition = _interopRequireDefault$5(Transition$1); - -function _interopRequireDefault$5(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -var reactTransitionGroup = { - Transition: _Transition.default, - TransitionGroup: _TransitionGroup.default, - ReplaceTransition: _ReplaceTransition.default, - CSSTransition: _CSSTransition.default -}; - -var AnimateGroupChild$1 = {}; - -function _typeof$t(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$t = function _typeof(obj) { return typeof obj; }; } else { _typeof$t = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$t(obj); } - -Object.defineProperty(AnimateGroupChild$1, "__esModule", { - value: true -}); -AnimateGroupChild$1.default = void 0; - -var _react$1 = _interopRequireWildcard$1(React__default); - -var _reactTransitionGroup$1 = reactTransitionGroup; - -var _propTypes$1 = _interopRequireDefault$4(propTypes.exports); - -var _Animate = _interopRequireDefault$4(Animate$2); - -function _interopRequireDefault$4(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -function _getRequireWildcardCache$1() { if (typeof WeakMap !== "function") return null; var cache = new WeakMap(); _getRequireWildcardCache$1 = function _getRequireWildcardCache() { return cache; }; return cache; } - -function _interopRequireWildcard$1(obj) { if (obj && obj.__esModule) { return obj; } if (obj === null || _typeof$t(obj) !== "object" && typeof obj !== "function") { return { default: obj }; } var cache = _getRequireWildcardCache$1(); if (cache && cache.has(obj)) { return cache.get(obj); } var newObj = {}; var hasPropertyDescriptor = Object.defineProperty && Object.getOwnPropertyDescriptor; for (var key in obj) { if (Object.prototype.hasOwnProperty.call(obj, key)) { var desc = hasPropertyDescriptor ? Object.getOwnPropertyDescriptor(obj, key) : null; if (desc && (desc.get || desc.set)) { Object.defineProperty(newObj, key, desc); } else { newObj[key] = obj[key]; } } } newObj.default = obj; if (cache) { cache.set(obj, newObj); } return newObj; } - -function _extends$v() { _extends$v = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$v.apply(this, arguments); } - -function _objectWithoutProperties$g(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$g(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$g(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function ownKeys$v(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$v(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$v(Object(source), true).forEach(function (key) { _defineProperty$w(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$v(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$w(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$s(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$s(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$s(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$s(Constructor.prototype, protoProps); if (staticProps) _defineProperties$s(Constructor, staticProps); return Constructor; } - -function _inherits$r(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$r(subClass, superClass); } - -function _setPrototypeOf$r(o, p) { _setPrototypeOf$r = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$r(o, p); } - -function _createSuper$r(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$r(); return function _createSuperInternal() { var Super = _getPrototypeOf$r(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$r(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$r(this, result); }; } - -function _possibleConstructorReturn$r(self, call) { if (call && (_typeof$t(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$r(self); } - -function _assertThisInitialized$r(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$r() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$r(o) { _getPrototypeOf$r = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$r(o); } - -if (Number.isFinite === undefined) { - Number.isFinite = function (value) { - return typeof value === 'number' && isFinite(value); - }; -} - -var parseDurationOfSingleTransition = function parseDurationOfSingleTransition() { - var options = arguments.length > 0 && arguments[0] !== undefined ? arguments[0] : {}; - var steps = options.steps, - duration = options.duration; - - if (steps && steps.length) { - return steps.reduce(function (result, entry) { - return result + (Number.isFinite(entry.duration) && entry.duration > 0 ? entry.duration : 0); - }, 0); - } - - if (Number.isFinite(duration)) { - return duration; - } - - return 0; -}; - -var AnimateGroupChild = /*#__PURE__*/function (_Component) { - _inherits$r(AnimateGroupChild, _Component); - - var _super = _createSuper$r(AnimateGroupChild); - - function AnimateGroupChild() { - var _this; - - _classCallCheck$s(this, AnimateGroupChild); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - isActive: false - }; - - _this.handleEnter = function (node, isAppearing) { - var _this$props = _this.props, - appearOptions = _this$props.appearOptions, - enterOptions = _this$props.enterOptions; - - _this.handleStyleActive(isAppearing ? appearOptions : enterOptions); - }; - - _this.handleExit = function () { - _this.handleStyleActive(_this.props.leaveOptions); - }; - - return _this; - } - - _createClass$s(AnimateGroupChild, [{ - key: "handleStyleActive", - value: function handleStyleActive(style) { - if (style) { - var onAnimationEnd = style.onAnimationEnd ? function () { - style.onAnimationEnd(); - } : null; - this.setState(_objectSpread$v(_objectSpread$v({}, style), {}, { - onAnimationEnd: onAnimationEnd, - isActive: true - })); - } - } - }, { - key: "parseTimeout", - value: function parseTimeout() { - var _this$props2 = this.props, - appearOptions = _this$props2.appearOptions, - enterOptions = _this$props2.enterOptions, - leaveOptions = _this$props2.leaveOptions; - return parseDurationOfSingleTransition(appearOptions) + parseDurationOfSingleTransition(enterOptions) + parseDurationOfSingleTransition(leaveOptions); - } - }, { - key: "render", - value: function render() { - var _this2 = this; - - var _this$props3 = this.props, - children = _this$props3.children; - _this$props3.appearOptions; - _this$props3.enterOptions; - _this$props3.leaveOptions; - var props = _objectWithoutProperties$g(_this$props3, ["children", "appearOptions", "enterOptions", "leaveOptions"]); - - return /*#__PURE__*/_react$1.default.createElement(_reactTransitionGroup$1.Transition, _extends$v({}, props, { - onEnter: this.handleEnter, - onExit: this.handleExit, - timeout: this.parseTimeout() - }), function () { - return /*#__PURE__*/_react$1.default.createElement(_Animate.default, _this2.state, _react$1.Children.only(children)); - }); - } - }]); - - return AnimateGroupChild; -}(_react$1.Component); - -AnimateGroupChild.propTypes = { - appearOptions: _propTypes$1.default.object, - enterOptions: _propTypes$1.default.object, - leaveOptions: _propTypes$1.default.object, - children: _propTypes$1.default.element -}; -var _default$2 = AnimateGroupChild; -AnimateGroupChild$1.default = _default$2; - -function _typeof$s(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$s = function _typeof(obj) { return typeof obj; }; } else { _typeof$s = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$s(obj); } - -Object.defineProperty(AnimateGroup$1, "__esModule", { - value: true -}); -AnimateGroup$1.default = void 0; - -var _react = _interopRequireWildcard(React__default); - -var _reactTransitionGroup = reactTransitionGroup; - -var _propTypes = _interopRequireDefault$3(propTypes.exports); - -var _AnimateGroupChild = _interopRequireDefault$3(AnimateGroupChild$1); - -function _interopRequireDefault$3(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -function _getRequireWildcardCache() { if (typeof WeakMap !== "function") return null; var cache = new WeakMap(); _getRequireWildcardCache = function _getRequireWildcardCache() { return cache; }; return cache; } - -function _interopRequireWildcard(obj) { if (obj && obj.__esModule) { return obj; } if (obj === null || _typeof$s(obj) !== "object" && typeof obj !== "function") { return { default: obj }; } var cache = _getRequireWildcardCache(); if (cache && cache.has(obj)) { return cache.get(obj); } var newObj = {}; var hasPropertyDescriptor = Object.defineProperty && Object.getOwnPropertyDescriptor; for (var key in obj) { if (Object.prototype.hasOwnProperty.call(obj, key)) { var desc = hasPropertyDescriptor ? Object.getOwnPropertyDescriptor(obj, key) : null; if (desc && (desc.get || desc.set)) { Object.defineProperty(newObj, key, desc); } else { newObj[key] = obj[key]; } } } newObj.default = obj; if (cache) { cache.set(obj, newObj); } return newObj; } - -function AnimateGroup(props) { - var component = props.component, - children = props.children, - appear = props.appear, - enter = props.enter, - leave = props.leave; - return /*#__PURE__*/_react.default.createElement(_reactTransitionGroup.TransitionGroup, { - component: component - }, _react.Children.map(children, function (child, index) { - return /*#__PURE__*/_react.default.createElement(_AnimateGroupChild.default, { - appearOptions: appear, - enterOptions: enter, - leaveOptions: leave, - key: "child-".concat(index) // eslint-disable-line - - }, child); - })); -} - -AnimateGroup.propTypes = { - appear: _propTypes.default.object, - enter: _propTypes.default.object, - leave: _propTypes.default.object, - children: _propTypes.default.oneOfType([_propTypes.default.array, _propTypes.default.element]), - component: _propTypes.default.any -}; -AnimateGroup.defaultProps = { - component: 'span' -}; -var _default$1 = AnimateGroup; -AnimateGroup$1.default = _default$1; - -(function (exports) { - - Object.defineProperty(exports, "__esModule", { - value: true - }); - Object.defineProperty(exports, "configBezier", { - enumerable: true, - get: function get() { - return _easing.configBezier; - } - }); - Object.defineProperty(exports, "configSpring", { - enumerable: true, - get: function get() { - return _easing.configSpring; - } - }); - Object.defineProperty(exports, "translateStyle", { - enumerable: true, - get: function get() { - return _util.translateStyle; - } - }); - Object.defineProperty(exports, "AnimateGroup", { - enumerable: true, - get: function get() { - return _AnimateGroup.default; - } - }); - exports.default = void 0; - - var _Animate = _interopRequireDefault(Animate$2); - - var _easing = easing; - - var _util = util; - - var _AnimateGroup = _interopRequireDefault(AnimateGroup$1); - - function _interopRequireDefault(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - - var _default = _Animate.default; - exports.default = _default; -} (lib$2)); - -var Animate = /*@__PURE__*/getDefaultExportFromCjs(lib$2); - -var Symbol$4 = _Symbol$1, - isArguments = isArguments_1, - isArray$5 = isArray_1; - -/** Built-in value references. */ -var spreadableSymbol = Symbol$4 ? Symbol$4.isConcatSpreadable : undefined; - -/** - * Checks if `value` is a flattenable `arguments` object or array. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is flattenable, else `false`. - */ -function isFlattenable$1(value) { - return isArray$5(value) || isArguments(value) || - !!(spreadableSymbol && value && value[spreadableSymbol]); -} - -var _isFlattenable = isFlattenable$1; - -var arrayPush$1 = _arrayPush, - isFlattenable = _isFlattenable; - -/** - * The base implementation of `_.flatten` with support for restricting flattening. - * - * @private - * @param {Array} array The array to flatten. - * @param {number} depth The maximum recursion depth. - * @param {boolean} [predicate=isFlattenable] The function invoked per iteration. - * @param {boolean} [isStrict] Restrict to values that pass `predicate` checks. - * @param {Array} [result=[]] The initial result value. - * @returns {Array} Returns the new flattened array. - */ -function baseFlatten$3(array, depth, predicate, isStrict, result) { - var index = -1, - length = array.length; - - predicate || (predicate = isFlattenable); - result || (result = []); - - while (++index < length) { - var value = array[index]; - if (depth > 0 && predicate(value)) { - if (depth > 1) { - // Recursively flatten arrays (susceptible to call stack limits). - baseFlatten$3(value, depth - 1, predicate, isStrict, result); - } else { - arrayPush$1(result, value); - } - } else if (!isStrict) { - result[result.length] = value; - } - } - return result; -} - -var _baseFlatten = baseFlatten$3; - -/** - * Creates a base function for methods like `_.forIn` and `_.forOwn`. - * - * @private - * @param {boolean} [fromRight] Specify iterating from right to left. - * @returns {Function} Returns the new base function. - */ - -function createBaseFor$1(fromRight) { - return function(object, iteratee, keysFunc) { - var index = -1, - iterable = Object(object), - props = keysFunc(object), - length = props.length; - - while (length--) { - var key = props[fromRight ? length : ++index]; - if (iteratee(iterable[key], key, iterable) === false) { - break; - } - } - return object; - }; -} - -var _createBaseFor = createBaseFor$1; - -var createBaseFor = _createBaseFor; - -/** - * The base implementation of `baseForOwn` which iterates over `object` - * properties returned by `keysFunc` and invokes `iteratee` for each property. - * Iteratee functions may exit iteration early by explicitly returning `false`. - * - * @private - * @param {Object} object The object to iterate over. - * @param {Function} iteratee The function invoked per iteration. - * @param {Function} keysFunc The function to get the keys of `object`. - * @returns {Object} Returns `object`. - */ -var baseFor$1 = createBaseFor(); - -var _baseFor = baseFor$1; - -var baseFor = _baseFor, - keys$3 = keys_1; - -/** - * The base implementation of `_.forOwn` without support for iteratee shorthands. - * - * @private - * @param {Object} object The object to iterate over. - * @param {Function} iteratee The function invoked per iteration. - * @returns {Object} Returns `object`. - */ -function baseForOwn$2(object, iteratee) { - return object && baseFor(object, iteratee, keys$3); -} - -var _baseForOwn = baseForOwn$2; - -var isArrayLike$4 = isArrayLike_1; - -/** - * Creates a `baseEach` or `baseEachRight` function. - * - * @private - * @param {Function} eachFunc The function to iterate over a collection. - * @param {boolean} [fromRight] Specify iterating from right to left. - * @returns {Function} Returns the new base function. - */ -function createBaseEach$1(eachFunc, fromRight) { - return function(collection, iteratee) { - if (collection == null) { - return collection; - } - if (!isArrayLike$4(collection)) { - return eachFunc(collection, iteratee); - } - var length = collection.length, - index = fromRight ? length : -1, - iterable = Object(collection); - - while ((fromRight ? index-- : ++index < length)) { - if (iteratee(iterable[index], index, iterable) === false) { - break; - } - } - return collection; - }; -} - -var _createBaseEach = createBaseEach$1; - -var baseForOwn$1 = _baseForOwn, - createBaseEach = _createBaseEach; - -/** - * The base implementation of `_.forEach` without support for iteratee shorthands. - * - * @private - * @param {Array|Object} collection The collection to iterate over. - * @param {Function} iteratee The function invoked per iteration. - * @returns {Array|Object} Returns `collection`. - */ -var baseEach$3 = createBaseEach(baseForOwn$1); - -var _baseEach = baseEach$3; - -var baseEach$2 = _baseEach, - isArrayLike$3 = isArrayLike_1; - -/** - * The base implementation of `_.map` without support for iteratee shorthands. - * - * @private - * @param {Array|Object} collection The collection to iterate over. - * @param {Function} iteratee The function invoked per iteration. - * @returns {Array} Returns the new mapped array. - */ -function baseMap$2(collection, iteratee) { - var index = -1, - result = isArrayLike$3(collection) ? Array(collection.length) : []; - - baseEach$2(collection, function(value, key, collection) { - result[++index] = iteratee(value, key, collection); - }); - return result; -} - -var _baseMap = baseMap$2; - -/** - * The base implementation of `_.sortBy` which uses `comparer` to define the - * sort order of `array` and replaces criteria objects with their corresponding - * values. - * - * @private - * @param {Array} array The array to sort. - * @param {Function} comparer The function to define sort order. - * @returns {Array} Returns `array`. - */ - -function baseSortBy$1(array, comparer) { - var length = array.length; - - array.sort(comparer); - while (length--) { - array[length] = array[length].value; - } - return array; -} - -var _baseSortBy = baseSortBy$1; - -var isSymbol$4 = isSymbol_1$1; - -/** - * Compares values to sort them in ascending order. - * - * @private - * @param {*} value The value to compare. - * @param {*} other The other value to compare. - * @returns {number} Returns the sort order indicator for `value`. - */ -function compareAscending$1(value, other) { - if (value !== other) { - var valIsDefined = value !== undefined, - valIsNull = value === null, - valIsReflexive = value === value, - valIsSymbol = isSymbol$4(value); - - var othIsDefined = other !== undefined, - othIsNull = other === null, - othIsReflexive = other === other, - othIsSymbol = isSymbol$4(other); - - if ((!othIsNull && !othIsSymbol && !valIsSymbol && value > other) || - (valIsSymbol && othIsDefined && othIsReflexive && !othIsNull && !othIsSymbol) || - (valIsNull && othIsDefined && othIsReflexive) || - (!valIsDefined && othIsReflexive) || - !valIsReflexive) { - return 1; - } - if ((!valIsNull && !valIsSymbol && !othIsSymbol && value < other) || - (othIsSymbol && valIsDefined && valIsReflexive && !valIsNull && !valIsSymbol) || - (othIsNull && valIsDefined && valIsReflexive) || - (!othIsDefined && valIsReflexive) || - !othIsReflexive) { - return -1; - } - } - return 0; -} - -var _compareAscending = compareAscending$1; - -var compareAscending = _compareAscending; - -/** - * Used by `_.orderBy` to compare multiple properties of a value to another - * and stable sort them. - * - * If `orders` is unspecified, all values are sorted in ascending order. Otherwise, - * specify an order of "desc" for descending or "asc" for ascending sort order - * of corresponding values. - * - * @private - * @param {Object} object The object to compare. - * @param {Object} other The other object to compare. - * @param {boolean[]|string[]} orders The order to sort by for each property. - * @returns {number} Returns the sort order indicator for `object`. - */ -function compareMultiple$1(object, other, orders) { - var index = -1, - objCriteria = object.criteria, - othCriteria = other.criteria, - length = objCriteria.length, - ordersLength = orders.length; - - while (++index < length) { - var result = compareAscending(objCriteria[index], othCriteria[index]); - if (result) { - if (index >= ordersLength) { - return result; - } - var order = orders[index]; - return result * (order == 'desc' ? -1 : 1); - } - } - // Fixes an `Array#sort` bug in the JS engine embedded in Adobe applications - // that causes it, under certain circumstances, to provide the same value for - // `object` and `other`. See https://github.com/jashkenas/underscore/pull/1247 - // for more details. - // - // This also ensures a stable sort in V8 and other engines. - // See https://bugs.chromium.org/p/v8/issues/detail?id=90 for more details. - return object.index - other.index; -} - -var _compareMultiple = compareMultiple$1; - -var arrayMap$2 = _arrayMap, - baseGet$1 = _baseGet, - baseIteratee$9 = _baseIteratee, - baseMap$1 = _baseMap, - baseSortBy = _baseSortBy, - baseUnary$2 = _baseUnary, - compareMultiple = _compareMultiple, - identity$7 = identity_1, - isArray$4 = isArray_1; - -/** - * The base implementation of `_.orderBy` without param guards. - * - * @private - * @param {Array|Object} collection The collection to iterate over. - * @param {Function[]|Object[]|string[]} iteratees The iteratees to sort by. - * @param {string[]} orders The sort orders of `iteratees`. - * @returns {Array} Returns the new sorted array. - */ -function baseOrderBy$1(collection, iteratees, orders) { - if (iteratees.length) { - iteratees = arrayMap$2(iteratees, function(iteratee) { - if (isArray$4(iteratee)) { - return function(value) { - return baseGet$1(value, iteratee.length === 1 ? iteratee[0] : iteratee); - } - } - return iteratee; - }); - } else { - iteratees = [identity$7]; - } - - var index = -1; - iteratees = arrayMap$2(iteratees, baseUnary$2(baseIteratee$9)); - - var result = baseMap$1(collection, function(value, key, collection) { - var criteria = arrayMap$2(iteratees, function(iteratee) { - return iteratee(value); - }); - return { 'criteria': criteria, 'index': ++index, 'value': value }; - }); - - return baseSortBy(result, function(object, other) { - return compareMultiple(object, other, orders); - }); -} - -var _baseOrderBy = baseOrderBy$1; - -/** - * A faster alternative to `Function#apply`, this function invokes `func` - * with the `this` binding of `thisArg` and the arguments of `args`. - * - * @private - * @param {Function} func The function to invoke. - * @param {*} thisArg The `this` binding of `func`. - * @param {Array} args The arguments to invoke `func` with. - * @returns {*} Returns the result of `func`. - */ - -function apply$1(func, thisArg, args) { - switch (args.length) { - case 0: return func.call(thisArg); - case 1: return func.call(thisArg, args[0]); - case 2: return func.call(thisArg, args[0], args[1]); - case 3: return func.call(thisArg, args[0], args[1], args[2]); - } - return func.apply(thisArg, args); -} - -var _apply = apply$1; - -var apply = _apply; - -/* Built-in method references for those with the same name as other `lodash` methods. */ -var nativeMax$4 = Math.max; - -/** - * A specialized version of `baseRest` which transforms the rest array. - * - * @private - * @param {Function} func The function to apply a rest parameter to. - * @param {number} [start=func.length-1] The start position of the rest parameter. - * @param {Function} transform The rest array transform. - * @returns {Function} Returns the new function. - */ -function overRest$2(func, start, transform) { - start = nativeMax$4(start === undefined ? (func.length - 1) : start, 0); - return function() { - var args = arguments, - index = -1, - length = nativeMax$4(args.length - start, 0), - array = Array(length); - - while (++index < length) { - array[index] = args[start + index]; - } - index = -1; - var otherArgs = Array(start + 1); - while (++index < start) { - otherArgs[index] = args[index]; - } - otherArgs[start] = transform(array); - return apply(func, this, otherArgs); - }; -} - -var _overRest = overRest$2; - -/** - * Creates a function that returns `value`. - * - * @static - * @memberOf _ - * @since 2.4.0 - * @category Util - * @param {*} value The value to return from the new function. - * @returns {Function} Returns the new constant function. - * @example - * - * var objects = _.times(2, _.constant({ 'a': 1 })); - * - * console.log(objects); - * // => [{ 'a': 1 }, { 'a': 1 }] - * - * console.log(objects[0] === objects[1]); - * // => true - */ - -function constant$2(value) { - return function() { - return value; - }; -} - -var constant_1 = constant$2; - -var getNative = _getNative; - -var defineProperty$2 = (function() { - try { - var func = getNative(Object, 'defineProperty'); - func({}, '', {}); - return func; - } catch (e) {} -}()); - -var _defineProperty$v = defineProperty$2; - -var constant$1 = constant_1, - defineProperty$1 = _defineProperty$v, - identity$6 = identity_1; - -/** - * The base implementation of `setToString` without support for hot loop shorting. - * - * @private - * @param {Function} func The function to modify. - * @param {Function} string The `toString` result. - * @returns {Function} Returns `func`. - */ -var baseSetToString$1 = !defineProperty$1 ? identity$6 : function(func, string) { - return defineProperty$1(func, 'toString', { - 'configurable': true, - 'enumerable': false, - 'value': constant$1(string), - 'writable': true - }); -}; - -var _baseSetToString = baseSetToString$1; - -/** Used to detect hot functions by number of calls within a span of milliseconds. */ - -var HOT_COUNT = 800, - HOT_SPAN = 16; - -/* Built-in method references for those with the same name as other `lodash` methods. */ -var nativeNow = Date.now; - -/** - * Creates a function that'll short out and invoke `identity` instead - * of `func` when it's called `HOT_COUNT` or more times in `HOT_SPAN` - * milliseconds. - * - * @private - * @param {Function} func The function to restrict. - * @returns {Function} Returns the new shortable function. - */ -function shortOut$1(func) { - var count = 0, - lastCalled = 0; - - return function() { - var stamp = nativeNow(), - remaining = HOT_SPAN - (stamp - lastCalled); - - lastCalled = stamp; - if (remaining > 0) { - if (++count >= HOT_COUNT) { - return arguments[0]; - } - } else { - count = 0; - } - return func.apply(undefined, arguments); - }; -} - -var _shortOut = shortOut$1; - -var baseSetToString = _baseSetToString, - shortOut = _shortOut; - -/** - * Sets the `toString` method of `func` to return `string`. - * - * @private - * @param {Function} func The function to modify. - * @param {Function} string The `toString` result. - * @returns {Function} Returns `func`. - */ -var setToString$2 = shortOut(baseSetToString); - -var _setToString = setToString$2; - -var identity$5 = identity_1, - overRest$1 = _overRest, - setToString$1 = _setToString; - -/** - * The base implementation of `_.rest` which doesn't validate or coerce arguments. - * - * @private - * @param {Function} func The function to apply a rest parameter to. - * @param {number} [start=func.length-1] The start position of the rest parameter. - * @returns {Function} Returns the new function. - */ -function baseRest$1(func, start) { - return setToString$1(overRest$1(func, start, identity$5), func + ''); -} - -var _baseRest = baseRest$1; - -var eq$1 = eq_1, - isArrayLike$2 = isArrayLike_1, - isIndex = _isIndex, - isObject$7 = isObject_1$1; - -/** - * Checks if the given arguments are from an iteratee call. - * - * @private - * @param {*} value The potential iteratee value argument. - * @param {*} index The potential iteratee index or key argument. - * @param {*} object The potential iteratee object argument. - * @returns {boolean} Returns `true` if the arguments are from an iteratee call, - * else `false`. - */ -function isIterateeCall$4(value, index, object) { - if (!isObject$7(object)) { - return false; - } - var type = typeof index; - if (type == 'number' - ? (isArrayLike$2(object) && isIndex(index, object.length)) - : (type == 'string' && index in object) - ) { - return eq$1(object[index], value); - } - return false; -} - -var _isIterateeCall = isIterateeCall$4; - -var baseFlatten$2 = _baseFlatten, - baseOrderBy = _baseOrderBy, - baseRest = _baseRest, - isIterateeCall$3 = _isIterateeCall; - -/** - * Creates an array of elements, sorted in ascending order by the results of - * running each element in a collection thru each iteratee. This method - * performs a stable sort, that is, it preserves the original sort order of - * equal elements. The iteratees are invoked with one argument: (value). - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Collection - * @param {Array|Object} collection The collection to iterate over. - * @param {...(Function|Function[])} [iteratees=[_.identity]] - * The iteratees to sort by. - * @returns {Array} Returns the new sorted array. - * @example - * - * var users = [ - * { 'user': 'fred', 'age': 48 }, - * { 'user': 'barney', 'age': 36 }, - * { 'user': 'fred', 'age': 30 }, - * { 'user': 'barney', 'age': 34 } - * ]; - * - * _.sortBy(users, [function(o) { return o.user; }]); - * // => objects for [['barney', 36], ['barney', 34], ['fred', 48], ['fred', 30]] - * - * _.sortBy(users, ['user', 'age']); - * // => objects for [['barney', 34], ['barney', 36], ['fred', 30], ['fred', 48]] - */ -var sortBy = baseRest(function(collection, iteratees) { - if (collection == null) { - return []; - } - var length = iteratees.length; - if (length > 1 && isIterateeCall$3(collection, iteratees[0], iteratees[1])) { - iteratees = []; - } else if (length > 2 && isIterateeCall$3(iteratees[0], iteratees[1], iteratees[2])) { - iteratees = [iteratees[0]]; - } - return baseOrderBy(collection, baseFlatten$2(iteratees, 1), []); -}); - -var sortBy_1 = sortBy; - -function _typeof$r(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$r = function _typeof(obj) { return typeof obj; }; } else { _typeof$r = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$r(obj); } - -function _slicedToArray$8(arr, i) { return _arrayWithHoles$8(arr) || _iterableToArrayLimit$8(arr, i) || _unsupportedIterableToArray$f(arr, i) || _nonIterableRest$8(); } - -function _nonIterableRest$8() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$f(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$f(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$f(o, minLen); } - -function _arrayLikeToArray$f(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit$8(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$8(arr) { if (Array.isArray(arr)) return arr; } - -function ownKeys$u(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$u(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$u(Object(source), true).forEach(function (key) { _defineProperty$u(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$u(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$u(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$r(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$r(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$r(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$r(Constructor.prototype, protoProps); if (staticProps) _defineProperties$r(Constructor, staticProps); return Constructor; } - -function _inherits$q(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$q(subClass, superClass); } - -function _setPrototypeOf$q(o, p) { _setPrototypeOf$q = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$q(o, p); } - -function _createSuper$q(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$q(); return function _createSuperInternal() { var Super = _getPrototypeOf$q(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$q(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$q(this, result); }; } - -function _possibleConstructorReturn$q(self, call) { if (call && (_typeof$r(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$q(self); } - -function _assertThisInitialized$q(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$q() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$q(o) { _getPrototypeOf$q = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$q(o); } - -function defaultFormatter(value) { - return isArray_1(value) && isNumOrStr(value[0]) && isNumOrStr(value[1]) ? value.join(' ~ ') : value; -} - -var DefaultTooltipContent = /*#__PURE__*/function (_PureComponent) { - _inherits$q(DefaultTooltipContent, _PureComponent); - - var _super = _createSuper$q(DefaultTooltipContent); - - function DefaultTooltipContent() { - _classCallCheck$r(this, DefaultTooltipContent); - - return _super.apply(this, arguments); - } - - _createClass$r(DefaultTooltipContent, [{ - key: "renderContent", - value: function renderContent() { - var _this$props = this.props, - payload = _this$props.payload, - separator = _this$props.separator, - formatter = _this$props.formatter, - itemStyle = _this$props.itemStyle, - itemSorter = _this$props.itemSorter; - - if (payload && payload.length) { - var listStyle = { - padding: 0, - margin: 0 - }; - var items = (itemSorter ? sortBy_1(payload, itemSorter) : payload).map(function (entry, i) { - if (entry.type === 'none') { - return null; - } - - var finalItemStyle = _objectSpread$u({ - display: 'block', - paddingTop: 4, - paddingBottom: 4, - color: entry.color || '#000' - }, itemStyle); - - var finalFormatter = entry.formatter || formatter || defaultFormatter; - var value = entry.value, - name = entry.name; - - if (finalFormatter && value != null && name != null) { - var formatted = finalFormatter(value, name, entry, i, payload); - - if (Array.isArray(formatted)) { - var _ref = formatted; - - var _ref2 = _slicedToArray$8(_ref, 2); - - value = _ref2[0]; - name = _ref2[1]; - } else { - value = formatted; - } - } - - return ( - /*#__PURE__*/ - // eslint-disable-next-line react/no-array-index-key - React__default.createElement("li", { - className: "recharts-tooltip-item", - key: "tooltip-item-".concat(i), - style: finalItemStyle - }, isNumOrStr(name) ? /*#__PURE__*/React__default.createElement("span", { - className: "recharts-tooltip-item-name" - }, name) : null, isNumOrStr(name) ? /*#__PURE__*/React__default.createElement("span", { - className: "recharts-tooltip-item-separator" - }, separator) : null, /*#__PURE__*/React__default.createElement("span", { - className: "recharts-tooltip-item-value" - }, value), /*#__PURE__*/React__default.createElement("span", { - className: "recharts-tooltip-item-unit" - }, entry.unit || '')) - ); - }); - return /*#__PURE__*/React__default.createElement("ul", { - className: "recharts-tooltip-item-list", - style: listStyle - }, items); - } - - return null; - } - }, { - key: "render", - value: function render() { - var _this$props2 = this.props, - wrapperClassName = _this$props2.wrapperClassName, - contentStyle = _this$props2.contentStyle, - labelClassName = _this$props2.labelClassName, - labelStyle = _this$props2.labelStyle, - label = _this$props2.label, - labelFormatter = _this$props2.labelFormatter, - payload = _this$props2.payload; - - var finalStyle = _objectSpread$u({ - margin: 0, - padding: 10, - backgroundColor: '#fff', - border: '1px solid #ccc', - whiteSpace: 'nowrap' - }, contentStyle); - - var finalLabelStyle = _objectSpread$u({ - margin: 0 - }, labelStyle); - - var hasLabel = !isNil_1(label); - var finalLabel = hasLabel ? label : ''; - var wrapperCN = classNames('recharts-default-tooltip', wrapperClassName); - var labelCN = classNames('recharts-tooltip-label', labelClassName); - - if (hasLabel && labelFormatter && payload !== undefined && payload !== null) { - finalLabel = labelFormatter(label, payload); - } - - return /*#__PURE__*/React__default.createElement("div", { - className: wrapperCN, - style: finalStyle - }, /*#__PURE__*/React__default.createElement("p", { - className: labelCN, - style: finalLabelStyle - }, /*#__PURE__*/React__default.isValidElement(finalLabel) ? finalLabel : "".concat(finalLabel)), this.renderContent()); - } - }]); - - return DefaultTooltipContent; -}(PureComponent); -DefaultTooltipContent.displayName = 'DefaultTooltipContent'; -DefaultTooltipContent.defaultProps = { - separator: ' : ', - contentStyle: {}, - itemStyle: {}, - labelStyle: {} -}; - -var parseIsSsrByDefault = function parseIsSsrByDefault() { - return !(window.document && window.document.createElement && window.setTimeout); -}; - -var Global = { - isSsr: parseIsSsrByDefault(), - get: function get(key) { - return Global[key]; - }, - set: function set(key, value) { - if (typeof key === 'string') { - Global[key] = value; - } else { - var keys = Object.keys(key); - - if (keys && keys.length) { - keys.forEach(function (k) { - Global[k] = key[k]; - }); - } - } - } -}; - -function _typeof$q(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$q = function _typeof(obj) { return typeof obj; }; } else { _typeof$q = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$q(obj); } - -function ownKeys$t(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$t(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$t(Object(source), true).forEach(function (key) { _defineProperty$t(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$t(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$t(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$q(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$q(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$q(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$q(Constructor.prototype, protoProps); if (staticProps) _defineProperties$q(Constructor, staticProps); return Constructor; } - -function _inherits$p(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$p(subClass, superClass); } - -function _setPrototypeOf$p(o, p) { _setPrototypeOf$p = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$p(o, p); } - -function _createSuper$p(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$p(); return function _createSuperInternal() { var Super = _getPrototypeOf$p(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$p(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$p(this, result); }; } - -function _possibleConstructorReturn$p(self, call) { if (call && (_typeof$q(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$p(self); } - -function _assertThisInitialized$p(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$p() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$p(o) { _getPrototypeOf$p = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$p(o); } -var CLS_PREFIX = 'recharts-tooltip-wrapper'; -var EPS$1 = 1; - -function defaultUniqBy(entry) { - return entry.dataKey; -} - -function getUniqPayload(option, payload) { - if (option === true) { - return uniqBy_1(payload, defaultUniqBy); - } - - if (isFunction_1(option)) { - return uniqBy_1(payload, option); - } - - return payload; -} - -function renderContent(content, props) { - if ( /*#__PURE__*/React__default.isValidElement(content)) { - return /*#__PURE__*/React__default.cloneElement(content, props); - } - - if (isFunction_1(content)) { - return /*#__PURE__*/React__default.createElement(content, props); - } - - return /*#__PURE__*/React__default.createElement(DefaultTooltipContent, props); -} - -var Tooltip = /*#__PURE__*/function (_PureComponent) { - _inherits$p(Tooltip, _PureComponent); - - var _super = _createSuper$p(Tooltip); - - function Tooltip() { - var _this; - - _classCallCheck$q(this, Tooltip); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - boxWidth: -1, - boxHeight: -1, - dismissed: false, - dismissedAtCoordinate: { - x: 0, - y: 0 - } - }; - _this.wrapperNode = void 0; - - _this.getTranslate = function (_ref) { - var key = _ref.key, - tooltipDimension = _ref.tooltipDimension, - viewBoxDimension = _ref.viewBoxDimension; - var _this$props = _this.props, - allowEscapeViewBox = _this$props.allowEscapeViewBox, - coordinate = _this$props.coordinate, - offset = _this$props.offset, - position = _this$props.position, - viewBox = _this$props.viewBox; - - if (position && isNumber(position[key])) { - return position[key]; - } - - var restricted = coordinate[key] - tooltipDimension - offset; - var unrestricted = coordinate[key] + offset; - - if (allowEscapeViewBox[key]) { - return unrestricted; - } - - var tooltipBoundary = coordinate[key] + tooltipDimension + offset; - var viewBoxBoundary = viewBox[key] + viewBoxDimension; - - if (tooltipBoundary > viewBoxBoundary) { - return Math.max(restricted, viewBox[key]); - } - - return Math.max(unrestricted, viewBox[key]); - }; - - return _this; - } - - _createClass$q(Tooltip, [{ - key: "componentDidMount", - value: function componentDidMount() { - this.updateBBox(); - } - }, { - key: "componentDidUpdate", - value: function componentDidUpdate() { - this.updateBBox(); - } - }, { - key: "updateBBox", - value: function updateBBox() { - var _this$state = this.state, - boxWidth = _this$state.boxWidth, - boxHeight = _this$state.boxHeight, - dismissed = _this$state.dismissed; - - if (dismissed) { - this.wrapperNode.blur(); - - if (this.props.coordinate.x !== this.state.dismissedAtCoordinate.x || this.props.coordinate.y !== this.state.dismissedAtCoordinate.y) { - this.setState({ - dismissed: false - }); - } - } else { - this.wrapperNode.focus({ - preventScroll: true - }); - } - - if (this.wrapperNode && this.wrapperNode.getBoundingClientRect) { - var box = this.wrapperNode.getBoundingClientRect(); - - if (Math.abs(box.width - boxWidth) > EPS$1 || Math.abs(box.height - boxHeight) > EPS$1) { - this.setState({ - boxWidth: box.width, - boxHeight: box.height - }); - } - } else if (boxWidth !== -1 || boxHeight !== -1) { - this.setState({ - boxWidth: -1, - boxHeight: -1 - }); - } - } - }, { - key: "render", - value: function render() { - var _classNames, - _this2 = this; - - var _this$props2 = this.props, - payload = _this$props2.payload, - isAnimationActive = _this$props2.isAnimationActive, - animationDuration = _this$props2.animationDuration, - animationEasing = _this$props2.animationEasing, - filterNull = _this$props2.filterNull, - payloadUniqBy = _this$props2.payloadUniqBy; - var finalPayload = getUniqPayload(payloadUniqBy, filterNull && payload && payload.length ? payload.filter(function (entry) { - return !isNil_1(entry.value); - }) : payload); - var hasPayload = finalPayload && finalPayload.length; - var _this$props3 = this.props, - content = _this$props3.content, - viewBox = _this$props3.viewBox, - coordinate = _this$props3.coordinate, - position = _this$props3.position, - active = _this$props3.active, - wrapperStyle = _this$props3.wrapperStyle; - - var outerStyle = _objectSpread$t({ - pointerEvents: 'none', - visibility: !this.state.dismissed && active && hasPayload ? 'visible' : 'hidden', - position: 'absolute', - top: 0, - left: 0 - }, wrapperStyle); - - var translateX, translateY; - - if (position && isNumber(position.x) && isNumber(position.y)) { - translateX = position.x; - translateY = position.y; - } else { - var _this$state2 = this.state, - boxWidth = _this$state2.boxWidth, - boxHeight = _this$state2.boxHeight; - - if (boxWidth > 0 && boxHeight > 0 && coordinate) { - translateX = this.getTranslate({ - key: 'x', - tooltipDimension: boxWidth, - viewBoxDimension: viewBox.width - }); - translateY = this.getTranslate({ - key: 'y', - tooltipDimension: boxHeight, - viewBoxDimension: viewBox.height - }); - } else { - outerStyle.visibility = 'hidden'; - } - } - - outerStyle = _objectSpread$t(_objectSpread$t({}, lib$2.translateStyle({ - transform: this.props.useTranslate3d ? "translate3d(".concat(translateX, "px, ").concat(translateY, "px, 0)") : "translate(".concat(translateX, "px, ").concat(translateY, "px)") - })), outerStyle); - - if (isAnimationActive && active) { - outerStyle = _objectSpread$t(_objectSpread$t({}, lib$2.translateStyle({ - transition: "transform ".concat(animationDuration, "ms ").concat(animationEasing) - })), outerStyle); - } - - var cls = classNames(CLS_PREFIX, (_classNames = {}, _defineProperty$t(_classNames, "".concat(CLS_PREFIX, "-right"), isNumber(translateX) && coordinate && isNumber(coordinate.x) && translateX >= coordinate.x), _defineProperty$t(_classNames, "".concat(CLS_PREFIX, "-left"), isNumber(translateX) && coordinate && isNumber(coordinate.x) && translateX < coordinate.x), _defineProperty$t(_classNames, "".concat(CLS_PREFIX, "-bottom"), isNumber(translateY) && coordinate && isNumber(coordinate.y) && translateY >= coordinate.y), _defineProperty$t(_classNames, "".concat(CLS_PREFIX, "-top"), isNumber(translateY) && coordinate && isNumber(coordinate.y) && translateY < coordinate.y), _classNames)); - return ( - /*#__PURE__*/ - // ESLint is disabled to allow listening to the `Escape` key. Refer to - // https://github.com/recharts/recharts/pull/2925 - // eslint-disable-next-line jsx-a11y/no-noninteractive-element-interactions - React__default.createElement("div", { - tabIndex: -1, - role: "dialog", - onKeyDown: function onKeyDown(event) { - if (event.key === 'Escape') { - _this2.setState({ - dismissed: true, - dismissedAtCoordinate: _objectSpread$t(_objectSpread$t({}, _this2.state.dismissedAtCoordinate), {}, { - x: _this2.props.coordinate.x, - y: _this2.props.coordinate.y - }) - }); - } - }, - className: cls, - style: outerStyle, - ref: function ref(node) { - _this2.wrapperNode = node; - } - }, renderContent(content, _objectSpread$t(_objectSpread$t({}, this.props), {}, { - payload: finalPayload - }))) - ); - } - }]); - - return Tooltip; -}(PureComponent); -Tooltip.displayName = 'Tooltip'; -Tooltip.defaultProps = { - active: false, - allowEscapeViewBox: { - x: false, - y: false - }, - offset: 10, - viewBox: { - x1: 0, - x2: 0, - y1: 0, - y2: 0 - }, - coordinate: { - x: 0, - y: 0 - }, - cursorStyle: {}, - separator: ' : ', - wrapperStyle: {}, - contentStyle: {}, - itemStyle: {}, - labelStyle: {}, - cursor: true, - trigger: 'hover', - isAnimationActive: !Global.isSsr, - animationEasing: 'ease', - animationDuration: 400, - filterNull: true, - useTranslate3d: false -}; - -var root$3 = _root$1; - -/** - * Gets the timestamp of the number of milliseconds that have elapsed since - * the Unix epoch (1 January 1970 00:00:00 UTC). - * - * @static - * @memberOf _ - * @since 2.4.0 - * @category Date - * @returns {number} Returns the timestamp. - * @example - * - * _.defer(function(stamp) { - * console.log(_.now() - stamp); - * }, _.now()); - * // => Logs the number of milliseconds it took for the deferred invocation. - */ -var now$3 = function() { - return root$3.Date.now(); -}; - -var now_1$1 = now$3; - -/** Used to match a single whitespace character. */ - -var reWhitespace$1 = /\s/; - -/** - * Used by `_.trim` and `_.trimEnd` to get the index of the last non-whitespace - * character of `string`. - * - * @private - * @param {string} string The string to inspect. - * @returns {number} Returns the index of the last non-whitespace character. - */ -function trimmedEndIndex$3(string) { - var index = string.length; - - while (index-- && reWhitespace$1.test(string.charAt(index))) {} - return index; -} - -var _trimmedEndIndex$1 = trimmedEndIndex$3; - -var trimmedEndIndex$2 = _trimmedEndIndex$1; - -/** Used to match leading whitespace. */ -var reTrimStart$1 = /^\s+/; - -/** - * The base implementation of `_.trim`. - * - * @private - * @param {string} string The string to trim. - * @returns {string} Returns the trimmed string. - */ -function baseTrim$3(string) { - return string - ? string.slice(0, trimmedEndIndex$2(string) + 1).replace(reTrimStart$1, '') - : string; -} - -var _baseTrim$1 = baseTrim$3; - -var baseTrim$2 = _baseTrim$1, - isObject$6 = isObject_1$1, - isSymbol$3 = isSymbol_1$1; - -/** Used as references for various `Number` constants. */ -var NAN$1 = 0 / 0; - -/** Used to detect bad signed hexadecimal string values. */ -var reIsBadHex$1 = /^[-+]0x[0-9a-f]+$/i; - -/** Used to detect binary string values. */ -var reIsBinary$1 = /^0b[01]+$/i; - -/** Used to detect octal string values. */ -var reIsOctal$1 = /^0o[0-7]+$/i; - -/** Built-in method references without a dependency on `root`. */ -var freeParseInt$1 = parseInt; - -/** - * Converts `value` to a number. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to process. - * @returns {number} Returns the number. - * @example - * - * _.toNumber(3.2); - * // => 3.2 - * - * _.toNumber(Number.MIN_VALUE); - * // => 5e-324 - * - * _.toNumber(Infinity); - * // => Infinity - * - * _.toNumber('3.2'); - * // => 3.2 - */ -function toNumber$4(value) { - if (typeof value == 'number') { - return value; - } - if (isSymbol$3(value)) { - return NAN$1; - } - if (isObject$6(value)) { - var other = typeof value.valueOf == 'function' ? value.valueOf() : value; - value = isObject$6(other) ? (other + '') : other; - } - if (typeof value != 'string') { - return value === 0 ? value : +value; - } - value = baseTrim$2(value); - var isBinary = reIsBinary$1.test(value); - return (isBinary || reIsOctal$1.test(value)) - ? freeParseInt$1(value.slice(2), isBinary ? 2 : 8) - : (reIsBadHex$1.test(value) ? NAN$1 : +value); -} - -var toNumber_1$1 = toNumber$4; - -var isObject$5 = isObject_1$1, - now$2 = now_1$1, - toNumber$3 = toNumber_1$1; - -/** Error message constants. */ -var FUNC_ERROR_TEXT$3 = 'Expected a function'; - -/* Built-in method references for those with the same name as other `lodash` methods. */ -var nativeMax$3 = Math.max, - nativeMin$1 = Math.min; - -/** - * Creates a debounced function that delays invoking `func` until after `wait` - * milliseconds have elapsed since the last time the debounced function was - * invoked. The debounced function comes with a `cancel` method to cancel - * delayed `func` invocations and a `flush` method to immediately invoke them. - * Provide `options` to indicate whether `func` should be invoked on the - * leading and/or trailing edge of the `wait` timeout. The `func` is invoked - * with the last arguments provided to the debounced function. Subsequent - * calls to the debounced function return the result of the last `func` - * invocation. - * - * **Note:** If `leading` and `trailing` options are `true`, `func` is - * invoked on the trailing edge of the timeout only if the debounced function - * is invoked more than once during the `wait` timeout. - * - * If `wait` is `0` and `leading` is `false`, `func` invocation is deferred - * until to the next tick, similar to `setTimeout` with a timeout of `0`. - * - * See [David Corbacho's article](https://css-tricks.com/debouncing-throttling-explained-examples/) - * for details over the differences between `_.debounce` and `_.throttle`. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Function - * @param {Function} func The function to debounce. - * @param {number} [wait=0] The number of milliseconds to delay. - * @param {Object} [options={}] The options object. - * @param {boolean} [options.leading=false] - * Specify invoking on the leading edge of the timeout. - * @param {number} [options.maxWait] - * The maximum time `func` is allowed to be delayed before it's invoked. - * @param {boolean} [options.trailing=true] - * Specify invoking on the trailing edge of the timeout. - * @returns {Function} Returns the new debounced function. - * @example - * - * // Avoid costly calculations while the window size is in flux. - * jQuery(window).on('resize', _.debounce(calculateLayout, 150)); - * - * // Invoke `sendMail` when clicked, debouncing subsequent calls. - * jQuery(element).on('click', _.debounce(sendMail, 300, { - * 'leading': true, - * 'trailing': false - * })); - * - * // Ensure `batchLog` is invoked once after 1 second of debounced calls. - * var debounced = _.debounce(batchLog, 250, { 'maxWait': 1000 }); - * var source = new EventSource('/stream'); - * jQuery(source).on('message', debounced); - * - * // Cancel the trailing debounced invocation. - * jQuery(window).on('popstate', debounced.cancel); - */ -function debounce$3(func, wait, options) { - var lastArgs, - lastThis, - maxWait, - result, - timerId, - lastCallTime, - lastInvokeTime = 0, - leading = false, - maxing = false, - trailing = true; - - if (typeof func != 'function') { - throw new TypeError(FUNC_ERROR_TEXT$3); - } - wait = toNumber$3(wait) || 0; - if (isObject$5(options)) { - leading = !!options.leading; - maxing = 'maxWait' in options; - maxWait = maxing ? nativeMax$3(toNumber$3(options.maxWait) || 0, wait) : maxWait; - trailing = 'trailing' in options ? !!options.trailing : trailing; - } - - function invokeFunc(time) { - var args = lastArgs, - thisArg = lastThis; - - lastArgs = lastThis = undefined; - lastInvokeTime = time; - result = func.apply(thisArg, args); - return result; - } - - function leadingEdge(time) { - // Reset any `maxWait` timer. - lastInvokeTime = time; - // Start the timer for the trailing edge. - timerId = setTimeout(timerExpired, wait); - // Invoke the leading edge. - return leading ? invokeFunc(time) : result; - } - - function remainingWait(time) { - var timeSinceLastCall = time - lastCallTime, - timeSinceLastInvoke = time - lastInvokeTime, - timeWaiting = wait - timeSinceLastCall; - - return maxing - ? nativeMin$1(timeWaiting, maxWait - timeSinceLastInvoke) - : timeWaiting; - } - - function shouldInvoke(time) { - var timeSinceLastCall = time - lastCallTime, - timeSinceLastInvoke = time - lastInvokeTime; - - // Either this is the first call, activity has stopped and we're at the - // trailing edge, the system time has gone backwards and we're treating - // it as the trailing edge, or we've hit the `maxWait` limit. - return (lastCallTime === undefined || (timeSinceLastCall >= wait) || - (timeSinceLastCall < 0) || (maxing && timeSinceLastInvoke >= maxWait)); - } - - function timerExpired() { - var time = now$2(); - if (shouldInvoke(time)) { - return trailingEdge(time); - } - // Restart the timer. - timerId = setTimeout(timerExpired, remainingWait(time)); - } - - function trailingEdge(time) { - timerId = undefined; - - // Only invoke if we have `lastArgs` which means `func` has been - // debounced at least once. - if (trailing && lastArgs) { - return invokeFunc(time); - } - lastArgs = lastThis = undefined; - return result; - } - - function cancel() { - if (timerId !== undefined) { - clearTimeout(timerId); - } - lastInvokeTime = 0; - lastArgs = lastCallTime = lastThis = timerId = undefined; - } - - function flush() { - return timerId === undefined ? result : trailingEdge(now$2()); - } - - function debounced() { - var time = now$2(), - isInvoking = shouldInvoke(time); - - lastArgs = arguments; - lastThis = this; - lastCallTime = time; - - if (isInvoking) { - if (timerId === undefined) { - return leadingEdge(lastCallTime); - } - if (maxing) { - // Handle invocations in a tight loop. - clearTimeout(timerId); - timerId = setTimeout(timerExpired, wait); - return invokeFunc(lastCallTime); - } - } - if (timerId === undefined) { - timerId = setTimeout(timerExpired, wait); - } - return result; - } - debounced.cancel = cancel; - debounced.flush = flush; - return debounced; -} - -var debounce_1$1 = debounce$3; - -/* global Reflect, Promise */ - -var extendStatics = function(d, b) { - extendStatics = Object.setPrototypeOf || - ({ __proto__: [] } instanceof Array && function (d, b) { d.__proto__ = b; }) || - function (d, b) { for (var p in b) if (Object.prototype.hasOwnProperty.call(b, p)) d[p] = b[p]; }; - return extendStatics(d, b); -}; - -function __extends(d, b) { - if (typeof b !== "function" && b !== null) - throw new TypeError("Class extends value " + String(b) + " is not a constructor or null"); - extendStatics(d, b); - function __() { this.constructor = d; } - d.prototype = b === null ? Object.create(b) : (__.prototype = b.prototype, new __()); -} - -function __rest(s, e) { - var t = {}; - for (var p in s) if (Object.prototype.hasOwnProperty.call(s, p) && e.indexOf(p) < 0) - t[p] = s[p]; - if (s != null && typeof Object.getOwnPropertySymbols === "function") - for (var i = 0, p = Object.getOwnPropertySymbols(s); i < p.length; i++) { - if (e.indexOf(p[i]) < 0 && Object.prototype.propertyIsEnumerable.call(s, p[i])) - t[p[i]] = s[p[i]]; - } - return t; -}var commonjsGlobal = typeof globalThis !== 'undefined' ? globalThis : window ;/** - * Checks if `value` is the - * [language type](http://www.ecma-international.org/ecma-262/7.0/#sec-ecmascript-language-types) - * of `Object`. (e.g. arrays, functions, objects, regexes, `new Number(0)`, and `new String('')`) - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is an object, else `false`. - * @example - * - * _.isObject({}); - * // => true - * - * _.isObject([1, 2, 3]); - * // => true - * - * _.isObject(_.noop); - * // => true - * - * _.isObject(null); - * // => false - */ - -function isObject$3$1(value) { - var type = typeof value; - return value != null && (type == 'object' || type == 'function'); -} - -var isObject_1 = isObject$3$1;/** Detect free variable `global` from Node.js. */ - -var freeGlobal$1 = typeof commonjsGlobal == 'object' && commonjsGlobal && commonjsGlobal.Object === Object && commonjsGlobal; - -var _freeGlobal = freeGlobal$1;var freeGlobal = _freeGlobal; - -/** Detect free variable `self`. */ -var freeSelf = typeof self == 'object' && self && self.Object === Object && self; - -/** Used as a reference to the global object. */ -var root$2 = freeGlobal || freeSelf || Function('return this')(); - -var _root = root$2;var root$1 = _root; - -/** - * Gets the timestamp of the number of milliseconds that have elapsed since - * the Unix epoch (1 January 1970 00:00:00 UTC). - * - * @static - * @memberOf _ - * @since 2.4.0 - * @category Date - * @returns {number} Returns the timestamp. - * @example - * - * _.defer(function(stamp) { - * console.log(_.now() - stamp); - * }, _.now()); - * // => Logs the number of milliseconds it took for the deferred invocation. - */ -var now$1 = function() { - return root$1.Date.now(); -}; - -var now_1 = now$1;/** Used to match a single whitespace character. */ - -var reWhitespace = /\s/; - -/** - * Used by `_.trim` and `_.trimEnd` to get the index of the last non-whitespace - * character of `string`. - * - * @private - * @param {string} string The string to inspect. - * @returns {number} Returns the index of the last non-whitespace character. - */ -function trimmedEndIndex$1(string) { - var index = string.length; - - while (index-- && reWhitespace.test(string.charAt(index))) {} - return index; -} - -var _trimmedEndIndex = trimmedEndIndex$1;var trimmedEndIndex = _trimmedEndIndex; - -/** Used to match leading whitespace. */ -var reTrimStart = /^\s+/; - -/** - * The base implementation of `_.trim`. - * - * @private - * @param {string} string The string to trim. - * @returns {string} Returns the trimmed string. - */ -function baseTrim$1(string) { - return string - ? string.slice(0, trimmedEndIndex(string) + 1).replace(reTrimStart, '') - : string; -} - -var _baseTrim = baseTrim$1;var root = _root; - -/** Built-in value references. */ -var Symbol$2 = root.Symbol; - -var _Symbol = Symbol$2;var Symbol$1$1 = _Symbol; - -/** Used for built-in method references. */ -var objectProto$1$1 = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$4 = objectProto$1$1.hasOwnProperty; - -/** - * Used to resolve the - * [`toStringTag`](http://ecma-international.org/ecma-262/7.0/#sec-object.prototype.tostring) - * of values. - */ -var nativeObjectToString$1 = objectProto$1$1.toString; - -/** Built-in value references. */ -var symToStringTag$1 = Symbol$1$1 ? Symbol$1$1.toStringTag : undefined; - -/** - * A specialized version of `baseGetTag` which ignores `Symbol.toStringTag` values. - * - * @private - * @param {*} value The value to query. - * @returns {string} Returns the raw `toStringTag`. - */ -function getRawTag$1(value) { - var isOwn = hasOwnProperty$4.call(value, symToStringTag$1), - tag = value[symToStringTag$1]; - - try { - value[symToStringTag$1] = undefined; - var unmasked = true; - } catch (e) {} - - var result = nativeObjectToString$1.call(value); - if (unmasked) { - if (isOwn) { - value[symToStringTag$1] = tag; - } else { - delete value[symToStringTag$1]; - } - } - return result; -} - -var _getRawTag = getRawTag$1;/** Used for built-in method references. */ - -var objectProto$4 = Object.prototype; - -/** - * Used to resolve the - * [`toStringTag`](http://ecma-international.org/ecma-262/7.0/#sec-object.prototype.tostring) - * of values. - */ -var nativeObjectToString = objectProto$4.toString; - -/** - * Converts `value` to a string using `Object.prototype.toString`. - * - * @private - * @param {*} value The value to convert. - * @returns {string} Returns the converted string. - */ -function objectToString$1(value) { - return nativeObjectToString.call(value); -} - -var _objectToString = objectToString$1;var Symbol$3 = _Symbol, - getRawTag = _getRawTag, - objectToString = _objectToString; - -/** `Object#toString` result references. */ -var nullTag = '[object Null]', - undefinedTag = '[object Undefined]'; - -/** Built-in value references. */ -var symToStringTag = Symbol$3 ? Symbol$3.toStringTag : undefined; - -/** - * The base implementation of `getTag` without fallbacks for buggy environments. - * - * @private - * @param {*} value The value to query. - * @returns {string} Returns the `toStringTag`. - */ -function baseGetTag$1$1(value) { - if (value == null) { - return value === undefined ? undefinedTag : nullTag; - } - return (symToStringTag && symToStringTag in Object(value)) - ? getRawTag(value) - : objectToString(value); -} - -var _baseGetTag = baseGetTag$1$1;/** - * Checks if `value` is object-like. A value is object-like if it's not `null` - * and has a `typeof` result of "object". - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is object-like, else `false`. - * @example - * - * _.isObjectLike({}); - * // => true - * - * _.isObjectLike([1, 2, 3]); - * // => true - * - * _.isObjectLike(_.noop); - * // => false - * - * _.isObjectLike(null); - * // => false - */ - -function isObjectLike$1$1(value) { - return value != null && typeof value == 'object'; -} - -var isObjectLike_1 = isObjectLike$1$1;var baseGetTag$2 = _baseGetTag, - isObjectLike$4 = isObjectLike_1; - -/** `Object#toString` result references. */ -var symbolTag$2 = '[object Symbol]'; - -/** - * Checks if `value` is classified as a `Symbol` primitive or object. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a symbol, else `false`. - * @example - * - * _.isSymbol(Symbol.iterator); - * // => true - * - * _.isSymbol('abc'); - * // => false - */ -function isSymbol$1(value) { - return typeof value == 'symbol' || - (isObjectLike$4(value) && baseGetTag$2(value) == symbolTag$2); -} - -var isSymbol_1 = isSymbol$1;var baseTrim = _baseTrim, - isObject$2$1 = isObject_1, - isSymbol$2 = isSymbol_1; - -/** Used as references for various `Number` constants. */ -var NAN = 0 / 0; - -/** Used to detect bad signed hexadecimal string values. */ -var reIsBadHex = /^[-+]0x[0-9a-f]+$/i; - -/** Used to detect binary string values. */ -var reIsBinary = /^0b[01]+$/i; - -/** Used to detect octal string values. */ -var reIsOctal = /^0o[0-7]+$/i; - -/** Built-in method references without a dependency on `root`. */ -var freeParseInt = parseInt; - -/** - * Converts `value` to a number. - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to process. - * @returns {number} Returns the number. - * @example - * - * _.toNumber(3.2); - * // => 3.2 - * - * _.toNumber(Number.MIN_VALUE); - * // => 5e-324 - * - * _.toNumber(Infinity); - * // => Infinity - * - * _.toNumber('3.2'); - * // => 3.2 - */ -function toNumber$1(value) { - if (typeof value == 'number') { - return value; - } - if (isSymbol$2(value)) { - return NAN; - } - if (isObject$2$1(value)) { - var other = typeof value.valueOf == 'function' ? value.valueOf() : value; - value = isObject$2$1(other) ? (other + '') : other; - } - if (typeof value != 'string') { - return value === 0 ? value : +value; - } - value = baseTrim(value); - var isBinary = reIsBinary.test(value); - return (isBinary || reIsOctal.test(value)) - ? freeParseInt(value.slice(2), isBinary ? 2 : 8) - : (reIsBadHex.test(value) ? NAN : +value); -} - -var toNumber_1 = toNumber$1;var isObject$1$1 = isObject_1, - now = now_1, - toNumber$2 = toNumber_1; - -/** Error message constants. */ -var FUNC_ERROR_TEXT$1 = 'Expected a function'; - -/* Built-in method references for those with the same name as other `lodash` methods. */ -var nativeMax$2 = Math.max, - nativeMin = Math.min; - -/** - * Creates a debounced function that delays invoking `func` until after `wait` - * milliseconds have elapsed since the last time the debounced function was - * invoked. The debounced function comes with a `cancel` method to cancel - * delayed `func` invocations and a `flush` method to immediately invoke them. - * Provide `options` to indicate whether `func` should be invoked on the - * leading and/or trailing edge of the `wait` timeout. The `func` is invoked - * with the last arguments provided to the debounced function. Subsequent - * calls to the debounced function return the result of the last `func` - * invocation. - * - * **Note:** If `leading` and `trailing` options are `true`, `func` is - * invoked on the trailing edge of the timeout only if the debounced function - * is invoked more than once during the `wait` timeout. - * - * If `wait` is `0` and `leading` is `false`, `func` invocation is deferred - * until to the next tick, similar to `setTimeout` with a timeout of `0`. - * - * See [David Corbacho's article](https://css-tricks.com/debouncing-throttling-explained-examples/) - * for details over the differences between `_.debounce` and `_.throttle`. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Function - * @param {Function} func The function to debounce. - * @param {number} [wait=0] The number of milliseconds to delay. - * @param {Object} [options={}] The options object. - * @param {boolean} [options.leading=false] - * Specify invoking on the leading edge of the timeout. - * @param {number} [options.maxWait] - * The maximum time `func` is allowed to be delayed before it's invoked. - * @param {boolean} [options.trailing=true] - * Specify invoking on the trailing edge of the timeout. - * @returns {Function} Returns the new debounced function. - * @example - * - * // Avoid costly calculations while the window size is in flux. - * jQuery(window).on('resize', _.debounce(calculateLayout, 150)); - * - * // Invoke `sendMail` when clicked, debouncing subsequent calls. - * jQuery(element).on('click', _.debounce(sendMail, 300, { - * 'leading': true, - * 'trailing': false - * })); - * - * // Ensure `batchLog` is invoked once after 1 second of debounced calls. - * var debounced = _.debounce(batchLog, 250, { 'maxWait': 1000 }); - * var source = new EventSource('/stream'); - * jQuery(source).on('message', debounced); - * - * // Cancel the trailing debounced invocation. - * jQuery(window).on('popstate', debounced.cancel); - */ -function debounce$1(func, wait, options) { - var lastArgs, - lastThis, - maxWait, - result, - timerId, - lastCallTime, - lastInvokeTime = 0, - leading = false, - maxing = false, - trailing = true; - - if (typeof func != 'function') { - throw new TypeError(FUNC_ERROR_TEXT$1); - } - wait = toNumber$2(wait) || 0; - if (isObject$1$1(options)) { - leading = !!options.leading; - maxing = 'maxWait' in options; - maxWait = maxing ? nativeMax$2(toNumber$2(options.maxWait) || 0, wait) : maxWait; - trailing = 'trailing' in options ? !!options.trailing : trailing; - } - - function invokeFunc(time) { - var args = lastArgs, - thisArg = lastThis; - - lastArgs = lastThis = undefined; - lastInvokeTime = time; - result = func.apply(thisArg, args); - return result; - } - - function leadingEdge(time) { - // Reset any `maxWait` timer. - lastInvokeTime = time; - // Start the timer for the trailing edge. - timerId = setTimeout(timerExpired, wait); - // Invoke the leading edge. - return leading ? invokeFunc(time) : result; - } - - function remainingWait(time) { - var timeSinceLastCall = time - lastCallTime, - timeSinceLastInvoke = time - lastInvokeTime, - timeWaiting = wait - timeSinceLastCall; - - return maxing - ? nativeMin(timeWaiting, maxWait - timeSinceLastInvoke) - : timeWaiting; - } - - function shouldInvoke(time) { - var timeSinceLastCall = time - lastCallTime, - timeSinceLastInvoke = time - lastInvokeTime; - - // Either this is the first call, activity has stopped and we're at the - // trailing edge, the system time has gone backwards and we're treating - // it as the trailing edge, or we've hit the `maxWait` limit. - return (lastCallTime === undefined || (timeSinceLastCall >= wait) || - (timeSinceLastCall < 0) || (maxing && timeSinceLastInvoke >= maxWait)); - } - - function timerExpired() { - var time = now(); - if (shouldInvoke(time)) { - return trailingEdge(time); - } - // Restart the timer. - timerId = setTimeout(timerExpired, remainingWait(time)); - } - - function trailingEdge(time) { - timerId = undefined; - - // Only invoke if we have `lastArgs` which means `func` has been - // debounced at least once. - if (trailing && lastArgs) { - return invokeFunc(time); - } - lastArgs = lastThis = undefined; - return result; - } - - function cancel() { - if (timerId !== undefined) { - clearTimeout(timerId); - } - lastInvokeTime = 0; - lastArgs = lastCallTime = lastThis = timerId = undefined; - } - - function flush() { - return timerId === undefined ? result : trailingEdge(now()); - } - - function debounced() { - var time = now(), - isInvoking = shouldInvoke(time); - - lastArgs = arguments; - lastThis = this; - lastCallTime = time; - - if (isInvoking) { - if (timerId === undefined) { - return leadingEdge(lastCallTime); - } - if (maxing) { - // Handle invocations in a tight loop. - clearTimeout(timerId); - timerId = setTimeout(timerExpired, wait); - return invokeFunc(lastCallTime); - } - } - if (timerId === undefined) { - timerId = setTimeout(timerExpired, wait); - } - return result; - } - debounced.cancel = cancel; - debounced.flush = flush; - return debounced; -} - -var debounce_1 = debounce$1;var debounce$2 = debounce_1, - isObject$4 = isObject_1; - -/** Error message constants. */ -var FUNC_ERROR_TEXT$2 = 'Expected a function'; - -/** - * Creates a throttled function that only invokes `func` at most once per - * every `wait` milliseconds. The throttled function comes with a `cancel` - * method to cancel delayed `func` invocations and a `flush` method to - * immediately invoke them. Provide `options` to indicate whether `func` - * should be invoked on the leading and/or trailing edge of the `wait` - * timeout. The `func` is invoked with the last arguments provided to the - * throttled function. Subsequent calls to the throttled function return the - * result of the last `func` invocation. - * - * **Note:** If `leading` and `trailing` options are `true`, `func` is - * invoked on the trailing edge of the timeout only if the throttled function - * is invoked more than once during the `wait` timeout. - * - * If `wait` is `0` and `leading` is `false`, `func` invocation is deferred - * until to the next tick, similar to `setTimeout` with a timeout of `0`. - * - * See [David Corbacho's article](https://css-tricks.com/debouncing-throttling-explained-examples/) - * for details over the differences between `_.throttle` and `_.debounce`. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Function - * @param {Function} func The function to throttle. - * @param {number} [wait=0] The number of milliseconds to throttle invocations to. - * @param {Object} [options={}] The options object. - * @param {boolean} [options.leading=true] - * Specify invoking on the leading edge of the timeout. - * @param {boolean} [options.trailing=true] - * Specify invoking on the trailing edge of the timeout. - * @returns {Function} Returns the new throttled function. - * @example - * - * // Avoid excessively updating the position while scrolling. - * jQuery(window).on('scroll', _.throttle(updatePosition, 100)); - * - * // Invoke `renewToken` when the click event is fired, but not more than once every 5 minutes. - * var throttled = _.throttle(renewToken, 300000, { 'trailing': false }); - * jQuery(element).on('click', throttled); - * - * // Cancel the trailing throttled invocation. - * jQuery(window).on('popstate', throttled.cancel); - */ -function throttle$1(func, wait, options) { - var leading = true, - trailing = true; - - if (typeof func != 'function') { - throw new TypeError(FUNC_ERROR_TEXT$2); - } - if (isObject$4(options)) { - leading = 'leading' in options ? !!options.leading : leading; - trailing = 'trailing' in options ? !!options.trailing : trailing; - } - return debounce$2(func, wait, { - 'leading': leading, - 'maxWait': wait, - 'trailing': trailing - }); -} - -var throttle_1$1 = throttle$1;var patchResizeHandler = function (resizeCallback, refreshMode, refreshRate, refreshOptions) { - switch (refreshMode) { - case 'debounce': - return debounce_1(resizeCallback, refreshRate, refreshOptions); - case 'throttle': - return throttle_1$1(resizeCallback, refreshRate, refreshOptions); - default: - return resizeCallback; - } -}; -var isFunction = function (fn) { return typeof fn === 'function'; }; -var isSSR = function () { return "object" === 'undefined'; }; -var isDOMElement = function (element) { return element instanceof Element || element instanceof HTMLDocument; }; -var createNotifier = function (onResize, setSize, handleWidth, handleHeight) { - return function (_a) { - var width = _a.width, height = _a.height; - setSize(function (prev) { - if (prev.width === width && prev.height === height) { - // skip if dimensions haven't changed - return prev; - } - if ((prev.width === width && !handleHeight) || (prev.height === height && !handleWidth)) { - // process `handleHeight/handleWidth` props - return prev; - } - if (onResize && isFunction(onResize)) { - onResize(width, height); - } - return { width: width, height: height }; - }); - }; -};var ResizeDetector = /** @class */ (function (_super) { - __extends(ResizeDetector, _super); - function ResizeDetector(props) { - var _this = _super.call(this, props) || this; - _this.cancelHandler = function () { - if (_this.resizeHandler && _this.resizeHandler.cancel) { - // cancel debounced handler - _this.resizeHandler.cancel(); - _this.resizeHandler = null; - } - }; - _this.attachObserver = function () { - var _a = _this.props, targetRef = _a.targetRef, observerOptions = _a.observerOptions; - if (targetRef && targetRef.current) { - _this.targetRef.current = targetRef.current; - } - var element = _this.getElement(); - if (!element) { - // can't find element to observe - return; - } - if (_this.observableElement && _this.observableElement === element) { - // element is already observed - return; - } - _this.observableElement = element; - _this.resizeObserver.observe(element, observerOptions); - }; - _this.getElement = function () { - var _a = _this.props, querySelector = _a.querySelector, targetDomEl = _a.targetDomEl; - // in case we pass a querySelector - if (querySelector) - return document.querySelector(querySelector); - // in case we pass a DOM element - if (targetDomEl && isDOMElement(targetDomEl)) - return targetDomEl; - // in case we pass a React ref using React.createRef() - if (_this.targetRef && isDOMElement(_this.targetRef.current)) - return _this.targetRef.current; - // the worse case when we don't receive any information from the parent and the library doesn't add any wrappers - // we have to use a deprecated `findDOMNode` method in order to find a DOM element to attach to - var currentElement = findDOMNode(_this); - if (!currentElement) - return null; - var renderType = _this.getRenderType(); - switch (renderType) { - case 'renderProp': - return currentElement; - case 'childFunction': - return currentElement; - case 'child': - return currentElement; - case 'childArray': - return currentElement; - default: - return currentElement.parentElement; - } - }; - _this.createResizeHandler = function (entries) { - var _a = _this.props, _b = _a.handleWidth, handleWidth = _b === void 0 ? true : _b, _c = _a.handleHeight, handleHeight = _c === void 0 ? true : _c, onResize = _a.onResize; - if (!handleWidth && !handleHeight) - return; - var notifyResize = createNotifier(onResize, _this.setState.bind(_this), handleWidth, handleHeight); - entries.forEach(function (entry) { - var _a = (entry && entry.contentRect) || {}, width = _a.width, height = _a.height; - var shouldSetSize = !_this.skipOnMount && !isSSR(); - if (shouldSetSize) { - notifyResize({ width: width, height: height }); - } - _this.skipOnMount = false; - }); - }; - _this.getRenderType = function () { - var _a = _this.props, render = _a.render, children = _a.children; - if (isFunction(render)) { - // DEPRECATED. Use `Child Function Pattern` instead - return 'renderProp'; - } - if (isFunction(children)) { - return 'childFunction'; - } - if (isValidElement(children)) { - return 'child'; - } - if (Array.isArray(children)) { - // DEPRECATED. Wrap children with a single parent - return 'childArray'; - } - // DEPRECATED. Use `Child Function Pattern` instead - return 'parent'; - }; - var skipOnMount = props.skipOnMount, refreshMode = props.refreshMode, _a = props.refreshRate, refreshRate = _a === void 0 ? 1000 : _a, refreshOptions = props.refreshOptions; - _this.state = { - width: undefined, - height: undefined - }; - _this.skipOnMount = skipOnMount; - _this.targetRef = createRef(); - _this.observableElement = null; - _this.resizeHandler = patchResizeHandler(_this.createResizeHandler, refreshMode, refreshRate, refreshOptions); - _this.resizeObserver = new window.ResizeObserver(_this.resizeHandler); - return _this; - } - ResizeDetector.prototype.componentDidMount = function () { - this.attachObserver(); - }; - ResizeDetector.prototype.componentDidUpdate = function () { - this.attachObserver(); - }; - ResizeDetector.prototype.componentWillUnmount = function () { - this.observableElement = null; - this.resizeObserver.disconnect(); - this.cancelHandler(); - }; - ResizeDetector.prototype.render = function () { - var _a = this.props, render = _a.render, children = _a.children, _b = _a.nodeType, WrapperTag = _b === void 0 ? 'div' : _b; - var _c = this.state, width = _c.width, height = _c.height; - var childProps = { width: width, height: height, targetRef: this.targetRef }; - var renderType = this.getRenderType(); - var typedChildren; - switch (renderType) { - case 'renderProp': - return render && render(childProps); - case 'childFunction': - typedChildren = children; - return typedChildren(childProps); - case 'child': - // @TODO bug prone logic - typedChildren = children; - if (typedChildren.type && typeof typedChildren.type === 'string') { - // child is a native DOM elements such as div, span etc - childProps.targetRef; var nativeProps = __rest(childProps, ["targetRef"]); - return cloneElement(typedChildren, nativeProps); - } - // class or functional component otherwise - return cloneElement(typedChildren, childProps); - case 'childArray': - typedChildren = children; - return typedChildren.map(function (el) { return !!el && cloneElement(el, childProps); }); - default: - return React.createElement(WrapperTag, null); - } - }; - return ResizeDetector; -}(PureComponent)); - -/* eslint no-console: 0 */ -var warn = function warn(condition, format) { - for (var _len = arguments.length, args = new Array(_len > 2 ? _len - 2 : 0), _key = 2; _key < _len; _key++) { - args[_key - 2] = arguments[_key]; - } -}; - -function _extends$u() { _extends$u = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$u.apply(this, arguments); } - -function _slicedToArray$7(arr, i) { return _arrayWithHoles$7(arr) || _iterableToArrayLimit$7(arr, i) || _unsupportedIterableToArray$e(arr, i) || _nonIterableRest$7(); } - -function _nonIterableRest$7() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$e(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$e(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$e(o, minLen); } - -function _arrayLikeToArray$e(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit$7(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$7(arr) { if (Array.isArray(arr)) return arr; } -var ResponsiveContainer = /*#__PURE__*/forwardRef(function (_ref, ref) { - var aspect = _ref.aspect, - _ref$width = _ref.width, - width = _ref$width === void 0 ? '100%' : _ref$width, - _ref$height = _ref.height, - height = _ref$height === void 0 ? '100%' : _ref$height, - minWidth = _ref.minWidth, - minHeight = _ref.minHeight, - maxHeight = _ref.maxHeight, - children = _ref.children, - _ref$debounce = _ref.debounce, - debounce = _ref$debounce === void 0 ? 0 : _ref$debounce, - id = _ref.id, - className = _ref.className; - - var _useState = useState({ - containerWidth: -1, - containerHeight: -1 - }), - _useState2 = _slicedToArray$7(_useState, 2), - sizes = _useState2[0], - setSizes = _useState2[1]; - - var containerRef = useRef(null); - useImperativeHandle(ref, function () { - return containerRef; - }, [containerRef]); - - var _useState3 = useState(false), - _useState4 = _slicedToArray$7(_useState3, 2), - mounted = _useState4[0], - setMounted = _useState4[1]; - - var getContainerSize = function getContainerSize() { - if (!containerRef.current) { - return null; - } - - return { - containerWidth: containerRef.current.clientWidth, - containerHeight: containerRef.current.clientHeight - }; - }; - - var updateDimensionsImmediate = function updateDimensionsImmediate() { - if (!mounted) { - return; - } - - var newSize = getContainerSize(); - - if (newSize) { - var oldWidth = sizes.containerWidth, - oldHeight = sizes.containerHeight; - var containerWidth = newSize.containerWidth, - containerHeight = newSize.containerHeight; - - if (containerWidth !== oldWidth || containerHeight !== oldHeight) { - setSizes({ - containerWidth: containerWidth, - containerHeight: containerHeight - }); - } - } - }; - - var handleResize = debounce > 0 ? debounce_1$1(updateDimensionsImmediate, debounce) : updateDimensionsImmediate; - - var renderChart = function renderChart() { - var containerWidth = sizes.containerWidth, - containerHeight = sizes.containerHeight; - - if (containerWidth < 0 || containerHeight < 0) { - return null; - } - - warn(isPercent(width) || isPercent(height), "The width(%s) and height(%s) are both fixed numbers,\n maybe you don't need to use a ResponsiveContainer.", width, height); - warn(!aspect || aspect > 0, 'The aspect(%s) must be greater than zero.', aspect); - var calculatedWidth = isPercent(width) ? containerWidth : width; - var calculatedHeight = isPercent(height) ? containerHeight : height; - - if (aspect && aspect > 0) { - // Preserve the desired aspect ratio - if (calculatedWidth) { - // Will default to using width for aspect ratio - calculatedHeight = calculatedWidth / aspect; - } else if (calculatedHeight) { - // But we should also take height into consideration - calculatedWidth = calculatedHeight * aspect; - } // if maxHeight is set, overwrite if calculatedHeight is greater than maxHeight - - - if (maxHeight && calculatedHeight > maxHeight) { - calculatedHeight = maxHeight; - } - } - - warn(calculatedWidth > 0 || calculatedHeight > 0, "The width(%s) and height(%s) of chart should be greater than 0,\n please check the style of container, or the props width(%s) and height(%s),\n or add a minWidth(%s) or minHeight(%s) or use aspect(%s) to control the\n height and width.", calculatedWidth, calculatedHeight, width, height, minWidth, minHeight, aspect); - return /*#__PURE__*/cloneElement(children, { - width: calculatedWidth, - height: calculatedHeight - }); - }; - - useEffect(function () { - if (mounted) { - var size = getContainerSize(); - - if (size) { - setSizes(size); - } - } - }, [mounted]); - useEffect(function () { - setMounted(true); - }, []); - var style = { - width: width, - height: height, - minWidth: minWidth, - minHeight: minHeight, - maxHeight: maxHeight - }; - return /*#__PURE__*/React__default.createElement(ResizeDetector, { - handleWidth: true, - handleHeight: true, - onResize: handleResize, - targetRef: containerRef - }, /*#__PURE__*/React__default.createElement("div", _extends$u({}, id != null ? { - id: "".concat(id) - } : {}, { - className: classNames('recharts-responsive-container', className), - style: style, - ref: containerRef - }), renderChart())); -}); - -/** - * @fileOverview Cross - */ -var Cell = function Cell(_props) { - return null; -}; -Cell.displayName = 'Cell'; - -var dist = {exports: {}}; - -var openParentheses = "(".charCodeAt(0); -var closeParentheses = ")".charCodeAt(0); -var singleQuote = "'".charCodeAt(0); -var doubleQuote = '"'.charCodeAt(0); -var backslash = "\\".charCodeAt(0); -var slash = "/".charCodeAt(0); -var comma = ",".charCodeAt(0); -var colon = ":".charCodeAt(0); -var star = "*".charCodeAt(0); - -var parse$1 = function(input) { - var tokens = []; - var value = input; - - var next, quote, prev, token, escape, escapePos, whitespacePos; - var pos = 0; - var code = value.charCodeAt(pos); - var max = value.length; - var stack = [{ nodes: tokens }]; - var balanced = 0; - var parent; - - var name = ""; - var before = ""; - var after = ""; - - while (pos < max) { - // Whitespaces - if (code <= 32) { - next = pos; - do { - next += 1; - code = value.charCodeAt(next); - } while (code <= 32); - token = value.slice(pos, next); - - prev = tokens[tokens.length - 1]; - if (code === closeParentheses && balanced) { - after = token; - } else if (prev && prev.type === "div") { - prev.after = token; - } else if ( - code === comma || - code === colon || - (code === slash && value.charCodeAt(next + 1) !== star) - ) { - before = token; - } else { - tokens.push({ - type: "space", - sourceIndex: pos, - value: token - }); - } - - pos = next; - - // Quotes - } else if (code === singleQuote || code === doubleQuote) { - next = pos; - quote = code === singleQuote ? "'" : '"'; - token = { - type: "string", - sourceIndex: pos, - quote: quote - }; - do { - escape = false; - next = value.indexOf(quote, next + 1); - if (~next) { - escapePos = next; - while (value.charCodeAt(escapePos - 1) === backslash) { - escapePos -= 1; - escape = !escape; - } - } else { - value += quote; - next = value.length - 1; - token.unclosed = true; - } - } while (escape); - token.value = value.slice(pos + 1, next); - - tokens.push(token); - pos = next + 1; - code = value.charCodeAt(pos); - - // Comments - } else if (code === slash && value.charCodeAt(pos + 1) === star) { - token = { - type: "comment", - sourceIndex: pos - }; - - next = value.indexOf("*/", pos); - if (next === -1) { - token.unclosed = true; - next = value.length; - } - - token.value = value.slice(pos + 2, next); - tokens.push(token); - - pos = next + 2; - code = value.charCodeAt(pos); - - // Dividers - } else if (code === slash || code === comma || code === colon) { - token = value[pos]; - - tokens.push({ - type: "div", - sourceIndex: pos - before.length, - value: token, - before: before, - after: "" - }); - before = ""; - - pos += 1; - code = value.charCodeAt(pos); - - // Open parentheses - } else if (openParentheses === code) { - // Whitespaces after open parentheses - next = pos; - do { - next += 1; - code = value.charCodeAt(next); - } while (code <= 32); - token = { - type: "function", - sourceIndex: pos - name.length, - value: name, - before: value.slice(pos + 1, next) - }; - pos = next; - - if (name === "url" && code !== singleQuote && code !== doubleQuote) { - next -= 1; - do { - escape = false; - next = value.indexOf(")", next + 1); - if (~next) { - escapePos = next; - while (value.charCodeAt(escapePos - 1) === backslash) { - escapePos -= 1; - escape = !escape; - } - } else { - value += ")"; - next = value.length - 1; - token.unclosed = true; - } - } while (escape); - // Whitespaces before closed - whitespacePos = next; - do { - whitespacePos -= 1; - code = value.charCodeAt(whitespacePos); - } while (code <= 32); - if (pos !== whitespacePos + 1) { - token.nodes = [ - { - type: "word", - sourceIndex: pos, - value: value.slice(pos, whitespacePos + 1) - } - ]; - } else { - token.nodes = []; - } - if (token.unclosed && whitespacePos + 1 !== next) { - token.after = ""; - token.nodes.push({ - type: "space", - sourceIndex: whitespacePos + 1, - value: value.slice(whitespacePos + 1, next) - }); - } else { - token.after = value.slice(whitespacePos + 1, next); - } - pos = next + 1; - code = value.charCodeAt(pos); - tokens.push(token); - } else { - balanced += 1; - token.after = ""; - tokens.push(token); - stack.push(token); - tokens = token.nodes = []; - parent = token; - } - name = ""; - - // Close parentheses - } else if (closeParentheses === code && balanced) { - pos += 1; - code = value.charCodeAt(pos); - - parent.after = after; - after = ""; - balanced -= 1; - stack.pop(); - parent = stack[balanced]; - tokens = parent.nodes; - - // Words - } else { - next = pos; - do { - if (code === backslash) { - next += 1; - } - next += 1; - code = value.charCodeAt(next); - } while ( - next < max && - !( - code <= 32 || - code === singleQuote || - code === doubleQuote || - code === comma || - code === colon || - code === slash || - code === openParentheses || - (code === closeParentheses && balanced) - ) - ); - token = value.slice(pos, next); - - if (openParentheses === code) { - name = token; - } else { - tokens.push({ - type: "word", - sourceIndex: pos, - value: token - }); - } - - pos = next; - } - } - - for (pos = stack.length - 1; pos; pos -= 1) { - stack[pos].unclosed = true; - } - - return stack[0].nodes; -}; - -var walk$1 = function walk(nodes, cb, bubble) { - var i, max, node, result; - - for (i = 0, max = nodes.length; i < max; i += 1) { - node = nodes[i]; - if (!bubble) { - result = cb(node, i, nodes); - } - - if ( - result !== false && - node.type === "function" && - Array.isArray(node.nodes) - ) { - walk(node.nodes, cb, bubble); - } - - if (bubble) { - cb(node, i, nodes); - } - } -}; - -function stringifyNode(node, custom) { - var type = node.type; - var value = node.value; - var buf; - var customResult; - - if (custom && (customResult = custom(node)) !== undefined) { - return customResult; - } else if (type === "word" || type === "space") { - return value; - } else if (type === "string") { - buf = node.quote || ""; - return buf + value + (node.unclosed ? "" : buf); - } else if (type === "comment") { - return "/*" + value + (node.unclosed ? "" : "*/"); - } else if (type === "div") { - return (node.before || "") + value + (node.after || ""); - } else if (Array.isArray(node.nodes)) { - buf = stringify$1(node.nodes); - if (type !== "function") { - return buf; - } - return ( - value + - "(" + - (node.before || "") + - buf + - (node.after || "") + - (node.unclosed ? "" : ")") - ); - } - return value; -} - -function stringify$1(nodes, custom) { - var result, i; - - if (Array.isArray(nodes)) { - result = ""; - for (i = nodes.length - 1; ~i; i -= 1) { - result = stringifyNode(nodes[i], custom) + result; - } - return result; - } - return stringifyNode(nodes, custom); -} - -var stringify_1 = stringify$1; - -var unit$1; -var hasRequiredUnit; - -function requireUnit () { - if (hasRequiredUnit) return unit$1; - hasRequiredUnit = 1; - var minus = "-".charCodeAt(0); - var plus = "+".charCodeAt(0); - var dot = ".".charCodeAt(0); - var exp = "e".charCodeAt(0); - var EXP = "E".charCodeAt(0); - - unit$1 = function(value) { - var pos = 0; - var length = value.length; - var dotted = false; - var sciPos = -1; - var containsNumber = false; - var code; - - while (pos < length) { - code = value.charCodeAt(pos); - - if (code >= 48 && code <= 57) { - containsNumber = true; - } else if (code === exp || code === EXP) { - if (sciPos > -1) { - break; - } - sciPos = pos; - } else if (code === dot) { - if (dotted) { - break; - } - dotted = true; - } else if (code === plus || code === minus) { - if (pos !== 0) { - break; - } - } else { - break; - } - - pos += 1; - } - - if (sciPos + 1 === pos) pos--; - - return containsNumber - ? { - number: value.slice(0, pos), - unit: value.slice(pos) - } - : false; - }; - return unit$1; -} - -var parse = parse$1; -var walk = walk$1; -var stringify = stringify_1; - -function ValueParser(value) { - if (this instanceof ValueParser) { - this.nodes = parse(value); - return this; - } - return new ValueParser(value); -} - -ValueParser.prototype.toString = function() { - return Array.isArray(this.nodes) ? stringify(this.nodes) : ""; -}; - -ValueParser.prototype.walk = function(cb, bubble) { - walk(this.nodes, cb, bubble); - return this; -}; - -ValueParser.unit = requireUnit(); - -ValueParser.walk = walk; - -ValueParser.stringify = stringify; - -var lib$1 = ValueParser; - -function commonjsRequire(path) { - throw new Error('Could not dynamically require "' + path + '". Please configure the dynamicRequireTargets or/and ignoreDynamicRequires option of @rollup/plugin-commonjs appropriately for this require call to work.'); -} - -var parser = {}; - -(function (exports) { - /* parser generated by jison 0.6.1-215 */ - - /* - * Returns a Parser object of the following structure: - * - * Parser: { - * yy: {} The so-called "shared state" or rather the *source* of it; - * the real "shared state" `yy` passed around to - * the rule actions, etc. is a derivative/copy of this one, - * not a direct reference! - * } - * - * Parser.prototype: { - * yy: {}, - * EOF: 1, - * TERROR: 2, - * - * trace: function(errorMessage, ...), - * - * JisonParserError: function(msg, hash), - * - * quoteName: function(name), - * Helper function which can be overridden by user code later on: put suitable - * quotes around literal IDs in a description string. - * - * originalQuoteName: function(name), - * The basic quoteName handler provided by JISON. - * `cleanupAfterParse()` will clean up and reset `quoteName()` to reference this function - * at the end of the `parse()`. - * - * describeSymbol: function(symbol), - * Return a more-or-less human-readable description of the given symbol, when - * available, or the symbol itself, serving as its own 'description' for lack - * of something better to serve up. - * - * Return NULL when the symbol is unknown to the parser. - * - * symbols_: {associative list: name ==> number}, - * terminals_: {associative list: number ==> name}, - * nonterminals: {associative list: rule-name ==> {associative list: number ==> rule-alt}}, - * terminal_descriptions_: (if there are any) {associative list: number ==> description}, - * productions_: [...], - * - * performAction: function parser__performAction(yytext, yyleng, yylineno, yyloc, yystate, yysp, yyvstack, yylstack, yystack, yysstack), - * - * The function parameters and `this` have the following value/meaning: - * - `this` : reference to the `yyval` internal object, which has members (`$` and `_$`) - * to store/reference the rule value `$$` and location info `@$`. - * - * One important thing to note about `this` a.k.a. `yyval`: every *reduce* action gets - * to see the same object via the `this` reference, i.e. if you wish to carry custom - * data from one reduce action through to the next within a single parse run, then you - * may get nasty and use `yyval` a.k.a. `this` for storing you own semi-permanent data. - * - * `this.yy` is a direct reference to the `yy` shared state object. - * - * `%parse-param`-specified additional `parse()` arguments have been added to this `yy` - * object at `parse()` start and are therefore available to the action code via the - * same named `yy.xxxx` attributes (where `xxxx` represents a identifier name from - * the %parse-param` list. - * - * - `yytext` : reference to the lexer value which belongs to the last lexer token used - * to match this rule. This is *not* the look-ahead token, but the last token - * that's actually part of this rule. - * - * Formulated another way, `yytext` is the value of the token immediately preceeding - * the current look-ahead token. - * Caveats apply for rules which don't require look-ahead, such as epsilon rules. - * - * - `yyleng` : ditto as `yytext`, only now for the lexer.yyleng value. - * - * - `yylineno`: ditto as `yytext`, only now for the lexer.yylineno value. - * - * - `yyloc` : ditto as `yytext`, only now for the lexer.yylloc lexer token location info. - * - * WARNING: since jison 0.4.18-186 this entry may be NULL/UNDEFINED instead - * of an empty object when no suitable location info can be provided. - * - * - `yystate` : the current parser state number, used internally for dispatching and - * executing the action code chunk matching the rule currently being reduced. - * - * - `yysp` : the current state stack position (a.k.a. 'stack pointer') - * - * This one comes in handy when you are going to do advanced things to the parser - * stacks, all of which are accessible from your action code (see the next entries below). - * - * Also note that you can access this and other stack index values using the new double-hash - * syntax, i.e. `##$ === ##0 === yysp`, while `##1` is the stack index for all things - * related to the first rule term, just like you have `$1`, `@1` and `#1`. - * This is made available to write very advanced grammar action rules, e.g. when you want - * to investigate the parse state stack in your action code, which would, for example, - * be relevant when you wish to implement error diagnostics and reporting schemes similar - * to the work described here: - * - * + Pottier, F., 2016. Reachability and error diagnosis in LR(1) automata. - * In Journées Francophones des Languages Applicatifs. - * - * + Jeffery, C.L., 2003. Generating LR syntax error messages from examples. - * ACM Transactions on Programming Languages and Systems (TOPLAS), 25(5), pp.631–640. - * - * - `yyrulelength`: the current rule's term count, i.e. the number of entries occupied on the stack. - * - * This one comes in handy when you are going to do advanced things to the parser - * stacks, all of which are accessible from your action code (see the next entries below). - * - * - `yyvstack`: reference to the parser value stack. Also accessed via the `$1` etc. - * constructs. - * - * - `yylstack`: reference to the parser token location stack. Also accessed via - * the `@1` etc. constructs. - * - * WARNING: since jison 0.4.18-186 this array MAY contain slots which are - * UNDEFINED rather than an empty (location) object, when the lexer/parser - * action code did not provide a suitable location info object when such a - * slot was filled! - * - * - `yystack` : reference to the parser token id stack. Also accessed via the - * `#1` etc. constructs. - * - * Note: this is a bit of a **white lie** as we can statically decode any `#n` reference to - * its numeric token id value, hence that code wouldn't need the `yystack` but *you* might - * want access this array for your own purposes, such as error analysis as mentioned above! - * - * Note that this stack stores the current stack of *tokens*, that is the sequence of - * already parsed=reduced *nonterminals* (tokens representing rules) and *terminals* - * (lexer tokens *shifted* onto the stack until the rule they belong to is found and - * *reduced*. - * - * - `yysstack`: reference to the parser state stack. This one carries the internal parser - * *states* such as the one in `yystate`, which are used to represent - * the parser state machine in the *parse table*. *Very* *internal* stuff, - * what can I say? If you access this one, you're clearly doing wicked things - * - * - `...` : the extra arguments you specified in the `%parse-param` statement in your - * grammar definition file. - * - * table: [...], - * State transition table - * ---------------------- - * - * index levels are: - * - `state` --> hash table - * - `symbol` --> action (number or array) - * - * If the `action` is an array, these are the elements' meaning: - * - index [0]: 1 = shift, 2 = reduce, 3 = accept - * - index [1]: GOTO `state` - * - * If the `action` is a number, it is the GOTO `state` - * - * defaultActions: {...}, - * - * parseError: function(str, hash, ExceptionClass), - * yyError: function(str, ...), - * yyRecovering: function(), - * yyErrOk: function(), - * yyClearIn: function(), - * - * constructParseErrorInfo: function(error_message, exception_object, expected_token_set, is_recoverable), - * Helper function **which will be set up during the first invocation of the `parse()` method**. - * Produces a new errorInfo 'hash object' which can be passed into `parseError()`. - * See it's use in this parser kernel in many places; example usage: - * - * var infoObj = parser.constructParseErrorInfo('fail!', null, - * parser.collect_expected_token_set(state), true); - * var retVal = parser.parseError(infoObj.errStr, infoObj, parser.JisonParserError); - * - * originalParseError: function(str, hash, ExceptionClass), - * The basic `parseError` handler provided by JISON. - * `cleanupAfterParse()` will clean up and reset `parseError()` to reference this function - * at the end of the `parse()`. - * - * options: { ... parser %options ... }, - * - * parse: function(input[, args...]), - * Parse the given `input` and return the parsed value (or `true` when none was provided by - * the root action, in which case the parser is acting as a *matcher*). - * You MAY use the additional `args...` parameters as per `%parse-param` spec of this grammar: - * these extra `args...` are added verbatim to the `yy` object reference as member variables. - * - * WARNING: - * Parser's additional `args...` parameters (via `%parse-param`) MAY conflict with - * any attributes already added to `yy` by the jison run-time; - * when such a collision is detected an exception is thrown to prevent the generated run-time - * from silently accepting this confusing and potentially hazardous situation! - * - * The lexer MAY add its own set of additional parameters (via the `%parse-param` line in - * the lexer section of the grammar spec): these will be inserted in the `yy` shared state - * object and any collision with those will be reported by the lexer via a thrown exception. - * - * cleanupAfterParse: function(resultValue, invoke_post_methods, do_not_nuke_errorinfos), - * Helper function **which will be set up during the first invocation of the `parse()` method**. - * This helper API is invoked at the end of the `parse()` call, unless an exception was thrown - * and `%options no-try-catch` has been defined for this grammar: in that case this helper MAY - * be invoked by calling user code to ensure the `post_parse` callbacks are invoked and - * the internal parser gets properly garbage collected under these particular circumstances. - * - * yyMergeLocationInfo: function(first_index, last_index, first_yylloc, last_yylloc, dont_look_back), - * Helper function **which will be set up during the first invocation of the `parse()` method**. - * This helper API can be invoked to calculate a spanning `yylloc` location info object. - * - * Note: %epsilon rules MAY specify no `first_index` and `first_yylloc`, in which case - * this function will attempt to obtain a suitable location marker by inspecting the location stack - * backwards. - * - * For more info see the documentation comment further below, immediately above this function's - * implementation. - * - * lexer: { - * yy: {...}, A reference to the so-called "shared state" `yy` once - * received via a call to the `.setInput(input, yy)` lexer API. - * EOF: 1, - * ERROR: 2, - * JisonLexerError: function(msg, hash), - * parseError: function(str, hash, ExceptionClass), - * setInput: function(input, [yy]), - * input: function(), - * unput: function(str), - * more: function(), - * reject: function(), - * less: function(n), - * pastInput: function(n), - * upcomingInput: function(n), - * showPosition: function(), - * test_match: function(regex_match_array, rule_index, ...), - * next: function(...), - * lex: function(...), - * begin: function(condition), - * pushState: function(condition), - * popState: function(), - * topState: function(), - * _currentRules: function(), - * stateStackSize: function(), - * cleanupAfterLex: function() - * - * options: { ... lexer %options ... }, - * - * performAction: function(yy, yy_, $avoiding_name_collisions, YY_START, ...), - * rules: [...], - * conditions: {associative list: name ==> set}, - * } - * } - * - * - * token location info (@$, _$, etc.): { - * first_line: n, - * last_line: n, - * first_column: n, - * last_column: n, - * range: [start_number, end_number] - * (where the numbers are indexes into the input string, zero-based) - * } - * - * --- - * - * The `parseError` function receives a 'hash' object with these members for lexer and - * parser errors: - * - * { - * text: (matched text) - * token: (the produced terminal token, if any) - * token_id: (the produced terminal token numeric ID, if any) - * line: (yylineno) - * loc: (yylloc) - * } - * - * parser (grammar) errors will also provide these additional members: - * - * { - * expected: (array describing the set of expected tokens; - * may be UNDEFINED when we cannot easily produce such a set) - * state: (integer (or array when the table includes grammar collisions); - * represents the current internal state of the parser kernel. - * can, for example, be used to pass to the `collect_expected_token_set()` - * API to obtain the expected token set) - * action: (integer; represents the current internal action which will be executed) - * new_state: (integer; represents the next/planned internal state, once the current - * action has executed) - * recoverable: (boolean: TRUE when the parser MAY have an error recovery rule - * available for this particular error) - * state_stack: (array: the current parser LALR/LR internal state stack; this can be used, - * for instance, for advanced error analysis and reporting) - * value_stack: (array: the current parser LALR/LR internal `$$` value stack; this can be used, - * for instance, for advanced error analysis and reporting) - * location_stack: (array: the current parser LALR/LR internal location stack; this can be used, - * for instance, for advanced error analysis and reporting) - * yy: (object: the current parser internal "shared state" `yy` - * as is also available in the rule actions; this can be used, - * for instance, for advanced error analysis and reporting) - * lexer: (reference to the current lexer instance used by the parser) - * parser: (reference to the current parser instance) - * } - * - * while `this` will reference the current parser instance. - * - * When `parseError` is invoked by the lexer, `this` will still reference the related *parser* - * instance, while these additional `hash` fields will also be provided: - * - * { - * lexer: (reference to the current lexer instance which reported the error) - * } - * - * When `parseError` is invoked by the parser due to a **JavaScript exception** being fired - * from either the parser or lexer, `this` will still reference the related *parser* - * instance, while these additional `hash` fields will also be provided: - * - * { - * exception: (reference to the exception thrown) - * } - * - * Please do note that in the latter situation, the `expected` field will be omitted as - * this type of failure is assumed not to be due to *parse errors* but rather due to user - * action code in either parser or lexer failing unexpectedly. - * - * --- - * - * You can specify parser options by setting / modifying the `.yy` object of your Parser instance. - * These options are available: - * - * ### options which are global for all parser instances - * - * Parser.pre_parse: function(yy) - * optional: you can specify a pre_parse() function in the chunk following - * the grammar, i.e. after the last `%%`. - * Parser.post_parse: function(yy, retval, parseInfo) { return retval; } - * optional: you can specify a post_parse() function in the chunk following - * the grammar, i.e. after the last `%%`. When it does not return any value, - * the parser will return the original `retval`. - * - * ### options which can be set up per parser instance - * - * yy: { - * pre_parse: function(yy) - * optional: is invoked before the parse cycle starts (and before the first - * invocation of `lex()`) but immediately after the invocation of - * `parser.pre_parse()`). - * post_parse: function(yy, retval, parseInfo) { return retval; } - * optional: is invoked when the parse terminates due to success ('accept') - * or failure (even when exceptions are thrown). - * `retval` contains the return value to be produced by `Parser.parse()`; - * this function can override the return value by returning another. - * When it does not return any value, the parser will return the original - * `retval`. - * This function is invoked immediately before `parser.post_parse()`. - * - * parseError: function(str, hash, ExceptionClass) - * optional: overrides the default `parseError` function. - * quoteName: function(name), - * optional: overrides the default `quoteName` function. - * } - * - * parser.lexer.options: { - * pre_lex: function() - * optional: is invoked before the lexer is invoked to produce another token. - * `this` refers to the Lexer object. - * post_lex: function(token) { return token; } - * optional: is invoked when the lexer has produced a token `token`; - * this function can override the returned token value by returning another. - * When it does not return any (truthy) value, the lexer will return - * the original `token`. - * `this` refers to the Lexer object. - * - * ranges: boolean - * optional: `true` ==> token location info will include a .range[] member. - * flex: boolean - * optional: `true` ==> flex-like lexing behaviour where the rules are tested - * exhaustively to find the longest match. - * backtrack_lexer: boolean - * optional: `true` ==> lexer regexes are tested in order and for invoked; - * the lexer terminates the scan when a token is returned by the action code. - * xregexp: boolean - * optional: `true` ==> lexer rule regexes are "extended regex format" requiring the - * `XRegExp` library. When this `%option` has not been specified at compile time, all lexer - * rule regexes have been written as standard JavaScript RegExp expressions. - * } - */ - - - - var parser = (function () { - - - // See also: - // http://stackoverflow.com/questions/1382107/whats-a-good-way-to-extend-error-in-javascript/#35881508 - // but we keep the prototype.constructor and prototype.name assignment lines too for compatibility - // with userland code which might access the derived class in a 'classic' way. - function JisonParserError(msg, hash) { - Object.defineProperty(this, 'name', { - enumerable: false, - writable: false, - value: 'JisonParserError' - }); - - if (msg == null) msg = '???'; - - Object.defineProperty(this, 'message', { - enumerable: false, - writable: true, - value: msg - }); - - this.hash = hash; - - var stacktrace; - if (hash && hash.exception instanceof Error) { - var ex2 = hash.exception; - this.message = ex2.message || msg; - stacktrace = ex2.stack; - } - if (!stacktrace) { - if (Error.hasOwnProperty('captureStackTrace')) { // V8/Chrome engine - Error.captureStackTrace(this, this.constructor); - } else { - stacktrace = (new Error(msg)).stack; - } - } - if (stacktrace) { - Object.defineProperty(this, 'stack', { - enumerable: false, - writable: false, - value: stacktrace - }); - } - } - - if (typeof Object.setPrototypeOf === 'function') { - Object.setPrototypeOf(JisonParserError.prototype, Error.prototype); - } else { - JisonParserError.prototype = Object.create(Error.prototype); - } - JisonParserError.prototype.constructor = JisonParserError; - JisonParserError.prototype.name = 'JisonParserError'; - - - - - // helper: reconstruct the productions[] table - function bp(s) { - var rv = []; - var p = s.pop; - var r = s.rule; - for (var i = 0, l = p.length; i < l; i++) { - rv.push([ - p[i], - r[i] - ]); - } - return rv; - } - - - - // helper: reconstruct the defaultActions[] table - function bda(s) { - var rv = {}; - var d = s.idx; - var g = s.goto; - for (var i = 0, l = d.length; i < l; i++) { - var j = d[i]; - rv[j] = g[i]; - } - return rv; - } - - - - // helper: reconstruct the 'goto' table - function bt(s) { - var rv = []; - var d = s.len; - var y = s.symbol; - var t = s.type; - var a = s.state; - var m = s.mode; - var g = s.goto; - for (var i = 0, l = d.length; i < l; i++) { - var n = d[i]; - var q = {}; - for (var j = 0; j < n; j++) { - var z = y.shift(); - switch (t.shift()) { - case 2: - q[z] = [ - m.shift(), - g.shift() - ]; - break; - - case 0: - q[z] = a.shift(); - break; - - default: - // type === 1: accept - q[z] = [ - 3 - ]; - } - } - rv.push(q); - } - return rv; - } - - - - // helper: runlength encoding with increment step: code, length: step (default step = 0) - // `this` references an array - function s(c, l, a) { - a = a || 0; - for (var i = 0; i < l; i++) { - this.push(c); - c += a; - } - } - - // helper: duplicate sequence from *relative* offset and length. - // `this` references an array - function c(i, l) { - i = this.length - i; - for (l += i; i < l; i++) { - this.push(this[i]); - } - } - - // helper: unpack an array using helpers and data, all passed in an array argument 'a'. - function u(a) { - var rv = []; - for (var i = 0, l = a.length; i < l; i++) { - var e = a[i]; - // Is this entry a helper function? - if (typeof e === 'function') { - i++; - e.apply(rv, a[i]); - } else { - rv.push(e); - } - } - return rv; - } - - - var parser = { - // Code Generator Information Report - // --------------------------------- - // - // Options: - // - // default action mode: ............. ["classic","merge"] - // test-compile action mode: ........ "parser:*,lexer:*" - // try..catch: ...................... true - // default resolve on conflict: ..... true - // on-demand look-ahead: ............ false - // error recovery token skip maximum: 3 - // yyerror in parse actions is: ..... NOT recoverable, - // yyerror in lexer actions and other non-fatal lexer are: - // .................................. NOT recoverable, - // debug grammar/output: ............ false - // has partial LR conflict upgrade: true - // rudimentary token-stack support: false - // parser table compression mode: ... 2 - // export debug tables: ............. false - // export *all* tables: ............. false - // module type: ..................... commonjs - // parser engine type: .............. lalr - // output main() in the module: ..... true - // has user-specified main(): ....... false - // has user-specified require()/import modules for main(): - // .................................. false - // number of expected conflicts: .... 0 - // - // - // Parser Analysis flags: - // - // no significant actions (parser is a language matcher only): - // .................................. false - // uses yyleng: ..................... false - // uses yylineno: ................... false - // uses yytext: ..................... false - // uses yylloc: ..................... false - // uses ParseError API: ............. false - // uses YYERROR: .................... false - // uses YYRECOVERING: ............... false - // uses YYERROK: .................... false - // uses YYCLEARIN: .................. false - // tracks rule values: .............. true - // assigns rule values: ............. true - // uses location tracking: .......... false - // assigns location: ................ false - // uses yystack: .................... false - // uses yysstack: ................... false - // uses yysp: ....................... true - // uses yyrulelength: ............... false - // uses yyMergeLocationInfo API: .... false - // has error recovery: .............. false - // has error reporting: ............. false - // - // --------- END OF REPORT ----------- - - trace: function no_op_trace() { }, - JisonParserError: JisonParserError, - yy: {}, - options: { - type: "lalr", - hasPartialLrUpgradeOnConflict: true, - errorRecoveryTokenDiscardCount: 3 - }, - symbols_: { - "$accept": 0, - "$end": 1, - "ADD": 3, - "ANGLE": 16, - "CHS": 22, - "COMMA": 14, - "CSS_CPROP": 13, - "CSS_VAR": 12, - "DIV": 6, - "EMS": 20, - "EOF": 1, - "EXS": 21, - "FREQ": 18, - "LENGTH": 15, - "LPAREN": 7, - "MUL": 5, - "NESTED_CALC": 9, - "NUMBER": 11, - "PERCENTAGE": 28, - "PREFIX": 10, - "REMS": 23, - "RES": 19, - "RPAREN": 8, - "SUB": 4, - "TIME": 17, - "VHS": 24, - "VMAXS": 27, - "VMINS": 26, - "VWS": 25, - "css_value": 33, - "css_variable": 32, - "error": 2, - "expression": 29, - "math_expression": 30, - "value": 31 - }, - terminals_: { - 1: "EOF", - 2: "error", - 3: "ADD", - 4: "SUB", - 5: "MUL", - 6: "DIV", - 7: "LPAREN", - 8: "RPAREN", - 9: "NESTED_CALC", - 10: "PREFIX", - 11: "NUMBER", - 12: "CSS_VAR", - 13: "CSS_CPROP", - 14: "COMMA", - 15: "LENGTH", - 16: "ANGLE", - 17: "TIME", - 18: "FREQ", - 19: "RES", - 20: "EMS", - 21: "EXS", - 22: "CHS", - 23: "REMS", - 24: "VHS", - 25: "VWS", - 26: "VMINS", - 27: "VMAXS", - 28: "PERCENTAGE" - }, - TERROR: 2, - EOF: 1, - - // internals: defined here so the object *structure* doesn't get modified by parse() et al, - // thus helping JIT compilers like Chrome V8. - originalQuoteName: null, - originalParseError: null, - cleanupAfterParse: null, - constructParseErrorInfo: null, - yyMergeLocationInfo: null, - - __reentrant_call_depth: 0, // INTERNAL USE ONLY - __error_infos: [], // INTERNAL USE ONLY: the set of parseErrorInfo objects created since the last cleanup - __error_recovery_infos: [], // INTERNAL USE ONLY: the set of parseErrorInfo objects created since the last cleanup - - // APIs which will be set up depending on user action code analysis: - //yyRecovering: 0, - //yyErrOk: 0, - //yyClearIn: 0, - - // Helper APIs - // ----------- - - // Helper function which can be overridden by user code later on: put suitable quotes around - // literal IDs in a description string. - quoteName: function parser_quoteName(id_str) { - return '"' + id_str + '"'; - }, - - // Return the name of the given symbol (terminal or non-terminal) as a string, when available. - // - // Return NULL when the symbol is unknown to the parser. - getSymbolName: function parser_getSymbolName(symbol) { - if (this.terminals_[symbol]) { - return this.terminals_[symbol]; - } - - // Otherwise... this might refer to a RULE token i.e. a non-terminal: see if we can dig that one up. - // - // An example of this may be where a rule's action code contains a call like this: - // - // parser.getSymbolName(#$) - // - // to obtain a human-readable name of the current grammar rule. - var s = this.symbols_; - for (var key in s) { - if (s[key] === symbol) { - return key; - } - } - return null; - }, - - // Return a more-or-less human-readable description of the given symbol, when available, - // or the symbol itself, serving as its own 'description' for lack of something better to serve up. - // - // Return NULL when the symbol is unknown to the parser. - describeSymbol: function parser_describeSymbol(symbol) { - if (symbol !== this.EOF && this.terminal_descriptions_ && this.terminal_descriptions_[symbol]) { - return this.terminal_descriptions_[symbol]; - } - else if (symbol === this.EOF) { - return 'end of input'; - } - var id = this.getSymbolName(symbol); - if (id) { - return this.quoteName(id); - } - return null; - }, - - // Produce a (more or less) human-readable list of expected tokens at the point of failure. - // - // The produced list may contain token or token set descriptions instead of the tokens - // themselves to help turning this output into something that easier to read by humans - // unless `do_not_describe` parameter is set, in which case a list of the raw, *numeric*, - // expected terminals and nonterminals is produced. - // - // The returned list (array) will not contain any duplicate entries. - collect_expected_token_set: function parser_collect_expected_token_set(state, do_not_describe) { - var TERROR = this.TERROR; - var tokenset = []; - var check = {}; - // Has this (error?) state been outfitted with a custom expectations description text for human consumption? - // If so, use that one instead of the less palatable token set. - if (!do_not_describe && this.state_descriptions_ && this.state_descriptions_[state]) { - return [ - this.state_descriptions_[state] - ]; - } - for (var p in this.table[state]) { - p = +p; - if (p !== TERROR) { - var d = do_not_describe ? p : this.describeSymbol(p); - if (d && !check[d]) { - tokenset.push(d); - check[d] = true; // Mark this token description as already mentioned to prevent outputting duplicate entries. - } - } - } - return tokenset; - }, - productions_: bp({ - pop: u([ - 29, - s, - [30, 10], - 31, - 31, - 32, - 32, - s, - [33, 15] - ]), - rule: u([ - 2, - s, - [3, 5], - 4, - 7, - s, - [1, 4], - 2, - 4, - 6, - s, - [1, 14], - 2 - ]) - }), - performAction: function parser__PerformAction(yystate /* action[1] */, yysp, yyvstack) { - - /* this == yyval */ - - // the JS engine itself can go and remove these statements when `yy` turns out to be unused in any action code! - var yy = this.yy; - yy.parser; - yy.lexer; - - - - switch (yystate) { - case 0: - /*! Production:: $accept : expression $end */ - - // default action (generated by JISON mode classic/merge :: 1,VT,VA,-,-,-,-,-,-): - this.$ = yyvstack[yysp - 1]; - // END of default action (generated by JISON mode classic/merge :: 1,VT,VA,-,-,-,-,-,-) - break; - - case 1: - /*! Production:: expression : math_expression EOF */ - - // default action (generated by JISON mode classic/merge :: 2,VT,VA,-,-,-,-,-,-): - this.$ = yyvstack[yysp - 1]; - // END of default action (generated by JISON mode classic/merge :: 2,VT,VA,-,-,-,-,-,-) - - - return yyvstack[yysp - 1]; - - case 2: - /*! Production:: math_expression : math_expression ADD math_expression */ - case 3: - /*! Production:: math_expression : math_expression SUB math_expression */ - case 4: - /*! Production:: math_expression : math_expression MUL math_expression */ - case 5: - /*! Production:: math_expression : math_expression DIV math_expression */ - - this.$ = { type: 'MathExpression', operator: yyvstack[yysp - 1], left: yyvstack[yysp - 2], right: yyvstack[yysp] }; - break; - - case 6: - /*! Production:: math_expression : LPAREN math_expression RPAREN */ - - this.$ = yyvstack[yysp - 1]; - break; - - case 7: - /*! Production:: math_expression : NESTED_CALC LPAREN math_expression RPAREN */ - - this.$ = { type: 'Calc', value: yyvstack[yysp - 1] }; - break; - - case 8: - /*! Production:: math_expression : SUB PREFIX SUB NESTED_CALC LPAREN math_expression RPAREN */ - - this.$ = { type: 'Calc', value: yyvstack[yysp - 1], prefix: yyvstack[yysp - 5] }; - break; - - case 9: - /*! Production:: math_expression : css_variable */ - case 10: - /*! Production:: math_expression : css_value */ - case 11: - /*! Production:: math_expression : value */ - - this.$ = yyvstack[yysp]; - break; - - case 12: - /*! Production:: value : NUMBER */ - - this.$ = { type: 'Value', value: parseFloat(yyvstack[yysp]) }; - break; - - case 13: - /*! Production:: value : SUB NUMBER */ - - this.$ = { type: 'Value', value: parseFloat(yyvstack[yysp]) * -1 }; - break; - - case 14: - /*! Production:: css_variable : CSS_VAR LPAREN CSS_CPROP RPAREN */ - - this.$ = { type: 'CssVariable', value: yyvstack[yysp - 1] }; - break; - - case 15: - /*! Production:: css_variable : CSS_VAR LPAREN CSS_CPROP COMMA math_expression RPAREN */ - - this.$ = { type: 'CssVariable', value: yyvstack[yysp - 3], fallback: yyvstack[yysp - 1] }; - break; - - case 16: - /*! Production:: css_value : LENGTH */ - - this.$ = { type: 'LengthValue', value: parseFloat(yyvstack[yysp]), unit: /[a-z]+/.exec(yyvstack[yysp])[0] }; - break; - - case 17: - /*! Production:: css_value : ANGLE */ - - this.$ = { type: 'AngleValue', value: parseFloat(yyvstack[yysp]), unit: /[a-z]+/.exec(yyvstack[yysp])[0] }; - break; - - case 18: - /*! Production:: css_value : TIME */ - - this.$ = { type: 'TimeValue', value: parseFloat(yyvstack[yysp]), unit: /[a-z]+/.exec(yyvstack[yysp])[0] }; - break; - - case 19: - /*! Production:: css_value : FREQ */ - - this.$ = { type: 'FrequencyValue', value: parseFloat(yyvstack[yysp]), unit: /[a-z]+/.exec(yyvstack[yysp])[0] }; - break; - - case 20: - /*! Production:: css_value : RES */ - - this.$ = { type: 'ResolutionValue', value: parseFloat(yyvstack[yysp]), unit: /[a-z]+/.exec(yyvstack[yysp])[0] }; - break; - - case 21: - /*! Production:: css_value : EMS */ - - this.$ = { type: 'EmValue', value: parseFloat(yyvstack[yysp]), unit: 'em' }; - break; - - case 22: - /*! Production:: css_value : EXS */ - - this.$ = { type: 'ExValue', value: parseFloat(yyvstack[yysp]), unit: 'ex' }; - break; - - case 23: - /*! Production:: css_value : CHS */ - - this.$ = { type: 'ChValue', value: parseFloat(yyvstack[yysp]), unit: 'ch' }; - break; - - case 24: - /*! Production:: css_value : REMS */ - - this.$ = { type: 'RemValue', value: parseFloat(yyvstack[yysp]), unit: 'rem' }; - break; - - case 25: - /*! Production:: css_value : VHS */ - - this.$ = { type: 'VhValue', value: parseFloat(yyvstack[yysp]), unit: 'vh' }; - break; - - case 26: - /*! Production:: css_value : VWS */ - - this.$ = { type: 'VwValue', value: parseFloat(yyvstack[yysp]), unit: 'vw' }; - break; - - case 27: - /*! Production:: css_value : VMINS */ - - this.$ = { type: 'VminValue', value: parseFloat(yyvstack[yysp]), unit: 'vmin' }; - break; - - case 28: - /*! Production:: css_value : VMAXS */ - - this.$ = { type: 'VmaxValue', value: parseFloat(yyvstack[yysp]), unit: 'vmax' }; - break; - - case 29: - /*! Production:: css_value : PERCENTAGE */ - - this.$ = { type: 'PercentageValue', value: parseFloat(yyvstack[yysp]), unit: '%' }; - break; - - case 30: - /*! Production:: css_value : SUB css_value */ - - var prev = yyvstack[yysp]; prev.value *= -1; this.$ = prev; - break; - - } - }, - table: bt({ - len: u([ - 24, - 1, - 5, - 23, - 1, - 18, - s, - [0, 3], - 1, - s, - [0, 16], - s, - [23, 4], - c, - [28, 3], - 0, - 0, - 16, - 1, - 6, - 6, - s, - [0, 3], - 5, - 1, - 2, - c, - [37, 3], - c, - [20, 3], - 5, - 0, - 0 - ]), - symbol: u([ - 4, - 7, - 9, - 11, - 12, - s, - [15, 19, 1], - 1, - 1, - s, - [3, 4, 1], - c, - [30, 19], - c, - [29, 4], - 7, - 4, - 10, - 11, - c, - [22, 14], - c, - [19, 3], - c, - [43, 22], - c, - [23, 69], - c, - [139, 4], - 8, - c, - [51, 24], - 4, - c, - [138, 15], - 13, - c, - [186, 5], - 8, - c, - [6, 6], - c, - [5, 5], - 9, - 8, - 14, - c, - [159, 47], - c, - [60, 10] - ]), - type: u([ - s, - [2, 19], - s, - [0, 5], - 1, - s, - [2, 24], - s, - [0, 4], - c, - [22, 19], - c, - [43, 42], - c, - [23, 70], - c, - [28, 25], - c, - [45, 25], - c, - [113, 54] - ]), - state: u([ - 1, - 2, - 8, - 6, - 7, - 30, - c, - [4, 3], - 33, - 37, - c, - [5, 3], - 38, - c, - [4, 3], - 39, - c, - [4, 3], - 40, - c, - [4, 3], - 42, - c, - [21, 4], - 50, - c, - [5, 3], - 51, - c, - [4, 3] - ]), - mode: u([ - s, - [1, 179], - s, - [2, 3], - c, - [5, 5], - c, - [6, 4], - s, - [1, 57] - ]), - goto: u([ - 5, - 3, - 4, - 24, - s, - [9, 15, 1], - s, - [25, 5, 1], - c, - [24, 19], - 31, - 35, - 32, - 34, - c, - [18, 14], - 36, - c, - [38, 19], - c, - [19, 57], - c, - [118, 4], - 41, - c, - [24, 19], - 43, - 35, - c, - [16, 14], - 44, - s, - [2, 3], - 28, - 29, - 2, - s, - [3, 3], - 28, - 29, - 3, - c, - [53, 4], - s, - [45, 5, 1], - c, - [100, 42], - 52, - c, - [5, 4], - 53 - ]) - }), - defaultActions: bda({ - idx: u([ - 6, - 7, - 8, - s, - [10, 16, 1], - 33, - 34, - 39, - 40, - 41, - 45, - 47, - 52, - 53 - ]), - goto: u([ - 9, - 10, - 11, - s, - [16, 14, 1], - 12, - 1, - 30, - 13, - s, - [4, 4, 1], - 14, - 15, - 8 - ]) - }), - parseError: function parseError(str, hash, ExceptionClass) { - if (hash.recoverable) { - if (typeof this.trace === 'function') { - this.trace(str); - } - hash.destroy(); // destroy... well, *almost*! - } else { - if (typeof this.trace === 'function') { - this.trace(str); - } - if (!ExceptionClass) { - ExceptionClass = this.JisonParserError; - } - throw new ExceptionClass(str, hash); - } - }, - parse: function parse(input) { - var self = this; - var stack = new Array(128); // token stack: stores token which leads to state at the same index (column storage) - var sstack = new Array(128); // state stack: stores states (column storage) - - var vstack = new Array(128); // semantic value stack - - var table = this.table; - var sp = 0; // 'stack pointer': index into the stacks - - - - - - var symbol = 0; - - - - this.TERROR; - var EOF = this.EOF; - (this.options.errorRecoveryTokenDiscardCount | 0) || 3; - var NO_ACTION = [0, 54 /* === table.length :: ensures that anyone using this new state will fail dramatically! */]; - - var lexer; - if (this.__lexer__) { - lexer = this.__lexer__; - } else { - lexer = this.__lexer__ = Object.create(this.lexer); - } - - var sharedState_yy = { - parseError: undefined, - quoteName: undefined, - lexer: undefined, - parser: undefined, - pre_parse: undefined, - post_parse: undefined, - pre_lex: undefined, - post_lex: undefined // WARNING: must be written this way for the code expanders to work correctly in both ES5 and ES6 modes! - }; - if (typeof assert !== 'function') ; else { - assert; - } - - this.yyGetSharedState = function yyGetSharedState() { - return sharedState_yy; - }; - - - - - - - - - function shallow_copy_noclobber(dst, src) { - for (var k in src) { - if (typeof dst[k] === 'undefined' && Object.prototype.hasOwnProperty.call(src, k)) { - dst[k] = src[k]; - } - } - } - - // copy state - shallow_copy_noclobber(sharedState_yy, this.yy); - - sharedState_yy.lexer = lexer; - sharedState_yy.parser = this; - - - - - - - // Does the shared state override the default `parseError` that already comes with this instance? - if (typeof sharedState_yy.parseError === 'function') { - this.parseError = function parseErrorAlt(str, hash, ExceptionClass) { - if (!ExceptionClass) { - ExceptionClass = this.JisonParserError; - } - return sharedState_yy.parseError.call(this, str, hash, ExceptionClass); - }; - } else { - this.parseError = this.originalParseError; - } - - // Does the shared state override the default `quoteName` that already comes with this instance? - if (typeof sharedState_yy.quoteName === 'function') { - this.quoteName = function quoteNameAlt(id_str) { - return sharedState_yy.quoteName.call(this, id_str); - }; - } else { - this.quoteName = this.originalQuoteName; - } - - // set up the cleanup function; make it an API so that external code can re-use this one in case of - // calamities or when the `%options no-try-catch` option has been specified for the grammar, in which - // case this parse() API method doesn't come with a `finally { ... }` block any more! - // - // NOTE: as this API uses parse() as a closure, it MUST be set again on every parse() invocation, - // or else your `sharedState`, etc. references will be *wrong*! - this.cleanupAfterParse = function parser_cleanupAfterParse(resultValue, invoke_post_methods, do_not_nuke_errorinfos) { - var rv; - - if (invoke_post_methods) { - var hash; - - if (sharedState_yy.post_parse || this.post_parse) { - // create an error hash info instance: we re-use this API in a **non-error situation** - // as this one delivers all parser internals ready for access by userland code. - hash = this.constructParseErrorInfo(null /* no error! */, null /* no exception! */, null, false); - } - - if (sharedState_yy.post_parse) { - rv = sharedState_yy.post_parse.call(this, sharedState_yy, resultValue, hash); - if (typeof rv !== 'undefined') resultValue = rv; - } - if (this.post_parse) { - rv = this.post_parse.call(this, sharedState_yy, resultValue, hash); - if (typeof rv !== 'undefined') resultValue = rv; - } - - // cleanup: - if (hash && hash.destroy) { - hash.destroy(); - } - } - - if (this.__reentrant_call_depth > 1) return resultValue; // do not (yet) kill the sharedState when this is a reentrant run. - - // clean up the lingering lexer structures as well: - if (lexer.cleanupAfterLex) { - lexer.cleanupAfterLex(do_not_nuke_errorinfos); - } - - // prevent lingering circular references from causing memory leaks: - if (sharedState_yy) { - sharedState_yy.lexer = undefined; - sharedState_yy.parser = undefined; - if (lexer.yy === sharedState_yy) { - lexer.yy = undefined; - } - } - sharedState_yy = undefined; - this.parseError = this.originalParseError; - this.quoteName = this.originalQuoteName; - - // nuke the vstack[] array at least as that one will still reference obsoleted user values. - // To be safe, we nuke the other internal stack columns as well... - stack.length = 0; // fastest way to nuke an array without overly bothering the GC - sstack.length = 0; - - vstack.length = 0; - sp = 0; - - // nuke the error hash info instances created during this run. - // Userland code must COPY any data/references - // in the error hash instance(s) it is more permanently interested in. - if (!do_not_nuke_errorinfos) { - for (var i = this.__error_infos.length - 1; i >= 0; i--) { - var el = this.__error_infos[i]; - if (el && typeof el.destroy === 'function') { - el.destroy(); - } - } - this.__error_infos.length = 0; - - - } - - return resultValue; - }; - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - // NOTE: as this API uses parse() as a closure, it MUST be set again on every parse() invocation, - // or else your `lexer`, `sharedState`, etc. references will be *wrong*! - this.constructParseErrorInfo = function parser_constructParseErrorInfo(msg, ex, expected, recoverable) { - var pei = { - errStr: msg, - exception: ex, - text: lexer.match, - value: lexer.yytext, - token: this.describeSymbol(symbol) || symbol, - token_id: symbol, - line: lexer.yylineno, - - expected: expected, - recoverable: recoverable, - state: state, - action: action, - new_state: newState, - symbol_stack: stack, - state_stack: sstack, - value_stack: vstack, - - stack_pointer: sp, - yy: sharedState_yy, - lexer: lexer, - parser: this, - - // and make sure the error info doesn't stay due to potential - // ref cycle via userland code manipulations. - // These would otherwise all be memory leak opportunities! - // - // Note that only array and object references are nuked as those - // constitute the set of elements which can produce a cyclic ref. - // The rest of the members is kept intact as they are harmless. - destroy: function destructParseErrorInfo() { - // remove cyclic references added to error info: - // info.yy = null; - // info.lexer = null; - // info.value = null; - // info.value_stack = null; - // ... - var rec = !!this.recoverable; - for (var key in this) { - if (this.hasOwnProperty(key) && typeof key === 'object') { - this[key] = undefined; - } - } - this.recoverable = rec; - } - }; - // track this instance so we can `destroy()` it once we deem it superfluous and ready for garbage collection! - this.__error_infos.push(pei); - return pei; - }; - - - function stdLex() { - var token = lexer.lex(); - // if token isn't its numeric value, convert - if (typeof token !== 'number') { - token = self.symbols_[token] || token; - } - - return token || EOF; - } - - function fastLex() { - var token = lexer.fastLex(); - // if token isn't its numeric value, convert - if (typeof token !== 'number') { - token = self.symbols_[token] || token; - } - - return token || EOF; - } - - var lex = stdLex; - - - var state, action, r, t; - var yyval = { - $: true, - _$: undefined, - yy: sharedState_yy - }; - var p; - var yyrulelen; - var this_production; - var newState; - var retval = false; - - - try { - this.__reentrant_call_depth++; - - lexer.setInput(input, sharedState_yy); - - // NOTE: we *assume* no lexer pre/post handlers are set up *after* - // this initial `setInput()` call: hence we can now check and decide - // whether we'll go with the standard, slower, lex() API or the - // `fast_lex()` one: - if (typeof lexer.canIUse === 'function') { - var lexerInfo = lexer.canIUse(); - if (lexerInfo.fastLex && typeof fastLex === 'function') { - lex = fastLex; - } - } - - - - vstack[sp] = null; - sstack[sp] = 0; - stack[sp] = 0; - ++sp; - - - - - - if (this.pre_parse) { - this.pre_parse.call(this, sharedState_yy); - } - if (sharedState_yy.pre_parse) { - sharedState_yy.pre_parse.call(this, sharedState_yy); - } - - newState = sstack[sp - 1]; - for (;;) { - // retrieve state number from top of stack - state = newState; // sstack[sp - 1]; - - // use default actions if available - if (this.defaultActions[state]) { - action = 2; - newState = this.defaultActions[state]; - } else { - // The single `==` condition below covers both these `===` comparisons in a single - // operation: - // - // if (symbol === null || typeof symbol === 'undefined') ... - if (!symbol) { - symbol = lex(); - } - // read action for current state and first input - t = (table[state] && table[state][symbol]) || NO_ACTION; - newState = t[1]; - action = t[0]; - - - - - - - - - - - - // handle parse error - if (!action) { - var errStr; - var errSymbolDescr = (this.describeSymbol(symbol) || symbol); - var expected = this.collect_expected_token_set(state); - - // Report error - if (typeof lexer.yylineno === 'number') { - errStr = 'Parse error on line ' + (lexer.yylineno + 1) + ': '; - } else { - errStr = 'Parse error: '; - } - if (typeof lexer.showPosition === 'function') { - errStr += '\n' + lexer.showPosition(79 - 10, 10) + '\n'; - } - if (expected.length) { - errStr += 'Expecting ' + expected.join(', ') + ', got unexpected ' + errSymbolDescr; - } else { - errStr += 'Unexpected ' + errSymbolDescr; - } - // we cannot recover from the error! - p = this.constructParseErrorInfo(errStr, null, expected, false); - r = this.parseError(p.errStr, p, this.JisonParserError); - if (typeof r !== 'undefined') { - retval = r; - } - break; - } - - - } - - - - - - - - - - - switch (action) { - // catch misc. parse failures: - default: - // this shouldn't happen, unless resolve defaults are off - if (action instanceof Array) { - p = this.constructParseErrorInfo('Parse Error: multiple actions possible at state: ' + state + ', token: ' + symbol, null, null, false); - r = this.parseError(p.errStr, p, this.JisonParserError); - if (typeof r !== 'undefined') { - retval = r; - } - break; - } - // Another case of better safe than sorry: in case state transitions come out of another error recovery process - // or a buggy LUT (LookUp Table): - p = this.constructParseErrorInfo('Parsing halted. No viable error recovery approach available due to internal system failure.', null, null, false); - r = this.parseError(p.errStr, p, this.JisonParserError); - if (typeof r !== 'undefined') { - retval = r; - } - break; - - // shift: - case 1: - stack[sp] = symbol; - vstack[sp] = lexer.yytext; - - sstack[sp] = newState; // push state - - ++sp; - symbol = 0; - - - - - // Pick up the lexer details for the current symbol as that one is not 'look-ahead' any more: - - - - - continue; - - // reduce: - case 2: - - - - this_production = this.productions_[newState - 1]; // `this.productions_[]` is zero-based indexed while states start from 1 upwards... - yyrulelen = this_production[1]; - - - - - - - - - - - r = this.performAction.call(yyval, newState, sp - 1, vstack); - - if (typeof r !== 'undefined') { - retval = r; - break; - } - - // pop off stack - sp -= yyrulelen; - - // don't overwrite the `symbol` variable: use a local var to speed things up: - var ntsymbol = this_production[0]; // push nonterminal (reduce) - stack[sp] = ntsymbol; - vstack[sp] = yyval.$; - - // goto new state = table[STATE][NONTERMINAL] - newState = table[sstack[sp - 1]][ntsymbol]; - sstack[sp] = newState; - ++sp; - - - - - - - - - - continue; - - // accept: - case 3: - if (sp !== -2) { - retval = true; - // Return the `$accept` rule's `$$` result, if available. - // - // Also note that JISON always adds this top-most `$accept` rule (with implicit, - // default, action): - // - // $accept: $end - // %{ $$ = $1; @$ = @1; %} - // - // which, combined with the parse kernel's `$accept` state behaviour coded below, - // will produce the `$$` value output of the rule as the parse result, - // IFF that result is *not* `undefined`. (See also the parser kernel code.) - // - // In code: - // - // %{ - // @$ = @1; // if location tracking support is included - // if (typeof $1 !== 'undefined') - // return $1; - // else - // return true; // the default parse result if the rule actions don't produce anything - // %} - sp--; - if (typeof vstack[sp] !== 'undefined') { - retval = vstack[sp]; - } - } - break; - } - - // break out of loop: we accept or fail with error - break; - } - } catch (ex) { - // report exceptions through the parseError callback too, but keep the exception intact - // if it is a known parser or lexer error which has been thrown by parseError() already: - if (ex instanceof this.JisonParserError) { - throw ex; - } - else if (lexer && typeof lexer.JisonLexerError === 'function' && ex instanceof lexer.JisonLexerError) { - throw ex; - } - - p = this.constructParseErrorInfo('Parsing aborted due to exception.', ex, null, false); - retval = false; - r = this.parseError(p.errStr, p, this.JisonParserError); - if (typeof r !== 'undefined') { - retval = r; - } - } finally { - retval = this.cleanupAfterParse(retval, true, true); - this.__reentrant_call_depth--; - } // /finally - - return retval; - } - }; - parser.originalParseError = parser.parseError; - parser.originalQuoteName = parser.quoteName; - /* lexer generated by jison-lex 0.6.1-215 */ - - /* - * Returns a Lexer object of the following structure: - * - * Lexer: { - * yy: {} The so-called "shared state" or rather the *source* of it; - * the real "shared state" `yy` passed around to - * the rule actions, etc. is a direct reference! - * - * This "shared context" object was passed to the lexer by way of - * the `lexer.setInput(str, yy)` API before you may use it. - * - * This "shared context" object is passed to the lexer action code in `performAction()` - * so userland code in the lexer actions may communicate with the outside world - * and/or other lexer rules' actions in more or less complex ways. - * - * } - * - * Lexer.prototype: { - * EOF: 1, - * ERROR: 2, - * - * yy: The overall "shared context" object reference. - * - * JisonLexerError: function(msg, hash), - * - * performAction: function lexer__performAction(yy, yyrulenumber, YY_START), - * - * The function parameters and `this` have the following value/meaning: - * - `this` : reference to the `lexer` instance. - * `yy_` is an alias for `this` lexer instance reference used internally. - * - * - `yy` : a reference to the `yy` "shared state" object which was passed to the lexer - * by way of the `lexer.setInput(str, yy)` API before. - * - * Note: - * The extra arguments you specified in the `%parse-param` statement in your - * **parser** grammar definition file are passed to the lexer via this object - * reference as member variables. - * - * - `yyrulenumber` : index of the matched lexer rule (regex), used internally. - * - * - `YY_START`: the current lexer "start condition" state. - * - * parseError: function(str, hash, ExceptionClass), - * - * constructLexErrorInfo: function(error_message, is_recoverable), - * Helper function. - * Produces a new errorInfo 'hash object' which can be passed into `parseError()`. - * See it's use in this lexer kernel in many places; example usage: - * - * var infoObj = lexer.constructParseErrorInfo('fail!', true); - * var retVal = lexer.parseError(infoObj.errStr, infoObj, lexer.JisonLexerError); - * - * options: { ... lexer %options ... }, - * - * lex: function(), - * Produce one token of lexed input, which was passed in earlier via the `lexer.setInput()` API. - * You MAY use the additional `args...` parameters as per `%parse-param` spec of the **lexer** grammar: - * these extra `args...` are added verbatim to the `yy` object reference as member variables. - * - * WARNING: - * Lexer's additional `args...` parameters (via lexer's `%parse-param`) MAY conflict with - * any attributes already added to `yy` by the **parser** or the jison run-time; - * when such a collision is detected an exception is thrown to prevent the generated run-time - * from silently accepting this confusing and potentially hazardous situation! - * - * cleanupAfterLex: function(do_not_nuke_errorinfos), - * Helper function. - * - * This helper API is invoked when the **parse process** has completed: it is the responsibility - * of the **parser** (or the calling userland code) to invoke this method once cleanup is desired. - * - * This helper may be invoked by user code to ensure the internal lexer gets properly garbage collected. - * - * setInput: function(input, [yy]), - * - * - * input: function(), - * - * - * unput: function(str), - * - * - * more: function(), - * - * - * reject: function(), - * - * - * less: function(n), - * - * - * pastInput: function(n), - * - * - * upcomingInput: function(n), - * - * - * showPosition: function(), - * - * - * test_match: function(regex_match_array, rule_index), - * - * - * next: function(), - * - * - * begin: function(condition), - * - * - * pushState: function(condition), - * - * - * popState: function(), - * - * - * topState: function(), - * - * - * _currentRules: function(), - * - * - * stateStackSize: function(), - * - * - * performAction: function(yy, yy_, yyrulenumber, YY_START), - * - * - * rules: [...], - * - * - * conditions: {associative list: name ==> set}, - * } - * - * - * token location info (`yylloc`): { - * first_line: n, - * last_line: n, - * first_column: n, - * last_column: n, - * range: [start_number, end_number] - * (where the numbers are indexes into the input string, zero-based) - * } - * - * --- - * - * The `parseError` function receives a 'hash' object with these members for lexer errors: - * - * { - * text: (matched text) - * token: (the produced terminal token, if any) - * token_id: (the produced terminal token numeric ID, if any) - * line: (yylineno) - * loc: (yylloc) - * recoverable: (boolean: TRUE when the parser MAY have an error recovery rule - * available for this particular error) - * yy: (object: the current parser internal "shared state" `yy` - * as is also available in the rule actions; this can be used, - * for instance, for advanced error analysis and reporting) - * lexer: (reference to the current lexer instance used by the parser) - * } - * - * while `this` will reference the current lexer instance. - * - * When `parseError` is invoked by the lexer, the default implementation will - * attempt to invoke `yy.parser.parseError()`; when this callback is not provided - * it will try to invoke `yy.parseError()` instead. When that callback is also not - * provided, a `JisonLexerError` exception will be thrown containing the error - * message and `hash`, as constructed by the `constructLexErrorInfo()` API. - * - * Note that the lexer's `JisonLexerError` error class is passed via the - * `ExceptionClass` argument, which is invoked to construct the exception - * instance to be thrown, so technically `parseError` will throw the object - * produced by the `new ExceptionClass(str, hash)` JavaScript expression. - * - * --- - * - * You can specify lexer options by setting / modifying the `.options` object of your Lexer instance. - * These options are available: - * - * (Options are permanent.) - * - * yy: { - * parseError: function(str, hash, ExceptionClass) - * optional: overrides the default `parseError` function. - * } - * - * lexer.options: { - * pre_lex: function() - * optional: is invoked before the lexer is invoked to produce another token. - * `this` refers to the Lexer object. - * post_lex: function(token) { return token; } - * optional: is invoked when the lexer has produced a token `token`; - * this function can override the returned token value by returning another. - * When it does not return any (truthy) value, the lexer will return - * the original `token`. - * `this` refers to the Lexer object. - * - * WARNING: the next set of options are not meant to be changed. They echo the abilities of - * the lexer as per when it was compiled! - * - * ranges: boolean - * optional: `true` ==> token location info will include a .range[] member. - * flex: boolean - * optional: `true` ==> flex-like lexing behaviour where the rules are tested - * exhaustively to find the longest match. - * backtrack_lexer: boolean - * optional: `true` ==> lexer regexes are tested in order and for invoked; - * the lexer terminates the scan when a token is returned by the action code. - * xregexp: boolean - * optional: `true` ==> lexer rule regexes are "extended regex format" requiring the - * `XRegExp` library. When this %option has not been specified at compile time, all lexer - * rule regexes have been written as standard JavaScript RegExp expressions. - * } - */ - - - var lexer = function() { - /** - * See also: - * http://stackoverflow.com/questions/1382107/whats-a-good-way-to-extend-error-in-javascript/#35881508 - * but we keep the prototype.constructor and prototype.name assignment lines too for compatibility - * with userland code which might access the derived class in a 'classic' way. - * - * @public - * @constructor - * @nocollapse - */ - function JisonLexerError(msg, hash) { - Object.defineProperty(this, 'name', { - enumerable: false, - writable: false, - value: 'JisonLexerError' - }); - - if (msg == null) - msg = '???'; - - Object.defineProperty(this, 'message', { - enumerable: false, - writable: true, - value: msg - }); - - this.hash = hash; - var stacktrace; - - if (hash && hash.exception instanceof Error) { - var ex2 = hash.exception; - this.message = ex2.message || msg; - stacktrace = ex2.stack; - } - - if (!stacktrace) { - if (Error.hasOwnProperty('captureStackTrace')) { - // V8 - Error.captureStackTrace(this, this.constructor); - } else { - stacktrace = new Error(msg).stack; - } - } - - if (stacktrace) { - Object.defineProperty(this, 'stack', { - enumerable: false, - writable: false, - value: stacktrace - }); - } - } - - if (typeof Object.setPrototypeOf === 'function') { - Object.setPrototypeOf(JisonLexerError.prototype, Error.prototype); - } else { - JisonLexerError.prototype = Object.create(Error.prototype); - } - - JisonLexerError.prototype.constructor = JisonLexerError; - JisonLexerError.prototype.name = 'JisonLexerError'; - - var lexer = { - - // Code Generator Information Report - // --------------------------------- - // - // Options: - // - // backtracking: .................... false - // location.ranges: ................. false - // location line+column tracking: ... true - // - // - // Forwarded Parser Analysis flags: - // - // uses yyleng: ..................... false - // uses yylineno: ................... false - // uses yytext: ..................... false - // uses yylloc: ..................... false - // uses lexer values: ............... true / true - // location tracking: ............... false - // location assignment: ............. false - // - // - // Lexer Analysis flags: - // - // uses yyleng: ..................... ??? - // uses yylineno: ................... ??? - // uses yytext: ..................... ??? - // uses yylloc: ..................... ??? - // uses ParseError API: ............. ??? - // uses yyerror: .................... ??? - // uses location tracking & editing: ??? - // uses more() API: ................. ??? - // uses unput() API: ................ ??? - // uses reject() API: ............... ??? - // uses less() API: ................. ??? - // uses display APIs pastInput(), upcomingInput(), showPosition(): - // ............................. ??? - // uses describeYYLLOC() API: ....... ??? - // - // --------- END OF REPORT ----------- - - EOF: 1, - ERROR: 2, - - // JisonLexerError: JisonLexerError, /// <-- injected by the code generator - - // options: {}, /// <-- injected by the code generator - - // yy: ..., /// <-- injected by setInput() - - __currentRuleSet__: null, /// INTERNAL USE ONLY: internal rule set cache for the current lexer state - - __error_infos: [], /// INTERNAL USE ONLY: the set of lexErrorInfo objects created since the last cleanup - __decompressed: false, /// INTERNAL USE ONLY: mark whether the lexer instance has been 'unfolded' completely and is now ready for use - done: false, /// INTERNAL USE ONLY - _backtrack: false, /// INTERNAL USE ONLY - _input: '', /// INTERNAL USE ONLY - _more: false, /// INTERNAL USE ONLY - _signaled_error_token: false, /// INTERNAL USE ONLY - conditionStack: [], /// INTERNAL USE ONLY; managed via `pushState()`, `popState()`, `topState()` and `stateStackSize()` - match: '', /// READ-ONLY EXTERNAL ACCESS - ADVANCED USE ONLY: tracks input which has been matched so far for the lexer token under construction. `match` is identical to `yytext` except that this one still contains the matched input string after `lexer.performAction()` has been invoked, where userland code MAY have changed/replaced the `yytext` value entirely! - matched: '', /// READ-ONLY EXTERNAL ACCESS - ADVANCED USE ONLY: tracks entire input which has been matched so far - matches: false, /// READ-ONLY EXTERNAL ACCESS - ADVANCED USE ONLY: tracks RE match result for last (successful) match attempt - yytext: '', /// ADVANCED USE ONLY: tracks input which has been matched so far for the lexer token under construction; this value is transferred to the parser as the 'token value' when the parser consumes the lexer token produced through a call to the `lex()` API. - offset: 0, /// READ-ONLY EXTERNAL ACCESS - ADVANCED USE ONLY: tracks the 'cursor position' in the input string, i.e. the number of characters matched so far - yyleng: 0, /// READ-ONLY EXTERNAL ACCESS - ADVANCED USE ONLY: length of matched input for the token under construction (`yytext`) - yylineno: 0, /// READ-ONLY EXTERNAL ACCESS - ADVANCED USE ONLY: 'line number' at which the token under construction is located - yylloc: null, /// READ-ONLY EXTERNAL ACCESS - ADVANCED USE ONLY: tracks location info (lines + columns) for the token under construction - - /** - * INTERNAL USE: construct a suitable error info hash object instance for `parseError`. - * - * @public - * @this {RegExpLexer} - */ - constructLexErrorInfo: function lexer_constructLexErrorInfo(msg, recoverable, show_input_position) { - msg = '' + msg; - - // heuristic to determine if the error message already contains a (partial) source code dump - // as produced by either `showPosition()` or `prettyPrintRange()`: - if (show_input_position == undefined) { - show_input_position = !(msg.indexOf('\n') > 0 && msg.indexOf('^') > 0); - } - - if (this.yylloc && show_input_position) { - if (typeof this.prettyPrintRange === 'function') { - this.prettyPrintRange(this.yylloc); - - if (!/\n\s*$/.test(msg)) { - msg += '\n'; - } - - msg += '\n Erroneous area:\n' + this.prettyPrintRange(this.yylloc); - } else if (typeof this.showPosition === 'function') { - var pos_str = this.showPosition(); - - if (pos_str) { - if (msg.length && msg[msg.length - 1] !== '\n' && pos_str[0] !== '\n') { - msg += '\n' + pos_str; - } else { - msg += pos_str; - } - } - } - } - - /** @constructor */ - var pei = { - errStr: msg, - recoverable: !!recoverable, - text: this.match, // This one MAY be empty; userland code should use the `upcomingInput` API to obtain more text which follows the 'lexer cursor position'... - token: null, - line: this.yylineno, - loc: this.yylloc, - yy: this.yy, - lexer: this, - - /** - * and make sure the error info doesn't stay due to potential - * ref cycle via userland code manipulations. - * These would otherwise all be memory leak opportunities! - * - * Note that only array and object references are nuked as those - * constitute the set of elements which can produce a cyclic ref. - * The rest of the members is kept intact as they are harmless. - * - * @public - * @this {LexErrorInfo} - */ - destroy: function destructLexErrorInfo() { - // remove cyclic references added to error info: - // info.yy = null; - // info.lexer = null; - // ... - var rec = !!this.recoverable; - - for (var key in this) { - if (this.hasOwnProperty(key) && typeof key === 'object') { - this[key] = undefined; - } - } - - this.recoverable = rec; - } - }; - - // track this instance so we can `destroy()` it once we deem it superfluous and ready for garbage collection! - this.__error_infos.push(pei); - - return pei; - }, - - /** - * handler which is invoked when a lexer error occurs. - * - * @public - * @this {RegExpLexer} - */ - parseError: function lexer_parseError(str, hash, ExceptionClass) { - if (!ExceptionClass) { - ExceptionClass = this.JisonLexerError; - } - - if (this.yy) { - if (this.yy.parser && typeof this.yy.parser.parseError === 'function') { - return this.yy.parser.parseError.call(this, str, hash, ExceptionClass) || this.ERROR; - } else if (typeof this.yy.parseError === 'function') { - return this.yy.parseError.call(this, str, hash, ExceptionClass) || this.ERROR; - } - } - - throw new ExceptionClass(str, hash); - }, - - /** - * method which implements `yyerror(str, ...args)` functionality for use inside lexer actions. - * - * @public - * @this {RegExpLexer} - */ - yyerror: function yyError(str /*, ...args */) { - var lineno_msg = ''; - - if (this.yylloc) { - lineno_msg = ' on line ' + (this.yylineno + 1); - } - - var p = this.constructLexErrorInfo( - 'Lexical error' + lineno_msg + ': ' + str, - this.options.lexerErrorsAreRecoverable - ); - - // Add any extra args to the hash under the name `extra_error_attributes`: - var args = Array.prototype.slice.call(arguments, 1); - - if (args.length) { - p.extra_error_attributes = args; - } - - return this.parseError(p.errStr, p, this.JisonLexerError) || this.ERROR; - }, - - /** - * final cleanup function for when we have completed lexing the input; - * make it an API so that external code can use this one once userland - * code has decided it's time to destroy any lingering lexer error - * hash object instances and the like: this function helps to clean - * up these constructs, which *may* carry cyclic references which would - * otherwise prevent the instances from being properly and timely - * garbage-collected, i.e. this function helps prevent memory leaks! - * - * @public - * @this {RegExpLexer} - */ - cleanupAfterLex: function lexer_cleanupAfterLex(do_not_nuke_errorinfos) { - // prevent lingering circular references from causing memory leaks: - this.setInput('', {}); - - // nuke the error hash info instances created during this run. - // Userland code must COPY any data/references - // in the error hash instance(s) it is more permanently interested in. - if (!do_not_nuke_errorinfos) { - for (var i = this.__error_infos.length - 1; i >= 0; i--) { - var el = this.__error_infos[i]; - - if (el && typeof el.destroy === 'function') { - el.destroy(); - } - } - - this.__error_infos.length = 0; - } - - return this; - }, - - /** - * clear the lexer token context; intended for internal use only - * - * @public - * @this {RegExpLexer} - */ - clear: function lexer_clear() { - this.yytext = ''; - this.yyleng = 0; - this.match = ''; - - // - DO NOT reset `this.matched` - this.matches = false; - - this._more = false; - this._backtrack = false; - var col = (this.yylloc ? this.yylloc.last_column : 0); - - this.yylloc = { - first_line: this.yylineno + 1, - first_column: col, - last_line: this.yylineno + 1, - last_column: col, - range: [this.offset, this.offset] - }; - }, - - /** - * resets the lexer, sets new input - * - * @public - * @this {RegExpLexer} - */ - setInput: function lexer_setInput(input, yy) { - this.yy = yy || this.yy || {}; - - // also check if we've fully initialized the lexer instance, - // including expansion work to be done to go from a loaded - // lexer to a usable lexer: - if (!this.__decompressed) { - // step 1: decompress the regex list: - var rules = this.rules; - - for (var i = 0, len = rules.length; i < len; i++) { - var rule_re = rules[i]; - - // compression: is the RE an xref to another RE slot in the rules[] table? - if (typeof rule_re === 'number') { - rules[i] = rules[rule_re]; - } - } - - // step 2: unfold the conditions[] set to make these ready for use: - var conditions = this.conditions; - - for (var k in conditions) { - var spec = conditions[k]; - var rule_ids = spec.rules; - var len = rule_ids.length; - var rule_regexes = new Array(len + 1); // slot 0 is unused; we use a 1-based index approach here to keep the hottest code in `lexer_next()` fast and simple! - var rule_new_ids = new Array(len + 1); - - for (var i = 0; i < len; i++) { - var idx = rule_ids[i]; - var rule_re = rules[idx]; - rule_regexes[i + 1] = rule_re; - rule_new_ids[i + 1] = idx; - } - - spec.rules = rule_new_ids; - spec.__rule_regexes = rule_regexes; - spec.__rule_count = len; - } - - this.__decompressed = true; - } - - this._input = input || ''; - this.clear(); - this._signaled_error_token = false; - this.done = false; - this.yylineno = 0; - this.matched = ''; - this.conditionStack = ['INITIAL']; - this.__currentRuleSet__ = null; - - this.yylloc = { - first_line: 1, - first_column: 0, - last_line: 1, - last_column: 0, - range: [0, 0] - }; - - this.offset = 0; - return this; - }, - - /** - * edit the remaining input via user-specified callback. - * This can be used to forward-adjust the input-to-parse, - * e.g. inserting macro expansions and alike in the - * input which has yet to be lexed. - * The behaviour of this API contrasts the `unput()` et al - * APIs as those act on the *consumed* input, while this - * one allows one to manipulate the future, without impacting - * the current `yyloc` cursor location or any history. - * - * Use this API to help implement C-preprocessor-like - * `#include` statements, etc. - * - * The provided callback must be synchronous and is - * expected to return the edited input (string). - * - * The `cpsArg` argument value is passed to the callback - * as-is. - * - * `callback` interface: - * `function callback(input, cpsArg)` - * - * - `input` will carry the remaining-input-to-lex string - * from the lexer. - * - `cpsArg` is `cpsArg` passed into this API. - * - * The `this` reference for the callback will be set to - * reference this lexer instance so that userland code - * in the callback can easily and quickly access any lexer - * API. - * - * When the callback returns a non-string-type falsey value, - * we assume the callback did not edit the input and we - * will using the input as-is. - * - * When the callback returns a non-string-type value, it - * is converted to a string for lexing via the `"" + retval` - * operation. (See also why: http://2ality.com/2012/03/converting-to-string.html - * -- that way any returned object's `toValue()` and `toString()` - * methods will be invoked in a proper/desirable order.) - * - * @public - * @this {RegExpLexer} - */ - editRemainingInput: function lexer_editRemainingInput(callback, cpsArg) { - var rv = callback.call(this, this._input, cpsArg); - - if (typeof rv !== 'string') { - if (rv) { - this._input = '' + rv; - } - // else: keep `this._input` as is. - } else { - this._input = rv; - } - - return this; - }, - - /** - * consumes and returns one char from the input - * - * @public - * @this {RegExpLexer} - */ - input: function lexer_input() { - if (!this._input) { - //this.done = true; -- don't set `done` as we want the lex()/next() API to be able to produce one custom EOF token match after this anyhow. (lexer can match special <> tokens and perform user action code for a <> match, but only does so *once*) - return null; - } - - var ch = this._input[0]; - this.yytext += ch; - this.yyleng++; - this.offset++; - this.match += ch; - this.matched += ch; - - // Count the linenumber up when we hit the LF (or a stand-alone CR). - // On CRLF, the linenumber is incremented when you fetch the CR or the CRLF combo - // and we advance immediately past the LF as well, returning both together as if - // it was all a single 'character' only. - var slice_len = 1; - - var lines = false; - - if (ch === '\n') { - lines = true; - } else if (ch === '\r') { - lines = true; - var ch2 = this._input[1]; - - if (ch2 === '\n') { - slice_len++; - ch += ch2; - this.yytext += ch2; - this.yyleng++; - this.offset++; - this.match += ch2; - this.matched += ch2; - this.yylloc.range[1]++; - } - } - - if (lines) { - this.yylineno++; - this.yylloc.last_line++; - this.yylloc.last_column = 0; - } else { - this.yylloc.last_column++; - } - - this.yylloc.range[1]++; - this._input = this._input.slice(slice_len); - return ch; - }, - - /** - * unshifts one char (or an entire string) into the input - * - * @public - * @this {RegExpLexer} - */ - unput: function lexer_unput(ch) { - var len = ch.length; - var lines = ch.split(/(?:\r\n?|\n)/g); - this._input = ch + this._input; - this.yytext = this.yytext.substr(0, this.yytext.length - len); - this.yyleng = this.yytext.length; - this.offset -= len; - this.match = this.match.substr(0, this.match.length - len); - this.matched = this.matched.substr(0, this.matched.length - len); - - if (lines.length > 1) { - this.yylineno -= lines.length - 1; - this.yylloc.last_line = this.yylineno + 1; - - // Get last entirely matched line into the `pre_lines[]` array's - // last index slot; we don't mind when other previously - // matched lines end up in the array too. - var pre = this.match; - - var pre_lines = pre.split(/(?:\r\n?|\n)/g); - - if (pre_lines.length === 1) { - pre = this.matched; - pre_lines = pre.split(/(?:\r\n?|\n)/g); - } - - this.yylloc.last_column = pre_lines[pre_lines.length - 1].length; - } else { - this.yylloc.last_column -= len; - } - - this.yylloc.range[1] = this.yylloc.range[0] + this.yyleng; - this.done = false; - return this; - }, - - /** - * cache matched text and append it on next action - * - * @public - * @this {RegExpLexer} - */ - more: function lexer_more() { - this._more = true; - return this; - }, - - /** - * signal the lexer that this rule fails to match the input, so the - * next matching rule (regex) should be tested instead. - * - * @public - * @this {RegExpLexer} - */ - reject: function lexer_reject() { - if (this.options.backtrack_lexer) { - this._backtrack = true; - } else { - // when the `parseError()` call returns, we MUST ensure that the error is registered. - // We accomplish this by signaling an 'error' token to be produced for the current - // `.lex()` run. - var lineno_msg = ''; - - if (this.yylloc) { - lineno_msg = ' on line ' + (this.yylineno + 1); - } - - var p = this.constructLexErrorInfo( - 'Lexical error' + lineno_msg + ': You can only invoke reject() in the lexer when the lexer is of the backtracking persuasion (options.backtrack_lexer = true).', - false - ); - - this._signaled_error_token = this.parseError(p.errStr, p, this.JisonLexerError) || this.ERROR; - } - - return this; - }, - - /** - * retain first n characters of the match - * - * @public - * @this {RegExpLexer} - */ - less: function lexer_less(n) { - return this.unput(this.match.slice(n)); - }, - - /** - * return (part of the) already matched input, i.e. for error - * messages. - * - * Limit the returned string length to `maxSize` (default: 20). - * - * Limit the returned string to the `maxLines` number of lines of - * input (default: 1). - * - * Negative limit values equal *unlimited*. - * - * @public - * @this {RegExpLexer} - */ - pastInput: function lexer_pastInput(maxSize, maxLines) { - var past = this.matched.substring(0, this.matched.length - this.match.length); - - if (maxSize < 0) - maxSize = past.length; - else if (!maxSize) - maxSize = 20; - - if (maxLines < 0) - maxLines = past.length; // can't ever have more input lines than this! - else if (!maxLines) - maxLines = 1; - - // `substr` anticipation: treat \r\n as a single character and take a little - // more than necessary so that we can still properly check against maxSize - // after we've transformed and limited the newLines in here: - past = past.substr(-maxSize * 2 - 2); - - // now that we have a significantly reduced string to process, transform the newlines - // and chop them, then limit them: - var a = past.replace(/\r\n|\r/g, '\n').split('\n'); - - a = a.slice(-maxLines); - past = a.join('\n'); - - // When, after limiting to maxLines, we still have too much to return, - // do add an ellipsis prefix... - if (past.length > maxSize) { - past = '...' + past.substr(-maxSize); - } - - return past; - }, - - /** - * return (part of the) upcoming input, i.e. for error messages. - * - * Limit the returned string length to `maxSize` (default: 20). - * - * Limit the returned string to the `maxLines` number of lines of input (default: 1). - * - * Negative limit values equal *unlimited*. - * - * > ### NOTE ### - * > - * > *"upcoming input"* is defined as the whole of the both - * > the *currently lexed* input, together with any remaining input - * > following that. *"currently lexed"* input is the input - * > already recognized by the lexer but not yet returned with - * > the lexer token. This happens when you are invoking this API - * > from inside any lexer rule action code block. - * > - * - * @public - * @this {RegExpLexer} - */ - upcomingInput: function lexer_upcomingInput(maxSize, maxLines) { - var next = this.match; - - if (maxSize < 0) - maxSize = next.length + this._input.length; - else if (!maxSize) - maxSize = 20; - - if (maxLines < 0) - maxLines = maxSize; // can't ever have more input lines than this! - else if (!maxLines) - maxLines = 1; - - // `substring` anticipation: treat \r\n as a single character and take a little - // more than necessary so that we can still properly check against maxSize - // after we've transformed and limited the newLines in here: - if (next.length < maxSize * 2 + 2) { - next += this._input.substring(0, maxSize * 2 + 2); // substring is faster on Chrome/V8 - } - - // now that we have a significantly reduced string to process, transform the newlines - // and chop them, then limit them: - var a = next.replace(/\r\n|\r/g, '\n').split('\n'); - - a = a.slice(0, maxLines); - next = a.join('\n'); - - // When, after limiting to maxLines, we still have too much to return, - // do add an ellipsis postfix... - if (next.length > maxSize) { - next = next.substring(0, maxSize) + '...'; - } - - return next; - }, - - /** - * return a string which displays the character position where the - * lexing error occurred, i.e. for error messages - * - * @public - * @this {RegExpLexer} - */ - showPosition: function lexer_showPosition(maxPrefix, maxPostfix) { - var pre = this.pastInput(maxPrefix).replace(/\s/g, ' '); - var c = new Array(pre.length + 1).join('-'); - return pre + this.upcomingInput(maxPostfix).replace(/\s/g, ' ') + '\n' + c + '^'; - }, - - /** - * return an YYLLOC info object derived off the given context (actual, preceding, following, current). - * Use this method when the given `actual` location is not guaranteed to exist (i.e. when - * it MAY be NULL) and you MUST have a valid location info object anyway: - * then we take the given context of the `preceding` and `following` locations, IFF those are available, - * and reconstruct the `actual` location info from those. - * If this fails, the heuristic is to take the `current` location, IFF available. - * If this fails as well, we assume the sought location is at/around the current lexer position - * and then produce that one as a response. DO NOTE that these heuristic/derived location info - * values MAY be inaccurate! - * - * NOTE: `deriveLocationInfo()` ALWAYS produces a location info object *copy* of `actual`, not just - * a *reference* hence all input location objects can be assumed to be 'constant' (function has no side-effects). - * - * @public - * @this {RegExpLexer} - */ - deriveLocationInfo: function lexer_deriveYYLLOC(actual, preceding, following, current) { - var loc = { - first_line: 1, - first_column: 0, - last_line: 1, - last_column: 0, - range: [0, 0] - }; - - if (actual) { - loc.first_line = actual.first_line | 0; - loc.last_line = actual.last_line | 0; - loc.first_column = actual.first_column | 0; - loc.last_column = actual.last_column | 0; - - if (actual.range) { - loc.range[0] = actual.range[0] | 0; - loc.range[1] = actual.range[1] | 0; - } - } - - if (loc.first_line <= 0 || loc.last_line < loc.first_line) { - // plan B: heuristic using preceding and following: - if (loc.first_line <= 0 && preceding) { - loc.first_line = preceding.last_line | 0; - loc.first_column = preceding.last_column | 0; - - if (preceding.range) { - loc.range[0] = actual.range[1] | 0; - } - } - - if ((loc.last_line <= 0 || loc.last_line < loc.first_line) && following) { - loc.last_line = following.first_line | 0; - loc.last_column = following.first_column | 0; - - if (following.range) { - loc.range[1] = actual.range[0] | 0; - } - } - - // plan C?: see if the 'current' location is useful/sane too: - if (loc.first_line <= 0 && current && (loc.last_line <= 0 || current.last_line <= loc.last_line)) { - loc.first_line = current.first_line | 0; - loc.first_column = current.first_column | 0; - - if (current.range) { - loc.range[0] = current.range[0] | 0; - } - } - - if (loc.last_line <= 0 && current && (loc.first_line <= 0 || current.first_line >= loc.first_line)) { - loc.last_line = current.last_line | 0; - loc.last_column = current.last_column | 0; - - if (current.range) { - loc.range[1] = current.range[1] | 0; - } - } - } - - // sanitize: fix last_line BEFORE we fix first_line as we use the 'raw' value of the latter - // or plan D heuristics to produce a 'sensible' last_line value: - if (loc.last_line <= 0) { - if (loc.first_line <= 0) { - loc.first_line = this.yylloc.first_line; - loc.last_line = this.yylloc.last_line; - loc.first_column = this.yylloc.first_column; - loc.last_column = this.yylloc.last_column; - loc.range[0] = this.yylloc.range[0]; - loc.range[1] = this.yylloc.range[1]; - } else { - loc.last_line = this.yylloc.last_line; - loc.last_column = this.yylloc.last_column; - loc.range[1] = this.yylloc.range[1]; - } - } - - if (loc.first_line <= 0) { - loc.first_line = loc.last_line; - loc.first_column = 0; // loc.last_column; - loc.range[1] = loc.range[0]; - } - - if (loc.first_column < 0) { - loc.first_column = 0; - } - - if (loc.last_column < 0) { - loc.last_column = (loc.first_column > 0 ? loc.first_column : 80); - } - - return loc; - }, - - /** - * return a string which displays the lines & columns of input which are referenced - * by the given location info range, plus a few lines of context. - * - * This function pretty-prints the indicated section of the input, with line numbers - * and everything! - * - * This function is very useful to provide highly readable error reports, while - * the location range may be specified in various flexible ways: - * - * - `loc` is the location info object which references the area which should be - * displayed and 'marked up': these lines & columns of text are marked up by `^` - * characters below each character in the entire input range. - * - * - `context_loc` is the *optional* location info object which instructs this - * pretty-printer how much *leading* context should be displayed alongside - * the area referenced by `loc`. This can help provide context for the displayed - * error, etc. - * - * When this location info is not provided, a default context of 3 lines is - * used. - * - * - `context_loc2` is another *optional* location info object, which serves - * a similar purpose to `context_loc`: it specifies the amount of *trailing* - * context lines to display in the pretty-print output. - * - * When this location info is not provided, a default context of 1 line only is - * used. - * - * Special Notes: - * - * - when the `loc`-indicated range is very large (about 5 lines or more), then - * only the first and last few lines of this block are printed while a - * `...continued...` message will be printed between them. - * - * This serves the purpose of not printing a huge amount of text when the `loc` - * range happens to be huge: this way a manageable & readable output results - * for arbitrary large ranges. - * - * - this function can display lines of input which whave not yet been lexed. - * `prettyPrintRange()` can access the entire input! - * - * @public - * @this {RegExpLexer} - */ - prettyPrintRange: function lexer_prettyPrintRange(loc, context_loc, context_loc2) { - loc = this.deriveLocationInfo(loc, context_loc, context_loc2); - const CONTEXT = 3; - const CONTEXT_TAIL = 1; - const MINIMUM_VISIBLE_NONEMPTY_LINE_COUNT = 2; - var input = this.matched + this._input; - var lines = input.split('\n'); - var l0 = Math.max(1, (context_loc ? context_loc.first_line : loc.first_line - CONTEXT)); - var l1 = Math.max(1, (context_loc2 ? context_loc2.last_line : loc.last_line + CONTEXT_TAIL)); - var lineno_display_width = 1 + Math.log10(l1 | 1) | 0; - var ws_prefix = new Array(lineno_display_width).join(' '); - var nonempty_line_indexes = []; - - var rv = lines.slice(l0 - 1, l1 + 1).map(function injectLineNumber(line, index) { - var lno = index + l0; - var lno_pfx = (ws_prefix + lno).substr(-lineno_display_width); - var rv = lno_pfx + ': ' + line; - var errpfx = new Array(lineno_display_width + 1).join('^'); - var offset = 2 + 1; - var len = 0; - - if (lno === loc.first_line) { - offset += loc.first_column; - - len = Math.max( - 2, - ((lno === loc.last_line ? loc.last_column : line.length)) - loc.first_column + 1 - ); - } else if (lno === loc.last_line) { - len = Math.max(2, loc.last_column + 1); - } else if (lno > loc.first_line && lno < loc.last_line) { - len = Math.max(2, line.length + 1); - } - - if (len) { - var lead = new Array(offset).join('.'); - var mark = new Array(len).join('^'); - rv += '\n' + errpfx + lead + mark; - - if (line.trim().length > 0) { - nonempty_line_indexes.push(index); - } - } - - rv = rv.replace(/\t/g, ' '); - return rv; - }); - - // now make sure we don't print an overly large amount of error area: limit it - // to the top and bottom line count: - if (nonempty_line_indexes.length > 2 * MINIMUM_VISIBLE_NONEMPTY_LINE_COUNT) { - var clip_start = nonempty_line_indexes[MINIMUM_VISIBLE_NONEMPTY_LINE_COUNT - 1] + 1; - var clip_end = nonempty_line_indexes[nonempty_line_indexes.length - MINIMUM_VISIBLE_NONEMPTY_LINE_COUNT] - 1; - var intermediate_line = new Array(lineno_display_width + 1).join(' ') + ' (...continued...)'; - intermediate_line += '\n' + new Array(lineno_display_width + 1).join('-') + ' (---------------)'; - rv.splice(clip_start, clip_end - clip_start + 1, intermediate_line); - } - - return rv.join('\n'); - }, - - /** - * helper function, used to produce a human readable description as a string, given - * the input `yylloc` location object. - * - * Set `display_range_too` to TRUE to include the string character index position(s) - * in the description if the `yylloc.range` is available. - * - * @public - * @this {RegExpLexer} - */ - describeYYLLOC: function lexer_describe_yylloc(yylloc, display_range_too) { - var l1 = yylloc.first_line; - var l2 = yylloc.last_line; - var c1 = yylloc.first_column; - var c2 = yylloc.last_column; - var dl = l2 - l1; - var dc = c2 - c1; - var rv; - - if (dl === 0) { - rv = 'line ' + l1 + ', '; - - if (dc <= 1) { - rv += 'column ' + c1; - } else { - rv += 'columns ' + c1 + ' .. ' + c2; - } - } else { - rv = 'lines ' + l1 + '(column ' + c1 + ') .. ' + l2 + '(column ' + c2 + ')'; - } - - if (yylloc.range && display_range_too) { - var r1 = yylloc.range[0]; - var r2 = yylloc.range[1] - 1; - - if (r2 <= r1) { - rv += ' {String Offset: ' + r1 + '}'; - } else { - rv += ' {String Offset range: ' + r1 + ' .. ' + r2 + '}'; - } - } - - return rv; - }, - - /** - * test the lexed token: return FALSE when not a match, otherwise return token. - * - * `match` is supposed to be an array coming out of a regex match, i.e. `match[0]` - * contains the actually matched text string. - * - * Also move the input cursor forward and update the match collectors: - * - * - `yytext` - * - `yyleng` - * - `match` - * - `matches` - * - `yylloc` - * - `offset` - * - * @public - * @this {RegExpLexer} - */ - test_match: function lexer_test_match(match, indexed_rule) { - var token, lines, backup, match_str, match_str_len; - - if (this.options.backtrack_lexer) { - // save context - backup = { - yylineno: this.yylineno, - - yylloc: { - first_line: this.yylloc.first_line, - last_line: this.yylloc.last_line, - first_column: this.yylloc.first_column, - last_column: this.yylloc.last_column, - range: this.yylloc.range.slice(0) - }, - - yytext: this.yytext, - match: this.match, - matches: this.matches, - matched: this.matched, - yyleng: this.yyleng, - offset: this.offset, - _more: this._more, - _input: this._input, - - //_signaled_error_token: this._signaled_error_token, - yy: this.yy, - - conditionStack: this.conditionStack.slice(0), - done: this.done - }; - } - - match_str = match[0]; - match_str_len = match_str.length; - - // if (match_str.indexOf('\n') !== -1 || match_str.indexOf('\r') !== -1) { - lines = match_str.split(/(?:\r\n?|\n)/g); - - if (lines.length > 1) { - this.yylineno += lines.length - 1; - this.yylloc.last_line = this.yylineno + 1; - this.yylloc.last_column = lines[lines.length - 1].length; - } else { - this.yylloc.last_column += match_str_len; - } - - // } - this.yytext += match_str; - - this.match += match_str; - this.matched += match_str; - this.matches = match; - this.yyleng = this.yytext.length; - this.yylloc.range[1] += match_str_len; - - // previous lex rules MAY have invoked the `more()` API rather than producing a token: - // those rules will already have moved this `offset` forward matching their match lengths, - // hence we must only add our own match length now: - this.offset += match_str_len; - - this._more = false; - this._backtrack = false; - this._input = this._input.slice(match_str_len); - - // calling this method: - // - // function lexer__performAction(yy, yyrulenumber, YY_START) {...} - token = this.performAction.call( - this, - this.yy, - indexed_rule, - this.conditionStack[this.conditionStack.length - 1] /* = YY_START */ - ); - - // otherwise, when the action codes are all simple return token statements: - //token = this.simpleCaseActionClusters[indexed_rule]; - - if (this.done && this._input) { - this.done = false; - } - - if (token) { - return token; - } else if (this._backtrack) { - // recover context - for (var k in backup) { - this[k] = backup[k]; - } - - this.__currentRuleSet__ = null; - return false; // rule action called reject() implying the next rule should be tested instead. - } else if (this._signaled_error_token) { - // produce one 'error' token as `.parseError()` in `reject()` - // did not guarantee a failure signal by throwing an exception! - token = this._signaled_error_token; - - this._signaled_error_token = false; - return token; - } - - return false; - }, - - /** - * return next match in input - * - * @public - * @this {RegExpLexer} - */ - next: function lexer_next() { - if (this.done) { - this.clear(); - return this.EOF; - } - - if (!this._input) { - this.done = true; - } - - var token, match, tempMatch, index; - - if (!this._more) { - this.clear(); - } - - var spec = this.__currentRuleSet__; - - if (!spec) { - // Update the ruleset cache as we apparently encountered a state change or just started lexing. - // The cache is set up for fast lookup -- we assume a lexer will switch states much less often than it will - // invoke the `lex()` token-producing API and related APIs, hence caching the set for direct access helps - // speed up those activities a tiny bit. - spec = this.__currentRuleSet__ = this._currentRules(); - - // Check whether a *sane* condition has been pushed before: this makes the lexer robust against - // user-programmer bugs such as https://github.com/zaach/jison-lex/issues/19 - if (!spec || !spec.rules) { - var lineno_msg = ''; - - if (this.options.trackPosition) { - lineno_msg = ' on line ' + (this.yylineno + 1); - } - - var p = this.constructLexErrorInfo( - 'Internal lexer engine error' + lineno_msg + ': The lex grammar programmer pushed a non-existing condition name "' + this.topState() + '"; this is a fatal error and should be reported to the application programmer team!', - false - ); - - // produce one 'error' token until this situation has been resolved, most probably by parse termination! - return this.parseError(p.errStr, p, this.JisonLexerError) || this.ERROR; - } - } - - var rule_ids = spec.rules; - var regexes = spec.__rule_regexes; - var len = spec.__rule_count; - - // Note: the arrays are 1-based, while `len` itself is a valid index, - // hence the non-standard less-or-equal check in the next loop condition! - for (var i = 1; i <= len; i++) { - tempMatch = this._input.match(regexes[i]); - - if (tempMatch && (!match || tempMatch[0].length > match[0].length)) { - match = tempMatch; - index = i; - - if (this.options.backtrack_lexer) { - token = this.test_match(tempMatch, rule_ids[i]); - - if (token !== false) { - return token; - } else if (this._backtrack) { - match = undefined; - continue; // rule action called reject() implying a rule MISmatch. - } else { - // else: this is a lexer rule which consumes input without producing a token (e.g. whitespace) - return false; - } - } else if (!this.options.flex) { - break; - } - } - } - - if (match) { - token = this.test_match(match, rule_ids[index]); - - if (token !== false) { - return token; - } - - // else: this is a lexer rule which consumes input without producing a token (e.g. whitespace) - return false; - } - - if (!this._input) { - this.done = true; - this.clear(); - return this.EOF; - } else { - var lineno_msg = ''; - - if (this.options.trackPosition) { - lineno_msg = ' on line ' + (this.yylineno + 1); - } - - var p = this.constructLexErrorInfo( - 'Lexical error' + lineno_msg + ': Unrecognized text.', - this.options.lexerErrorsAreRecoverable - ); - - var pendingInput = this._input; - var activeCondition = this.topState(); - var conditionStackDepth = this.conditionStack.length; - token = this.parseError(p.errStr, p, this.JisonLexerError) || this.ERROR; - - if (token === this.ERROR) { - // we can try to recover from a lexer error that `parseError()` did not 'recover' for us - // by moving forward at least one character at a time IFF the (user-specified?) `parseError()` - // has not consumed/modified any pending input or changed state in the error handler: - if (!this.matches && // and make sure the input has been modified/consumed ... - pendingInput === this._input && // ...or the lexer state has been modified significantly enough - // to merit a non-consuming error handling action right now. - activeCondition === this.topState() && conditionStackDepth === this.conditionStack.length) { - this.input(); - } - } - - return token; - } - }, - - /** - * return next match that has a token - * - * @public - * @this {RegExpLexer} - */ - lex: function lexer_lex() { - var r; - - // allow the PRE/POST handlers set/modify the return token for maximum flexibility of the generated lexer: - if (typeof this.pre_lex === 'function') { - r = this.pre_lex.call(this, 0); - } - - if (typeof this.options.pre_lex === 'function') { - // (also account for a userdef function which does not return any value: keep the token as is) - r = this.options.pre_lex.call(this, r) || r; - } - - if (this.yy && typeof this.yy.pre_lex === 'function') { - // (also account for a userdef function which does not return any value: keep the token as is) - r = this.yy.pre_lex.call(this, r) || r; - } - - while (!r) { - r = this.next(); - } - - if (this.yy && typeof this.yy.post_lex === 'function') { - // (also account for a userdef function which does not return any value: keep the token as is) - r = this.yy.post_lex.call(this, r) || r; - } - - if (typeof this.options.post_lex === 'function') { - // (also account for a userdef function which does not return any value: keep the token as is) - r = this.options.post_lex.call(this, r) || r; - } - - if (typeof this.post_lex === 'function') { - // (also account for a userdef function which does not return any value: keep the token as is) - r = this.post_lex.call(this, r) || r; - } - - return r; - }, - - /** - * return next match that has a token. Identical to the `lex()` API but does not invoke any of the - * `pre_lex()` nor any of the `post_lex()` callbacks. - * - * @public - * @this {RegExpLexer} - */ - fastLex: function lexer_fastLex() { - var r; - - while (!r) { - r = this.next(); - } - - return r; - }, - - /** - * return info about the lexer state that can help a parser or other lexer API user to use the - * most efficient means available. This API is provided to aid run-time performance for larger - * systems which employ this lexer. - * - * @public - * @this {RegExpLexer} - */ - canIUse: function lexer_canIUse() { - var rv = { - fastLex: !(typeof this.pre_lex === 'function' || typeof this.options.pre_lex === 'function' || this.yy && typeof this.yy.pre_lex === 'function' || this.yy && typeof this.yy.post_lex === 'function' || typeof this.options.post_lex === 'function' || typeof this.post_lex === 'function') && typeof this.fastLex === 'function' - }; - - return rv; - }, - - /** - * backwards compatible alias for `pushState()`; - * the latter is symmetrical with `popState()` and we advise to use - * those APIs in any modern lexer code, rather than `begin()`. - * - * @public - * @this {RegExpLexer} - */ - begin: function lexer_begin(condition) { - return this.pushState(condition); - }, - - /** - * activates a new lexer condition state (pushes the new lexer - * condition state onto the condition stack) - * - * @public - * @this {RegExpLexer} - */ - pushState: function lexer_pushState(condition) { - this.conditionStack.push(condition); - this.__currentRuleSet__ = null; - return this; - }, - - /** - * pop the previously active lexer condition state off the condition - * stack - * - * @public - * @this {RegExpLexer} - */ - popState: function lexer_popState() { - var n = this.conditionStack.length - 1; - - if (n > 0) { - this.__currentRuleSet__ = null; - return this.conditionStack.pop(); - } else { - return this.conditionStack[0]; - } - }, - - /** - * return the currently active lexer condition state; when an index - * argument is provided it produces the N-th previous condition state, - * if available - * - * @public - * @this {RegExpLexer} - */ - topState: function lexer_topState(n) { - n = this.conditionStack.length - 1 - Math.abs(n || 0); - - if (n >= 0) { - return this.conditionStack[n]; - } else { - return 'INITIAL'; - } - }, - - /** - * (internal) determine the lexer rule set which is active for the - * currently active lexer condition state - * - * @public - * @this {RegExpLexer} - */ - _currentRules: function lexer__currentRules() { - if (this.conditionStack.length && this.conditionStack[this.conditionStack.length - 1]) { - return this.conditions[this.conditionStack[this.conditionStack.length - 1]]; - } else { - return this.conditions['INITIAL']; - } - }, - - /** - * return the number of states currently on the stack - * - * @public - * @this {RegExpLexer} - */ - stateStackSize: function lexer_stateStackSize() { - return this.conditionStack.length; - }, - - options: { - trackPosition: true - }, - - JisonLexerError: JisonLexerError, - - performAction: function lexer__performAction(yy, yyrulenumber, YY_START) { - - switch (yyrulenumber) { - case 1: - /*! Conditions:: INITIAL */ - /*! Rule:: \s+ */ - /* skip whitespace */ - break; - - default: - return this.simpleCaseActionClusters[yyrulenumber]; - } - }, - - simpleCaseActionClusters: { - /*! Conditions:: INITIAL */ - /*! Rule:: (--[0-9a-z-A-Z-]*) */ - 0: 13, - - /*! Conditions:: INITIAL */ - /*! Rule:: \* */ - 2: 5, - - /*! Conditions:: INITIAL */ - /*! Rule:: \/ */ - 3: 6, - - /*! Conditions:: INITIAL */ - /*! Rule:: \+ */ - 4: 3, - - /*! Conditions:: INITIAL */ - /*! Rule:: - */ - 5: 4, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)px\b */ - 6: 15, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)cm\b */ - 7: 15, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)mm\b */ - 8: 15, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)in\b */ - 9: 15, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)pt\b */ - 10: 15, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)pc\b */ - 11: 15, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)deg\b */ - 12: 16, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)grad\b */ - 13: 16, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)rad\b */ - 14: 16, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)turn\b */ - 15: 16, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)s\b */ - 16: 17, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)ms\b */ - 17: 17, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)Hz\b */ - 18: 18, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)kHz\b */ - 19: 18, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)dpi\b */ - 20: 19, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)dpcm\b */ - 21: 19, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)dppx\b */ - 22: 19, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)em\b */ - 23: 20, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)ex\b */ - 24: 21, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)ch\b */ - 25: 22, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)rem\b */ - 26: 23, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)vw\b */ - 27: 25, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)vh\b */ - 28: 24, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)vmin\b */ - 29: 26, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)vmax\b */ - 30: 27, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)% */ - 31: 28, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([0-9]+(\.[0-9]*)?|\.[0-9]+)\b */ - 32: 11, - - /*! Conditions:: INITIAL */ - /*! Rule:: (calc) */ - 33: 9, - - /*! Conditions:: INITIAL */ - /*! Rule:: (var) */ - 34: 12, - - /*! Conditions:: INITIAL */ - /*! Rule:: ([a-z]+) */ - 35: 10, - - /*! Conditions:: INITIAL */ - /*! Rule:: \( */ - 36: 7, - - /*! Conditions:: INITIAL */ - /*! Rule:: \) */ - 37: 8, - - /*! Conditions:: INITIAL */ - /*! Rule:: , */ - 38: 14, - - /*! Conditions:: INITIAL */ - /*! Rule:: $ */ - 39: 1 - }, - - rules: [ - /* 0: */ /^(?:(--[\d\-A-Za-z]*))/, - /* 1: */ /^(?:\s+)/, - /* 2: */ /^(?:\*)/, - /* 3: */ /^(?:\/)/, - /* 4: */ /^(?:\+)/, - /* 5: */ /^(?:-)/, - /* 6: */ /^(?:(\d+(\.\d*)?|\.\d+)px\b)/, - /* 7: */ /^(?:(\d+(\.\d*)?|\.\d+)cm\b)/, - /* 8: */ /^(?:(\d+(\.\d*)?|\.\d+)mm\b)/, - /* 9: */ /^(?:(\d+(\.\d*)?|\.\d+)in\b)/, - /* 10: */ /^(?:(\d+(\.\d*)?|\.\d+)pt\b)/, - /* 11: */ /^(?:(\d+(\.\d*)?|\.\d+)pc\b)/, - /* 12: */ /^(?:(\d+(\.\d*)?|\.\d+)deg\b)/, - /* 13: */ /^(?:(\d+(\.\d*)?|\.\d+)grad\b)/, - /* 14: */ /^(?:(\d+(\.\d*)?|\.\d+)rad\b)/, - /* 15: */ /^(?:(\d+(\.\d*)?|\.\d+)turn\b)/, - /* 16: */ /^(?:(\d+(\.\d*)?|\.\d+)s\b)/, - /* 17: */ /^(?:(\d+(\.\d*)?|\.\d+)ms\b)/, - /* 18: */ /^(?:(\d+(\.\d*)?|\.\d+)Hz\b)/, - /* 19: */ /^(?:(\d+(\.\d*)?|\.\d+)kHz\b)/, - /* 20: */ /^(?:(\d+(\.\d*)?|\.\d+)dpi\b)/, - /* 21: */ /^(?:(\d+(\.\d*)?|\.\d+)dpcm\b)/, - /* 22: */ /^(?:(\d+(\.\d*)?|\.\d+)dppx\b)/, - /* 23: */ /^(?:(\d+(\.\d*)?|\.\d+)em\b)/, - /* 24: */ /^(?:(\d+(\.\d*)?|\.\d+)ex\b)/, - /* 25: */ /^(?:(\d+(\.\d*)?|\.\d+)ch\b)/, - /* 26: */ /^(?:(\d+(\.\d*)?|\.\d+)rem\b)/, - /* 27: */ /^(?:(\d+(\.\d*)?|\.\d+)vw\b)/, - /* 28: */ /^(?:(\d+(\.\d*)?|\.\d+)vh\b)/, - /* 29: */ /^(?:(\d+(\.\d*)?|\.\d+)vmin\b)/, - /* 30: */ /^(?:(\d+(\.\d*)?|\.\d+)vmax\b)/, - /* 31: */ /^(?:(\d+(\.\d*)?|\.\d+)%)/, - /* 32: */ /^(?:(\d+(\.\d*)?|\.\d+)\b)/, - /* 33: */ /^(?:(calc))/, - /* 34: */ /^(?:(var))/, - /* 35: */ /^(?:([a-z]+))/, - /* 36: */ /^(?:\()/, - /* 37: */ /^(?:\))/, - /* 38: */ /^(?:,)/, - /* 39: */ /^(?:$)/ - ], - - conditions: { - 'INITIAL': { - rules: [ - 0, - 1, - 2, - 3, - 4, - 5, - 6, - 7, - 8, - 9, - 10, - 11, - 12, - 13, - 14, - 15, - 16, - 17, - 18, - 19, - 20, - 21, - 22, - 23, - 24, - 25, - 26, - 27, - 28, - 29, - 30, - 31, - 32, - 33, - 34, - 35, - 36, - 37, - 38, - 39 - ], - - inclusive: true - } - } - }; - - return lexer; - }(); - parser.lexer = lexer; - - - - function Parser() { - this.yy = {}; - } - Parser.prototype = parser; - parser.Parser = Parser; - - return new Parser(); - })(); - - - - - if (typeof commonjsRequire !== 'undefined' && 'object' !== 'undefined') { - exports.parser = parser; - exports.Parser = parser.Parser; - exports.parse = function () { - return parser.parse.apply(parser, arguments); - }; - - } -} (parser)); - -var reducer = {}; - -var convert = {exports: {}}; - -var conversions = { - // length - 'px': { - 'px': 1, - 'cm': 96.0/2.54, - 'mm': 96.0/25.4, - 'in': 96, - 'pt': 96.0/72.0, - 'pc': 16 - }, - 'cm': { - 'px': 2.54/96.0, - 'cm': 1, - 'mm': 0.1, - 'in': 2.54, - 'pt': 2.54/72.0, - 'pc': 2.54/6.0 - }, - 'mm': { - 'px': 25.4/96.0, - 'cm': 10, - 'mm': 1, - 'in': 25.4, - 'pt': 25.4/72.0, - 'pc': 25.4/6.0 - }, - 'in': { - 'px': 1.0/96.0, - 'cm': 1.0/2.54, - 'mm': 1.0/25.4, - 'in': 1, - 'pt': 1.0/72.0, - 'pc': 1.0/6.0 - }, - 'pt': { - 'px': 0.75, - 'cm': 72.0/2.54, - 'mm': 72.0/25.4, - 'in': 72, - 'pt': 1, - 'pc': 12 - }, - 'pc': { - 'px': 6.0/96.0, - 'cm': 6.0/2.54, - 'mm': 6.0/25.4, - 'in': 6, - 'pt': 6.0/72.0, - 'pc': 1 - }, - // angle - 'deg': { - 'deg': 1, - 'grad': 0.9, - 'rad': 180/Math.PI, - 'turn': 360 - }, - 'grad': { - 'deg': 400/360, - 'grad': 1, - 'rad': 200/Math.PI, - 'turn': 400 - }, - 'rad': { - 'deg': Math.PI/180, - 'grad': Math.PI/200, - 'rad': 1, - 'turn': Math.PI*2 - }, - 'turn': { - 'deg': 1/360, - 'grad': 1/400, - 'rad': 0.5/Math.PI, - 'turn': 1 - }, - // time - 's': { - 's': 1, - 'ms': 1/1000 - }, - 'ms': { - 's': 1000, - 'ms': 1 - }, - // frequency - 'Hz': { - 'Hz': 1, - 'kHz': 1000 - }, - 'kHz': { - 'Hz': 1/1000, - 'kHz': 1 - }, - // resolution - 'dpi': { - 'dpi': 1, - 'dpcm': 1.0/2.54, - 'dppx': 1/96 - }, - 'dpcm': { - 'dpi': 2.54, - 'dpcm': 1, - 'dppx': 2.54/96.0 - }, - 'dppx': { - 'dpi': 96, - 'dpcm': 96.0/2.54, - 'dppx': 1 - } -}; - -var cssUnitConverter = function (value, sourceUnit, targetUnit, precision) { - if (!conversions.hasOwnProperty(targetUnit)) - throw new Error("Cannot convert to " + targetUnit); - - if (!conversions[targetUnit].hasOwnProperty(sourceUnit)) - throw new Error("Cannot convert from " + sourceUnit + " to " + targetUnit); - - var converted = conversions[targetUnit][sourceUnit] * value; - - if (precision !== false) { - precision = Math.pow(10, parseInt(precision) || 5); - return Math.round(converted * precision) / precision; - } - - return converted; -}; - -(function (module, exports) { - - Object.defineProperty(exports, "__esModule", { - value: true - }); - - var _cssUnitConverter = cssUnitConverter; - - var _cssUnitConverter2 = _interopRequireDefault(_cssUnitConverter); - - function _interopRequireDefault(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - - function convertNodes(left, right, precision) { - switch (left.type) { - case 'LengthValue': - case 'AngleValue': - case 'TimeValue': - case 'FrequencyValue': - case 'ResolutionValue': - return convertAbsoluteLength(left, right, precision); - default: - return { left: left, right: right }; - } - } - - function convertAbsoluteLength(left, right, precision) { - if (right.type === left.type) { - right = { - type: left.type, - value: (0, _cssUnitConverter2.default)(right.value, right.unit, left.unit, precision), - unit: left.unit - }; - } - return { left: left, right: right }; - } - - exports.default = convertNodes; - module.exports = exports['default']; -} (convert, convert.exports)); - -Object.defineProperty(reducer, "__esModule", { - value: true -}); -reducer.flip = flip; - -var _convert = convert.exports; - -var _convert2 = _interopRequireDefault$2(_convert); - -function _interopRequireDefault$2(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -function reduce(node, precision) { - if (node.type === "MathExpression") return reduceMathExpression(node, precision); - if (node.type === "Calc") return reduce(node.value, precision); - - return node; -} - -function isEqual$1(left, right) { - return left.type === right.type && left.value === right.value; -} - -function isValueType(type) { - switch (type) { - case 'LengthValue': - case 'AngleValue': - case 'TimeValue': - case 'FrequencyValue': - case 'ResolutionValue': - case 'EmValue': - case 'ExValue': - case 'ChValue': - case 'RemValue': - case 'VhValue': - case 'VwValue': - case 'VminValue': - case 'VmaxValue': - case 'PercentageValue': - case 'Value': - return true; - } - return false; -} - -function convertMathExpression(node, precision) { - var nodes = (0, _convert2.default)(node.left, node.right, precision); - var left = reduce(nodes.left, precision); - var right = reduce(nodes.right, precision); - - if (left.type === "MathExpression" && right.type === "MathExpression") { - - if (left.operator === '/' && right.operator === '*' || left.operator === '-' && right.operator === '+' || left.operator === '*' && right.operator === '/' || left.operator === '+' && right.operator === '-') { - - if (isEqual$1(left.right, right.right)) nodes = (0, _convert2.default)(left.left, right.left, precision);else if (isEqual$1(left.right, right.left)) nodes = (0, _convert2.default)(left.left, right.right, precision); - - left = reduce(nodes.left, precision); - right = reduce(nodes.right, precision); - } - } - - node.left = left; - node.right = right; - return node; -} - -function flip(operator) { - return operator === '+' ? '-' : '+'; -} - -function flipValue(node) { - if (isValueType(node.type)) node.value = -node.value;else if (node.type == 'MathExpression') { - node.left = flipValue(node.left); - node.right = flipValue(node.right); - } - return node; -} - -function reduceAddSubExpression(node, precision) { - var _node = node, - left = _node.left, - right = _node.right, - op = _node.operator; - - - if (left.type === 'CssVariable' || right.type === 'CssVariable') return node; - - // something + 0 => something - // something - 0 => something - if (right.value === 0) return left; - - // 0 + something => something - if (left.value === 0 && op === "+") return right; - - // 0 - something => -something - if (left.value === 0 && op === "-") return flipValue(right); - - // value + value - // value - value - if (left.type === right.type && isValueType(left.type)) { - node = Object.assign({}, left); - if (op === "+") node.value = left.value + right.value;else node.value = left.value - right.value; - } - - // value (expr) - if (isValueType(left.type) && (right.operator === '+' || right.operator === '-') && right.type === 'MathExpression') { - // value + (value + something) => (value + value) + something - // value + (value - something) => (value + value) - something - // value - (value + something) => (value - value) - something - // value - (value - something) => (value - value) + something - if (left.type === right.left.type) { - node = Object.assign({}, node); - node.left = reduce({ - type: 'MathExpression', - operator: op, - left: left, - right: right.left - }, precision); - node.right = right.right; - node.operator = op === '-' ? flip(right.operator) : right.operator; - return reduce(node, precision); - } - // value + (something + value) => (value + value) + something - // value + (something - value) => (value - value) + something - // value - (something + value) => (value - value) - something - // value - (something - value) => (value + value) - something - else if (left.type === right.right.type) { - node = Object.assign({}, node); - node.left = reduce({ - type: 'MathExpression', - operator: op === '-' ? flip(right.operator) : right.operator, - left: left, - right: right.right - }, precision); - node.right = right.left; - return reduce(node, precision); - } - } - - // (expr) value - if (left.type === 'MathExpression' && (left.operator === '+' || left.operator === '-') && isValueType(right.type)) { - // (value + something) + value => (value + value) + something - // (value - something) + value => (value + value) - something - // (value + something) - value => (value - value) + something - // (value - something) - value => (value - value) - something - if (right.type === left.left.type) { - node = Object.assign({}, left); - node.left = reduce({ - type: 'MathExpression', - operator: op, - left: left.left, - right: right - }, precision); - return reduce(node, precision); - } - // (something + value) + value => something + (value + value) - // (something - value1) + value2 => something - (value2 - value1) - // (something + value) - value => something + (value - value) - // (something - value) - value => something - (value + value) - else if (right.type === left.right.type) { - node = Object.assign({}, left); - if (left.operator === '-') { - node.right = reduce({ - type: 'MathExpression', - operator: op === '-' ? '+' : '-', - left: right, - right: left.right - }, precision); - node.operator = op === '-' ? '-' : '+'; - } else { - node.right = reduce({ - type: 'MathExpression', - operator: op, - left: left.right, - right: right - }, precision); - } - if (node.right.value < 0) { - node.right.value *= -1; - node.operator = node.operator === '-' ? '+' : '-'; - } - return reduce(node, precision); - } - } - return node; -} - -function reduceDivisionExpression(node, precision) { - if (!isValueType(node.right.type)) return node; - - if (node.right.type !== 'Value') throw new Error("Cannot divide by \"" + node.right.unit + "\", number expected"); - - if (node.right.value === 0) throw new Error('Cannot divide by zero'); - - // (expr) / value - if (node.left.type === 'MathExpression') { - if (isValueType(node.left.left.type) && isValueType(node.left.right.type)) { - node.left.left.value /= node.right.value; - node.left.right.value /= node.right.value; - return reduce(node.left, precision); - } - return node; - } - // something / value - else if (isValueType(node.left.type)) { - node.left.value /= node.right.value; - return node.left; - } - return node; -} - -function reduceMultiplicationExpression(node) { - // (expr) * value - if (node.left.type === 'MathExpression' && node.right.type === 'Value') { - if (isValueType(node.left.left.type) && isValueType(node.left.right.type)) { - node.left.left.value *= node.right.value; - node.left.right.value *= node.right.value; - return node.left; - } - } - // something * value - else if (isValueType(node.left.type) && node.right.type === 'Value') { - node.left.value *= node.right.value; - return node.left; - } - // value * (expr) - else if (node.left.type === 'Value' && node.right.type === 'MathExpression') { - if (isValueType(node.right.left.type) && isValueType(node.right.right.type)) { - node.right.left.value *= node.left.value; - node.right.right.value *= node.left.value; - return node.right; - } - } - // value * something - else if (node.left.type === 'Value' && isValueType(node.right.type)) { - node.right.value *= node.left.value; - return node.right; - } - return node; -} - -function reduceMathExpression(node, precision) { - node = convertMathExpression(node, precision); - - switch (node.operator) { - case "+": - case "-": - return reduceAddSubExpression(node, precision); - case "/": - return reduceDivisionExpression(node, precision); - case "*": - return reduceMultiplicationExpression(node); - } - return node; -} - -reducer.default = reduce; - -var stringifier = {exports: {}}; - -(function (module, exports) { - - Object.defineProperty(exports, "__esModule", { - value: true - }); - - exports.default = function (calc, node, precision) { - var str = stringify(node, precision); - - if (node.type === "MathExpression") { - // if calc expression couldn't be resolved to a single value, re-wrap it as - // a calc() - str = calc + "(" + str + ")"; - } - return str; - }; - - var _reducer = reducer; - - var order = { - "*": 0, - "/": 0, - "+": 1, - "-": 1 - }; - - function round(value, prec) { - if (prec !== false) { - var precision = Math.pow(10, prec); - return Math.round(value * precision) / precision; - } - return value; - } - - function stringify(node, prec) { - switch (node.type) { - case "MathExpression": - { - var left = node.left, - right = node.right, - op = node.operator; - - var str = ""; - - if (left.type === 'MathExpression' && order[op] < order[left.operator]) str += "(" + stringify(left, prec) + ")";else str += stringify(left, prec); - - str += " " + node.operator + " "; - - if (right.type === 'MathExpression' && order[op] < order[right.operator]) { - str += "(" + stringify(right, prec) + ")"; - } else if (right.type === 'MathExpression' && op === "-" && ["+", "-"].includes(right.operator)) { - // fix #52 : a-(b+c) = a-b-c - right.operator = (0, _reducer.flip)(right.operator); - str += stringify(right, prec); - } else { - str += stringify(right, prec); - } - - return str; - } - case "Value": - return round(node.value, prec); - case 'CssVariable': - if (node.fallback) { - return "var(" + node.value + ", " + stringify(node.fallback, prec) + ")"; - } - return "var(" + node.value + ")"; - case 'Calc': - if (node.prefix) { - return "-" + node.prefix + "-calc(" + stringify(node.value, prec) + ")"; - } - return "calc(" + stringify(node.value, prec) + ")"; - default: - return round(node.value, prec) + node.unit; - } - } - - module.exports = exports["default"]; -} (stringifier, stringifier.exports)); - -(function (module, exports) { - - Object.defineProperty(exports, "__esModule", { - value: true - }); - - var _postcssValueParser = lib$1; - - var _postcssValueParser2 = _interopRequireDefault(_postcssValueParser); - - var _parser = parser; - - var _reducer = reducer; - - var _reducer2 = _interopRequireDefault(_reducer); - - var _stringifier = stringifier.exports; - - var _stringifier2 = _interopRequireDefault(_stringifier); - - function _interopRequireDefault(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - - // eslint-disable-line - var MATCH_CALC = /((?:\-[a-z]+\-)?calc)/; - - exports.default = function (value) { - var precision = arguments.length > 1 && arguments[1] !== undefined ? arguments[1] : 5; - - return (0, _postcssValueParser2.default)(value).walk(function (node) { - // skip anything which isn't a calc() function - if (node.type !== 'function' || !MATCH_CALC.test(node.value)) return; - - // stringify calc expression and produce an AST - var contents = _postcssValueParser2.default.stringify(node.nodes); - - // skip constant() and env() - if (contents.indexOf('constant') >= 0 || contents.indexOf('env') >= 0) return; - - var ast = _parser.parser.parse(contents); - - // reduce AST to its simplest form, that is, either to a single value - // or a simplified calc expression - var reducedAst = (0, _reducer2.default)(ast, precision); - - // stringify AST and write it back - node.type = 'word'; - node.value = (0, _stringifier2.default)(node.value, reducedAst, precision); - }, true).toString(); - }; - - module.exports = exports['default']; -} (dist, dist.exports)); - -var reduceCSSCalc = /*@__PURE__*/getDefaultExportFromCjs(dist.exports); - -function ownKeys$s(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$s(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$s(Object(source), true).forEach(function (key) { _defineProperty$s(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$s(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$s(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _toConsumableArray$8(arr) { return _arrayWithoutHoles$8(arr) || _iterableToArray$8(arr) || _unsupportedIterableToArray$d(arr) || _nonIterableSpread$8(); } - -function _nonIterableSpread$8() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$d(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$d(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$d(o, minLen); } - -function _iterableToArray$8(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$8(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$d(arr); } - -function _arrayLikeToArray$d(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } -var stringCache = { - widthCache: {}, - cacheCount: 0 -}; -var MAX_CACHE_NUM = 2000; -var SPAN_STYLE = { - position: 'absolute', - top: '-20000px', - left: 0, - padding: 0, - margin: 0, - border: 'none', - whiteSpace: 'pre' -}; -var STYLE_LIST = ['minWidth', 'maxWidth', 'width', 'minHeight', 'maxHeight', 'height', 'top', 'left', 'fontSize', 'lineHeight', 'padding', 'margin', 'paddingLeft', 'paddingRight', 'paddingTop', 'paddingBottom', 'marginLeft', 'marginRight', 'marginTop', 'marginBottom']; -var MEASUREMENT_SPAN_ID = 'recharts_measurement_span'; - -function autoCompleteStyle(name, value) { - if (STYLE_LIST.indexOf(name) >= 0 && value === +value) { - return "".concat(value, "px"); - } - - return value; -} - -function camelToMiddleLine(text) { - var strs = text.split(''); - var formatStrs = strs.reduce(function (result, entry) { - if (entry === entry.toUpperCase()) { - return [].concat(_toConsumableArray$8(result), ['-', entry.toLowerCase()]); - } - - return [].concat(_toConsumableArray$8(result), [entry]); - }, []); - return formatStrs.join(''); -} - -var getStyleString = function getStyleString(style) { - return Object.keys(style).reduce(function (result, s) { - return "".concat(result).concat(camelToMiddleLine(s), ":").concat(autoCompleteStyle(s, style[s]), ";"); - }, ''); -}; -var getStringSize = function getStringSize(text) { - var style = arguments.length > 1 && arguments[1] !== undefined ? arguments[1] : {}; - - if (text === undefined || text === null || Global.isSsr) { - return { - width: 0, - height: 0 - }; - } - - var str = "".concat(text); - var styleString = getStyleString(style); - var cacheKey = "".concat(str, "-").concat(styleString); - - if (stringCache.widthCache[cacheKey]) { - return stringCache.widthCache[cacheKey]; - } - - try { - var measurementSpan = document.getElementById(MEASUREMENT_SPAN_ID); - - if (!measurementSpan) { - measurementSpan = document.createElement('span'); - measurementSpan.setAttribute('id', MEASUREMENT_SPAN_ID); - measurementSpan.setAttribute('aria-hidden', 'true'); - document.body.appendChild(measurementSpan); - } // Need to use CSS Object Model (CSSOM) to be able to comply with Content Security Policy (CSP) - // https://en.wikipedia.org/wiki/Content_Security_Policy - - - var measurementSpanStyle = _objectSpread$s(_objectSpread$s({}, SPAN_STYLE), style); - - Object.keys(measurementSpanStyle).map(function (styleKey) { - measurementSpan.style[styleKey] = measurementSpanStyle[styleKey]; - return styleKey; - }); - measurementSpan.textContent = str; - var rect = measurementSpan.getBoundingClientRect(); - var result = { - width: rect.width, - height: rect.height - }; - stringCache.widthCache[cacheKey] = result; - - if (++stringCache.cacheCount > MAX_CACHE_NUM) { - stringCache.cacheCount = 0; - stringCache.widthCache = {}; - } - - return result; - } catch (e) { - return { - width: 0, - height: 0 - }; - } -}; -var getOffset = function getOffset(el) { - var html = el.ownerDocument.documentElement; - var box = { - top: 0, - left: 0 - }; // If we don't have gBCR, just use 0,0 rather than error - // BlackBerry 5, iOS 3 (original iPhone) - - if (typeof el.getBoundingClientRect !== 'undefined') { - box = el.getBoundingClientRect(); - } - - return { - top: box.top + window.pageYOffset - html.clientTop, - left: box.left + window.pageXOffset - html.clientLeft - }; -}; -/** - * Calculate coordinate of cursor in chart - * @param {Object} event Event object - * @param {Object} offset The offset of main part in the svg element - * @return {Object} {chartX, chartY} - */ - -var calculateChartCoordinate = function calculateChartCoordinate(event, offset) { - return { - chartX: Math.round(event.pageX - offset.left), - chartY: Math.round(event.pageY - offset.top) - }; -}; - -function _typeof$p(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$p = function _typeof(obj) { return typeof obj; }; } else { _typeof$p = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$p(obj); } - -function _extends$t() { _extends$t = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$t.apply(this, arguments); } - -function _objectWithoutProperties$f(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$f(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$f(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _classCallCheck$p(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$p(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$p(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$p(Constructor.prototype, protoProps); if (staticProps) _defineProperties$p(Constructor, staticProps); return Constructor; } - -function _inherits$o(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$o(subClass, superClass); } - -function _setPrototypeOf$o(o, p) { _setPrototypeOf$o = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$o(o, p); } - -function _createSuper$o(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$o(); return function _createSuperInternal() { var Super = _getPrototypeOf$o(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$o(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$o(this, result); }; } - -function _possibleConstructorReturn$o(self, call) { if (call && (_typeof$p(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$o(self); } - -function _assertThisInitialized$o(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$o() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$o(o) { _getPrototypeOf$o = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$o(o); } - -function _slicedToArray$6(arr, i) { return _arrayWithHoles$6(arr) || _iterableToArrayLimit$6(arr, i) || _unsupportedIterableToArray$c(arr, i) || _nonIterableRest$6(); } - -function _nonIterableRest$6() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$c(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$c(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$c(o, minLen); } - -function _arrayLikeToArray$c(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit$6(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$6(arr) { if (Array.isArray(arr)) return arr; } - -function ownKeys$r(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$r(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$r(Object(source), true).forEach(function (key) { _defineProperty$r(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$r(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$r(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } -var BREAKING_SPACES = /[ \f\n\r\t\v\u2028\u2029]+/; - -var calculateWordWidths = function calculateWordWidths(props) { - try { - var words = []; - - if (!isNil_1(props.children)) { - if (props.breakAll) { - words = props.children.toString().split(''); - } else { - words = props.children.toString().split(BREAKING_SPACES); - } - } - - var wordsWithComputedWidth = words.map(function (word) { - return { - word: word, - width: getStringSize(word, props.style).width - }; - }); - var spaceWidth = props.breakAll ? 0 : getStringSize("\xA0", props.style).width; - return { - wordsWithComputedWidth: wordsWithComputedWidth, - spaceWidth: spaceWidth - }; - } catch (e) { - return null; - } -}; - -var calculateWordsByLines = function calculateWordsByLines(props, initialWordsWithComputedWith, spaceWidth, lineWidth, scaleToFit) { - var shouldLimitLines = isNumber(props.maxLines); - var text = props.children; - - var calculate = function calculate() { - var words = arguments.length > 0 && arguments[0] !== undefined ? arguments[0] : []; - return words.reduce(function (result, _ref) { - var word = _ref.word, - width = _ref.width; - var currentLine = result[result.length - 1]; - - if (currentLine && (lineWidth == null || scaleToFit || currentLine.width + width + spaceWidth < lineWidth)) { - // Word can be added to an existing line - currentLine.words.push(word); - currentLine.width += width + spaceWidth; - } else { - // Add first word to line or word is too long to scaleToFit on existing line - var newLine = { - words: [word], - width: width - }; - result.push(newLine); - } - - return result; - }, []); - }; - - var originalResult = calculate(initialWordsWithComputedWith); - - var findLongestLine = function findLongestLine(words) { - return words.reduce(function (a, b) { - return a.width > b.width ? a : b; - }); - }; - - if (!shouldLimitLines) { - return originalResult; - } - - var suffix = '…'; - - var checkOverflow = function checkOverflow(index) { - var tempText = text.slice(0, index); - var words = calculateWordWidths(_objectSpread$r(_objectSpread$r({}, props), {}, { - children: tempText + suffix - })).wordsWithComputedWidth; - var result = calculate(words); - var doesOverflow = result.length > props.maxLines || findLongestLine(result).width > lineWidth; - return [doesOverflow, result]; - }; - - var start = 0; - var end = text.length - 1; - var iterations = 0; - var trimmedResult; - - while (start <= end && iterations <= text.length - 1) { - var middle = Math.floor((start + end) / 2); - var prev = middle - 1; - - var _checkOverflow = checkOverflow(prev), - _checkOverflow2 = _slicedToArray$6(_checkOverflow, 2), - doesPrevOverflow = _checkOverflow2[0], - result = _checkOverflow2[1]; - - var _checkOverflow3 = checkOverflow(middle), - _checkOverflow4 = _slicedToArray$6(_checkOverflow3, 1), - doesMiddleOverflow = _checkOverflow4[0]; - - if (!doesPrevOverflow && !doesMiddleOverflow) { - start = middle + 1; - } - - if (doesPrevOverflow && doesMiddleOverflow) { - end = middle - 1; - } - - if (!doesPrevOverflow && doesMiddleOverflow) { - trimmedResult = result; - break; - } - - iterations++; - } // Fallback to originalResult (result without trimming) if we cannot find the - // where to trim. This should not happen :tm: - - - return trimmedResult || originalResult; -}; - -var getWordsWithoutCalculate = function getWordsWithoutCalculate(children) { - var words = !isNil_1(children) ? children.toString().split(BREAKING_SPACES) : []; - return [{ - words: words - }]; -}; - -var getWordsByLines = function getWordsByLines(props, needCalculate) { - // Only perform calculations if using features that require them (multiline, scaleToFit) - if ((props.width || props.scaleToFit) && !Global.isSsr) { - var wordsWithComputedWidth, spaceWidth; - - if (needCalculate) { - var wordWidths = calculateWordWidths(props); - - if (wordWidths) { - var wcw = wordWidths.wordsWithComputedWidth, - sw = wordWidths.spaceWidth; - wordsWithComputedWidth = wcw; - spaceWidth = sw; - } else { - return getWordsWithoutCalculate(props.children); - } - - return calculateWordsByLines(props, wordsWithComputedWidth, spaceWidth, props.width, props.scaleToFit); - } - } - - return getWordsWithoutCalculate(props.children); -}; - -var Text = /*#__PURE__*/function (_Component) { - _inherits$o(Text, _Component); - - var _super = _createSuper$o(Text); - - function Text() { - var _this; - - _classCallCheck$p(this, Text); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = {}; - return _this; - } - - _createClass$p(Text, [{ - key: "render", - value: function render() { - var _this$props = this.props, - dx = _this$props.dx, - dy = _this$props.dy, - textAnchor = _this$props.textAnchor, - verticalAnchor = _this$props.verticalAnchor, - scaleToFit = _this$props.scaleToFit, - angle = _this$props.angle, - lineHeight = _this$props.lineHeight, - capHeight = _this$props.capHeight, - className = _this$props.className, - breakAll = _this$props.breakAll, - textProps = _objectWithoutProperties$f(_this$props, ["dx", "dy", "textAnchor", "verticalAnchor", "scaleToFit", "angle", "lineHeight", "capHeight", "className", "breakAll"]); - - var wordsByLines = this.state.wordsByLines; - - if (!isNumOrStr(textProps.x) || !isNumOrStr(textProps.y)) { - return null; - } - - var x = textProps.x + (isNumber(dx) ? dx : 0); - var y = textProps.y + (isNumber(dy) ? dy : 0); - var startDy; - - switch (verticalAnchor) { - case 'start': - startDy = reduceCSSCalc("calc(".concat(capHeight, ")")); - break; - - case 'middle': - startDy = reduceCSSCalc("calc(".concat((wordsByLines.length - 1) / 2, " * -").concat(lineHeight, " + (").concat(capHeight, " / 2))")); - break; - - default: - startDy = reduceCSSCalc("calc(".concat(wordsByLines.length - 1, " * -").concat(lineHeight, ")")); - break; - } - - var transforms = []; - - if (scaleToFit) { - var lineWidth = wordsByLines[0].width; - var width = this.props.width; - transforms.push("scale(".concat((isNumber(width) ? width / lineWidth : 1) / lineWidth, ")")); - } - - if (angle) { - transforms.push("rotate(".concat(angle, ", ").concat(x, ", ").concat(y, ")")); - } - - if (transforms.length) { - textProps.transform = transforms.join(' '); - } - - return /*#__PURE__*/React__default.createElement("text", _extends$t({}, filterProps(textProps, true), { - x: x, - y: y, - className: classNames('recharts-text', className), - textAnchor: textAnchor - }), wordsByLines.map(function (line, index) { - return ( - /*#__PURE__*/ - // eslint-disable-next-line react/no-array-index-key - React__default.createElement("tspan", { - x: x, - dy: index === 0 ? startDy : lineHeight, - key: index - }, line.words.join(breakAll ? '' : ' ')) - ); - })); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.width !== prevState.prevWidth || nextProps.scaleToFit !== prevState.prevScaleToFit || nextProps.children !== prevState.prevChildren || nextProps.style !== prevState.prevStyle || nextProps.breakAll !== prevState.prevBreakAll) { - var needCalculate = nextProps.children !== prevState.prevChildren || nextProps.style !== prevState.prevStyle || nextProps.breakAll !== prevState.prevBreakAll; - return { - prevWidth: nextProps.width, - prevScaleToFit: nextProps.scaleToFit, - prevChildren: nextProps.children, - prevStyle: nextProps.style, - wordsByLines: getWordsByLines(nextProps, needCalculate) - }; - } - - return null; - } - }]); - - return Text; -}(Component); -Text.defaultProps = { - x: 0, - y: 0, - lineHeight: '1em', - capHeight: '0.71em', - // Magic number from d3 - scaleToFit: false, - textAnchor: 'start', - verticalAnchor: 'end' // Maintain compat with existing charts / default SVG behavior - -}; - -var reactIs = {exports: {}}; - -var reactIs_production_min = {}; - -/** @license React v16.13.1 - * react-is.production.min.js - * - * Copyright (c) Facebook, Inc. and its affiliates. - * - * This source code is licensed under the MIT license found in the - * LICENSE file in the root directory of this source tree. - */ -var b="function"===typeof Symbol&&Symbol.for,c=b?Symbol.for("react.element"):60103,d=b?Symbol.for("react.portal"):60106,e=b?Symbol.for("react.fragment"):60107,f=b?Symbol.for("react.strict_mode"):60108,g=b?Symbol.for("react.profiler"):60114,h=b?Symbol.for("react.provider"):60109,k=b?Symbol.for("react.context"):60110,l=b?Symbol.for("react.async_mode"):60111,m=b?Symbol.for("react.concurrent_mode"):60111,n=b?Symbol.for("react.forward_ref"):60112,p=b?Symbol.for("react.suspense"):60113,q=b? -Symbol.for("react.suspense_list"):60120,r=b?Symbol.for("react.memo"):60115,t=b?Symbol.for("react.lazy"):60116,v=b?Symbol.for("react.block"):60121,w=b?Symbol.for("react.fundamental"):60117,x=b?Symbol.for("react.responder"):60118,y=b?Symbol.for("react.scope"):60119; -function z(a){if("object"===typeof a&&null!==a){var u=a.$$typeof;switch(u){case c:switch(a=a.type,a){case l:case m:case e:case g:case f:case p:return a;default:switch(a=a&&a.$$typeof,a){case k:case n:case t:case r:case h:return a;default:return u}}case d:return u}}}function A$1(a){return z(a)===m}reactIs_production_min.AsyncMode=l;reactIs_production_min.ConcurrentMode=m;reactIs_production_min.ContextConsumer=k;reactIs_production_min.ContextProvider=h;reactIs_production_min.Element=c;reactIs_production_min.ForwardRef=n;reactIs_production_min.Fragment=e;reactIs_production_min.Lazy=t;reactIs_production_min.Memo=r;reactIs_production_min.Portal=d; -reactIs_production_min.Profiler=g;reactIs_production_min.StrictMode=f;reactIs_production_min.Suspense=p;reactIs_production_min.isAsyncMode=function(a){return A$1(a)||z(a)===l};reactIs_production_min.isConcurrentMode=A$1;reactIs_production_min.isContextConsumer=function(a){return z(a)===k};reactIs_production_min.isContextProvider=function(a){return z(a)===h};reactIs_production_min.isElement=function(a){return "object"===typeof a&&null!==a&&a.$$typeof===c};reactIs_production_min.isForwardRef=function(a){return z(a)===n};reactIs_production_min.isFragment=function(a){return z(a)===e};reactIs_production_min.isLazy=function(a){return z(a)===t}; -reactIs_production_min.isMemo=function(a){return z(a)===r};reactIs_production_min.isPortal=function(a){return z(a)===d};reactIs_production_min.isProfiler=function(a){return z(a)===g};reactIs_production_min.isStrictMode=function(a){return z(a)===f};reactIs_production_min.isSuspense=function(a){return z(a)===p}; -reactIs_production_min.isValidElementType=function(a){return "string"===typeof a||"function"===typeof a||a===e||a===m||a===g||a===f||a===p||a===q||"object"===typeof a&&null!==a&&(a.$$typeof===t||a.$$typeof===r||a.$$typeof===h||a.$$typeof===k||a.$$typeof===n||a.$$typeof===w||a.$$typeof===x||a.$$typeof===y||a.$$typeof===v)};reactIs_production_min.typeOf=z; - -(function (module) { - - { - module.exports = reactIs_production_min; - } -} (reactIs)); - -function shallowEqual(a, b) { - /* eslint-disable no-restricted-syntax */ - for (var key in a) { - if ({}.hasOwnProperty.call(a, key) && (!{}.hasOwnProperty.call(b, key) || a[key] !== b[key])) { - return false; - } - } - - for (var _key in b) { - if ({}.hasOwnProperty.call(b, _key) && !{}.hasOwnProperty.call(a, _key)) { - return false; - } - } - - return true; -} - -function _objectWithoutProperties$e(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$e(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$e(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } -var REACT_BROWSER_EVENT_MAP = { - click: 'onClick', - mousedown: 'onMouseDown', - mouseup: 'onMouseUp', - mouseover: 'onMouseOver', - mousemove: 'onMouseMove', - mouseout: 'onMouseOut', - mouseenter: 'onMouseEnter', - mouseleave: 'onMouseLeave', - touchcancel: 'onTouchCancel', - touchend: 'onTouchEnd', - touchmove: 'onTouchMove', - touchstart: 'onTouchStart' -}; -/** - * Get the display name of a component - * @param {Object} Comp Specified Component - * @return {String} Display name of Component - */ - -var getDisplayName = function getDisplayName(Comp) { - if (typeof Comp === 'string') { - return Comp; - } - - if (!Comp) { - return ''; - } - - return Comp.displayName || Comp.name || 'Component'; -}; // `toArray` gets called multiple times during the render -// so we can memoize last invocation (since reference to `children` is the same) - -var lastChildren = null; -var lastResult = null; -var toArray = function toArray(children) { - if (children === lastChildren && isArray_1(lastResult)) { - return lastResult; - } - - var result = []; - Children.forEach(children, function (child) { - if (isNil_1(child)) return; - - if (reactIs.exports.isFragment(child)) { - result = result.concat(toArray(child.props.children)); - } else { - result.push(child); - } - }); - lastResult = result; - lastChildren = children; - return result; -}; -/* - * Find and return all matched children by type. `type` can be a React element class or - * string - */ - -var findAllByType = function findAllByType(children, type) { - var result = []; - var types = []; - - if (isArray_1(type)) { - types = type.map(function (t) { - return getDisplayName(t); - }); - } else { - types = [getDisplayName(type)]; - } - - toArray(children).forEach(function (child) { - var childType = get_1(child, 'type.displayName') || get_1(child, 'type.name'); - - if (types.indexOf(childType) !== -1) { - result.push(child); - } - }); - return result; -}; -/* - * Return the first matched child by type, return null otherwise. - * `type` can be a React element class or string. - */ - -var findChildByType = function findChildByType(children, type) { - var result = findAllByType(children, type); - return result && result[0]; -}; -/** - * validate the width and height props of a chart element - * @param {Object} el A chart element - * @return {Boolean} true If the props width and height are number, and greater than 0 - */ - -var validateWidthHeight = function validateWidthHeight(el) { - if (!el || !el.props) { - return false; - } - - var _el$props = el.props, - width = _el$props.width, - height = _el$props.height; - - if (!isNumber(width) || width <= 0 || !isNumber(height) || height <= 0) { - return false; - } - - return true; -}; -var SVG_TAGS = ['a', 'altGlyph', 'altGlyphDef', 'altGlyphItem', 'animate', 'animateColor', 'animateMotion', 'animateTransform', 'circle', 'clipPath', 'color-profile', 'cursor', 'defs', 'desc', 'ellipse', 'feBlend', 'feColormatrix', 'feComponentTransfer', 'feComposite', 'feConvolveMatrix', 'feDiffuseLighting', 'feDisplacementMap', 'feDistantLight', 'feFlood', 'feFuncA', 'feFuncB', 'feFuncG', 'feFuncR', 'feGaussianBlur', 'feImage', 'feMerge', 'feMergeNode', 'feMorphology', 'feOffset', 'fePointLight', 'feSpecularLighting', 'feSpotLight', 'feTile', 'feTurbulence', 'filter', 'font', 'font-face', 'font-face-format', 'font-face-name', 'font-face-url', 'foreignObject', 'g', 'glyph', 'glyphRef', 'hkern', 'image', 'line', 'lineGradient', 'marker', 'mask', 'metadata', 'missing-glyph', 'mpath', 'path', 'pattern', 'polygon', 'polyline', 'radialGradient', 'rect', 'script', 'set', 'stop', 'style', 'svg', 'switch', 'symbol', 'text', 'textPath', 'title', 'tref', 'tspan', 'use', 'view', 'vkern']; - -var isSvgElement = function isSvgElement(child) { - return child && child.type && isString_1(child.type) && SVG_TAGS.indexOf(child.type) >= 0; -}; -/** - * Filter all the svg elements of children - * @param {Array} children The children of a react element - * @return {Array} All the svg elements - */ - - -var filterSvgElements = function filterSvgElements(children) { - var svgElements = []; - toArray(children).forEach(function (entry) { - if (isSvgElement(entry)) { - svgElements.push(entry); - } - }); - return svgElements; -}; -/** - * Wether props of children changed - * @param {Object} nextChildren The latest children - * @param {Object} prevChildren The prev children - * @return {Boolean} equal or not - */ - -var isChildrenEqual = function isChildrenEqual(nextChildren, prevChildren) { - if (nextChildren === prevChildren) { - return true; - } - - var count = Children.count(nextChildren); - - if (count !== Children.count(prevChildren)) { - return false; - } - - if (count === 0) { - return true; - } - - if (count === 1) { - // eslint-disable-next-line no-use-before-define,@typescript-eslint/no-use-before-define - return isSingleChildEqual(isArray_1(nextChildren) ? nextChildren[0] : nextChildren, isArray_1(prevChildren) ? prevChildren[0] : prevChildren); - } - - for (var i = 0; i < count; i++) { - var nextChild = nextChildren[i]; - var prevChild = prevChildren[i]; - - if (isArray_1(nextChild) || isArray_1(prevChild)) { - if (!isChildrenEqual(nextChild, prevChild)) { - return false; - } // eslint-disable-next-line no-use-before-define,@typescript-eslint/no-use-before-define - - } else if (!isSingleChildEqual(nextChild, prevChild)) { - return false; - } - } - - return true; -}; -var isSingleChildEqual = function isSingleChildEqual(nextChild, prevChild) { - if (isNil_1(nextChild) && isNil_1(prevChild)) { - return true; - } - - if (!isNil_1(nextChild) && !isNil_1(prevChild)) { - var _ref = nextChild.props || {}, - nextChildren = _ref.children, - nextProps = _objectWithoutProperties$e(_ref, ["children"]); - - var _ref2 = prevChild.props || {}, - prevChildren = _ref2.children, - prevProps = _objectWithoutProperties$e(_ref2, ["children"]); - - if (nextChildren && prevChildren) { - // eslint-disable-next-line no-use-before-define - return shallowEqual(nextProps, prevProps) && isChildrenEqual(nextChildren, prevChildren); - } - - if (!nextChildren && !prevChildren) { - return shallowEqual(nextProps, prevProps); - } - - return false; - } - - return false; -}; -var renderByOrder = function renderByOrder(children, renderMap) { - var elements = []; - var record = {}; - toArray(children).forEach(function (child, index) { - if (isSvgElement(child)) { - elements.push(child); - } else if (child) { - var displayName = getDisplayName(child.type); - - var _ref3 = renderMap[displayName] || {}, - handler = _ref3.handler, - once = _ref3.once; - - if (handler && (!once || !record[displayName])) { - var results = handler(child, displayName, index); - elements.push(results); - record[displayName] = true; - } - } - }); - return elements; -}; -var getReactEventByType = function getReactEventByType(e) { - var type = e && e.type; - - if (type && REACT_BROWSER_EVENT_MAP[type]) { - return REACT_BROWSER_EVENT_MAP[type]; - } - - return null; -}; -var parseChildIndex = function parseChildIndex(child, children) { - return toArray(children).indexOf(child); -}; - -var baseIsEqual = _baseIsEqual; - -/** - * Performs a deep comparison between two values to determine if they are - * equivalent. - * - * **Note:** This method supports comparing arrays, array buffers, booleans, - * date objects, error objects, maps, numbers, `Object` objects, regexes, - * sets, strings, symbols, and typed arrays. `Object` objects are compared - * by their own, not inherited, enumerable properties. Functions and DOM - * nodes are compared by strict equality, i.e. `===`. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to compare. - * @param {*} other The other value to compare. - * @returns {boolean} Returns `true` if the values are equivalent, else `false`. - * @example - * - * var object = { 'a': 1 }; - * var other = { 'a': 1 }; - * - * _.isEqual(object, other); - * // => true - * - * object === other; - * // => false - */ -function isEqual(value, other) { - return baseIsEqual(value, other); -} - -var isEqual_1 = isEqual; - -var isSymbol = isSymbol_1$1; - -/** - * The base implementation of methods like `_.max` and `_.min` which accepts a - * `comparator` to determine the extremum value. - * - * @private - * @param {Array} array The array to iterate over. - * @param {Function} iteratee The iteratee invoked per iteration. - * @param {Function} comparator The comparator used to compare values. - * @returns {*} Returns the extremum value. - */ -function baseExtremum$4(array, iteratee, comparator) { - var index = -1, - length = array.length; - - while (++index < length) { - var value = array[index], - current = iteratee(value); - - if (current != null && (computed === undefined - ? (current === current && !isSymbol(current)) - : comparator(current, computed) - )) { - var computed = current, - result = value; - } - } - return result; -} - -var _baseExtremum = baseExtremum$4; - -/** - * The base implementation of `_.gt` which doesn't coerce arguments. - * - * @private - * @param {*} value The value to compare. - * @param {*} other The other value to compare. - * @returns {boolean} Returns `true` if `value` is greater than `other`, - * else `false`. - */ - -function baseGt$2(value, other) { - return value > other; -} - -var _baseGt = baseGt$2; - -var baseExtremum$3 = _baseExtremum, - baseGt$1 = _baseGt, - identity$4 = identity_1; - -/** - * Computes the maximum value of `array`. If `array` is empty or falsey, - * `undefined` is returned. - * - * @static - * @since 0.1.0 - * @memberOf _ - * @category Math - * @param {Array} array The array to iterate over. - * @returns {*} Returns the maximum value. - * @example - * - * _.max([4, 2, 8, 6]); - * // => 8 - * - * _.max([]); - * // => undefined - */ -function max(array) { - return (array && array.length) - ? baseExtremum$3(array, identity$4, baseGt$1) - : undefined; -} - -var max_1 = max; - -/** - * The base implementation of `_.lt` which doesn't coerce arguments. - * - * @private - * @param {*} value The value to compare. - * @param {*} other The other value to compare. - * @returns {boolean} Returns `true` if `value` is less than `other`, - * else `false`. - */ - -function baseLt$2(value, other) { - return value < other; -} - -var _baseLt = baseLt$2; - -var baseExtremum$2 = _baseExtremum, - baseLt$1 = _baseLt, - identity$3 = identity_1; - -/** - * Computes the minimum value of `array`. If `array` is empty or falsey, - * `undefined` is returned. - * - * @static - * @since 0.1.0 - * @memberOf _ - * @category Math - * @param {Array} array The array to iterate over. - * @returns {*} Returns the minimum value. - * @example - * - * _.min([4, 2, 8, 6]); - * // => 2 - * - * _.min([]); - * // => undefined - */ -function min(array) { - return (array && array.length) - ? baseExtremum$2(array, identity$3, baseLt$1) - : undefined; -} - -var min_1 = min; - -var arrayMap$1 = _arrayMap, - baseIteratee$8 = _baseIteratee, - baseMap = _baseMap, - isArray$3 = isArray_1; - -/** - * Creates an array of values by running each element in `collection` thru - * `iteratee`. The iteratee is invoked with three arguments: - * (value, index|key, collection). - * - * Many lodash methods are guarded to work as iteratees for methods like - * `_.every`, `_.filter`, `_.map`, `_.mapValues`, `_.reject`, and `_.some`. - * - * The guarded methods are: - * `ary`, `chunk`, `curry`, `curryRight`, `drop`, `dropRight`, `every`, - * `fill`, `invert`, `parseInt`, `random`, `range`, `rangeRight`, `repeat`, - * `sampleSize`, `slice`, `some`, `sortBy`, `split`, `take`, `takeRight`, - * `template`, `trim`, `trimEnd`, `trimStart`, and `words` - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Collection - * @param {Array|Object} collection The collection to iterate over. - * @param {Function} [iteratee=_.identity] The function invoked per iteration. - * @returns {Array} Returns the new mapped array. - * @example - * - * function square(n) { - * return n * n; - * } - * - * _.map([4, 8], square); - * // => [16, 64] - * - * _.map({ 'a': 4, 'b': 8 }, square); - * // => [16, 64] (iteration order is not guaranteed) - * - * var users = [ - * { 'user': 'barney' }, - * { 'user': 'fred' } - * ]; - * - * // The `_.property` iteratee shorthand. - * _.map(users, 'user'); - * // => ['barney', 'fred'] - */ -function map$2(collection, iteratee) { - var func = isArray$3(collection) ? arrayMap$1 : baseMap; - return func(collection, baseIteratee$8(iteratee)); -} - -var map_1 = map$2; - -var baseFlatten$1 = _baseFlatten, - map$1 = map_1; - -/** - * Creates a flattened array of values by running each element in `collection` - * thru `iteratee` and flattening the mapped results. The iteratee is invoked - * with three arguments: (value, index|key, collection). - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Collection - * @param {Array|Object} collection The collection to iterate over. - * @param {Function} [iteratee=_.identity] The function invoked per iteration. - * @returns {Array} Returns the new flattened array. - * @example - * - * function duplicate(n) { - * return [n, n]; - * } - * - * _.flatMap([1, 2], duplicate); - * // => [1, 1, 2, 2] - */ -function flatMap(collection, iteratee) { - return baseFlatten$1(map$1(collection, iteratee), 1); -} - -var flatMap_1 = flatMap; - -var lib = {}; - -var getNiceTickValues$1 = {}; - -var decimal = {exports: {}}; - -/*! decimal.js-light v2.5.1 https://github.com/MikeMcl/decimal.js-light/LICENCE */ - -(function (module) { -(function (globalScope) { - - - /* - * decimal.js-light v2.5.1 - * An arbitrary-precision Decimal type for JavaScript. - * https://github.com/MikeMcl/decimal.js-light - * Copyright (c) 2020 Michael Mclaughlin - * MIT Expat Licence - */ - - - // ----------------------------------- EDITABLE DEFAULTS ------------------------------------ // - - - // The limit on the value of `precision`, and on the value of the first argument to - // `toDecimalPlaces`, `toExponential`, `toFixed`, `toPrecision` and `toSignificantDigits`. - var MAX_DIGITS = 1e9, // 0 to 1e9 - - - // The initial configuration properties of the Decimal constructor. - Decimal = { - - // These values must be integers within the stated ranges (inclusive). - // Most of these values can be changed during run-time using `Decimal.config`. - - // The maximum number of significant digits of the result of a calculation or base conversion. - // E.g. `Decimal.config({ precision: 20 });` - precision: 20, // 1 to MAX_DIGITS - - // The rounding mode used by default by `toInteger`, `toDecimalPlaces`, `toExponential`, - // `toFixed`, `toPrecision` and `toSignificantDigits`. - // - // ROUND_UP 0 Away from zero. - // ROUND_DOWN 1 Towards zero. - // ROUND_CEIL 2 Towards +Infinity. - // ROUND_FLOOR 3 Towards -Infinity. - // ROUND_HALF_UP 4 Towards nearest neighbour. If equidistant, up. - // ROUND_HALF_DOWN 5 Towards nearest neighbour. If equidistant, down. - // ROUND_HALF_EVEN 6 Towards nearest neighbour. If equidistant, towards even neighbour. - // ROUND_HALF_CEIL 7 Towards nearest neighbour. If equidistant, towards +Infinity. - // ROUND_HALF_FLOOR 8 Towards nearest neighbour. If equidistant, towards -Infinity. - // - // E.g. - // `Decimal.rounding = 4;` - // `Decimal.rounding = Decimal.ROUND_HALF_UP;` - rounding: 4, // 0 to 8 - - // The exponent value at and beneath which `toString` returns exponential notation. - // JavaScript numbers: -7 - toExpNeg: -7, // 0 to -MAX_E - - // The exponent value at and above which `toString` returns exponential notation. - // JavaScript numbers: 21 - toExpPos: 21, // 0 to MAX_E - - // The natural logarithm of 10. - // 115 digits - LN10: '2.302585092994045684017991454684364207601101488628772976033327900967572609677352480235997205089598298341967784042286' - }, - - - // ----------------------------------- END OF EDITABLE DEFAULTS ------------------------------- // - - - external = true, - - decimalError = '[DecimalError] ', - invalidArgument = decimalError + 'Invalid argument: ', - exponentOutOfRange = decimalError + 'Exponent out of range: ', - - mathfloor = Math.floor, - mathpow = Math.pow, - - isDecimal = /^(\d+(\.\d*)?|\.\d+)(e[+-]?\d+)?$/i, - - ONE, - BASE = 1e7, - LOG_BASE = 7, - MAX_SAFE_INTEGER = 9007199254740991, - MAX_E = mathfloor(MAX_SAFE_INTEGER / LOG_BASE), // 1286742750677284 - - // Decimal.prototype object - P = {}; - - - // Decimal prototype methods - - - /* - * absoluteValue abs - * comparedTo cmp - * decimalPlaces dp - * dividedBy div - * dividedToIntegerBy idiv - * equals eq - * exponent - * greaterThan gt - * greaterThanOrEqualTo gte - * isInteger isint - * isNegative isneg - * isPositive ispos - * isZero - * lessThan lt - * lessThanOrEqualTo lte - * logarithm log - * minus sub - * modulo mod - * naturalExponential exp - * naturalLogarithm ln - * negated neg - * plus add - * precision sd - * squareRoot sqrt - * times mul - * toDecimalPlaces todp - * toExponential - * toFixed - * toInteger toint - * toNumber - * toPower pow - * toPrecision - * toSignificantDigits tosd - * toString - * valueOf val - */ - - - /* - * Return a new Decimal whose value is the absolute value of this Decimal. - * - */ - P.absoluteValue = P.abs = function () { - var x = new this.constructor(this); - if (x.s) x.s = 1; - return x; - }; - - - /* - * Return - * 1 if the value of this Decimal is greater than the value of `y`, - * -1 if the value of this Decimal is less than the value of `y`, - * 0 if they have the same value - * - */ - P.comparedTo = P.cmp = function (y) { - var i, j, xdL, ydL, - x = this; - - y = new x.constructor(y); - - // Signs differ? - if (x.s !== y.s) return x.s || -y.s; - - // Compare exponents. - if (x.e !== y.e) return x.e > y.e ^ x.s < 0 ? 1 : -1; - - xdL = x.d.length; - ydL = y.d.length; - - // Compare digit by digit. - for (i = 0, j = xdL < ydL ? xdL : ydL; i < j; ++i) { - if (x.d[i] !== y.d[i]) return x.d[i] > y.d[i] ^ x.s < 0 ? 1 : -1; - } - - // Compare lengths. - return xdL === ydL ? 0 : xdL > ydL ^ x.s < 0 ? 1 : -1; - }; - - - /* - * Return the number of decimal places of the value of this Decimal. - * - */ - P.decimalPlaces = P.dp = function () { - var x = this, - w = x.d.length - 1, - dp = (w - x.e) * LOG_BASE; - - // Subtract the number of trailing zeros of the last word. - w = x.d[w]; - if (w) for (; w % 10 == 0; w /= 10) dp--; - - return dp < 0 ? 0 : dp; - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal divided by `y`, truncated to - * `precision` significant digits. - * - */ - P.dividedBy = P.div = function (y) { - return divide(this, new this.constructor(y)); - }; - - - /* - * Return a new Decimal whose value is the integer part of dividing the value of this Decimal - * by the value of `y`, truncated to `precision` significant digits. - * - */ - P.dividedToIntegerBy = P.idiv = function (y) { - var x = this, - Ctor = x.constructor; - return round(divide(x, new Ctor(y), 0, 1), Ctor.precision); - }; - - - /* - * Return true if the value of this Decimal is equal to the value of `y`, otherwise return false. - * - */ - P.equals = P.eq = function (y) { - return !this.cmp(y); - }; - - - /* - * Return the (base 10) exponent value of this Decimal (this.e is the base 10000000 exponent). - * - */ - P.exponent = function () { - return getBase10Exponent(this); - }; - - - /* - * Return true if the value of this Decimal is greater than the value of `y`, otherwise return - * false. - * - */ - P.greaterThan = P.gt = function (y) { - return this.cmp(y) > 0; - }; - - - /* - * Return true if the value of this Decimal is greater than or equal to the value of `y`, - * otherwise return false. - * - */ - P.greaterThanOrEqualTo = P.gte = function (y) { - return this.cmp(y) >= 0; - }; - - - /* - * Return true if the value of this Decimal is an integer, otherwise return false. - * - */ - P.isInteger = P.isint = function () { - return this.e > this.d.length - 2; - }; - - - /* - * Return true if the value of this Decimal is negative, otherwise return false. - * - */ - P.isNegative = P.isneg = function () { - return this.s < 0; - }; - - - /* - * Return true if the value of this Decimal is positive, otherwise return false. - * - */ - P.isPositive = P.ispos = function () { - return this.s > 0; - }; - - - /* - * Return true if the value of this Decimal is 0, otherwise return false. - * - */ - P.isZero = function () { - return this.s === 0; - }; - - - /* - * Return true if the value of this Decimal is less than `y`, otherwise return false. - * - */ - P.lessThan = P.lt = function (y) { - return this.cmp(y) < 0; - }; - - - /* - * Return true if the value of this Decimal is less than or equal to `y`, otherwise return false. - * - */ - P.lessThanOrEqualTo = P.lte = function (y) { - return this.cmp(y) < 1; - }; - - - /* - * Return the logarithm of the value of this Decimal to the specified base, truncated to - * `precision` significant digits. - * - * If no base is specified, return log[10](x). - * - * log[base](x) = ln(x) / ln(base) - * - * The maximum error of the result is 1 ulp (unit in the last place). - * - * [base] {number|string|Decimal} The base of the logarithm. - * - */ - P.logarithm = P.log = function (base) { - var r, - x = this, - Ctor = x.constructor, - pr = Ctor.precision, - wpr = pr + 5; - - // Default base is 10. - if (base === void 0) { - base = new Ctor(10); - } else { - base = new Ctor(base); - - // log[-b](x) = NaN - // log[0](x) = NaN - // log[1](x) = NaN - if (base.s < 1 || base.eq(ONE)) throw Error(decimalError + 'NaN'); - } - - // log[b](-x) = NaN - // log[b](0) = -Infinity - if (x.s < 1) throw Error(decimalError + (x.s ? 'NaN' : '-Infinity')); - - // log[b](1) = 0 - if (x.eq(ONE)) return new Ctor(0); - - external = false; - r = divide(ln(x, wpr), ln(base, wpr), wpr); - external = true; - - return round(r, pr); - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal minus `y`, truncated to - * `precision` significant digits. - * - */ - P.minus = P.sub = function (y) { - var x = this; - y = new x.constructor(y); - return x.s == y.s ? subtract(x, y) : add(x, (y.s = -y.s, y)); - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal modulo `y`, truncated to - * `precision` significant digits. - * - */ - P.modulo = P.mod = function (y) { - var q, - x = this, - Ctor = x.constructor, - pr = Ctor.precision; - - y = new Ctor(y); - - // x % 0 = NaN - if (!y.s) throw Error(decimalError + 'NaN'); - - // Return x if x is 0. - if (!x.s) return round(new Ctor(x), pr); - - // Prevent rounding of intermediate calculations. - external = false; - q = divide(x, y, 0, 1).times(y); - external = true; - - return x.minus(q); - }; - - - /* - * Return a new Decimal whose value is the natural exponential of the value of this Decimal, - * i.e. the base e raised to the power the value of this Decimal, truncated to `precision` - * significant digits. - * - */ - P.naturalExponential = P.exp = function () { - return exp(this); - }; - - - /* - * Return a new Decimal whose value is the natural logarithm of the value of this Decimal, - * truncated to `precision` significant digits. - * - */ - P.naturalLogarithm = P.ln = function () { - return ln(this); - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal negated, i.e. as if multiplied by - * -1. - * - */ - P.negated = P.neg = function () { - var x = new this.constructor(this); - x.s = -x.s || 0; - return x; - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal plus `y`, truncated to - * `precision` significant digits. - * - */ - P.plus = P.add = function (y) { - var x = this; - y = new x.constructor(y); - return x.s == y.s ? add(x, y) : subtract(x, (y.s = -y.s, y)); - }; - - - /* - * Return the number of significant digits of the value of this Decimal. - * - * [z] {boolean|number} Whether to count integer-part trailing zeros: true, false, 1 or 0. - * - */ - P.precision = P.sd = function (z) { - var e, sd, w, - x = this; - - if (z !== void 0 && z !== !!z && z !== 1 && z !== 0) throw Error(invalidArgument + z); - - e = getBase10Exponent(x) + 1; - w = x.d.length - 1; - sd = w * LOG_BASE + 1; - w = x.d[w]; - - // If non-zero... - if (w) { - - // Subtract the number of trailing zeros of the last word. - for (; w % 10 == 0; w /= 10) sd--; - - // Add the number of digits of the first word. - for (w = x.d[0]; w >= 10; w /= 10) sd++; - } - - return z && e > sd ? e : sd; - }; - - - /* - * Return a new Decimal whose value is the square root of this Decimal, truncated to `precision` - * significant digits. - * - */ - P.squareRoot = P.sqrt = function () { - var e, n, pr, r, s, t, wpr, - x = this, - Ctor = x.constructor; - - // Negative or zero? - if (x.s < 1) { - if (!x.s) return new Ctor(0); - - // sqrt(-x) = NaN - throw Error(decimalError + 'NaN'); - } - - e = getBase10Exponent(x); - external = false; - - // Initial estimate. - s = Math.sqrt(+x); - - // Math.sqrt underflow/overflow? - // Pass x to Math.sqrt as integer, then adjust the exponent of the result. - if (s == 0 || s == 1 / 0) { - n = digitsToString(x.d); - if ((n.length + e) % 2 == 0) n += '0'; - s = Math.sqrt(n); - e = mathfloor((e + 1) / 2) - (e < 0 || e % 2); - - if (s == 1 / 0) { - n = '5e' + e; - } else { - n = s.toExponential(); - n = n.slice(0, n.indexOf('e') + 1) + e; - } - - r = new Ctor(n); - } else { - r = new Ctor(s.toString()); - } - - pr = Ctor.precision; - s = wpr = pr + 3; - - // Newton-Raphson iteration. - for (;;) { - t = r; - r = t.plus(divide(x, t, wpr + 2)).times(0.5); - - if (digitsToString(t.d).slice(0, wpr) === (n = digitsToString(r.d)).slice(0, wpr)) { - n = n.slice(wpr - 3, wpr + 1); - - // The 4th rounding digit may be in error by -1 so if the 4 rounding digits are 9999 or - // 4999, i.e. approaching a rounding boundary, continue the iteration. - if (s == wpr && n == '4999') { - - // On the first iteration only, check to see if rounding up gives the exact result as the - // nines may infinitely repeat. - round(t, pr + 1, 0); - - if (t.times(t).eq(x)) { - r = t; - break; - } - } else if (n != '9999') { - break; - } - - wpr += 4; - } - } - - external = true; - - return round(r, pr); - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal times `y`, truncated to - * `precision` significant digits. - * - */ - P.times = P.mul = function (y) { - var carry, e, i, k, r, rL, t, xdL, ydL, - x = this, - Ctor = x.constructor, - xd = x.d, - yd = (y = new Ctor(y)).d; - - // Return 0 if either is 0. - if (!x.s || !y.s) return new Ctor(0); - - y.s *= x.s; - e = x.e + y.e; - xdL = xd.length; - ydL = yd.length; - - // Ensure xd points to the longer array. - if (xdL < ydL) { - r = xd; - xd = yd; - yd = r; - rL = xdL; - xdL = ydL; - ydL = rL; - } - - // Initialise the result array with zeros. - r = []; - rL = xdL + ydL; - for (i = rL; i--;) r.push(0); - - // Multiply! - for (i = ydL; --i >= 0;) { - carry = 0; - for (k = xdL + i; k > i;) { - t = r[k] + yd[i] * xd[k - i - 1] + carry; - r[k--] = t % BASE | 0; - carry = t / BASE | 0; - } - - r[k] = (r[k] + carry) % BASE | 0; - } - - // Remove trailing zeros. - for (; !r[--rL];) r.pop(); - - if (carry) ++e; - else r.shift(); - - y.d = r; - y.e = e; - - return external ? round(y, Ctor.precision) : y; - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal rounded to a maximum of `dp` - * decimal places using rounding mode `rm` or `rounding` if `rm` is omitted. - * - * If `dp` is omitted, return a new Decimal whose value is the value of this Decimal. - * - * [dp] {number} Decimal places. Integer, 0 to MAX_DIGITS inclusive. - * [rm] {number} Rounding mode. Integer, 0 to 8 inclusive. - * - */ - P.toDecimalPlaces = P.todp = function (dp, rm) { - var x = this, - Ctor = x.constructor; - - x = new Ctor(x); - if (dp === void 0) return x; - - checkInt32(dp, 0, MAX_DIGITS); - - if (rm === void 0) rm = Ctor.rounding; - else checkInt32(rm, 0, 8); - - return round(x, dp + getBase10Exponent(x) + 1, rm); - }; - - - /* - * Return a string representing the value of this Decimal in exponential notation rounded to - * `dp` fixed decimal places using rounding mode `rounding`. - * - * [dp] {number} Decimal places. Integer, 0 to MAX_DIGITS inclusive. - * [rm] {number} Rounding mode. Integer, 0 to 8 inclusive. - * - */ - P.toExponential = function (dp, rm) { - var str, - x = this, - Ctor = x.constructor; - - if (dp === void 0) { - str = toString(x, true); - } else { - checkInt32(dp, 0, MAX_DIGITS); - - if (rm === void 0) rm = Ctor.rounding; - else checkInt32(rm, 0, 8); - - x = round(new Ctor(x), dp + 1, rm); - str = toString(x, true, dp + 1); - } - - return str; - }; - - - /* - * Return a string representing the value of this Decimal in normal (fixed-point) notation to - * `dp` fixed decimal places and rounded using rounding mode `rm` or `rounding` if `rm` is - * omitted. - * - * As with JavaScript numbers, (-0).toFixed(0) is '0', but e.g. (-0.00001).toFixed(0) is '-0'. - * - * [dp] {number} Decimal places. Integer, 0 to MAX_DIGITS inclusive. - * [rm] {number} Rounding mode. Integer, 0 to 8 inclusive. - * - * (-0).toFixed(0) is '0', but (-0.1).toFixed(0) is '-0'. - * (-0).toFixed(1) is '0.0', but (-0.01).toFixed(1) is '-0.0'. - * (-0).toFixed(3) is '0.000'. - * (-0.5).toFixed(0) is '-0'. - * - */ - P.toFixed = function (dp, rm) { - var str, y, - x = this, - Ctor = x.constructor; - - if (dp === void 0) return toString(x); - - checkInt32(dp, 0, MAX_DIGITS); - - if (rm === void 0) rm = Ctor.rounding; - else checkInt32(rm, 0, 8); - - y = round(new Ctor(x), dp + getBase10Exponent(x) + 1, rm); - str = toString(y.abs(), false, dp + getBase10Exponent(y) + 1); - - // To determine whether to add the minus sign look at the value before it was rounded, - // i.e. look at `x` rather than `y`. - return x.isneg() && !x.isZero() ? '-' + str : str; - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal rounded to a whole number using - * rounding mode `rounding`. - * - */ - P.toInteger = P.toint = function () { - var x = this, - Ctor = x.constructor; - return round(new Ctor(x), getBase10Exponent(x) + 1, Ctor.rounding); - }; - - - /* - * Return the value of this Decimal converted to a number primitive. - * - */ - P.toNumber = function () { - return +this; - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal raised to the power `y`, - * truncated to `precision` significant digits. - * - * For non-integer or very large exponents pow(x, y) is calculated using - * - * x^y = exp(y*ln(x)) - * - * The maximum error is 1 ulp (unit in last place). - * - * y {number|string|Decimal} The power to which to raise this Decimal. - * - */ - P.toPower = P.pow = function (y) { - var e, k, pr, r, sign, yIsInt, - x = this, - Ctor = x.constructor, - guard = 12, - yn = +(y = new Ctor(y)); - - // pow(x, 0) = 1 - if (!y.s) return new Ctor(ONE); - - x = new Ctor(x); - - // pow(0, y > 0) = 0 - // pow(0, y < 0) = Infinity - if (!x.s) { - if (y.s < 1) throw Error(decimalError + 'Infinity'); - return x; - } - - // pow(1, y) = 1 - if (x.eq(ONE)) return x; - - pr = Ctor.precision; - - // pow(x, 1) = x - if (y.eq(ONE)) return round(x, pr); - - e = y.e; - k = y.d.length - 1; - yIsInt = e >= k; - sign = x.s; - - if (!yIsInt) { - - // pow(x < 0, y non-integer) = NaN - if (sign < 0) throw Error(decimalError + 'NaN'); - - // If y is a small integer use the 'exponentiation by squaring' algorithm. - } else if ((k = yn < 0 ? -yn : yn) <= MAX_SAFE_INTEGER) { - r = new Ctor(ONE); - - // Max k of 9007199254740991 takes 53 loop iterations. - // Maximum digits array length; leaves [28, 34] guard digits. - e = Math.ceil(pr / LOG_BASE + 4); - - external = false; - - for (;;) { - if (k % 2) { - r = r.times(x); - truncate(r.d, e); - } - - k = mathfloor(k / 2); - if (k === 0) break; - - x = x.times(x); - truncate(x.d, e); - } - - external = true; - - return y.s < 0 ? new Ctor(ONE).div(r) : round(r, pr); - } - - // Result is negative if x is negative and the last digit of integer y is odd. - sign = sign < 0 && y.d[Math.max(e, k)] & 1 ? -1 : 1; - - x.s = 1; - external = false; - r = y.times(ln(x, pr + guard)); - external = true; - r = exp(r); - r.s = sign; - - return r; - }; - - - /* - * Return a string representing the value of this Decimal rounded to `sd` significant digits - * using rounding mode `rounding`. - * - * Return exponential notation if `sd` is less than the number of digits necessary to represent - * the integer part of the value in normal notation. - * - * [sd] {number} Significant digits. Integer, 1 to MAX_DIGITS inclusive. - * [rm] {number} Rounding mode. Integer, 0 to 8 inclusive. - * - */ - P.toPrecision = function (sd, rm) { - var e, str, - x = this, - Ctor = x.constructor; - - if (sd === void 0) { - e = getBase10Exponent(x); - str = toString(x, e <= Ctor.toExpNeg || e >= Ctor.toExpPos); - } else { - checkInt32(sd, 1, MAX_DIGITS); - - if (rm === void 0) rm = Ctor.rounding; - else checkInt32(rm, 0, 8); - - x = round(new Ctor(x), sd, rm); - e = getBase10Exponent(x); - str = toString(x, sd <= e || e <= Ctor.toExpNeg, sd); - } - - return str; - }; - - - /* - * Return a new Decimal whose value is the value of this Decimal rounded to a maximum of `sd` - * significant digits using rounding mode `rm`, or to `precision` and `rounding` respectively if - * omitted. - * - * [sd] {number} Significant digits. Integer, 1 to MAX_DIGITS inclusive. - * [rm] {number} Rounding mode. Integer, 0 to 8 inclusive. - * - */ - P.toSignificantDigits = P.tosd = function (sd, rm) { - var x = this, - Ctor = x.constructor; - - if (sd === void 0) { - sd = Ctor.precision; - rm = Ctor.rounding; - } else { - checkInt32(sd, 1, MAX_DIGITS); - - if (rm === void 0) rm = Ctor.rounding; - else checkInt32(rm, 0, 8); - } - - return round(new Ctor(x), sd, rm); - }; - - - /* - * Return a string representing the value of this Decimal. - * - * Return exponential notation if this Decimal has a positive exponent equal to or greater than - * `toExpPos`, or a negative exponent equal to or less than `toExpNeg`. - * - */ - P.toString = P.valueOf = P.val = P.toJSON = function () { - var x = this, - e = getBase10Exponent(x), - Ctor = x.constructor; - - return toString(x, e <= Ctor.toExpNeg || e >= Ctor.toExpPos); - }; - - - // Helper functions for Decimal.prototype (P) and/or Decimal methods, and their callers. - - - /* - * add P.minus, P.plus - * checkInt32 P.todp, P.toExponential, P.toFixed, P.toPrecision, P.tosd - * digitsToString P.log, P.sqrt, P.pow, toString, exp, ln - * divide P.div, P.idiv, P.log, P.mod, P.sqrt, exp, ln - * exp P.exp, P.pow - * getBase10Exponent P.exponent, P.sd, P.toint, P.sqrt, P.todp, P.toFixed, P.toPrecision, - * P.toString, divide, round, toString, exp, ln - * getLn10 P.log, ln - * getZeroString digitsToString, toString - * ln P.log, P.ln, P.pow, exp - * parseDecimal Decimal - * round P.abs, P.idiv, P.log, P.minus, P.mod, P.neg, P.plus, P.toint, P.sqrt, - * P.times, P.todp, P.toExponential, P.toFixed, P.pow, P.toPrecision, P.tosd, - * divide, getLn10, exp, ln - * subtract P.minus, P.plus - * toString P.toExponential, P.toFixed, P.toPrecision, P.toString, P.valueOf - * truncate P.pow - * - * Throws: P.log, P.mod, P.sd, P.sqrt, P.pow, checkInt32, divide, round, - * getLn10, exp, ln, parseDecimal, Decimal, config - */ - - - function add(x, y) { - var carry, d, e, i, k, len, xd, yd, - Ctor = x.constructor, - pr = Ctor.precision; - - // If either is zero... - if (!x.s || !y.s) { - - // Return x if y is zero. - // Return y if y is non-zero. - if (!y.s) y = new Ctor(x); - return external ? round(y, pr) : y; - } - - xd = x.d; - yd = y.d; - - // x and y are finite, non-zero numbers with the same sign. - - k = x.e; - e = y.e; - xd = xd.slice(); - i = k - e; - - // If base 1e7 exponents differ... - if (i) { - if (i < 0) { - d = xd; - i = -i; - len = yd.length; - } else { - d = yd; - e = k; - len = xd.length; - } - - // Limit number of zeros prepended to max(ceil(pr / LOG_BASE), len) + 1. - k = Math.ceil(pr / LOG_BASE); - len = k > len ? k + 1 : len + 1; - - if (i > len) { - i = len; - d.length = 1; - } - - // Prepend zeros to equalise exponents. Note: Faster to use reverse then do unshifts. - d.reverse(); - for (; i--;) d.push(0); - d.reverse(); - } - - len = xd.length; - i = yd.length; - - // If yd is longer than xd, swap xd and yd so xd points to the longer array. - if (len - i < 0) { - i = len; - d = yd; - yd = xd; - xd = d; - } - - // Only start adding at yd.length - 1 as the further digits of xd can be left as they are. - for (carry = 0; i;) { - carry = (xd[--i] = xd[i] + yd[i] + carry) / BASE | 0; - xd[i] %= BASE; - } - - if (carry) { - xd.unshift(carry); - ++e; - } - - // Remove trailing zeros. - // No need to check for zero, as +x + +y != 0 && -x + -y != 0 - for (len = xd.length; xd[--len] == 0;) xd.pop(); - - y.d = xd; - y.e = e; - - return external ? round(y, pr) : y; - } - - - function checkInt32(i, min, max) { - if (i !== ~~i || i < min || i > max) { - throw Error(invalidArgument + i); - } - } - - - function digitsToString(d) { - var i, k, ws, - indexOfLastWord = d.length - 1, - str = '', - w = d[0]; - - if (indexOfLastWord > 0) { - str += w; - for (i = 1; i < indexOfLastWord; i++) { - ws = d[i] + ''; - k = LOG_BASE - ws.length; - if (k) str += getZeroString(k); - str += ws; - } - - w = d[i]; - ws = w + ''; - k = LOG_BASE - ws.length; - if (k) str += getZeroString(k); - } else if (w === 0) { - return '0'; - } - - // Remove trailing zeros of last w. - for (; w % 10 === 0;) w /= 10; - - return str + w; - } - - - var divide = (function () { - - // Assumes non-zero x and k, and hence non-zero result. - function multiplyInteger(x, k) { - var temp, - carry = 0, - i = x.length; - - for (x = x.slice(); i--;) { - temp = x[i] * k + carry; - x[i] = temp % BASE | 0; - carry = temp / BASE | 0; - } - - if (carry) x.unshift(carry); - - return x; - } - - function compare(a, b, aL, bL) { - var i, r; - - if (aL != bL) { - r = aL > bL ? 1 : -1; - } else { - for (i = r = 0; i < aL; i++) { - if (a[i] != b[i]) { - r = a[i] > b[i] ? 1 : -1; - break; - } - } - } - - return r; - } - - function subtract(a, b, aL) { - var i = 0; - - // Subtract b from a. - for (; aL--;) { - a[aL] -= i; - i = a[aL] < b[aL] ? 1 : 0; - a[aL] = i * BASE + a[aL] - b[aL]; - } - - // Remove leading zeros. - for (; !a[0] && a.length > 1;) a.shift(); - } - - return function (x, y, pr, dp) { - var cmp, e, i, k, prod, prodL, q, qd, rem, remL, rem0, sd, t, xi, xL, yd0, yL, yz, - Ctor = x.constructor, - sign = x.s == y.s ? 1 : -1, - xd = x.d, - yd = y.d; - - // Either 0? - if (!x.s) return new Ctor(x); - if (!y.s) throw Error(decimalError + 'Division by zero'); - - e = x.e - y.e; - yL = yd.length; - xL = xd.length; - q = new Ctor(sign); - qd = q.d = []; - - // Result exponent may be one less than e. - for (i = 0; yd[i] == (xd[i] || 0); ) ++i; - if (yd[i] > (xd[i] || 0)) --e; - - if (pr == null) { - sd = pr = Ctor.precision; - } else if (dp) { - sd = pr + (getBase10Exponent(x) - getBase10Exponent(y)) + 1; - } else { - sd = pr; - } - - if (sd < 0) return new Ctor(0); - - // Convert precision in number of base 10 digits to base 1e7 digits. - sd = sd / LOG_BASE + 2 | 0; - i = 0; - - // divisor < 1e7 - if (yL == 1) { - k = 0; - yd = yd[0]; - sd++; - - // k is the carry. - for (; (i < xL || k) && sd--; i++) { - t = k * BASE + (xd[i] || 0); - qd[i] = t / yd | 0; - k = t % yd | 0; - } - - // divisor >= 1e7 - } else { - - // Normalise xd and yd so highest order digit of yd is >= BASE/2 - k = BASE / (yd[0] + 1) | 0; - - if (k > 1) { - yd = multiplyInteger(yd, k); - xd = multiplyInteger(xd, k); - yL = yd.length; - xL = xd.length; - } - - xi = yL; - rem = xd.slice(0, yL); - remL = rem.length; - - // Add zeros to make remainder as long as divisor. - for (; remL < yL;) rem[remL++] = 0; - - yz = yd.slice(); - yz.unshift(0); - yd0 = yd[0]; - - if (yd[1] >= BASE / 2) ++yd0; - - do { - k = 0; - - // Compare divisor and remainder. - cmp = compare(yd, rem, yL, remL); - - // If divisor < remainder. - if (cmp < 0) { - - // Calculate trial digit, k. - rem0 = rem[0]; - if (yL != remL) rem0 = rem0 * BASE + (rem[1] || 0); - - // k will be how many times the divisor goes into the current remainder. - k = rem0 / yd0 | 0; - - // Algorithm: - // 1. product = divisor * trial digit (k) - // 2. if product > remainder: product -= divisor, k-- - // 3. remainder -= product - // 4. if product was < remainder at 2: - // 5. compare new remainder and divisor - // 6. If remainder > divisor: remainder -= divisor, k++ - - if (k > 1) { - if (k >= BASE) k = BASE - 1; - - // product = divisor * trial digit. - prod = multiplyInteger(yd, k); - prodL = prod.length; - remL = rem.length; - - // Compare product and remainder. - cmp = compare(prod, rem, prodL, remL); - - // product > remainder. - if (cmp == 1) { - k--; - - // Subtract divisor from product. - subtract(prod, yL < prodL ? yz : yd, prodL); - } - } else { - - // cmp is -1. - // If k is 0, there is no need to compare yd and rem again below, so change cmp to 1 - // to avoid it. If k is 1 there is a need to compare yd and rem again below. - if (k == 0) cmp = k = 1; - prod = yd.slice(); - } - - prodL = prod.length; - if (prodL < remL) prod.unshift(0); - - // Subtract product from remainder. - subtract(rem, prod, remL); - - // If product was < previous remainder. - if (cmp == -1) { - remL = rem.length; - - // Compare divisor and new remainder. - cmp = compare(yd, rem, yL, remL); - - // If divisor < new remainder, subtract divisor from remainder. - if (cmp < 1) { - k++; - - // Subtract divisor from remainder. - subtract(rem, yL < remL ? yz : yd, remL); - } - } - - remL = rem.length; - } else if (cmp === 0) { - k++; - rem = [0]; - } // if cmp === 1, k will be 0 - - // Add the next digit, k, to the result array. - qd[i++] = k; - - // Update the remainder. - if (cmp && rem[0]) { - rem[remL++] = xd[xi] || 0; - } else { - rem = [xd[xi]]; - remL = 1; - } - - } while ((xi++ < xL || rem[0] !== void 0) && sd--); - } - - // Leading zero? - if (!qd[0]) qd.shift(); - - q.e = e; - - return round(q, dp ? pr + getBase10Exponent(q) + 1 : pr); - }; - })(); - - - /* - * Return a new Decimal whose value is the natural exponential of `x` truncated to `sd` - * significant digits. - * - * Taylor/Maclaurin series. - * - * exp(x) = x^0/0! + x^1/1! + x^2/2! + x^3/3! + ... - * - * Argument reduction: - * Repeat x = x / 32, k += 5, until |x| < 0.1 - * exp(x) = exp(x / 2^k)^(2^k) - * - * Previously, the argument was initially reduced by - * exp(x) = exp(r) * 10^k where r = x - k * ln10, k = floor(x / ln10) - * to first put r in the range [0, ln10], before dividing by 32 until |x| < 0.1, but this was - * found to be slower than just dividing repeatedly by 32 as above. - * - * (Math object integer min/max: Math.exp(709) = 8.2e+307, Math.exp(-745) = 5e-324) - * - * exp(x) is non-terminating for any finite, non-zero x. - * - */ - function exp(x, sd) { - var denominator, guard, pow, sum, t, wpr, - i = 0, - k = 0, - Ctor = x.constructor, - pr = Ctor.precision; - - if (getBase10Exponent(x) > 16) throw Error(exponentOutOfRange + getBase10Exponent(x)); - - // exp(0) = 1 - if (!x.s) return new Ctor(ONE); - - if (sd == null) { - external = false; - wpr = pr; - } else { - wpr = sd; - } - - t = new Ctor(0.03125); - - while (x.abs().gte(0.1)) { - x = x.times(t); // x = x / 2^5 - k += 5; - } - - // Estimate the precision increase necessary to ensure the first 4 rounding digits are correct. - guard = Math.log(mathpow(2, k)) / Math.LN10 * 2 + 5 | 0; - wpr += guard; - denominator = pow = sum = new Ctor(ONE); - Ctor.precision = wpr; - - for (;;) { - pow = round(pow.times(x), wpr); - denominator = denominator.times(++i); - t = sum.plus(divide(pow, denominator, wpr)); - - if (digitsToString(t.d).slice(0, wpr) === digitsToString(sum.d).slice(0, wpr)) { - while (k--) sum = round(sum.times(sum), wpr); - Ctor.precision = pr; - return sd == null ? (external = true, round(sum, pr)) : sum; - } - - sum = t; - } - } - - - // Calculate the base 10 exponent from the base 1e7 exponent. - function getBase10Exponent(x) { - var e = x.e * LOG_BASE, - w = x.d[0]; - - // Add the number of digits of the first word of the digits array. - for (; w >= 10; w /= 10) e++; - return e; - } - - - function getLn10(Ctor, sd, pr) { - - if (sd > Ctor.LN10.sd()) { - - - // Reset global state in case the exception is caught. - external = true; - if (pr) Ctor.precision = pr; - throw Error(decimalError + 'LN10 precision limit exceeded'); - } - - return round(new Ctor(Ctor.LN10), sd); - } - - - function getZeroString(k) { - var zs = ''; - for (; k--;) zs += '0'; - return zs; - } - - - /* - * Return a new Decimal whose value is the natural logarithm of `x` truncated to `sd` significant - * digits. - * - * ln(n) is non-terminating (n != 1) - * - */ - function ln(y, sd) { - var c, c0, denominator, e, numerator, sum, t, wpr, x2, - n = 1, - guard = 10, - x = y, - xd = x.d, - Ctor = x.constructor, - pr = Ctor.precision; - - // ln(-x) = NaN - // ln(0) = -Infinity - if (x.s < 1) throw Error(decimalError + (x.s ? 'NaN' : '-Infinity')); - - // ln(1) = 0 - if (x.eq(ONE)) return new Ctor(0); - - if (sd == null) { - external = false; - wpr = pr; - } else { - wpr = sd; - } - - if (x.eq(10)) { - if (sd == null) external = true; - return getLn10(Ctor, wpr); - } - - wpr += guard; - Ctor.precision = wpr; - c = digitsToString(xd); - c0 = c.charAt(0); - e = getBase10Exponent(x); - - if (Math.abs(e) < 1.5e15) { - - // Argument reduction. - // The series converges faster the closer the argument is to 1, so using - // ln(a^b) = b * ln(a), ln(a) = ln(a^b) / b - // multiply the argument by itself until the leading digits of the significand are 7, 8, 9, - // 10, 11, 12 or 13, recording the number of multiplications so the sum of the series can - // later be divided by this number, then separate out the power of 10 using - // ln(a*10^b) = ln(a) + b*ln(10). - - // max n is 21 (gives 0.9, 1.0 or 1.1) (9e15 / 21 = 4.2e14). - //while (c0 < 9 && c0 != 1 || c0 == 1 && c.charAt(1) > 1) { - // max n is 6 (gives 0.7 - 1.3) - while (c0 < 7 && c0 != 1 || c0 == 1 && c.charAt(1) > 3) { - x = x.times(y); - c = digitsToString(x.d); - c0 = c.charAt(0); - n++; - } - - e = getBase10Exponent(x); - - if (c0 > 1) { - x = new Ctor('0.' + c); - e++; - } else { - x = new Ctor(c0 + '.' + c.slice(1)); - } - } else { - - // The argument reduction method above may result in overflow if the argument y is a massive - // number with exponent >= 1500000000000000 (9e15 / 6 = 1.5e15), so instead recall this - // function using ln(x*10^e) = ln(x) + e*ln(10). - t = getLn10(Ctor, wpr + 2, pr).times(e + ''); - x = ln(new Ctor(c0 + '.' + c.slice(1)), wpr - guard).plus(t); - - Ctor.precision = pr; - return sd == null ? (external = true, round(x, pr)) : x; - } - - // x is reduced to a value near 1. - - // Taylor series. - // ln(y) = ln((1 + x)/(1 - x)) = 2(x + x^3/3 + x^5/5 + x^7/7 + ...) - // where x = (y - 1)/(y + 1) (|x| < 1) - sum = numerator = x = divide(x.minus(ONE), x.plus(ONE), wpr); - x2 = round(x.times(x), wpr); - denominator = 3; - - for (;;) { - numerator = round(numerator.times(x2), wpr); - t = sum.plus(divide(numerator, new Ctor(denominator), wpr)); - - if (digitsToString(t.d).slice(0, wpr) === digitsToString(sum.d).slice(0, wpr)) { - sum = sum.times(2); - - // Reverse the argument reduction. - if (e !== 0) sum = sum.plus(getLn10(Ctor, wpr + 2, pr).times(e + '')); - sum = divide(sum, new Ctor(n), wpr); - - Ctor.precision = pr; - return sd == null ? (external = true, round(sum, pr)) : sum; - } - - sum = t; - denominator += 2; - } - } - - - /* - * Parse the value of a new Decimal `x` from string `str`. - */ - function parseDecimal(x, str) { - var e, i, len; - - // Decimal point? - if ((e = str.indexOf('.')) > -1) str = str.replace('.', ''); - - // Exponential form? - if ((i = str.search(/e/i)) > 0) { - - // Determine exponent. - if (e < 0) e = i; - e += +str.slice(i + 1); - str = str.substring(0, i); - } else if (e < 0) { - - // Integer. - e = str.length; - } - - // Determine leading zeros. - for (i = 0; str.charCodeAt(i) === 48;) ++i; - - // Determine trailing zeros. - for (len = str.length; str.charCodeAt(len - 1) === 48;) --len; - str = str.slice(i, len); - - if (str) { - len -= i; - e = e - i - 1; - x.e = mathfloor(e / LOG_BASE); - x.d = []; - - // Transform base - - // e is the base 10 exponent. - // i is where to slice str to get the first word of the digits array. - i = (e + 1) % LOG_BASE; - if (e < 0) i += LOG_BASE; - - if (i < len) { - if (i) x.d.push(+str.slice(0, i)); - for (len -= LOG_BASE; i < len;) x.d.push(+str.slice(i, i += LOG_BASE)); - str = str.slice(i); - i = LOG_BASE - str.length; - } else { - i -= len; - } - - for (; i--;) str += '0'; - x.d.push(+str); - - if (external && (x.e > MAX_E || x.e < -MAX_E)) throw Error(exponentOutOfRange + e); - } else { - - // Zero. - x.s = 0; - x.e = 0; - x.d = [0]; - } - - return x; - } - - - /* - * Round `x` to `sd` significant digits, using rounding mode `rm` if present (truncate otherwise). - */ - function round(x, sd, rm) { - var i, j, k, n, rd, doRound, w, xdi, - xd = x.d; - - // rd: the rounding digit, i.e. the digit after the digit that may be rounded up. - // w: the word of xd which contains the rounding digit, a base 1e7 number. - // xdi: the index of w within xd. - // n: the number of digits of w. - // i: what would be the index of rd within w if all the numbers were 7 digits long (i.e. if - // they had leading zeros) - // j: if > 0, the actual index of rd within w (if < 0, rd is a leading zero). - - // Get the length of the first word of the digits array xd. - for (n = 1, k = xd[0]; k >= 10; k /= 10) n++; - i = sd - n; - - // Is the rounding digit in the first word of xd? - if (i < 0) { - i += LOG_BASE; - j = sd; - w = xd[xdi = 0]; - } else { - xdi = Math.ceil((i + 1) / LOG_BASE); - k = xd.length; - if (xdi >= k) return x; - w = k = xd[xdi]; - - // Get the number of digits of w. - for (n = 1; k >= 10; k /= 10) n++; - - // Get the index of rd within w. - i %= LOG_BASE; - - // Get the index of rd within w, adjusted for leading zeros. - // The number of leading zeros of w is given by LOG_BASE - n. - j = i - LOG_BASE + n; - } - - if (rm !== void 0) { - k = mathpow(10, n - j - 1); - - // Get the rounding digit at index j of w. - rd = w / k % 10 | 0; - - // Are there any non-zero digits after the rounding digit? - doRound = sd < 0 || xd[xdi + 1] !== void 0 || w % k; - - // The expression `w % mathpow(10, n - j - 1)` returns all the digits of w to the right of the - // digit at (left-to-right) index j, e.g. if w is 908714 and j is 2, the expression will give - // 714. - - doRound = rm < 4 - ? (rd || doRound) && (rm == 0 || rm == (x.s < 0 ? 3 : 2)) - : rd > 5 || rd == 5 && (rm == 4 || doRound || rm == 6 && - - // Check whether the digit to the left of the rounding digit is odd. - ((i > 0 ? j > 0 ? w / mathpow(10, n - j) : 0 : xd[xdi - 1]) % 10) & 1 || - rm == (x.s < 0 ? 8 : 7)); - } - - if (sd < 1 || !xd[0]) { - if (doRound) { - k = getBase10Exponent(x); - xd.length = 1; - - // Convert sd to decimal places. - sd = sd - k - 1; - - // 1, 0.1, 0.01, 0.001, 0.0001 etc. - xd[0] = mathpow(10, (LOG_BASE - sd % LOG_BASE) % LOG_BASE); - x.e = mathfloor(-sd / LOG_BASE) || 0; - } else { - xd.length = 1; - - // Zero. - xd[0] = x.e = x.s = 0; - } - - return x; - } - - // Remove excess digits. - if (i == 0) { - xd.length = xdi; - k = 1; - xdi--; - } else { - xd.length = xdi + 1; - k = mathpow(10, LOG_BASE - i); - - // E.g. 56700 becomes 56000 if 7 is the rounding digit. - // j > 0 means i > number of leading zeros of w. - xd[xdi] = j > 0 ? (w / mathpow(10, n - j) % mathpow(10, j) | 0) * k : 0; - } - - if (doRound) { - for (;;) { - - // Is the digit to be rounded up in the first word of xd? - if (xdi == 0) { - if ((xd[0] += k) == BASE) { - xd[0] = 1; - ++x.e; - } - - break; - } else { - xd[xdi] += k; - if (xd[xdi] != BASE) break; - xd[xdi--] = 0; - k = 1; - } - } - } - - // Remove trailing zeros. - for (i = xd.length; xd[--i] === 0;) xd.pop(); - - if (external && (x.e > MAX_E || x.e < -MAX_E)) { - throw Error(exponentOutOfRange + getBase10Exponent(x)); - } - - return x; - } - - - function subtract(x, y) { - var d, e, i, j, k, len, xd, xe, xLTy, yd, - Ctor = x.constructor, - pr = Ctor.precision; - - // Return y negated if x is zero. - // Return x if y is zero and x is non-zero. - if (!x.s || !y.s) { - if (y.s) y.s = -y.s; - else y = new Ctor(x); - return external ? round(y, pr) : y; - } - - xd = x.d; - yd = y.d; - - // x and y are non-zero numbers with the same sign. - - e = y.e; - xe = x.e; - xd = xd.slice(); - k = xe - e; - - // If exponents differ... - if (k) { - xLTy = k < 0; - - if (xLTy) { - d = xd; - k = -k; - len = yd.length; - } else { - d = yd; - e = xe; - len = xd.length; - } - - // Numbers with massively different exponents would result in a very high number of zeros - // needing to be prepended, but this can be avoided while still ensuring correct rounding by - // limiting the number of zeros to `Math.ceil(pr / LOG_BASE) + 2`. - i = Math.max(Math.ceil(pr / LOG_BASE), len) + 2; - - if (k > i) { - k = i; - d.length = 1; - } - - // Prepend zeros to equalise exponents. - d.reverse(); - for (i = k; i--;) d.push(0); - d.reverse(); - - // Base 1e7 exponents equal. - } else { - - // Check digits to determine which is the bigger number. - - i = xd.length; - len = yd.length; - xLTy = i < len; - if (xLTy) len = i; - - for (i = 0; i < len; i++) { - if (xd[i] != yd[i]) { - xLTy = xd[i] < yd[i]; - break; - } - } - - k = 0; - } - - if (xLTy) { - d = xd; - xd = yd; - yd = d; - y.s = -y.s; - } - - len = xd.length; - - // Append zeros to xd if shorter. - // Don't add zeros to yd if shorter as subtraction only needs to start at yd length. - for (i = yd.length - len; i > 0; --i) xd[len++] = 0; - - // Subtract yd from xd. - for (i = yd.length; i > k;) { - if (xd[--i] < yd[i]) { - for (j = i; j && xd[--j] === 0;) xd[j] = BASE - 1; - --xd[j]; - xd[i] += BASE; - } - - xd[i] -= yd[i]; - } - - // Remove trailing zeros. - for (; xd[--len] === 0;) xd.pop(); - - // Remove leading zeros and adjust exponent accordingly. - for (; xd[0] === 0; xd.shift()) --e; - - // Zero? - if (!xd[0]) return new Ctor(0); - - y.d = xd; - y.e = e; - - //return external && xd.length >= pr / LOG_BASE ? round(y, pr) : y; - return external ? round(y, pr) : y; - } - - - function toString(x, isExp, sd) { - var k, - e = getBase10Exponent(x), - str = digitsToString(x.d), - len = str.length; - - if (isExp) { - if (sd && (k = sd - len) > 0) { - str = str.charAt(0) + '.' + str.slice(1) + getZeroString(k); - } else if (len > 1) { - str = str.charAt(0) + '.' + str.slice(1); - } - - str = str + (e < 0 ? 'e' : 'e+') + e; - } else if (e < 0) { - str = '0.' + getZeroString(-e - 1) + str; - if (sd && (k = sd - len) > 0) str += getZeroString(k); - } else if (e >= len) { - str += getZeroString(e + 1 - len); - if (sd && (k = sd - e - 1) > 0) str = str + '.' + getZeroString(k); - } else { - if ((k = e + 1) < len) str = str.slice(0, k) + '.' + str.slice(k); - if (sd && (k = sd - len) > 0) { - if (e + 1 === len) str += '.'; - str += getZeroString(k); - } - } - - return x.s < 0 ? '-' + str : str; - } - - - // Does not strip trailing zeros. - function truncate(arr, len) { - if (arr.length > len) { - arr.length = len; - return true; - } - } - - - // Decimal methods - - - /* - * clone - * config/set - */ - - - /* - * Create and return a Decimal constructor with the same configuration properties as this Decimal - * constructor. - * - */ - function clone(obj) { - var i, p, ps; - - /* - * The Decimal constructor and exported function. - * Return a new Decimal instance. - * - * value {number|string|Decimal} A numeric value. - * - */ - function Decimal(value) { - var x = this; - - // Decimal called without new. - if (!(x instanceof Decimal)) return new Decimal(value); - - // Retain a reference to this Decimal constructor, and shadow Decimal.prototype.constructor - // which points to Object. - x.constructor = Decimal; - - // Duplicate. - if (value instanceof Decimal) { - x.s = value.s; - x.e = value.e; - x.d = (value = value.d) ? value.slice() : value; - return; - } - - if (typeof value === 'number') { - - // Reject Infinity/NaN. - if (value * 0 !== 0) { - throw Error(invalidArgument + value); - } - - if (value > 0) { - x.s = 1; - } else if (value < 0) { - value = -value; - x.s = -1; - } else { - x.s = 0; - x.e = 0; - x.d = [0]; - return; - } - - // Fast path for small integers. - if (value === ~~value && value < 1e7) { - x.e = 0; - x.d = [value]; - return; - } - - return parseDecimal(x, value.toString()); - } else if (typeof value !== 'string') { - throw Error(invalidArgument + value); - } - - // Minus sign? - if (value.charCodeAt(0) === 45) { - value = value.slice(1); - x.s = -1; - } else { - x.s = 1; - } - - if (isDecimal.test(value)) parseDecimal(x, value); - else throw Error(invalidArgument + value); - } - - Decimal.prototype = P; - - Decimal.ROUND_UP = 0; - Decimal.ROUND_DOWN = 1; - Decimal.ROUND_CEIL = 2; - Decimal.ROUND_FLOOR = 3; - Decimal.ROUND_HALF_UP = 4; - Decimal.ROUND_HALF_DOWN = 5; - Decimal.ROUND_HALF_EVEN = 6; - Decimal.ROUND_HALF_CEIL = 7; - Decimal.ROUND_HALF_FLOOR = 8; - - Decimal.clone = clone; - Decimal.config = Decimal.set = config; - - if (obj === void 0) obj = {}; - if (obj) { - ps = ['precision', 'rounding', 'toExpNeg', 'toExpPos', 'LN10']; - for (i = 0; i < ps.length;) if (!obj.hasOwnProperty(p = ps[i++])) obj[p] = this[p]; - } - - Decimal.config(obj); - - return Decimal; - } - - - /* - * Configure global settings for a Decimal constructor. - * - * `obj` is an object with one or more of the following properties, - * - * precision {number} - * rounding {number} - * toExpNeg {number} - * toExpPos {number} - * - * E.g. Decimal.config({ precision: 20, rounding: 4 }) - * - */ - function config(obj) { - if (!obj || typeof obj !== 'object') { - throw Error(decimalError + 'Object expected'); - } - var i, p, v, - ps = [ - 'precision', 1, MAX_DIGITS, - 'rounding', 0, 8, - 'toExpNeg', -1 / 0, 0, - 'toExpPos', 0, 1 / 0 - ]; - - for (i = 0; i < ps.length; i += 3) { - if ((v = obj[p = ps[i]]) !== void 0) { - if (mathfloor(v) === v && v >= ps[i + 1] && v <= ps[i + 2]) this[p] = v; - else throw Error(invalidArgument + p + ': ' + v); - } - } - - if ((v = obj[p = 'LN10']) !== void 0) { - if (v == Math.LN10) this[p] = new this(v); - else throw Error(invalidArgument + p + ': ' + v); - } - - return this; - } - - - // Create and configure initial Decimal constructor. - Decimal = clone(Decimal); - - Decimal['default'] = Decimal.Decimal = Decimal; - - // Internal constant. - ONE = new Decimal(1); - - - // Export. - - - // AMD. - if (module.exports) { - module.exports = Decimal; - - // Browser. - } else { - if (!globalScope) { - globalScope = typeof self != 'undefined' && self && self.self == self - ? self : Function('return this')(); - } - - globalScope.Decimal = Decimal; - } - })(commonjsGlobal$1); -} (decimal)); - -var utils = {}; - -Object.defineProperty(utils, "__esModule", { - value: true -}); -utils.memoize = utils.reverse = utils.compose = utils.map = utils.range = utils.curry = utils.PLACE_HOLDER = void 0; - -function _toConsumableArray$7(arr) { return _arrayWithoutHoles$7(arr) || _iterableToArray$7(arr) || _unsupportedIterableToArray$b(arr) || _nonIterableSpread$7(); } - -function _nonIterableSpread$7() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$b(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$b(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$b(o, minLen); } - -function _iterableToArray$7(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$7(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$b(arr); } - -function _arrayLikeToArray$b(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -var identity$2 = function identity(i) { - return i; -}; - -var PLACE_HOLDER = { - '@@functional/placeholder': true -}; -utils.PLACE_HOLDER = PLACE_HOLDER; - -var isPlaceHolder = function isPlaceHolder(val) { - return val === PLACE_HOLDER; -}; - -var curry0 = function curry0(fn) { - return function _curried() { - if (arguments.length === 0 || arguments.length === 1 && isPlaceHolder(arguments.length <= 0 ? undefined : arguments[0])) { - return _curried; - } - - return fn.apply(void 0, arguments); - }; -}; - -var curryN = function curryN(n, fn) { - if (n === 1) { - return fn; - } - - return curry0(function () { - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - var argsLength = args.filter(function (arg) { - return arg !== PLACE_HOLDER; - }).length; - - if (argsLength >= n) { - return fn.apply(void 0, args); - } - - return curryN(n - argsLength, curry0(function () { - for (var _len2 = arguments.length, restArgs = new Array(_len2), _key2 = 0; _key2 < _len2; _key2++) { - restArgs[_key2] = arguments[_key2]; - } - - var newArgs = args.map(function (arg) { - return isPlaceHolder(arg) ? restArgs.shift() : arg; - }); - return fn.apply(void 0, _toConsumableArray$7(newArgs).concat(restArgs)); - })); - }); -}; - -var curry = function curry(fn) { - return curryN(fn.length, fn); -}; - -utils.curry = curry; - -var range$1 = function range(begin, end) { - var arr = []; - - for (var i = begin; i < end; ++i) { - arr[i - begin] = i; - } - - return arr; -}; - -utils.range = range$1; -var map = curry(function (fn, arr) { - if (Array.isArray(arr)) { - return arr.map(fn); - } - - return Object.keys(arr).map(function (key) { - return arr[key]; - }).map(fn); -}); -utils.map = map; - -var compose = function compose() { - for (var _len3 = arguments.length, args = new Array(_len3), _key3 = 0; _key3 < _len3; _key3++) { - args[_key3] = arguments[_key3]; - } - - if (!args.length) { - return identity$2; - } - - var fns = args.reverse(); // first function can receive multiply arguments - - var firstFn = fns[0]; - var tailsFn = fns.slice(1); - return function () { - return tailsFn.reduce(function (res, fn) { - return fn(res); - }, firstFn.apply(void 0, arguments)); - }; -}; - -utils.compose = compose; - -var reverse = function reverse(arr) { - if (Array.isArray(arr)) { - return arr.reverse(); - } // can be string - - - return arr.split('').reverse.join(''); -}; - -utils.reverse = reverse; - -var memoize = function memoize(fn) { - var lastArgs = null; - var lastResult = null; - return function () { - for (var _len4 = arguments.length, args = new Array(_len4), _key4 = 0; _key4 < _len4; _key4++) { - args[_key4] = arguments[_key4]; - } - - if (lastArgs && args.every(function (val, i) { - return val === lastArgs[i]; - })) { - return lastResult; - } - - lastArgs = args; - lastResult = fn.apply(void 0, args); - return lastResult; - }; -}; - -utils.memoize = memoize; - -var arithmetic = {}; - -Object.defineProperty(arithmetic, "__esModule", { - value: true -}); -arithmetic.default = void 0; - -var _decimal$1 = _interopRequireDefault$1(decimal.exports); - -var _utils$1 = utils; - -function _interopRequireDefault$1(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -/** - * @fileOverview 一些公用的运算方法 - * @author xile611 - * @date 2015-09-17 - */ - -/** - * 获取数值的位数 - * 其中绝对值属于区间[0.1, 1), 得到的值为0 - * 绝对值属于区间[0.01, 0.1),得到的位数为 -1 - * 绝对值属于区间[0.001, 0.01),得到的位数为 -2 - * - * @param {Number} value 数值 - * @return {Integer} 位数 - */ -function getDigitCount(value) { - var result; - - if (value === 0) { - result = 1; - } else { - result = Math.floor(new _decimal$1.default(value).abs().log(10).toNumber()) + 1; - } - - return result; -} -/** - * 按照固定的步长获取[start, end)这个区间的数据 - * 并且需要处理js计算精度的问题 - * - * @param {Decimal} start 起点 - * @param {Decimal} end 终点,不包含该值 - * @param {Decimal} step 步长 - * @return {Array} 若干数值 - */ - - -function rangeStep(start, end, step) { - var num = new _decimal$1.default(start); - var i = 0; - var result = []; // magic number to prevent infinite loop - - while (num.lt(end) && i < 100000) { - result.push(num.toNumber()); - num = num.add(step); - i++; - } - - return result; -} -/** - * 对数值进行线性插值 - * - * @param {Number} a 定义域的极点 - * @param {Number} b 定义域的极点 - * @param {Number} t [0, 1]内的某个值 - * @return {Number} 定义域内的某个值 - */ - - -var interpolateNumber$1 = (0, _utils$1.curry)(function (a, b, t) { - var newA = +a; - var newB = +b; - return newA + t * (newB - newA); -}); -/** - * 线性插值的逆运算 - * - * @param {Number} a 定义域的极点 - * @param {Number} b 定义域的极点 - * @param {Number} x 可以认为是插值后的一个输出值 - * @return {Number} 当x在 a ~ b这个范围内时,返回值属于[0, 1] - */ - -var uninterpolateNumber = (0, _utils$1.curry)(function (a, b, x) { - var diff = b - +a; - diff = diff || Infinity; - return (x - a) / diff; -}); -/** - * 线性插值的逆运算,并且有截断的操作 - * - * @param {Number} a 定义域的极点 - * @param {Number} b 定义域的极点 - * @param {Number} x 可以认为是插值后的一个输出值 - * @return {Number} 当x在 a ~ b这个区间内时,返回值属于[0, 1], - * 当x不在 a ~ b这个区间时,会截断到 a ~ b 这个区间 - */ - -var uninterpolateTruncation = (0, _utils$1.curry)(function (a, b, x) { - var diff = b - +a; - diff = diff || Infinity; - return Math.max(0, Math.min(1, (x - a) / diff)); -}); -var _default = { - rangeStep: rangeStep, - getDigitCount: getDigitCount, - interpolateNumber: interpolateNumber$1, - uninterpolateNumber: uninterpolateNumber, - uninterpolateTruncation: uninterpolateTruncation -}; -arithmetic.default = _default; - -Object.defineProperty(getNiceTickValues$1, "__esModule", { - value: true -}); -getNiceTickValues$1.getTickValuesFixedDomain = getNiceTickValues$1.getTickValues = getNiceTickValues$1.getNiceTickValues = void 0; - -var _decimal = _interopRequireDefault(decimal.exports); - -var _utils = utils; - -var _arithmetic = _interopRequireDefault(arithmetic); - -function _interopRequireDefault(obj) { return obj && obj.__esModule ? obj : { default: obj }; } - -function _toConsumableArray$6(arr) { return _arrayWithoutHoles$6(arr) || _iterableToArray$6(arr) || _unsupportedIterableToArray$a(arr) || _nonIterableSpread$6(); } - -function _nonIterableSpread$6() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _iterableToArray$6(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$6(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$a(arr); } - -function _slicedToArray$5(arr, i) { return _arrayWithHoles$5(arr) || _iterableToArrayLimit$5(arr, i) || _unsupportedIterableToArray$a(arr, i) || _nonIterableRest$5(); } - -function _nonIterableRest$5() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$a(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$a(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$a(o, minLen); } - -function _arrayLikeToArray$a(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit$5(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$5(arr) { if (Array.isArray(arr)) return arr; } - -/** - * Calculate a interval of a minimum value and a maximum value - * - * @param {Number} min The minimum value - * @param {Number} max The maximum value - * @return {Array} An interval - */ -function getValidInterval(_ref) { - var _ref2 = _slicedToArray$5(_ref, 2), - min = _ref2[0], - max = _ref2[1]; - - var validMin = min, - validMax = max; // exchange - - if (min > max) { - validMin = max; - validMax = min; - } - - return [validMin, validMax]; -} -/** - * Calculate the step which is easy to understand between ticks, like 10, 20, 25 - * - * @param {Decimal} roughStep The rough step calculated by deviding the - * difference by the tickCount - * @param {Boolean} allowDecimals Allow the ticks to be decimals or not - * @param {Integer} correctionFactor A correction factor - * @return {Decimal} The step which is easy to understand between two ticks - */ - - -function getFormatStep(roughStep, allowDecimals, correctionFactor) { - if (roughStep.lte(0)) { - return new _decimal.default(0); - } - - var digitCount = _arithmetic.default.getDigitCount(roughStep.toNumber()); // The ratio between the rough step and the smallest number which has a bigger - // order of magnitudes than the rough step - - - var digitCountValue = new _decimal.default(10).pow(digitCount); - var stepRatio = roughStep.div(digitCountValue); // When an integer and a float multiplied, the accuracy of result may be wrong - - var stepRatioScale = digitCount !== 1 ? 0.05 : 0.1; - var amendStepRatio = new _decimal.default(Math.ceil(stepRatio.div(stepRatioScale).toNumber())).add(correctionFactor).mul(stepRatioScale); - var formatStep = amendStepRatio.mul(digitCountValue); - return allowDecimals ? formatStep : new _decimal.default(Math.ceil(formatStep)); -} -/** - * calculate the ticks when the minimum value equals to the maximum value - * - * @param {Number} value The minimum valuue which is also the maximum value - * @param {Integer} tickCount The count of ticks - * @param {Boolean} allowDecimals Allow the ticks to be decimals or not - * @return {Array} ticks - */ - - -function getTickOfSingleValue(value, tickCount, allowDecimals) { - var step = 1; // calculate the middle value of ticks - - var middle = new _decimal.default(value); - - if (!middle.isint() && allowDecimals) { - var absVal = Math.abs(value); - - if (absVal < 1) { - // The step should be a float number when the difference is smaller than 1 - step = new _decimal.default(10).pow(_arithmetic.default.getDigitCount(value) - 1); - middle = new _decimal.default(Math.floor(middle.div(step).toNumber())).mul(step); - } else if (absVal > 1) { - // Return the maximum integer which is smaller than 'value' when 'value' is greater than 1 - middle = new _decimal.default(Math.floor(value)); - } - } else if (value === 0) { - middle = new _decimal.default(Math.floor((tickCount - 1) / 2)); - } else if (!allowDecimals) { - middle = new _decimal.default(Math.floor(value)); - } - - var middleIndex = Math.floor((tickCount - 1) / 2); - var fn = (0, _utils.compose)((0, _utils.map)(function (n) { - return middle.add(new _decimal.default(n - middleIndex).mul(step)).toNumber(); - }), _utils.range); - return fn(0, tickCount); -} -/** - * Calculate the step - * - * @param {Number} min The minimum value of an interval - * @param {Number} max The maximum value of an interval - * @param {Integer} tickCount The count of ticks - * @param {Boolean} allowDecimals Allow the ticks to be decimals or not - * @param {Number} correctionFactor A correction factor - * @return {Object} The step, minimum value of ticks, maximum value of ticks - */ - - -function calculateStep(min, max, tickCount, allowDecimals) { - var correctionFactor = arguments.length > 4 && arguments[4] !== undefined ? arguments[4] : 0; - - // dirty hack (for recharts' test) - if (!Number.isFinite((max - min) / (tickCount - 1))) { - return { - step: new _decimal.default(0), - tickMin: new _decimal.default(0), - tickMax: new _decimal.default(0) - }; - } // The step which is easy to understand between two ticks - - - var step = getFormatStep(new _decimal.default(max).sub(min).div(tickCount - 1), allowDecimals, correctionFactor); // A medial value of ticks - - var middle; // When 0 is inside the interval, 0 should be a tick - - if (min <= 0 && max >= 0) { - middle = new _decimal.default(0); - } else { - // calculate the middle value - middle = new _decimal.default(min).add(max).div(2); // minus modulo value - - middle = middle.sub(new _decimal.default(middle).mod(step)); - } - - var belowCount = Math.ceil(middle.sub(min).div(step).toNumber()); - var upCount = Math.ceil(new _decimal.default(max).sub(middle).div(step).toNumber()); - var scaleCount = belowCount + upCount + 1; - - if (scaleCount > tickCount) { - // When more ticks need to cover the interval, step should be bigger. - return calculateStep(min, max, tickCount, allowDecimals, correctionFactor + 1); - } - - if (scaleCount < tickCount) { - // When less ticks can cover the interval, we should add some additional ticks - upCount = max > 0 ? upCount + (tickCount - scaleCount) : upCount; - belowCount = max > 0 ? belowCount : belowCount + (tickCount - scaleCount); - } - - return { - step: step, - tickMin: middle.sub(new _decimal.default(belowCount).mul(step)), - tickMax: middle.add(new _decimal.default(upCount).mul(step)) - }; -} -/** - * Calculate the ticks of an interval, the count of ticks will be guraranteed - * - * @param {Number} min, max min: The minimum value, max: The maximum value - * @param {Integer} tickCount The count of ticks - * @param {Boolean} allowDecimals Allow the ticks to be decimals or not - * @return {Array} ticks - */ - - -function getNiceTickValuesFn(_ref3) { - var _ref4 = _slicedToArray$5(_ref3, 2), - min = _ref4[0], - max = _ref4[1]; - - var tickCount = arguments.length > 1 && arguments[1] !== undefined ? arguments[1] : 6; - var allowDecimals = arguments.length > 2 && arguments[2] !== undefined ? arguments[2] : true; - // More than two ticks should be return - var count = Math.max(tickCount, 2); - - var _getValidInterval = getValidInterval([min, max]), - _getValidInterval2 = _slicedToArray$5(_getValidInterval, 2), - cormin = _getValidInterval2[0], - cormax = _getValidInterval2[1]; - - if (cormin === -Infinity || cormax === Infinity) { - var _values = cormax === Infinity ? [cormin].concat(_toConsumableArray$6((0, _utils.range)(0, tickCount - 1).map(function () { - return Infinity; - }))) : [].concat(_toConsumableArray$6((0, _utils.range)(0, tickCount - 1).map(function () { - return -Infinity; - })), [cormax]); - - return min > max ? (0, _utils.reverse)(_values) : _values; - } - - if (cormin === cormax) { - return getTickOfSingleValue(cormin, tickCount, allowDecimals); - } // Get the step between two ticks - - - var _calculateStep = calculateStep(cormin, cormax, count, allowDecimals), - step = _calculateStep.step, - tickMin = _calculateStep.tickMin, - tickMax = _calculateStep.tickMax; - - var values = _arithmetic.default.rangeStep(tickMin, tickMax.add(new _decimal.default(0.1).mul(step)), step); - - return min > max ? (0, _utils.reverse)(values) : values; -} -/** - * Calculate the ticks of an interval, the count of ticks won't be guraranteed - * - * @param {Number} min, max min: The minimum value, max: The maximum value - * @param {Integer} tickCount The count of ticks - * @param {Boolean} allowDecimals Allow the ticks to be decimals or not - * @return {Array} ticks - */ - - -function getTickValuesFn(_ref5) { - var _ref6 = _slicedToArray$5(_ref5, 2), - min = _ref6[0], - max = _ref6[1]; - - var tickCount = arguments.length > 1 && arguments[1] !== undefined ? arguments[1] : 6; - var allowDecimals = arguments.length > 2 && arguments[2] !== undefined ? arguments[2] : true; - // More than two ticks should be return - var count = Math.max(tickCount, 2); - - var _getValidInterval3 = getValidInterval([min, max]), - _getValidInterval4 = _slicedToArray$5(_getValidInterval3, 2), - cormin = _getValidInterval4[0], - cormax = _getValidInterval4[1]; - - if (cormin === -Infinity || cormax === Infinity) { - return [min, max]; - } - - if (cormin === cormax) { - return getTickOfSingleValue(cormin, tickCount, allowDecimals); - } - - var step = getFormatStep(new _decimal.default(cormax).sub(cormin).div(count - 1), allowDecimals, 0); - var fn = (0, _utils.compose)((0, _utils.map)(function (n) { - return new _decimal.default(cormin).add(new _decimal.default(n).mul(step)).toNumber(); - }), _utils.range); - var values = fn(0, count).filter(function (entry) { - return entry >= cormin && entry <= cormax; - }); - return min > max ? (0, _utils.reverse)(values) : values; -} -/** - * Calculate the ticks of an interval, the count of ticks won't be guraranteed, - * but the domain will be guaranteed - * - * @param {Number} min, max min: The minimum value, max: The maximum value - * @param {Integer} tickCount The count of ticks - * @param {Boolean} allowDecimals Allow the ticks to be decimals or not - * @return {Array} ticks - */ - - -function getTickValuesFixedDomainFn(_ref7, tickCount) { - var _ref8 = _slicedToArray$5(_ref7, 2), - min = _ref8[0], - max = _ref8[1]; - - var allowDecimals = arguments.length > 2 && arguments[2] !== undefined ? arguments[2] : true; - - // More than two ticks should be return - var _getValidInterval5 = getValidInterval([min, max]), - _getValidInterval6 = _slicedToArray$5(_getValidInterval5, 2), - cormin = _getValidInterval6[0], - cormax = _getValidInterval6[1]; - - if (cormin === -Infinity || cormax === Infinity) { - return [min, max]; - } - - if (cormin === cormax) { - return [cormin]; - } - - var count = Math.max(tickCount, 2); - var step = getFormatStep(new _decimal.default(cormax).sub(cormin).div(count - 1), allowDecimals, 0); - var values = [].concat(_toConsumableArray$6(_arithmetic.default.rangeStep(new _decimal.default(cormin), new _decimal.default(cormax).sub(new _decimal.default(0.99).mul(step)), step)), [cormax]); - return min > max ? (0, _utils.reverse)(values) : values; -} - -var getNiceTickValues = (0, _utils.memoize)(getNiceTickValuesFn); -getNiceTickValues$1.getNiceTickValues = getNiceTickValues; -var getTickValues = (0, _utils.memoize)(getTickValuesFn); -getNiceTickValues$1.getTickValues = getTickValues; -var getTickValuesFixedDomain = (0, _utils.memoize)(getTickValuesFixedDomainFn); -getNiceTickValues$1.getTickValuesFixedDomain = getTickValuesFixedDomain; - -(function (exports) { - - Object.defineProperty(exports, "__esModule", { - value: true - }); - Object.defineProperty(exports, "getTickValues", { - enumerable: true, - get: function get() { - return _getNiceTickValues.getTickValues; - } - }); - Object.defineProperty(exports, "getNiceTickValues", { - enumerable: true, - get: function get() { - return _getNiceTickValues.getNiceTickValues; - } - }); - Object.defineProperty(exports, "getTickValuesFixedDomain", { - enumerable: true, - get: function get() { - return _getNiceTickValues.getTickValuesFixedDomain; - } - }); - - var _getNiceTickValues = getNiceTickValues$1; -} (lib)); - -var d3Array = {exports: {}}; - -(function (module, exports) { - // https://d3js.org/d3-array/ v2.12.1 Copyright 2021 Mike Bostock - (function (global, factory) { - factory(exports) ; - }(commonjsGlobal$1, (function (exports) { - function ascending(a, b) { - return a < b ? -1 : a > b ? 1 : a >= b ? 0 : NaN; - } - - function bisector(f) { - let delta = f; - let compare = f; - - if (f.length === 1) { - delta = (d, x) => f(d) - x; - compare = ascendingComparator(f); - } - - function left(a, x, lo, hi) { - if (lo == null) lo = 0; - if (hi == null) hi = a.length; - while (lo < hi) { - const mid = (lo + hi) >>> 1; - if (compare(a[mid], x) < 0) lo = mid + 1; - else hi = mid; - } - return lo; - } - - function right(a, x, lo, hi) { - if (lo == null) lo = 0; - if (hi == null) hi = a.length; - while (lo < hi) { - const mid = (lo + hi) >>> 1; - if (compare(a[mid], x) > 0) hi = mid; - else lo = mid + 1; - } - return lo; - } - - function center(a, x, lo, hi) { - if (lo == null) lo = 0; - if (hi == null) hi = a.length; - const i = left(a, x, lo, hi - 1); - return i > lo && delta(a[i - 1], x) > -delta(a[i], x) ? i - 1 : i; - } - - return {left, center, right}; - } - - function ascendingComparator(f) { - return (d, x) => ascending(f(d), x); - } - - function number(x) { - return x === null ? NaN : +x; - } - - function* numbers(values, valueof) { - if (valueof === undefined) { - for (let value of values) { - if (value != null && (value = +value) >= value) { - yield value; - } - } - } else { - let index = -1; - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null && (value = +value) >= value) { - yield value; - } - } - } - } - - const ascendingBisect = bisector(ascending); - const bisectRight = ascendingBisect.right; - const bisectLeft = ascendingBisect.left; - const bisectCenter = bisector(number).center; - - function count(values, valueof) { - let count = 0; - if (valueof === undefined) { - for (let value of values) { - if (value != null && (value = +value) >= value) { - ++count; - } - } - } else { - let index = -1; - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null && (value = +value) >= value) { - ++count; - } - } - } - return count; - } - - function length$1(array) { - return array.length | 0; - } - - function empty(length) { - return !(length > 0); - } - - function arrayify(values) { - return typeof values !== "object" || "length" in values ? values : Array.from(values); - } - - function reducer(reduce) { - return values => reduce(...values); - } - - function cross(...values) { - const reduce = typeof values[values.length - 1] === "function" && reducer(values.pop()); - values = values.map(arrayify); - const lengths = values.map(length$1); - const j = values.length - 1; - const index = new Array(j + 1).fill(0); - const product = []; - if (j < 0 || lengths.some(empty)) return product; - while (true) { - product.push(index.map((j, i) => values[i][j])); - let i = j; - while (++index[i] === lengths[i]) { - if (i === 0) return reduce ? product.map(reduce) : product; - index[i--] = 0; - } - } - } - - function cumsum(values, valueof) { - var sum = 0, index = 0; - return Float64Array.from(values, valueof === undefined - ? v => (sum += +v || 0) - : v => (sum += +valueof(v, index++, values) || 0)); - } - - function descending(a, b) { - return b < a ? -1 : b > a ? 1 : b >= a ? 0 : NaN; - } - - function variance(values, valueof) { - let count = 0; - let delta; - let mean = 0; - let sum = 0; - if (valueof === undefined) { - for (let value of values) { - if (value != null && (value = +value) >= value) { - delta = value - mean; - mean += delta / ++count; - sum += delta * (value - mean); - } - } - } else { - let index = -1; - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null && (value = +value) >= value) { - delta = value - mean; - mean += delta / ++count; - sum += delta * (value - mean); - } - } - } - if (count > 1) return sum / (count - 1); - } - - function deviation(values, valueof) { - const v = variance(values, valueof); - return v ? Math.sqrt(v) : v; - } - - function extent(values, valueof) { - let min; - let max; - if (valueof === undefined) { - for (const value of values) { - if (value != null) { - if (min === undefined) { - if (value >= value) min = max = value; - } else { - if (min > value) min = value; - if (max < value) max = value; - } - } - } - } else { - let index = -1; - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null) { - if (min === undefined) { - if (value >= value) min = max = value; - } else { - if (min > value) min = value; - if (max < value) max = value; - } - } - } - } - return [min, max]; - } - - // https://github.com/python/cpython/blob/a74eea238f5baba15797e2e8b570d153bc8690a7/Modules/mathmodule.c#L1423 - class Adder { - constructor() { - this._partials = new Float64Array(32); - this._n = 0; - } - add(x) { - const p = this._partials; - let i = 0; - for (let j = 0; j < this._n && j < 32; j++) { - const y = p[j], - hi = x + y, - lo = Math.abs(x) < Math.abs(y) ? x - (hi - y) : y - (hi - x); - if (lo) p[i++] = lo; - x = hi; - } - p[i] = x; - this._n = i + 1; - return this; - } - valueOf() { - const p = this._partials; - let n = this._n, x, y, lo, hi = 0; - if (n > 0) { - hi = p[--n]; - while (n > 0) { - x = hi; - y = p[--n]; - hi = x + y; - lo = y - (hi - x); - if (lo) break; - } - if (n > 0 && ((lo < 0 && p[n - 1] < 0) || (lo > 0 && p[n - 1] > 0))) { - y = lo * 2; - x = hi + y; - if (y == x - hi) hi = x; - } - } - return hi; - } - } - - function fsum(values, valueof) { - const adder = new Adder(); - if (valueof === undefined) { - for (let value of values) { - if (value = +value) { - adder.add(value); - } - } - } else { - let index = -1; - for (let value of values) { - if (value = +valueof(value, ++index, values)) { - adder.add(value); - } - } - } - return +adder; - } - - function fcumsum(values, valueof) { - const adder = new Adder(); - let index = -1; - return Float64Array.from(values, valueof === undefined - ? v => adder.add(+v || 0) - : v => adder.add(+valueof(v, ++index, values) || 0) - ); - } - - class InternMap extends Map { - constructor(entries, key = keyof) { - super(); - Object.defineProperties(this, {_intern: {value: new Map()}, _key: {value: key}}); - if (entries != null) for (const [key, value] of entries) this.set(key, value); - } - get(key) { - return super.get(intern_get(this, key)); - } - has(key) { - return super.has(intern_get(this, key)); - } - set(key, value) { - return super.set(intern_set(this, key), value); - } - delete(key) { - return super.delete(intern_delete(this, key)); - } - } - - class InternSet extends Set { - constructor(values, key = keyof) { - super(); - Object.defineProperties(this, {_intern: {value: new Map()}, _key: {value: key}}); - if (values != null) for (const value of values) this.add(value); - } - has(value) { - return super.has(intern_get(this, value)); - } - add(value) { - return super.add(intern_set(this, value)); - } - delete(value) { - return super.delete(intern_delete(this, value)); - } - } - - function intern_get({_intern, _key}, value) { - const key = _key(value); - return _intern.has(key) ? _intern.get(key) : value; - } - - function intern_set({_intern, _key}, value) { - const key = _key(value); - if (_intern.has(key)) return _intern.get(key); - _intern.set(key, value); - return value; - } - - function intern_delete({_intern, _key}, value) { - const key = _key(value); - if (_intern.has(key)) { - value = _intern.get(value); - _intern.delete(key); - } - return value; - } - - function keyof(value) { - return value !== null && typeof value === "object" ? value.valueOf() : value; - } - - function identity(x) { - return x; - } - - function group(values, ...keys) { - return nest(values, identity, identity, keys); - } - - function groups(values, ...keys) { - return nest(values, Array.from, identity, keys); - } - - function rollup(values, reduce, ...keys) { - return nest(values, identity, reduce, keys); - } - - function rollups(values, reduce, ...keys) { - return nest(values, Array.from, reduce, keys); - } - - function index(values, ...keys) { - return nest(values, identity, unique, keys); - } - - function indexes(values, ...keys) { - return nest(values, Array.from, unique, keys); - } - - function unique(values) { - if (values.length !== 1) throw new Error("duplicate key"); - return values[0]; - } - - function nest(values, map, reduce, keys) { - return (function regroup(values, i) { - if (i >= keys.length) return reduce(values); - const groups = new InternMap(); - const keyof = keys[i++]; - let index = -1; - for (const value of values) { - const key = keyof(value, ++index, values); - const group = groups.get(key); - if (group) group.push(value); - else groups.set(key, [value]); - } - for (const [key, values] of groups) { - groups.set(key, regroup(values, i)); - } - return map(groups); - })(values, 0); - } - - function permute(source, keys) { - return Array.from(keys, key => source[key]); - } - - function sort(values, ...F) { - if (typeof values[Symbol.iterator] !== "function") throw new TypeError("values is not iterable"); - values = Array.from(values); - let [f = ascending] = F; - if (f.length === 1 || F.length > 1) { - const index = Uint32Array.from(values, (d, i) => i); - if (F.length > 1) { - F = F.map(f => values.map(f)); - index.sort((i, j) => { - for (const f of F) { - const c = ascending(f[i], f[j]); - if (c) return c; - } - }); - } else { - f = values.map(f); - index.sort((i, j) => ascending(f[i], f[j])); - } - return permute(values, index); - } - return values.sort(f); - } - - function groupSort(values, reduce, key) { - return (reduce.length === 1 - ? sort(rollup(values, reduce, key), (([ak, av], [bk, bv]) => ascending(av, bv) || ascending(ak, bk))) - : sort(group(values, key), (([ak, av], [bk, bv]) => reduce(av, bv) || ascending(ak, bk)))) - .map(([key]) => key); - } - - var array = Array.prototype; - - var slice = array.slice; - - function constant(x) { - return function() { - return x; - }; - } - - var e10 = Math.sqrt(50), - e5 = Math.sqrt(10), - e2 = Math.sqrt(2); - - function ticks(start, stop, count) { - var reverse, - i = -1, - n, - ticks, - step; - - stop = +stop, start = +start, count = +count; - if (start === stop && count > 0) return [start]; - if (reverse = stop < start) n = start, start = stop, stop = n; - if ((step = tickIncrement(start, stop, count)) === 0 || !isFinite(step)) return []; - - if (step > 0) { - let r0 = Math.round(start / step), r1 = Math.round(stop / step); - if (r0 * step < start) ++r0; - if (r1 * step > stop) --r1; - ticks = new Array(n = r1 - r0 + 1); - while (++i < n) ticks[i] = (r0 + i) * step; - } else { - step = -step; - let r0 = Math.round(start * step), r1 = Math.round(stop * step); - if (r0 / step < start) ++r0; - if (r1 / step > stop) --r1; - ticks = new Array(n = r1 - r0 + 1); - while (++i < n) ticks[i] = (r0 + i) / step; - } - - if (reverse) ticks.reverse(); - - return ticks; - } - - function tickIncrement(start, stop, count) { - var step = (stop - start) / Math.max(0, count), - power = Math.floor(Math.log(step) / Math.LN10), - error = step / Math.pow(10, power); - return power >= 0 - ? (error >= e10 ? 10 : error >= e5 ? 5 : error >= e2 ? 2 : 1) * Math.pow(10, power) - : -Math.pow(10, -power) / (error >= e10 ? 10 : error >= e5 ? 5 : error >= e2 ? 2 : 1); - } - - function tickStep(start, stop, count) { - var step0 = Math.abs(stop - start) / Math.max(0, count), - step1 = Math.pow(10, Math.floor(Math.log(step0) / Math.LN10)), - error = step0 / step1; - if (error >= e10) step1 *= 10; - else if (error >= e5) step1 *= 5; - else if (error >= e2) step1 *= 2; - return stop < start ? -step1 : step1; - } - - function nice(start, stop, count) { - let prestep; - while (true) { - const step = tickIncrement(start, stop, count); - if (step === prestep || step === 0 || !isFinite(step)) { - return [start, stop]; - } else if (step > 0) { - start = Math.floor(start / step) * step; - stop = Math.ceil(stop / step) * step; - } else if (step < 0) { - start = Math.ceil(start * step) / step; - stop = Math.floor(stop * step) / step; - } - prestep = step; - } - } - - function sturges(values) { - return Math.ceil(Math.log(count(values)) / Math.LN2) + 1; - } - - function bin() { - var value = identity, - domain = extent, - threshold = sturges; - - function histogram(data) { - if (!Array.isArray(data)) data = Array.from(data); - - var i, - n = data.length, - x, - values = new Array(n); - - for (i = 0; i < n; ++i) { - values[i] = value(data[i], i, data); - } - - var xz = domain(values), - x0 = xz[0], - x1 = xz[1], - tz = threshold(values, x0, x1); - - // Convert number of thresholds into uniform thresholds, and nice the - // default domain accordingly. - if (!Array.isArray(tz)) { - const max = x1, tn = +tz; - if (domain === extent) [x0, x1] = nice(x0, x1, tn); - tz = ticks(x0, x1, tn); - - // If the last threshold is coincident with the domain’s upper bound, the - // last bin will be zero-width. If the default domain is used, and this - // last threshold is coincident with the maximum input value, we can - // extend the niced upper bound by one tick to ensure uniform bin widths; - // otherwise, we simply remove the last threshold. Note that we don’t - // coerce values or the domain to numbers, and thus must be careful to - // compare order (>=) rather than strict equality (===)! - if (tz[tz.length - 1] >= x1) { - if (max >= x1 && domain === extent) { - const step = tickIncrement(x0, x1, tn); - if (isFinite(step)) { - if (step > 0) { - x1 = (Math.floor(x1 / step) + 1) * step; - } else if (step < 0) { - x1 = (Math.ceil(x1 * -step) + 1) / -step; - } - } - } else { - tz.pop(); - } - } - } - - // Remove any thresholds outside the domain. - var m = tz.length; - while (tz[0] <= x0) tz.shift(), --m; - while (tz[m - 1] > x1) tz.pop(), --m; - - var bins = new Array(m + 1), - bin; - - // Initialize bins. - for (i = 0; i <= m; ++i) { - bin = bins[i] = []; - bin.x0 = i > 0 ? tz[i - 1] : x0; - bin.x1 = i < m ? tz[i] : x1; - } - - // Assign data to bins by value, ignoring any outside the domain. - for (i = 0; i < n; ++i) { - x = values[i]; - if (x0 <= x && x <= x1) { - bins[bisectRight(tz, x, 0, m)].push(data[i]); - } - } - - return bins; - } - - histogram.value = function(_) { - return arguments.length ? (value = typeof _ === "function" ? _ : constant(_), histogram) : value; - }; - - histogram.domain = function(_) { - return arguments.length ? (domain = typeof _ === "function" ? _ : constant([_[0], _[1]]), histogram) : domain; - }; - - histogram.thresholds = function(_) { - return arguments.length ? (threshold = typeof _ === "function" ? _ : Array.isArray(_) ? constant(slice.call(_)) : constant(_), histogram) : threshold; - }; - - return histogram; - } - - function max(values, valueof) { - let max; - if (valueof === undefined) { - for (const value of values) { - if (value != null - && (max < value || (max === undefined && value >= value))) { - max = value; - } - } - } else { - let index = -1; - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null - && (max < value || (max === undefined && value >= value))) { - max = value; - } - } - } - return max; - } - - function min(values, valueof) { - let min; - if (valueof === undefined) { - for (const value of values) { - if (value != null - && (min > value || (min === undefined && value >= value))) { - min = value; - } - } - } else { - let index = -1; - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null - && (min > value || (min === undefined && value >= value))) { - min = value; - } - } - } - return min; - } - - // Based on https://github.com/mourner/quickselect - // ISC license, Copyright 2018 Vladimir Agafonkin. - function quickselect(array, k, left = 0, right = array.length - 1, compare = ascending) { - while (right > left) { - if (right - left > 600) { - const n = right - left + 1; - const m = k - left + 1; - const z = Math.log(n); - const s = 0.5 * Math.exp(2 * z / 3); - const sd = 0.5 * Math.sqrt(z * s * (n - s) / n) * (m - n / 2 < 0 ? -1 : 1); - const newLeft = Math.max(left, Math.floor(k - m * s / n + sd)); - const newRight = Math.min(right, Math.floor(k + (n - m) * s / n + sd)); - quickselect(array, k, newLeft, newRight, compare); - } - - const t = array[k]; - let i = left; - let j = right; - - swap(array, left, k); - if (compare(array[right], t) > 0) swap(array, left, right); - - while (i < j) { - swap(array, i, j), ++i, --j; - while (compare(array[i], t) < 0) ++i; - while (compare(array[j], t) > 0) --j; - } - - if (compare(array[left], t) === 0) swap(array, left, j); - else ++j, swap(array, j, right); - - if (j <= k) left = j + 1; - if (k <= j) right = j - 1; - } - return array; - } - - function swap(array, i, j) { - const t = array[i]; - array[i] = array[j]; - array[j] = t; - } - - function quantile(values, p, valueof) { - values = Float64Array.from(numbers(values, valueof)); - if (!(n = values.length)) return; - if ((p = +p) <= 0 || n < 2) return min(values); - if (p >= 1) return max(values); - var n, - i = (n - 1) * p, - i0 = Math.floor(i), - value0 = max(quickselect(values, i0).subarray(0, i0 + 1)), - value1 = min(values.subarray(i0 + 1)); - return value0 + (value1 - value0) * (i - i0); - } - - function quantileSorted(values, p, valueof = number) { - if (!(n = values.length)) return; - if ((p = +p) <= 0 || n < 2) return +valueof(values[0], 0, values); - if (p >= 1) return +valueof(values[n - 1], n - 1, values); - var n, - i = (n - 1) * p, - i0 = Math.floor(i), - value0 = +valueof(values[i0], i0, values), - value1 = +valueof(values[i0 + 1], i0 + 1, values); - return value0 + (value1 - value0) * (i - i0); - } - - function freedmanDiaconis(values, min, max) { - return Math.ceil((max - min) / (2 * (quantile(values, 0.75) - quantile(values, 0.25)) * Math.pow(count(values), -1 / 3))); - } - - function scott(values, min, max) { - return Math.ceil((max - min) / (3.5 * deviation(values) * Math.pow(count(values), -1 / 3))); - } - - function maxIndex(values, valueof) { - let max; - let maxIndex = -1; - let index = -1; - if (valueof === undefined) { - for (const value of values) { - ++index; - if (value != null - && (max < value || (max === undefined && value >= value))) { - max = value, maxIndex = index; - } - } - } else { - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null - && (max < value || (max === undefined && value >= value))) { - max = value, maxIndex = index; - } - } - } - return maxIndex; - } - - function mean(values, valueof) { - let count = 0; - let sum = 0; - if (valueof === undefined) { - for (let value of values) { - if (value != null && (value = +value) >= value) { - ++count, sum += value; - } - } - } else { - let index = -1; - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null && (value = +value) >= value) { - ++count, sum += value; - } - } - } - if (count) return sum / count; - } - - function median(values, valueof) { - return quantile(values, 0.5, valueof); - } - - function* flatten(arrays) { - for (const array of arrays) { - yield* array; - } - } - - function merge(arrays) { - return Array.from(flatten(arrays)); - } - - function minIndex(values, valueof) { - let min; - let minIndex = -1; - let index = -1; - if (valueof === undefined) { - for (const value of values) { - ++index; - if (value != null - && (min > value || (min === undefined && value >= value))) { - min = value, minIndex = index; - } - } - } else { - for (let value of values) { - if ((value = valueof(value, ++index, values)) != null - && (min > value || (min === undefined && value >= value))) { - min = value, minIndex = index; - } - } - } - return minIndex; - } - - function pairs(values, pairof = pair) { - const pairs = []; - let previous; - let first = false; - for (const value of values) { - if (first) pairs.push(pairof(previous, value)); - previous = value; - first = true; - } - return pairs; - } - - function pair(a, b) { - return [a, b]; - } - - function range(start, stop, step) { - start = +start, stop = +stop, step = (n = arguments.length) < 2 ? (stop = start, start = 0, 1) : n < 3 ? 1 : +step; - - var i = -1, - n = Math.max(0, Math.ceil((stop - start) / step)) | 0, - range = new Array(n); - - while (++i < n) { - range[i] = start + i * step; - } - - return range; - } - - function least(values, compare = ascending) { - let min; - let defined = false; - if (compare.length === 1) { - let minValue; - for (const element of values) { - const value = compare(element); - if (defined - ? ascending(value, minValue) < 0 - : ascending(value, value) === 0) { - min = element; - minValue = value; - defined = true; - } - } - } else { - for (const value of values) { - if (defined - ? compare(value, min) < 0 - : compare(value, value) === 0) { - min = value; - defined = true; - } - } - } - return min; - } - - function leastIndex(values, compare = ascending) { - if (compare.length === 1) return minIndex(values, compare); - let minValue; - let min = -1; - let index = -1; - for (const value of values) { - ++index; - if (min < 0 - ? compare(value, value) === 0 - : compare(value, minValue) < 0) { - minValue = value; - min = index; - } - } - return min; - } - - function greatest(values, compare = ascending) { - let max; - let defined = false; - if (compare.length === 1) { - let maxValue; - for (const element of values) { - const value = compare(element); - if (defined - ? ascending(value, maxValue) > 0 - : ascending(value, value) === 0) { - max = element; - maxValue = value; - defined = true; - } - } - } else { - for (const value of values) { - if (defined - ? compare(value, max) > 0 - : compare(value, value) === 0) { - max = value; - defined = true; - } - } - } - return max; - } - - function greatestIndex(values, compare = ascending) { - if (compare.length === 1) return maxIndex(values, compare); - let maxValue; - let max = -1; - let index = -1; - for (const value of values) { - ++index; - if (max < 0 - ? compare(value, value) === 0 - : compare(value, maxValue) > 0) { - maxValue = value; - max = index; - } - } - return max; - } - - function scan(values, compare) { - const index = leastIndex(values, compare); - return index < 0 ? undefined : index; - } - - var shuffle = shuffler(Math.random); - - function shuffler(random) { - return function shuffle(array, i0 = 0, i1 = array.length) { - let m = i1 - (i0 = +i0); - while (m) { - const i = random() * m-- | 0, t = array[m + i0]; - array[m + i0] = array[i + i0]; - array[i + i0] = t; - } - return array; - }; - } - - function sum(values, valueof) { - let sum = 0; - if (valueof === undefined) { - for (let value of values) { - if (value = +value) { - sum += value; - } - } - } else { - let index = -1; - for (let value of values) { - if (value = +valueof(value, ++index, values)) { - sum += value; - } - } - } - return sum; - } - - function transpose(matrix) { - if (!(n = matrix.length)) return []; - for (var i = -1, m = min(matrix, length), transpose = new Array(m); ++i < m;) { - for (var j = -1, n, row = transpose[i] = new Array(n); ++j < n;) { - row[j] = matrix[j][i]; - } - } - return transpose; - } - - function length(d) { - return d.length; - } - - function zip() { - return transpose(arguments); - } - - function every(values, test) { - if (typeof test !== "function") throw new TypeError("test is not a function"); - let index = -1; - for (const value of values) { - if (!test(value, ++index, values)) { - return false; - } - } - return true; - } - - function some(values, test) { - if (typeof test !== "function") throw new TypeError("test is not a function"); - let index = -1; - for (const value of values) { - if (test(value, ++index, values)) { - return true; - } - } - return false; - } - - function filter(values, test) { - if (typeof test !== "function") throw new TypeError("test is not a function"); - const array = []; - let index = -1; - for (const value of values) { - if (test(value, ++index, values)) { - array.push(value); - } - } - return array; - } - - function map(values, mapper) { - if (typeof values[Symbol.iterator] !== "function") throw new TypeError("values is not iterable"); - if (typeof mapper !== "function") throw new TypeError("mapper is not a function"); - return Array.from(values, (value, index) => mapper(value, index, values)); - } - - function reduce(values, reducer, value) { - if (typeof reducer !== "function") throw new TypeError("reducer is not a function"); - const iterator = values[Symbol.iterator](); - let done, next, index = -1; - if (arguments.length < 3) { - ({done, value} = iterator.next()); - if (done) return; - ++index; - } - while (({done, value: next} = iterator.next()), !done) { - value = reducer(value, next, ++index, values); - } - return value; - } - - function reverse(values) { - if (typeof values[Symbol.iterator] !== "function") throw new TypeError("values is not iterable"); - return Array.from(values).reverse(); - } - - function difference(values, ...others) { - values = new Set(values); - for (const other of others) { - for (const value of other) { - values.delete(value); - } - } - return values; - } - - function disjoint(values, other) { - const iterator = other[Symbol.iterator](), set = new Set(); - for (const v of values) { - if (set.has(v)) return false; - let value, done; - while (({value, done} = iterator.next())) { - if (done) break; - if (Object.is(v, value)) return false; - set.add(value); - } - } - return true; - } - - function set(values) { - return values instanceof Set ? values : new Set(values); - } - - function intersection(values, ...others) { - values = new Set(values); - others = others.map(set); - out: for (const value of values) { - for (const other of others) { - if (!other.has(value)) { - values.delete(value); - continue out; - } - } - } - return values; - } - - function superset(values, other) { - const iterator = values[Symbol.iterator](), set = new Set(); - for (const o of other) { - if (set.has(o)) continue; - let value, done; - while (({value, done} = iterator.next())) { - if (done) return false; - set.add(value); - if (Object.is(o, value)) break; - } - } - return true; - } - - function subset(values, other) { - return superset(other, values); - } - - function union(...others) { - const set = new Set(); - for (const other of others) { - for (const o of other) { - set.add(o); - } - } - return set; - } - - exports.Adder = Adder; - exports.InternMap = InternMap; - exports.InternSet = InternSet; - exports.ascending = ascending; - exports.bin = bin; - exports.bisect = bisectRight; - exports.bisectCenter = bisectCenter; - exports.bisectLeft = bisectLeft; - exports.bisectRight = bisectRight; - exports.bisector = bisector; - exports.count = count; - exports.cross = cross; - exports.cumsum = cumsum; - exports.descending = descending; - exports.deviation = deviation; - exports.difference = difference; - exports.disjoint = disjoint; - exports.every = every; - exports.extent = extent; - exports.fcumsum = fcumsum; - exports.filter = filter; - exports.fsum = fsum; - exports.greatest = greatest; - exports.greatestIndex = greatestIndex; - exports.group = group; - exports.groupSort = groupSort; - exports.groups = groups; - exports.histogram = bin; - exports.index = index; - exports.indexes = indexes; - exports.intersection = intersection; - exports.least = least; - exports.leastIndex = leastIndex; - exports.map = map; - exports.max = max; - exports.maxIndex = maxIndex; - exports.mean = mean; - exports.median = median; - exports.merge = merge; - exports.min = min; - exports.minIndex = minIndex; - exports.nice = nice; - exports.pairs = pairs; - exports.permute = permute; - exports.quantile = quantile; - exports.quantileSorted = quantileSorted; - exports.quickselect = quickselect; - exports.range = range; - exports.reduce = reduce; - exports.reverse = reverse; - exports.rollup = rollup; - exports.rollups = rollups; - exports.scan = scan; - exports.shuffle = shuffle; - exports.shuffler = shuffler; - exports.some = some; - exports.sort = sort; - exports.subset = subset; - exports.sum = sum; - exports.superset = superset; - exports.thresholdFreedmanDiaconis = freedmanDiaconis; - exports.thresholdScott = scott; - exports.thresholdSturges = sturges; - exports.tickIncrement = tickIncrement; - exports.tickStep = tickStep; - exports.ticks = ticks; - exports.transpose = transpose; - exports.union = union; - exports.variance = variance; - exports.zip = zip; - - Object.defineProperty(exports, '__esModule', { value: true }); - - }))); -} (d3Array, d3Array.exports)); - -function initRange(domain, range) { - switch (arguments.length) { - case 0: break; - case 1: this.range(domain); break; - default: this.range(range).domain(domain); break; - } - return this; -} - -function initInterpolator(domain, interpolator) { - switch (arguments.length) { - case 0: break; - case 1: { - if (typeof domain === "function") this.interpolator(domain); - else this.range(domain); - break; - } - default: { - this.domain(domain); - if (typeof interpolator === "function") this.interpolator(interpolator); - else this.range(interpolator); - break; - } - } - return this; -} - -const implicit = Symbol("implicit"); - -function ordinal() { - var index = new Map(), - domain = [], - range = [], - unknown = implicit; - - function scale(d) { - var key = d + "", i = index.get(key); - if (!i) { - if (unknown !== implicit) return unknown; - index.set(key, i = domain.push(d)); - } - return range[(i - 1) % range.length]; - } - - scale.domain = function(_) { - if (!arguments.length) return domain.slice(); - domain = [], index = new Map(); - for (const value of _) { - const key = value + ""; - if (index.has(key)) continue; - index.set(key, domain.push(value)); - } - return scale; - }; - - scale.range = function(_) { - return arguments.length ? (range = Array.from(_), scale) : range.slice(); - }; - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : unknown; - }; - - scale.copy = function() { - return ordinal(domain, range).unknown(unknown); - }; - - initRange.apply(scale, arguments); - - return scale; -} - -function band() { - var scale = ordinal().unknown(undefined), - domain = scale.domain, - ordinalRange = scale.range, - r0 = 0, - r1 = 1, - step, - bandwidth, - round = false, - paddingInner = 0, - paddingOuter = 0, - align = 0.5; - - delete scale.unknown; - - function rescale() { - var n = domain().length, - reverse = r1 < r0, - start = reverse ? r1 : r0, - stop = reverse ? r0 : r1; - step = (stop - start) / Math.max(1, n - paddingInner + paddingOuter * 2); - if (round) step = Math.floor(step); - start += (stop - start - step * (n - paddingInner)) * align; - bandwidth = step * (1 - paddingInner); - if (round) start = Math.round(start), bandwidth = Math.round(bandwidth); - var values = d3Array.exports.range(n).map(function(i) { return start + step * i; }); - return ordinalRange(reverse ? values.reverse() : values); - } - - scale.domain = function(_) { - return arguments.length ? (domain(_), rescale()) : domain(); - }; - - scale.range = function(_) { - return arguments.length ? ([r0, r1] = _, r0 = +r0, r1 = +r1, rescale()) : [r0, r1]; - }; - - scale.rangeRound = function(_) { - return [r0, r1] = _, r0 = +r0, r1 = +r1, round = true, rescale(); - }; - - scale.bandwidth = function() { - return bandwidth; - }; - - scale.step = function() { - return step; - }; - - scale.round = function(_) { - return arguments.length ? (round = !!_, rescale()) : round; - }; - - scale.padding = function(_) { - return arguments.length ? (paddingInner = Math.min(1, paddingOuter = +_), rescale()) : paddingInner; - }; - - scale.paddingInner = function(_) { - return arguments.length ? (paddingInner = Math.min(1, _), rescale()) : paddingInner; - }; - - scale.paddingOuter = function(_) { - return arguments.length ? (paddingOuter = +_, rescale()) : paddingOuter; - }; - - scale.align = function(_) { - return arguments.length ? (align = Math.max(0, Math.min(1, _)), rescale()) : align; - }; - - scale.copy = function() { - return band(domain(), [r0, r1]) - .round(round) - .paddingInner(paddingInner) - .paddingOuter(paddingOuter) - .align(align); - }; - - return initRange.apply(rescale(), arguments); -} - -function pointish(scale) { - var copy = scale.copy; - - scale.padding = scale.paddingOuter; - delete scale.paddingInner; - delete scale.paddingOuter; - - scale.copy = function() { - return pointish(copy()); - }; - - return scale; -} - -function point() { - return pointish(band.apply(null, arguments).paddingInner(1)); -} - -function define(constructor, factory, prototype) { - constructor.prototype = factory.prototype = prototype; - prototype.constructor = constructor; -} - -function extend(parent, definition) { - var prototype = Object.create(parent.prototype); - for (var key in definition) prototype[key] = definition[key]; - return prototype; -} - -function Color() {} - -var darker = 0.7; -var brighter = 1 / darker; - -var reI = "\\s*([+-]?\\d+)\\s*", - reN = "\\s*([+-]?\\d*\\.?\\d+(?:[eE][+-]?\\d+)?)\\s*", - reP = "\\s*([+-]?\\d*\\.?\\d+(?:[eE][+-]?\\d+)?)%\\s*", - reHex = /^#([0-9a-f]{3,8})$/, - reRgbInteger = new RegExp("^rgb\\(" + [reI, reI, reI] + "\\)$"), - reRgbPercent = new RegExp("^rgb\\(" + [reP, reP, reP] + "\\)$"), - reRgbaInteger = new RegExp("^rgba\\(" + [reI, reI, reI, reN] + "\\)$"), - reRgbaPercent = new RegExp("^rgba\\(" + [reP, reP, reP, reN] + "\\)$"), - reHslPercent = new RegExp("^hsl\\(" + [reN, reP, reP] + "\\)$"), - reHslaPercent = new RegExp("^hsla\\(" + [reN, reP, reP, reN] + "\\)$"); - -var named = { - aliceblue: 0xf0f8ff, - antiquewhite: 0xfaebd7, - aqua: 0x00ffff, - aquamarine: 0x7fffd4, - azure: 0xf0ffff, - beige: 0xf5f5dc, - bisque: 0xffe4c4, - black: 0x000000, - blanchedalmond: 0xffebcd, - blue: 0x0000ff, - blueviolet: 0x8a2be2, - brown: 0xa52a2a, - burlywood: 0xdeb887, - cadetblue: 0x5f9ea0, - chartreuse: 0x7fff00, - chocolate: 0xd2691e, - coral: 0xff7f50, - cornflowerblue: 0x6495ed, - cornsilk: 0xfff8dc, - crimson: 0xdc143c, - cyan: 0x00ffff, - darkblue: 0x00008b, - darkcyan: 0x008b8b, - darkgoldenrod: 0xb8860b, - darkgray: 0xa9a9a9, - darkgreen: 0x006400, - darkgrey: 0xa9a9a9, - darkkhaki: 0xbdb76b, - darkmagenta: 0x8b008b, - darkolivegreen: 0x556b2f, - darkorange: 0xff8c00, - darkorchid: 0x9932cc, - darkred: 0x8b0000, - darksalmon: 0xe9967a, - darkseagreen: 0x8fbc8f, - darkslateblue: 0x483d8b, - darkslategray: 0x2f4f4f, - darkslategrey: 0x2f4f4f, - darkturquoise: 0x00ced1, - darkviolet: 0x9400d3, - deeppink: 0xff1493, - deepskyblue: 0x00bfff, - dimgray: 0x696969, - dimgrey: 0x696969, - dodgerblue: 0x1e90ff, - firebrick: 0xb22222, - floralwhite: 0xfffaf0, - forestgreen: 0x228b22, - fuchsia: 0xff00ff, - gainsboro: 0xdcdcdc, - ghostwhite: 0xf8f8ff, - gold: 0xffd700, - goldenrod: 0xdaa520, - gray: 0x808080, - green: 0x008000, - greenyellow: 0xadff2f, - grey: 0x808080, - honeydew: 0xf0fff0, - hotpink: 0xff69b4, - indianred: 0xcd5c5c, - indigo: 0x4b0082, - ivory: 0xfffff0, - khaki: 0xf0e68c, - lavender: 0xe6e6fa, - lavenderblush: 0xfff0f5, - lawngreen: 0x7cfc00, - lemonchiffon: 0xfffacd, - lightblue: 0xadd8e6, - lightcoral: 0xf08080, - lightcyan: 0xe0ffff, - lightgoldenrodyellow: 0xfafad2, - lightgray: 0xd3d3d3, - lightgreen: 0x90ee90, - lightgrey: 0xd3d3d3, - lightpink: 0xffb6c1, - lightsalmon: 0xffa07a, - lightseagreen: 0x20b2aa, - lightskyblue: 0x87cefa, - lightslategray: 0x778899, - lightslategrey: 0x778899, - lightsteelblue: 0xb0c4de, - lightyellow: 0xffffe0, - lime: 0x00ff00, - limegreen: 0x32cd32, - linen: 0xfaf0e6, - magenta: 0xff00ff, - maroon: 0x800000, - mediumaquamarine: 0x66cdaa, - mediumblue: 0x0000cd, - mediumorchid: 0xba55d3, - mediumpurple: 0x9370db, - mediumseagreen: 0x3cb371, - mediumslateblue: 0x7b68ee, - mediumspringgreen: 0x00fa9a, - mediumturquoise: 0x48d1cc, - mediumvioletred: 0xc71585, - midnightblue: 0x191970, - mintcream: 0xf5fffa, - mistyrose: 0xffe4e1, - moccasin: 0xffe4b5, - navajowhite: 0xffdead, - navy: 0x000080, - oldlace: 0xfdf5e6, - olive: 0x808000, - olivedrab: 0x6b8e23, - orange: 0xffa500, - orangered: 0xff4500, - orchid: 0xda70d6, - palegoldenrod: 0xeee8aa, - palegreen: 0x98fb98, - paleturquoise: 0xafeeee, - palevioletred: 0xdb7093, - papayawhip: 0xffefd5, - peachpuff: 0xffdab9, - peru: 0xcd853f, - pink: 0xffc0cb, - plum: 0xdda0dd, - powderblue: 0xb0e0e6, - purple: 0x800080, - rebeccapurple: 0x663399, - red: 0xff0000, - rosybrown: 0xbc8f8f, - royalblue: 0x4169e1, - saddlebrown: 0x8b4513, - salmon: 0xfa8072, - sandybrown: 0xf4a460, - seagreen: 0x2e8b57, - seashell: 0xfff5ee, - sienna: 0xa0522d, - silver: 0xc0c0c0, - skyblue: 0x87ceeb, - slateblue: 0x6a5acd, - slategray: 0x708090, - slategrey: 0x708090, - snow: 0xfffafa, - springgreen: 0x00ff7f, - steelblue: 0x4682b4, - tan: 0xd2b48c, - teal: 0x008080, - thistle: 0xd8bfd8, - tomato: 0xff6347, - turquoise: 0x40e0d0, - violet: 0xee82ee, - wheat: 0xf5deb3, - white: 0xffffff, - whitesmoke: 0xf5f5f5, - yellow: 0xffff00, - yellowgreen: 0x9acd32 -}; - -define(Color, color, { - copy: function(channels) { - return Object.assign(new this.constructor, this, channels); - }, - displayable: function() { - return this.rgb().displayable(); - }, - hex: color_formatHex, // Deprecated! Use color.formatHex. - formatHex: color_formatHex, - formatHsl: color_formatHsl, - formatRgb: color_formatRgb, - toString: color_formatRgb -}); - -function color_formatHex() { - return this.rgb().formatHex(); -} - -function color_formatHsl() { - return hslConvert(this).formatHsl(); -} - -function color_formatRgb() { - return this.rgb().formatRgb(); -} - -function color(format) { - var m, l; - format = (format + "").trim().toLowerCase(); - return (m = reHex.exec(format)) ? (l = m[1].length, m = parseInt(m[1], 16), l === 6 ? rgbn(m) // #ff0000 - : l === 3 ? new Rgb((m >> 8 & 0xf) | (m >> 4 & 0xf0), (m >> 4 & 0xf) | (m & 0xf0), ((m & 0xf) << 4) | (m & 0xf), 1) // #f00 - : l === 8 ? rgba(m >> 24 & 0xff, m >> 16 & 0xff, m >> 8 & 0xff, (m & 0xff) / 0xff) // #ff000000 - : l === 4 ? rgba((m >> 12 & 0xf) | (m >> 8 & 0xf0), (m >> 8 & 0xf) | (m >> 4 & 0xf0), (m >> 4 & 0xf) | (m & 0xf0), (((m & 0xf) << 4) | (m & 0xf)) / 0xff) // #f000 - : null) // invalid hex - : (m = reRgbInteger.exec(format)) ? new Rgb(m[1], m[2], m[3], 1) // rgb(255, 0, 0) - : (m = reRgbPercent.exec(format)) ? new Rgb(m[1] * 255 / 100, m[2] * 255 / 100, m[3] * 255 / 100, 1) // rgb(100%, 0%, 0%) - : (m = reRgbaInteger.exec(format)) ? rgba(m[1], m[2], m[3], m[4]) // rgba(255, 0, 0, 1) - : (m = reRgbaPercent.exec(format)) ? rgba(m[1] * 255 / 100, m[2] * 255 / 100, m[3] * 255 / 100, m[4]) // rgb(100%, 0%, 0%, 1) - : (m = reHslPercent.exec(format)) ? hsla(m[1], m[2] / 100, m[3] / 100, 1) // hsl(120, 50%, 50%) - : (m = reHslaPercent.exec(format)) ? hsla(m[1], m[2] / 100, m[3] / 100, m[4]) // hsla(120, 50%, 50%, 1) - : named.hasOwnProperty(format) ? rgbn(named[format]) // eslint-disable-line no-prototype-builtins - : format === "transparent" ? new Rgb(NaN, NaN, NaN, 0) - : null; -} - -function rgbn(n) { - return new Rgb(n >> 16 & 0xff, n >> 8 & 0xff, n & 0xff, 1); -} - -function rgba(r, g, b, a) { - if (a <= 0) r = g = b = NaN; - return new Rgb(r, g, b, a); -} - -function rgbConvert(o) { - if (!(o instanceof Color)) o = color(o); - if (!o) return new Rgb; - o = o.rgb(); - return new Rgb(o.r, o.g, o.b, o.opacity); -} - -function rgb$1(r, g, b, opacity) { - return arguments.length === 1 ? rgbConvert(r) : new Rgb(r, g, b, opacity == null ? 1 : opacity); -} - -function Rgb(r, g, b, opacity) { - this.r = +r; - this.g = +g; - this.b = +b; - this.opacity = +opacity; -} - -define(Rgb, rgb$1, extend(Color, { - brighter: function(k) { - k = k == null ? brighter : Math.pow(brighter, k); - return new Rgb(this.r * k, this.g * k, this.b * k, this.opacity); - }, - darker: function(k) { - k = k == null ? darker : Math.pow(darker, k); - return new Rgb(this.r * k, this.g * k, this.b * k, this.opacity); - }, - rgb: function() { - return this; - }, - displayable: function() { - return (-0.5 <= this.r && this.r < 255.5) - && (-0.5 <= this.g && this.g < 255.5) - && (-0.5 <= this.b && this.b < 255.5) - && (0 <= this.opacity && this.opacity <= 1); - }, - hex: rgb_formatHex, // Deprecated! Use color.formatHex. - formatHex: rgb_formatHex, - formatRgb: rgb_formatRgb, - toString: rgb_formatRgb -})); - -function rgb_formatHex() { - return "#" + hex(this.r) + hex(this.g) + hex(this.b); -} - -function rgb_formatRgb() { - var a = this.opacity; a = isNaN(a) ? 1 : Math.max(0, Math.min(1, a)); - return (a === 1 ? "rgb(" : "rgba(") - + Math.max(0, Math.min(255, Math.round(this.r) || 0)) + ", " - + Math.max(0, Math.min(255, Math.round(this.g) || 0)) + ", " - + Math.max(0, Math.min(255, Math.round(this.b) || 0)) - + (a === 1 ? ")" : ", " + a + ")"); -} - -function hex(value) { - value = Math.max(0, Math.min(255, Math.round(value) || 0)); - return (value < 16 ? "0" : "") + value.toString(16); -} - -function hsla(h, s, l, a) { - if (a <= 0) h = s = l = NaN; - else if (l <= 0 || l >= 1) h = s = NaN; - else if (s <= 0) h = NaN; - return new Hsl(h, s, l, a); -} - -function hslConvert(o) { - if (o instanceof Hsl) return new Hsl(o.h, o.s, o.l, o.opacity); - if (!(o instanceof Color)) o = color(o); - if (!o) return new Hsl; - if (o instanceof Hsl) return o; - o = o.rgb(); - var r = o.r / 255, - g = o.g / 255, - b = o.b / 255, - min = Math.min(r, g, b), - max = Math.max(r, g, b), - h = NaN, - s = max - min, - l = (max + min) / 2; - if (s) { - if (r === max) h = (g - b) / s + (g < b) * 6; - else if (g === max) h = (b - r) / s + 2; - else h = (r - g) / s + 4; - s /= l < 0.5 ? max + min : 2 - max - min; - h *= 60; - } else { - s = l > 0 && l < 1 ? 0 : h; - } - return new Hsl(h, s, l, o.opacity); -} - -function hsl(h, s, l, opacity) { - return arguments.length === 1 ? hslConvert(h) : new Hsl(h, s, l, opacity == null ? 1 : opacity); -} - -function Hsl(h, s, l, opacity) { - this.h = +h; - this.s = +s; - this.l = +l; - this.opacity = +opacity; -} - -define(Hsl, hsl, extend(Color, { - brighter: function(k) { - k = k == null ? brighter : Math.pow(brighter, k); - return new Hsl(this.h, this.s, this.l * k, this.opacity); - }, - darker: function(k) { - k = k == null ? darker : Math.pow(darker, k); - return new Hsl(this.h, this.s, this.l * k, this.opacity); - }, - rgb: function() { - var h = this.h % 360 + (this.h < 0) * 360, - s = isNaN(h) || isNaN(this.s) ? 0 : this.s, - l = this.l, - m2 = l + (l < 0.5 ? l : 1 - l) * s, - m1 = 2 * l - m2; - return new Rgb( - hsl2rgb(h >= 240 ? h - 240 : h + 120, m1, m2), - hsl2rgb(h, m1, m2), - hsl2rgb(h < 120 ? h + 240 : h - 120, m1, m2), - this.opacity - ); - }, - displayable: function() { - return (0 <= this.s && this.s <= 1 || isNaN(this.s)) - && (0 <= this.l && this.l <= 1) - && (0 <= this.opacity && this.opacity <= 1); - }, - formatHsl: function() { - var a = this.opacity; a = isNaN(a) ? 1 : Math.max(0, Math.min(1, a)); - return (a === 1 ? "hsl(" : "hsla(") - + (this.h || 0) + ", " - + (this.s || 0) * 100 + "%, " - + (this.l || 0) * 100 + "%" - + (a === 1 ? ")" : ", " + a + ")"); - } -})); - -/* From FvD 13.37, CSS Color Module Level 3 */ -function hsl2rgb(h, m1, m2) { - return (h < 60 ? m1 + (m2 - m1) * h / 60 - : h < 180 ? m2 - : h < 240 ? m1 + (m2 - m1) * (240 - h) / 60 - : m1) * 255; -} - -const radians = Math.PI / 180; -const degrees = 180 / Math.PI; - -// https://observablehq.com/@mbostock/lab-and-rgb -const K = 18, - Xn = 0.96422, - Yn = 1, - Zn = 0.82521, - t0$1 = 4 / 29, - t1$1 = 6 / 29, - t2 = 3 * t1$1 * t1$1, - t3 = t1$1 * t1$1 * t1$1; - -function labConvert(o) { - if (o instanceof Lab) return new Lab(o.l, o.a, o.b, o.opacity); - if (o instanceof Hcl) return hcl2lab(o); - if (!(o instanceof Rgb)) o = rgbConvert(o); - var r = rgb2lrgb(o.r), - g = rgb2lrgb(o.g), - b = rgb2lrgb(o.b), - y = xyz2lab((0.2225045 * r + 0.7168786 * g + 0.0606169 * b) / Yn), x, z; - if (r === g && g === b) x = z = y; else { - x = xyz2lab((0.4360747 * r + 0.3850649 * g + 0.1430804 * b) / Xn); - z = xyz2lab((0.0139322 * r + 0.0971045 * g + 0.7141733 * b) / Zn); - } - return new Lab(116 * y - 16, 500 * (x - y), 200 * (y - z), o.opacity); -} - -function lab(l, a, b, opacity) { - return arguments.length === 1 ? labConvert(l) : new Lab(l, a, b, opacity == null ? 1 : opacity); -} - -function Lab(l, a, b, opacity) { - this.l = +l; - this.a = +a; - this.b = +b; - this.opacity = +opacity; -} - -define(Lab, lab, extend(Color, { - brighter: function(k) { - return new Lab(this.l + K * (k == null ? 1 : k), this.a, this.b, this.opacity); - }, - darker: function(k) { - return new Lab(this.l - K * (k == null ? 1 : k), this.a, this.b, this.opacity); - }, - rgb: function() { - var y = (this.l + 16) / 116, - x = isNaN(this.a) ? y : y + this.a / 500, - z = isNaN(this.b) ? y : y - this.b / 200; - x = Xn * lab2xyz(x); - y = Yn * lab2xyz(y); - z = Zn * lab2xyz(z); - return new Rgb( - lrgb2rgb( 3.1338561 * x - 1.6168667 * y - 0.4906146 * z), - lrgb2rgb(-0.9787684 * x + 1.9161415 * y + 0.0334540 * z), - lrgb2rgb( 0.0719453 * x - 0.2289914 * y + 1.4052427 * z), - this.opacity - ); - } -})); - -function xyz2lab(t) { - return t > t3 ? Math.pow(t, 1 / 3) : t / t2 + t0$1; -} - -function lab2xyz(t) { - return t > t1$1 ? t * t * t : t2 * (t - t0$1); -} - -function lrgb2rgb(x) { - return 255 * (x <= 0.0031308 ? 12.92 * x : 1.055 * Math.pow(x, 1 / 2.4) - 0.055); -} - -function rgb2lrgb(x) { - return (x /= 255) <= 0.04045 ? x / 12.92 : Math.pow((x + 0.055) / 1.055, 2.4); -} - -function hclConvert(o) { - if (o instanceof Hcl) return new Hcl(o.h, o.c, o.l, o.opacity); - if (!(o instanceof Lab)) o = labConvert(o); - if (o.a === 0 && o.b === 0) return new Hcl(NaN, 0 < o.l && o.l < 100 ? 0 : NaN, o.l, o.opacity); - var h = Math.atan2(o.b, o.a) * degrees; - return new Hcl(h < 0 ? h + 360 : h, Math.sqrt(o.a * o.a + o.b * o.b), o.l, o.opacity); -} - -function hcl(h, c, l, opacity) { - return arguments.length === 1 ? hclConvert(h) : new Hcl(h, c, l, opacity == null ? 1 : opacity); -} - -function Hcl(h, c, l, opacity) { - this.h = +h; - this.c = +c; - this.l = +l; - this.opacity = +opacity; -} - -function hcl2lab(o) { - if (isNaN(o.h)) return new Lab(o.l, 0, 0, o.opacity); - var h = o.h * radians; - return new Lab(o.l, Math.cos(h) * o.c, Math.sin(h) * o.c, o.opacity); -} - -define(Hcl, hcl, extend(Color, { - brighter: function(k) { - return new Hcl(this.h, this.c, this.l + K * (k == null ? 1 : k), this.opacity); - }, - darker: function(k) { - return new Hcl(this.h, this.c, this.l - K * (k == null ? 1 : k), this.opacity); - }, - rgb: function() { - return hcl2lab(this).rgb(); - } -})); - -var A = -0.14861, - B = +1.78277, - C = -0.29227, - D = -0.90649, - E = +1.97294, - ED = E * D, - EB = E * B, - BC_DA = B * C - D * A; - -function cubehelixConvert(o) { - if (o instanceof Cubehelix) return new Cubehelix(o.h, o.s, o.l, o.opacity); - if (!(o instanceof Rgb)) o = rgbConvert(o); - var r = o.r / 255, - g = o.g / 255, - b = o.b / 255, - l = (BC_DA * b + ED * r - EB * g) / (BC_DA + ED - EB), - bl = b - l, - k = (E * (g - l) - C * bl) / D, - s = Math.sqrt(k * k + bl * bl) / (E * l * (1 - l)), // NaN if l=0 or l=1 - h = s ? Math.atan2(k, bl) * degrees - 120 : NaN; - return new Cubehelix(h < 0 ? h + 360 : h, s, l, o.opacity); -} - -function cubehelix$1(h, s, l, opacity) { - return arguments.length === 1 ? cubehelixConvert(h) : new Cubehelix(h, s, l, opacity == null ? 1 : opacity); -} - -function Cubehelix(h, s, l, opacity) { - this.h = +h; - this.s = +s; - this.l = +l; - this.opacity = +opacity; -} - -define(Cubehelix, cubehelix$1, extend(Color, { - brighter: function(k) { - k = k == null ? brighter : Math.pow(brighter, k); - return new Cubehelix(this.h, this.s, this.l * k, this.opacity); - }, - darker: function(k) { - k = k == null ? darker : Math.pow(darker, k); - return new Cubehelix(this.h, this.s, this.l * k, this.opacity); - }, - rgb: function() { - var h = isNaN(this.h) ? 0 : (this.h + 120) * radians, - l = +this.l, - a = isNaN(this.s) ? 0 : this.s * l * (1 - l), - cosh = Math.cos(h), - sinh = Math.sin(h); - return new Rgb( - 255 * (l + a * (A * cosh + B * sinh)), - 255 * (l + a * (C * cosh + D * sinh)), - 255 * (l + a * (E * cosh)), - this.opacity - ); - } -})); - -var constant = x => () => x; - -function linear$1(a, d) { - return function(t) { - return a + t * d; - }; -} - -function exponential(a, b, y) { - return a = Math.pow(a, y), b = Math.pow(b, y) - a, y = 1 / y, function(t) { - return Math.pow(a + t * b, y); - }; -} - -function hue(a, b) { - var d = b - a; - return d ? linear$1(a, d > 180 || d < -180 ? d - 360 * Math.round(d / 360) : d) : constant(isNaN(a) ? b : a); -} - -function gamma(y) { - return (y = +y) === 1 ? nogamma : function(a, b) { - return b - a ? exponential(a, b, y) : constant(isNaN(a) ? b : a); - }; -} - -function nogamma(a, b) { - var d = b - a; - return d ? linear$1(a, d) : constant(isNaN(a) ? b : a); -} - -var rgb = (function rgbGamma(y) { - var color = gamma(y); - - function rgb(start, end) { - var r = color((start = rgb$1(start)).r, (end = rgb$1(end)).r), - g = color(start.g, end.g), - b = color(start.b, end.b), - opacity = nogamma(start.opacity, end.opacity); - return function(t) { - start.r = r(t); - start.g = g(t); - start.b = b(t); - start.opacity = opacity(t); - return start + ""; - }; - } - - rgb.gamma = rgbGamma; - - return rgb; -})(1); - -function numberArray(a, b) { - if (!b) b = []; - var n = a ? Math.min(b.length, a.length) : 0, - c = b.slice(), - i; - return function(t) { - for (i = 0; i < n; ++i) c[i] = a[i] * (1 - t) + b[i] * t; - return c; - }; -} - -function isNumberArray(x) { - return ArrayBuffer.isView(x) && !(x instanceof DataView); -} - -function genericArray(a, b) { - var nb = b ? b.length : 0, - na = a ? Math.min(nb, a.length) : 0, - x = new Array(na), - c = new Array(nb), - i; - - for (i = 0; i < na; ++i) x[i] = interpolate(a[i], b[i]); - for (; i < nb; ++i) c[i] = b[i]; - - return function(t) { - for (i = 0; i < na; ++i) c[i] = x[i](t); - return c; - }; -} - -function date$1(a, b) { - var d = new Date; - return a = +a, b = +b, function(t) { - return d.setTime(a * (1 - t) + b * t), d; - }; -} - -function interpolateNumber(a, b) { - return a = +a, b = +b, function(t) { - return a * (1 - t) + b * t; - }; -} - -function object(a, b) { - var i = {}, - c = {}, - k; - - if (a === null || typeof a !== "object") a = {}; - if (b === null || typeof b !== "object") b = {}; - - for (k in b) { - if (k in a) { - i[k] = interpolate(a[k], b[k]); - } else { - c[k] = b[k]; - } - } - - return function(t) { - for (k in i) c[k] = i[k](t); - return c; - }; -} - -var reA = /[-+]?(?:\d+\.?\d*|\.?\d+)(?:[eE][-+]?\d+)?/g, - reB = new RegExp(reA.source, "g"); - -function zero(b) { - return function() { - return b; - }; -} - -function one(b) { - return function(t) { - return b(t) + ""; - }; -} - -function string(a, b) { - var bi = reA.lastIndex = reB.lastIndex = 0, // scan index for next number in b - am, // current match in a - bm, // current match in b - bs, // string preceding current number in b, if any - i = -1, // index in s - s = [], // string constants and placeholders - q = []; // number interpolators - - // Coerce inputs to strings. - a = a + "", b = b + ""; - - // Interpolate pairs of numbers in a & b. - while ((am = reA.exec(a)) - && (bm = reB.exec(b))) { - if ((bs = bm.index) > bi) { // a string precedes the next number in b - bs = b.slice(bi, bs); - if (s[i]) s[i] += bs; // coalesce with previous string - else s[++i] = bs; - } - if ((am = am[0]) === (bm = bm[0])) { // numbers in a & b match - if (s[i]) s[i] += bm; // coalesce with previous string - else s[++i] = bm; - } else { // interpolate non-matching numbers - s[++i] = null; - q.push({i: i, x: interpolateNumber(am, bm)}); - } - bi = reB.lastIndex; - } - - // Add remains of b. - if (bi < b.length) { - bs = b.slice(bi); - if (s[i]) s[i] += bs; // coalesce with previous string - else s[++i] = bs; - } - - // Special optimization for only a single match. - // Otherwise, interpolate each of the numbers and rejoin the string. - return s.length < 2 ? (q[0] - ? one(q[0].x) - : zero(b)) - : (b = q.length, function(t) { - for (var i = 0, o; i < b; ++i) s[(o = q[i]).i] = o.x(t); - return s.join(""); - }); -} - -function interpolate(a, b) { - var t = typeof b, c; - return b == null || t === "boolean" ? constant(b) - : (t === "number" ? interpolateNumber - : t === "string" ? ((c = color(b)) ? (b = c, rgb) : string) - : b instanceof color ? rgb - : b instanceof Date ? date$1 - : isNumberArray(b) ? numberArray - : Array.isArray(b) ? genericArray - : typeof b.valueOf !== "function" && typeof b.toString !== "function" || isNaN(b) ? object - : interpolateNumber)(a, b); -} - -function interpolateRound(a, b) { - return a = +a, b = +b, function(t) { - return Math.round(a * (1 - t) + b * t); - }; -} - -var epsilon2 = 1e-12; - -function cosh(x) { - return ((x = Math.exp(x)) + 1 / x) / 2; -} - -function sinh(x) { - return ((x = Math.exp(x)) - 1 / x) / 2; -} - -function tanh(x) { - return ((x = Math.exp(2 * x)) - 1) / (x + 1); -} - -((function zoomRho(rho, rho2, rho4) { - - // p0 = [ux0, uy0, w0] - // p1 = [ux1, uy1, w1] - function zoom(p0, p1) { - var ux0 = p0[0], uy0 = p0[1], w0 = p0[2], - ux1 = p1[0], uy1 = p1[1], w1 = p1[2], - dx = ux1 - ux0, - dy = uy1 - uy0, - d2 = dx * dx + dy * dy, - i, - S; - - // Special case for u0 ≅ u1. - if (d2 < epsilon2) { - S = Math.log(w1 / w0) / rho; - i = function(t) { - return [ - ux0 + t * dx, - uy0 + t * dy, - w0 * Math.exp(rho * t * S) - ]; - }; - } - - // General case. - else { - var d1 = Math.sqrt(d2), - b0 = (w1 * w1 - w0 * w0 + rho4 * d2) / (2 * w0 * rho2 * d1), - b1 = (w1 * w1 - w0 * w0 - rho4 * d2) / (2 * w1 * rho2 * d1), - r0 = Math.log(Math.sqrt(b0 * b0 + 1) - b0), - r1 = Math.log(Math.sqrt(b1 * b1 + 1) - b1); - S = (r1 - r0) / rho; - i = function(t) { - var s = t * S, - coshr0 = cosh(r0), - u = w0 / (rho2 * d1) * (coshr0 * tanh(rho * s + r0) - sinh(r0)); - return [ - ux0 + u * dx, - uy0 + u * dy, - w0 * coshr0 / cosh(rho * s + r0) - ]; - }; - } - - i.duration = S * 1000 * rho / Math.SQRT2; - - return i; - } - - zoom.rho = function(_) { - var _1 = Math.max(1e-3, +_), _2 = _1 * _1, _4 = _2 * _2; - return zoomRho(_1, _2, _4); - }; - - return zoom; -}))(Math.SQRT2, 2, 4); - -function cubehelix(hue) { - return (function cubehelixGamma(y) { - y = +y; - - function cubehelix(start, end) { - var h = hue((start = cubehelix$1(start)).h, (end = cubehelix$1(end)).h), - s = nogamma(start.s, end.s), - l = nogamma(start.l, end.l), - opacity = nogamma(start.opacity, end.opacity); - return function(t) { - start.h = h(t); - start.s = s(t); - start.l = l(Math.pow(t, y)); - start.opacity = opacity(t); - return start + ""; - }; - } - - cubehelix.gamma = cubehelixGamma; - - return cubehelix; - })(1); -} - -cubehelix(hue); -cubehelix(nogamma); - -function piecewise(interpolate$1, values) { - if (values === undefined) values = interpolate$1, interpolate$1 = interpolate; - var i = 0, n = values.length - 1, v = values[0], I = new Array(n < 0 ? 0 : n); - while (i < n) I[i] = interpolate$1(v, v = values[++i]); - return function(t) { - var i = Math.max(0, Math.min(n - 1, Math.floor(t *= n))); - return I[i](t - i); - }; -} - -function constants(x) { - return function() { - return x; - }; -} - -function number$1(x) { - return +x; -} - -var unit = [0, 1]; - -function identity$1(x) { - return x; -} - -function normalize(a, b) { - return (b -= (a = +a)) - ? function(x) { return (x - a) / b; } - : constants(isNaN(b) ? NaN : 0.5); -} - -function clamper(a, b) { - var t; - if (a > b) t = a, a = b, b = t; - return function(x) { return Math.max(a, Math.min(b, x)); }; -} - -// normalize(a, b)(x) takes a domain value x in [a,b] and returns the corresponding parameter t in [0,1]. -// interpolate(a, b)(t) takes a parameter t in [0,1] and returns the corresponding range value x in [a,b]. -function bimap(domain, range, interpolate) { - var d0 = domain[0], d1 = domain[1], r0 = range[0], r1 = range[1]; - if (d1 < d0) d0 = normalize(d1, d0), r0 = interpolate(r1, r0); - else d0 = normalize(d0, d1), r0 = interpolate(r0, r1); - return function(x) { return r0(d0(x)); }; -} - -function polymap(domain, range, interpolate) { - var j = Math.min(domain.length, range.length) - 1, - d = new Array(j), - r = new Array(j), - i = -1; - - // Reverse descending domains. - if (domain[j] < domain[0]) { - domain = domain.slice().reverse(); - range = range.slice().reverse(); - } - - while (++i < j) { - d[i] = normalize(domain[i], domain[i + 1]); - r[i] = interpolate(range[i], range[i + 1]); - } - - return function(x) { - var i = d3Array.exports.bisect(domain, x, 1, j) - 1; - return r[i](d[i](x)); - }; -} - -function copy$1(source, target) { - return target - .domain(source.domain()) - .range(source.range()) - .interpolate(source.interpolate()) - .clamp(source.clamp()) - .unknown(source.unknown()); -} - -function transformer$2() { - var domain = unit, - range = unit, - interpolate$1 = interpolate, - transform, - untransform, - unknown, - clamp = identity$1, - piecewise, - output, - input; - - function rescale() { - var n = Math.min(domain.length, range.length); - if (clamp !== identity$1) clamp = clamper(domain[0], domain[n - 1]); - piecewise = n > 2 ? polymap : bimap; - output = input = null; - return scale; - } - - function scale(x) { - return x == null || isNaN(x = +x) ? unknown : (output || (output = piecewise(domain.map(transform), range, interpolate$1)))(transform(clamp(x))); - } - - scale.invert = function(y) { - return clamp(untransform((input || (input = piecewise(range, domain.map(transform), interpolateNumber)))(y))); - }; - - scale.domain = function(_) { - return arguments.length ? (domain = Array.from(_, number$1), rescale()) : domain.slice(); - }; - - scale.range = function(_) { - return arguments.length ? (range = Array.from(_), rescale()) : range.slice(); - }; - - scale.rangeRound = function(_) { - return range = Array.from(_), interpolate$1 = interpolateRound, rescale(); - }; - - scale.clamp = function(_) { - return arguments.length ? (clamp = _ ? true : identity$1, rescale()) : clamp !== identity$1; - }; - - scale.interpolate = function(_) { - return arguments.length ? (interpolate$1 = _, rescale()) : interpolate$1; - }; - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : unknown; - }; - - return function(t, u) { - transform = t, untransform = u; - return rescale(); - }; -} - -function continuous() { - return transformer$2()(identity$1, identity$1); -} - -var d3Format = {exports: {}}; - -(function (module, exports) { - // https://d3js.org/d3-format/ v2.0.0 Copyright 2020 Mike Bostock - (function (global, factory) { - factory(exports) ; - }(commonjsGlobal$1, (function (exports) { - function formatDecimal(x) { - return Math.abs(x = Math.round(x)) >= 1e21 - ? x.toLocaleString("en").replace(/,/g, "") - : x.toString(10); - } - - // Computes the decimal coefficient and exponent of the specified number x with - // significant digits p, where x is positive and p is in [1, 21] or undefined. - // For example, formatDecimalParts(1.23) returns ["123", 0]. - function formatDecimalParts(x, p) { - if ((i = (x = p ? x.toExponential(p - 1) : x.toExponential()).indexOf("e")) < 0) return null; // NaN, ±Infinity - var i, coefficient = x.slice(0, i); - - // The string returned by toExponential either has the form \d\.\d+e[-+]\d+ - // (e.g., 1.2e+3) or the form \de[-+]\d+ (e.g., 1e+3). - return [ - coefficient.length > 1 ? coefficient[0] + coefficient.slice(2) : coefficient, - +x.slice(i + 1) - ]; - } - - function exponent(x) { - return x = formatDecimalParts(Math.abs(x)), x ? x[1] : NaN; - } - - function formatGroup(grouping, thousands) { - return function(value, width) { - var i = value.length, - t = [], - j = 0, - g = grouping[0], - length = 0; - - while (i > 0 && g > 0) { - if (length + g + 1 > width) g = Math.max(1, width - length); - t.push(value.substring(i -= g, i + g)); - if ((length += g + 1) > width) break; - g = grouping[j = (j + 1) % grouping.length]; - } - - return t.reverse().join(thousands); - }; - } - - function formatNumerals(numerals) { - return function(value) { - return value.replace(/[0-9]/g, function(i) { - return numerals[+i]; - }); - }; - } - - // [[fill]align][sign][symbol][0][width][,][.precision][~][type] - var re = /^(?:(.)?([<>=^]))?([+\-( ])?([$#])?(0)?(\d+)?(,)?(\.\d+)?(~)?([a-z%])?$/i; - - function formatSpecifier(specifier) { - if (!(match = re.exec(specifier))) throw new Error("invalid format: " + specifier); - var match; - return new FormatSpecifier({ - fill: match[1], - align: match[2], - sign: match[3], - symbol: match[4], - zero: match[5], - width: match[6], - comma: match[7], - precision: match[8] && match[8].slice(1), - trim: match[9], - type: match[10] - }); - } - - formatSpecifier.prototype = FormatSpecifier.prototype; // instanceof - - function FormatSpecifier(specifier) { - this.fill = specifier.fill === undefined ? " " : specifier.fill + ""; - this.align = specifier.align === undefined ? ">" : specifier.align + ""; - this.sign = specifier.sign === undefined ? "-" : specifier.sign + ""; - this.symbol = specifier.symbol === undefined ? "" : specifier.symbol + ""; - this.zero = !!specifier.zero; - this.width = specifier.width === undefined ? undefined : +specifier.width; - this.comma = !!specifier.comma; - this.precision = specifier.precision === undefined ? undefined : +specifier.precision; - this.trim = !!specifier.trim; - this.type = specifier.type === undefined ? "" : specifier.type + ""; - } - - FormatSpecifier.prototype.toString = function() { - return this.fill - + this.align - + this.sign - + this.symbol - + (this.zero ? "0" : "") - + (this.width === undefined ? "" : Math.max(1, this.width | 0)) - + (this.comma ? "," : "") - + (this.precision === undefined ? "" : "." + Math.max(0, this.precision | 0)) - + (this.trim ? "~" : "") - + this.type; - }; - - // Trims insignificant zeros, e.g., replaces 1.2000k with 1.2k. - function formatTrim(s) { - out: for (var n = s.length, i = 1, i0 = -1, i1; i < n; ++i) { - switch (s[i]) { - case ".": i0 = i1 = i; break; - case "0": if (i0 === 0) i0 = i; i1 = i; break; - default: if (!+s[i]) break out; if (i0 > 0) i0 = 0; break; - } - } - return i0 > 0 ? s.slice(0, i0) + s.slice(i1 + 1) : s; - } - - var prefixExponent; - - function formatPrefixAuto(x, p) { - var d = formatDecimalParts(x, p); - if (!d) return x + ""; - var coefficient = d[0], - exponent = d[1], - i = exponent - (prefixExponent = Math.max(-8, Math.min(8, Math.floor(exponent / 3))) * 3) + 1, - n = coefficient.length; - return i === n ? coefficient - : i > n ? coefficient + new Array(i - n + 1).join("0") - : i > 0 ? coefficient.slice(0, i) + "." + coefficient.slice(i) - : "0." + new Array(1 - i).join("0") + formatDecimalParts(x, Math.max(0, p + i - 1))[0]; // less than 1y! - } - - function formatRounded(x, p) { - var d = formatDecimalParts(x, p); - if (!d) return x + ""; - var coefficient = d[0], - exponent = d[1]; - return exponent < 0 ? "0." + new Array(-exponent).join("0") + coefficient - : coefficient.length > exponent + 1 ? coefficient.slice(0, exponent + 1) + "." + coefficient.slice(exponent + 1) - : coefficient + new Array(exponent - coefficient.length + 2).join("0"); - } - - var formatTypes = { - "%": (x, p) => (x * 100).toFixed(p), - "b": (x) => Math.round(x).toString(2), - "c": (x) => x + "", - "d": formatDecimal, - "e": (x, p) => x.toExponential(p), - "f": (x, p) => x.toFixed(p), - "g": (x, p) => x.toPrecision(p), - "o": (x) => Math.round(x).toString(8), - "p": (x, p) => formatRounded(x * 100, p), - "r": formatRounded, - "s": formatPrefixAuto, - "X": (x) => Math.round(x).toString(16).toUpperCase(), - "x": (x) => Math.round(x).toString(16) - }; - - function identity(x) { - return x; - } - - var map = Array.prototype.map, - prefixes = ["y","z","a","f","p","n","µ","m","","k","M","G","T","P","E","Z","Y"]; - - function formatLocale(locale) { - var group = locale.grouping === undefined || locale.thousands === undefined ? identity : formatGroup(map.call(locale.grouping, Number), locale.thousands + ""), - currencyPrefix = locale.currency === undefined ? "" : locale.currency[0] + "", - currencySuffix = locale.currency === undefined ? "" : locale.currency[1] + "", - decimal = locale.decimal === undefined ? "." : locale.decimal + "", - numerals = locale.numerals === undefined ? identity : formatNumerals(map.call(locale.numerals, String)), - percent = locale.percent === undefined ? "%" : locale.percent + "", - minus = locale.minus === undefined ? "−" : locale.minus + "", - nan = locale.nan === undefined ? "NaN" : locale.nan + ""; - - function newFormat(specifier) { - specifier = formatSpecifier(specifier); - - var fill = specifier.fill, - align = specifier.align, - sign = specifier.sign, - symbol = specifier.symbol, - zero = specifier.zero, - width = specifier.width, - comma = specifier.comma, - precision = specifier.precision, - trim = specifier.trim, - type = specifier.type; - - // The "n" type is an alias for ",g". - if (type === "n") comma = true, type = "g"; - - // The "" type, and any invalid type, is an alias for ".12~g". - else if (!formatTypes[type]) precision === undefined && (precision = 12), trim = true, type = "g"; - - // If zero fill is specified, padding goes after sign and before digits. - if (zero || (fill === "0" && align === "=")) zero = true, fill = "0", align = "="; - - // Compute the prefix and suffix. - // For SI-prefix, the suffix is lazily computed. - var prefix = symbol === "$" ? currencyPrefix : symbol === "#" && /[boxX]/.test(type) ? "0" + type.toLowerCase() : "", - suffix = symbol === "$" ? currencySuffix : /[%p]/.test(type) ? percent : ""; - - // What format function should we use? - // Is this an integer type? - // Can this type generate exponential notation? - var formatType = formatTypes[type], - maybeSuffix = /[defgprs%]/.test(type); - - // Set the default precision if not specified, - // or clamp the specified precision to the supported range. - // For significant precision, it must be in [1, 21]. - // For fixed precision, it must be in [0, 20]. - precision = precision === undefined ? 6 - : /[gprs]/.test(type) ? Math.max(1, Math.min(21, precision)) - : Math.max(0, Math.min(20, precision)); - - function format(value) { - var valuePrefix = prefix, - valueSuffix = suffix, - i, n, c; - - if (type === "c") { - valueSuffix = formatType(value) + valueSuffix; - value = ""; - } else { - value = +value; - - // Determine the sign. -0 is not less than 0, but 1 / -0 is! - var valueNegative = value < 0 || 1 / value < 0; - - // Perform the initial formatting. - value = isNaN(value) ? nan : formatType(Math.abs(value), precision); - - // Trim insignificant zeros. - if (trim) value = formatTrim(value); - - // If a negative value rounds to zero after formatting, and no explicit positive sign is requested, hide the sign. - if (valueNegative && +value === 0 && sign !== "+") valueNegative = false; - - // Compute the prefix and suffix. - valuePrefix = (valueNegative ? (sign === "(" ? sign : minus) : sign === "-" || sign === "(" ? "" : sign) + valuePrefix; - valueSuffix = (type === "s" ? prefixes[8 + prefixExponent / 3] : "") + valueSuffix + (valueNegative && sign === "(" ? ")" : ""); - - // Break the formatted value into the integer “value” part that can be - // grouped, and fractional or exponential “suffix” part that is not. - if (maybeSuffix) { - i = -1, n = value.length; - while (++i < n) { - if (c = value.charCodeAt(i), 48 > c || c > 57) { - valueSuffix = (c === 46 ? decimal + value.slice(i + 1) : value.slice(i)) + valueSuffix; - value = value.slice(0, i); - break; - } - } - } - } - - // If the fill character is not "0", grouping is applied before padding. - if (comma && !zero) value = group(value, Infinity); - - // Compute the padding. - var length = valuePrefix.length + value.length + valueSuffix.length, - padding = length < width ? new Array(width - length + 1).join(fill) : ""; - - // If the fill character is "0", grouping is applied after padding. - if (comma && zero) value = group(padding + value, padding.length ? width - valueSuffix.length : Infinity), padding = ""; - - // Reconstruct the final output based on the desired alignment. - switch (align) { - case "<": value = valuePrefix + value + valueSuffix + padding; break; - case "=": value = valuePrefix + padding + value + valueSuffix; break; - case "^": value = padding.slice(0, length = padding.length >> 1) + valuePrefix + value + valueSuffix + padding.slice(length); break; - default: value = padding + valuePrefix + value + valueSuffix; break; - } - - return numerals(value); - } - - format.toString = function() { - return specifier + ""; - }; - - return format; - } - - function formatPrefix(specifier, value) { - var f = newFormat((specifier = formatSpecifier(specifier), specifier.type = "f", specifier)), - e = Math.max(-8, Math.min(8, Math.floor(exponent(value) / 3))) * 3, - k = Math.pow(10, -e), - prefix = prefixes[8 + e / 3]; - return function(value) { - return f(k * value) + prefix; - }; - } - - return { - format: newFormat, - formatPrefix: formatPrefix - }; - } - - var locale; - - defaultLocale({ - thousands: ",", - grouping: [3], - currency: ["$", ""] - }); - - function defaultLocale(definition) { - locale = formatLocale(definition); - exports.format = locale.format; - exports.formatPrefix = locale.formatPrefix; - return locale; - } - - function precisionFixed(step) { - return Math.max(0, -exponent(Math.abs(step))); - } - - function precisionPrefix(step, value) { - return Math.max(0, Math.max(-8, Math.min(8, Math.floor(exponent(value) / 3))) * 3 - exponent(Math.abs(step))); - } - - function precisionRound(step, max) { - step = Math.abs(step), max = Math.abs(max) - step; - return Math.max(0, exponent(max) - exponent(step)) + 1; - } - - exports.FormatSpecifier = FormatSpecifier; - exports.formatDefaultLocale = defaultLocale; - exports.formatLocale = formatLocale; - exports.formatSpecifier = formatSpecifier; - exports.precisionFixed = precisionFixed; - exports.precisionPrefix = precisionPrefix; - exports.precisionRound = precisionRound; - - Object.defineProperty(exports, '__esModule', { value: true }); - - }))); -} (d3Format, d3Format.exports)); - -function tickFormat(start, stop, count, specifier) { - var step = d3Array.exports.tickStep(start, stop, count), - precision; - specifier = d3Format.exports.formatSpecifier(specifier == null ? ",f" : specifier); - switch (specifier.type) { - case "s": { - var value = Math.max(Math.abs(start), Math.abs(stop)); - if (specifier.precision == null && !isNaN(precision = d3Format.exports.precisionPrefix(step, value))) specifier.precision = precision; - return d3Format.exports.formatPrefix(specifier, value); - } - case "": - case "e": - case "g": - case "p": - case "r": { - if (specifier.precision == null && !isNaN(precision = d3Format.exports.precisionRound(step, Math.max(Math.abs(start), Math.abs(stop))))) specifier.precision = precision - (specifier.type === "e"); - break; - } - case "f": - case "%": { - if (specifier.precision == null && !isNaN(precision = d3Format.exports.precisionFixed(step))) specifier.precision = precision - (specifier.type === "%") * 2; - break; - } - } - return d3Format.exports.format(specifier); -} - -function linearish(scale) { - var domain = scale.domain; - - scale.ticks = function(count) { - var d = domain(); - return d3Array.exports.ticks(d[0], d[d.length - 1], count == null ? 10 : count); - }; - - scale.tickFormat = function(count, specifier) { - var d = domain(); - return tickFormat(d[0], d[d.length - 1], count == null ? 10 : count, specifier); - }; - - scale.nice = function(count) { - if (count == null) count = 10; - - var d = domain(); - var i0 = 0; - var i1 = d.length - 1; - var start = d[i0]; - var stop = d[i1]; - var prestep; - var step; - var maxIter = 10; - - if (stop < start) { - step = start, start = stop, stop = step; - step = i0, i0 = i1, i1 = step; - } - - while (maxIter-- > 0) { - step = d3Array.exports.tickIncrement(start, stop, count); - if (step === prestep) { - d[i0] = start; - d[i1] = stop; - return domain(d); - } else if (step > 0) { - start = Math.floor(start / step) * step; - stop = Math.ceil(stop / step) * step; - } else if (step < 0) { - start = Math.ceil(start * step) / step; - stop = Math.floor(stop * step) / step; - } else { - break; - } - prestep = step; - } - - return scale; - }; - - return scale; -} - -function linear() { - var scale = continuous(); - - scale.copy = function() { - return copy$1(scale, linear()); - }; - - initRange.apply(scale, arguments); - - return linearish(scale); -} - -function identity(domain) { - var unknown; - - function scale(x) { - return x == null || isNaN(x = +x) ? unknown : x; - } - - scale.invert = scale; - - scale.domain = scale.range = function(_) { - return arguments.length ? (domain = Array.from(_, number$1), scale) : domain.slice(); - }; - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : unknown; - }; - - scale.copy = function() { - return identity(domain).unknown(unknown); - }; - - domain = arguments.length ? Array.from(domain, number$1) : [0, 1]; - - return linearish(scale); -} - -function nice(domain, interval) { - domain = domain.slice(); - - var i0 = 0, - i1 = domain.length - 1, - x0 = domain[i0], - x1 = domain[i1], - t; - - if (x1 < x0) { - t = i0, i0 = i1, i1 = t; - t = x0, x0 = x1, x1 = t; - } - - domain[i0] = interval.floor(x0); - domain[i1] = interval.ceil(x1); - return domain; -} - -function transformLog(x) { - return Math.log(x); -} - -function transformExp(x) { - return Math.exp(x); -} - -function transformLogn(x) { - return -Math.log(-x); -} - -function transformExpn(x) { - return -Math.exp(-x); -} - -function pow10(x) { - return isFinite(x) ? +("1e" + x) : x < 0 ? 0 : x; -} - -function powp(base) { - return base === 10 ? pow10 - : base === Math.E ? Math.exp - : function(x) { return Math.pow(base, x); }; -} - -function logp(base) { - return base === Math.E ? Math.log - : base === 10 && Math.log10 - || base === 2 && Math.log2 - || (base = Math.log(base), function(x) { return Math.log(x) / base; }); -} - -function reflect(f) { - return function(x) { - return -f(-x); - }; -} - -function loggish(transform) { - var scale = transform(transformLog, transformExp), - domain = scale.domain, - base = 10, - logs, - pows; - - function rescale() { - logs = logp(base), pows = powp(base); - if (domain()[0] < 0) { - logs = reflect(logs), pows = reflect(pows); - transform(transformLogn, transformExpn); - } else { - transform(transformLog, transformExp); - } - return scale; - } - - scale.base = function(_) { - return arguments.length ? (base = +_, rescale()) : base; - }; - - scale.domain = function(_) { - return arguments.length ? (domain(_), rescale()) : domain(); - }; - - scale.ticks = function(count) { - var d = domain(), - u = d[0], - v = d[d.length - 1], - r; - - if (r = v < u) i = u, u = v, v = i; - - var i = logs(u), - j = logs(v), - p, - k, - t, - n = count == null ? 10 : +count, - z = []; - - if (!(base % 1) && j - i < n) { - i = Math.floor(i), j = Math.ceil(j); - if (u > 0) for (; i <= j; ++i) { - for (k = 1, p = pows(i); k < base; ++k) { - t = p * k; - if (t < u) continue; - if (t > v) break; - z.push(t); - } - } else for (; i <= j; ++i) { - for (k = base - 1, p = pows(i); k >= 1; --k) { - t = p * k; - if (t < u) continue; - if (t > v) break; - z.push(t); - } - } - if (z.length * 2 < n) z = d3Array.exports.ticks(u, v, n); - } else { - z = d3Array.exports.ticks(i, j, Math.min(j - i, n)).map(pows); - } - - return r ? z.reverse() : z; - }; - - scale.tickFormat = function(count, specifier) { - if (specifier == null) specifier = base === 10 ? ".0e" : ","; - if (typeof specifier !== "function") specifier = d3Format.exports.format(specifier); - if (count === Infinity) return specifier; - if (count == null) count = 10; - var k = Math.max(1, base * count / scale.ticks().length); // TODO fast estimate? - return function(d) { - var i = d / pows(Math.round(logs(d))); - if (i * base < base - 0.5) i *= base; - return i <= k ? specifier(d) : ""; - }; - }; - - scale.nice = function() { - return domain(nice(domain(), { - floor: function(x) { return pows(Math.floor(logs(x))); }, - ceil: function(x) { return pows(Math.ceil(logs(x))); } - })); - }; - - return scale; -} - -function log() { - var scale = loggish(transformer$2()).domain([1, 10]); - - scale.copy = function() { - return copy$1(scale, log()).base(scale.base()); - }; - - initRange.apply(scale, arguments); - - return scale; -} - -function transformSymlog(c) { - return function(x) { - return Math.sign(x) * Math.log1p(Math.abs(x / c)); - }; -} - -function transformSymexp(c) { - return function(x) { - return Math.sign(x) * Math.expm1(Math.abs(x)) * c; - }; -} - -function symlogish(transform) { - var c = 1, scale = transform(transformSymlog(c), transformSymexp(c)); - - scale.constant = function(_) { - return arguments.length ? transform(transformSymlog(c = +_), transformSymexp(c)) : c; - }; - - return linearish(scale); -} - -function symlog() { - var scale = symlogish(transformer$2()); - - scale.copy = function() { - return copy$1(scale, symlog()).constant(scale.constant()); - }; - - return initRange.apply(scale, arguments); -} - -function transformPow(exponent) { - return function(x) { - return x < 0 ? -Math.pow(-x, exponent) : Math.pow(x, exponent); - }; -} - -function transformSqrt(x) { - return x < 0 ? -Math.sqrt(-x) : Math.sqrt(x); -} - -function transformSquare(x) { - return x < 0 ? -x * x : x * x; -} - -function powish(transform) { - var scale = transform(identity$1, identity$1), - exponent = 1; - - function rescale() { - return exponent === 1 ? transform(identity$1, identity$1) - : exponent === 0.5 ? transform(transformSqrt, transformSquare) - : transform(transformPow(exponent), transformPow(1 / exponent)); - } - - scale.exponent = function(_) { - return arguments.length ? (exponent = +_, rescale()) : exponent; - }; - - return linearish(scale); -} - -function pow() { - var scale = powish(transformer$2()); - - scale.copy = function() { - return copy$1(scale, pow()).exponent(scale.exponent()); - }; - - initRange.apply(scale, arguments); - - return scale; -} - -function sqrt() { - return pow.apply(null, arguments).exponent(0.5); -} - -function square(x) { - return Math.sign(x) * x * x; -} - -function unsquare(x) { - return Math.sign(x) * Math.sqrt(Math.abs(x)); -} - -function radial() { - var squared = continuous(), - range = [0, 1], - round = false, - unknown; - - function scale(x) { - var y = unsquare(squared(x)); - return isNaN(y) ? unknown : round ? Math.round(y) : y; - } - - scale.invert = function(y) { - return squared.invert(square(y)); - }; - - scale.domain = function(_) { - return arguments.length ? (squared.domain(_), scale) : squared.domain(); - }; - - scale.range = function(_) { - return arguments.length ? (squared.range((range = Array.from(_, number$1)).map(square)), scale) : range.slice(); - }; - - scale.rangeRound = function(_) { - return scale.range(_).round(true); - }; - - scale.round = function(_) { - return arguments.length ? (round = !!_, scale) : round; - }; - - scale.clamp = function(_) { - return arguments.length ? (squared.clamp(_), scale) : squared.clamp(); - }; - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : unknown; - }; - - scale.copy = function() { - return radial(squared.domain(), range) - .round(round) - .clamp(squared.clamp()) - .unknown(unknown); - }; - - initRange.apply(scale, arguments); - - return linearish(scale); -} - -function quantile() { - var domain = [], - range = [], - thresholds = [], - unknown; - - function rescale() { - var i = 0, n = Math.max(1, range.length); - thresholds = new Array(n - 1); - while (++i < n) thresholds[i - 1] = d3Array.exports.quantileSorted(domain, i / n); - return scale; - } - - function scale(x) { - return x == null || isNaN(x = +x) ? unknown : range[d3Array.exports.bisect(thresholds, x)]; - } - - scale.invertExtent = function(y) { - var i = range.indexOf(y); - return i < 0 ? [NaN, NaN] : [ - i > 0 ? thresholds[i - 1] : domain[0], - i < thresholds.length ? thresholds[i] : domain[domain.length - 1] - ]; - }; - - scale.domain = function(_) { - if (!arguments.length) return domain.slice(); - domain = []; - for (let d of _) if (d != null && !isNaN(d = +d)) domain.push(d); - domain.sort(d3Array.exports.ascending); - return rescale(); - }; - - scale.range = function(_) { - return arguments.length ? (range = Array.from(_), rescale()) : range.slice(); - }; - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : unknown; - }; - - scale.quantiles = function() { - return thresholds.slice(); - }; - - scale.copy = function() { - return quantile() - .domain(domain) - .range(range) - .unknown(unknown); - }; - - return initRange.apply(scale, arguments); -} - -function quantize() { - var x0 = 0, - x1 = 1, - n = 1, - domain = [0.5], - range = [0, 1], - unknown; - - function scale(x) { - return x != null && x <= x ? range[d3Array.exports.bisect(domain, x, 0, n)] : unknown; - } - - function rescale() { - var i = -1; - domain = new Array(n); - while (++i < n) domain[i] = ((i + 1) * x1 - (i - n) * x0) / (n + 1); - return scale; - } - - scale.domain = function(_) { - return arguments.length ? ([x0, x1] = _, x0 = +x0, x1 = +x1, rescale()) : [x0, x1]; - }; - - scale.range = function(_) { - return arguments.length ? (n = (range = Array.from(_)).length - 1, rescale()) : range.slice(); - }; - - scale.invertExtent = function(y) { - var i = range.indexOf(y); - return i < 0 ? [NaN, NaN] - : i < 1 ? [x0, domain[0]] - : i >= n ? [domain[n - 1], x1] - : [domain[i - 1], domain[i]]; - }; - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : scale; - }; - - scale.thresholds = function() { - return domain.slice(); - }; - - scale.copy = function() { - return quantize() - .domain([x0, x1]) - .range(range) - .unknown(unknown); - }; - - return initRange.apply(linearish(scale), arguments); -} - -function threshold() { - var domain = [0.5], - range = [0, 1], - unknown, - n = 1; - - function scale(x) { - return x != null && x <= x ? range[d3Array.exports.bisect(domain, x, 0, n)] : unknown; - } - - scale.domain = function(_) { - return arguments.length ? (domain = Array.from(_), n = Math.min(domain.length, range.length - 1), scale) : domain.slice(); - }; - - scale.range = function(_) { - return arguments.length ? (range = Array.from(_), n = Math.min(domain.length, range.length - 1), scale) : range.slice(); - }; - - scale.invertExtent = function(y) { - var i = range.indexOf(y); - return [domain[i - 1], domain[i]]; - }; - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : unknown; - }; - - scale.copy = function() { - return threshold() - .domain(domain) - .range(range) - .unknown(unknown); - }; - - return initRange.apply(scale, arguments); -} - -var t0 = new Date, - t1 = new Date; - -function newInterval(floori, offseti, count, field) { - - function interval(date) { - return floori(date = arguments.length === 0 ? new Date : new Date(+date)), date; - } - - interval.floor = function(date) { - return floori(date = new Date(+date)), date; - }; - - interval.ceil = function(date) { - return floori(date = new Date(date - 1)), offseti(date, 1), floori(date), date; - }; - - interval.round = function(date) { - var d0 = interval(date), - d1 = interval.ceil(date); - return date - d0 < d1 - date ? d0 : d1; - }; - - interval.offset = function(date, step) { - return offseti(date = new Date(+date), step == null ? 1 : Math.floor(step)), date; - }; - - interval.range = function(start, stop, step) { - var range = [], previous; - start = interval.ceil(start); - step = step == null ? 1 : Math.floor(step); - if (!(start < stop) || !(step > 0)) return range; // also handles Invalid Date - do range.push(previous = new Date(+start)), offseti(start, step), floori(start); - while (previous < start && start < stop); - return range; - }; - - interval.filter = function(test) { - return newInterval(function(date) { - if (date >= date) while (floori(date), !test(date)) date.setTime(date - 1); - }, function(date, step) { - if (date >= date) { - if (step < 0) while (++step <= 0) { - while (offseti(date, -1), !test(date)) {} // eslint-disable-line no-empty - } else while (--step >= 0) { - while (offseti(date, +1), !test(date)) {} // eslint-disable-line no-empty - } - } - }); - }; - - if (count) { - interval.count = function(start, end) { - t0.setTime(+start), t1.setTime(+end); - floori(t0), floori(t1); - return Math.floor(count(t0, t1)); - }; - - interval.every = function(step) { - step = Math.floor(step); - return !isFinite(step) || !(step > 0) ? null - : !(step > 1) ? interval - : interval.filter(field - ? function(d) { return field(d) % step === 0; } - : function(d) { return interval.count(0, d) % step === 0; }); - }; - } - - return interval; -} - -var millisecond = newInterval(function() { - // noop -}, function(date, step) { - date.setTime(+date + step); -}, function(start, end) { - return end - start; -}); - -// An optimized implementation for this simple case. -millisecond.every = function(k) { - k = Math.floor(k); - if (!isFinite(k) || !(k > 0)) return null; - if (!(k > 1)) return millisecond; - return newInterval(function(date) { - date.setTime(Math.floor(date / k) * k); - }, function(date, step) { - date.setTime(+date + step * k); - }, function(start, end) { - return (end - start) / k; - }); -}; -var milliseconds = millisecond.range; - -const durationSecond = 1000; -const durationMinute = durationSecond * 60; -const durationHour = durationMinute * 60; -const durationDay = durationHour * 24; -const durationWeek = durationDay * 7; -const durationMonth = durationDay * 30; -const durationYear = durationDay * 365; - -var second = newInterval(function(date) { - date.setTime(date - date.getMilliseconds()); -}, function(date, step) { - date.setTime(+date + step * durationSecond); -}, function(start, end) { - return (end - start) / durationSecond; -}, function(date) { - return date.getUTCSeconds(); -}); -var seconds = second.range; - -var minute = newInterval(function(date) { - date.setTime(date - date.getMilliseconds() - date.getSeconds() * durationSecond); -}, function(date, step) { - date.setTime(+date + step * durationMinute); -}, function(start, end) { - return (end - start) / durationMinute; -}, function(date) { - return date.getMinutes(); -}); -var minutes = minute.range; - -var hour = newInterval(function(date) { - date.setTime(date - date.getMilliseconds() - date.getSeconds() * durationSecond - date.getMinutes() * durationMinute); -}, function(date, step) { - date.setTime(+date + step * durationHour); -}, function(start, end) { - return (end - start) / durationHour; -}, function(date) { - return date.getHours(); -}); -var hours = hour.range; - -var day = newInterval( - date => date.setHours(0, 0, 0, 0), - (date, step) => date.setDate(date.getDate() + step), - (start, end) => (end - start - (end.getTimezoneOffset() - start.getTimezoneOffset()) * durationMinute) / durationDay, - date => date.getDate() - 1 -); -var days = day.range; - -function weekday(i) { - return newInterval(function(date) { - date.setDate(date.getDate() - (date.getDay() + 7 - i) % 7); - date.setHours(0, 0, 0, 0); - }, function(date, step) { - date.setDate(date.getDate() + step * 7); - }, function(start, end) { - return (end - start - (end.getTimezoneOffset() - start.getTimezoneOffset()) * durationMinute) / durationWeek; - }); -} - -var sunday = weekday(0); -var monday = weekday(1); -var tuesday = weekday(2); -var wednesday = weekday(3); -var thursday = weekday(4); -var friday = weekday(5); -var saturday = weekday(6); - -var sundays = sunday.range; -var mondays = monday.range; -var tuesdays = tuesday.range; -var wednesdays = wednesday.range; -var thursdays = thursday.range; -var fridays = friday.range; -var saturdays = saturday.range; - -var month = newInterval(function(date) { - date.setDate(1); - date.setHours(0, 0, 0, 0); -}, function(date, step) { - date.setMonth(date.getMonth() + step); -}, function(start, end) { - return end.getMonth() - start.getMonth() + (end.getFullYear() - start.getFullYear()) * 12; -}, function(date) { - return date.getMonth(); -}); -var months = month.range; - -var year = newInterval(function(date) { - date.setMonth(0, 1); - date.setHours(0, 0, 0, 0); -}, function(date, step) { - date.setFullYear(date.getFullYear() + step); -}, function(start, end) { - return end.getFullYear() - start.getFullYear(); -}, function(date) { - return date.getFullYear(); -}); - -// An optimized implementation for this simple case. -year.every = function(k) { - return !isFinite(k = Math.floor(k)) || !(k > 0) ? null : newInterval(function(date) { - date.setFullYear(Math.floor(date.getFullYear() / k) * k); - date.setMonth(0, 1); - date.setHours(0, 0, 0, 0); - }, function(date, step) { - date.setFullYear(date.getFullYear() + step * k); - }); -}; -var years = year.range; - -var utcMinute = newInterval(function(date) { - date.setUTCSeconds(0, 0); -}, function(date, step) { - date.setTime(+date + step * durationMinute); -}, function(start, end) { - return (end - start) / durationMinute; -}, function(date) { - return date.getUTCMinutes(); -}); -var utcMinutes = utcMinute.range; - -var utcHour = newInterval(function(date) { - date.setUTCMinutes(0, 0, 0); -}, function(date, step) { - date.setTime(+date + step * durationHour); -}, function(start, end) { - return (end - start) / durationHour; -}, function(date) { - return date.getUTCHours(); -}); -var utcHours = utcHour.range; - -var utcDay = newInterval(function(date) { - date.setUTCHours(0, 0, 0, 0); -}, function(date, step) { - date.setUTCDate(date.getUTCDate() + step); -}, function(start, end) { - return (end - start) / durationDay; -}, function(date) { - return date.getUTCDate() - 1; -}); -var utcDays = utcDay.range; - -function utcWeekday(i) { - return newInterval(function(date) { - date.setUTCDate(date.getUTCDate() - (date.getUTCDay() + 7 - i) % 7); - date.setUTCHours(0, 0, 0, 0); - }, function(date, step) { - date.setUTCDate(date.getUTCDate() + step * 7); - }, function(start, end) { - return (end - start) / durationWeek; - }); -} - -var utcSunday = utcWeekday(0); -var utcMonday = utcWeekday(1); -var utcTuesday = utcWeekday(2); -var utcWednesday = utcWeekday(3); -var utcThursday = utcWeekday(4); -var utcFriday = utcWeekday(5); -var utcSaturday = utcWeekday(6); - -var utcSundays = utcSunday.range; -var utcMondays = utcMonday.range; -var utcTuesdays = utcTuesday.range; -var utcWednesdays = utcWednesday.range; -var utcThursdays = utcThursday.range; -var utcFridays = utcFriday.range; -var utcSaturdays = utcSaturday.range; - -var utcMonth = newInterval(function(date) { - date.setUTCDate(1); - date.setUTCHours(0, 0, 0, 0); -}, function(date, step) { - date.setUTCMonth(date.getUTCMonth() + step); -}, function(start, end) { - return end.getUTCMonth() - start.getUTCMonth() + (end.getUTCFullYear() - start.getUTCFullYear()) * 12; -}, function(date) { - return date.getUTCMonth(); -}); -var utcMonths = utcMonth.range; - -var utcYear = newInterval(function(date) { - date.setUTCMonth(0, 1); - date.setUTCHours(0, 0, 0, 0); -}, function(date, step) { - date.setUTCFullYear(date.getUTCFullYear() + step); -}, function(start, end) { - return end.getUTCFullYear() - start.getUTCFullYear(); -}, function(date) { - return date.getUTCFullYear(); -}); - -// An optimized implementation for this simple case. -utcYear.every = function(k) { - return !isFinite(k = Math.floor(k)) || !(k > 0) ? null : newInterval(function(date) { - date.setUTCFullYear(Math.floor(date.getUTCFullYear() / k) * k); - date.setUTCMonth(0, 1); - date.setUTCHours(0, 0, 0, 0); - }, function(date, step) { - date.setUTCFullYear(date.getUTCFullYear() + step * k); - }); -}; -var utcYears = utcYear.range; - -function ticker(year, month, week, day, hour, minute) { - - const tickIntervals = [ - [second, 1, durationSecond], - [second, 5, 5 * durationSecond], - [second, 15, 15 * durationSecond], - [second, 30, 30 * durationSecond], - [minute, 1, durationMinute], - [minute, 5, 5 * durationMinute], - [minute, 15, 15 * durationMinute], - [minute, 30, 30 * durationMinute], - [ hour, 1, durationHour ], - [ hour, 3, 3 * durationHour ], - [ hour, 6, 6 * durationHour ], - [ hour, 12, 12 * durationHour ], - [ day, 1, durationDay ], - [ day, 2, 2 * durationDay ], - [ week, 1, durationWeek ], - [ month, 1, durationMonth ], - [ month, 3, 3 * durationMonth ], - [ year, 1, durationYear ] - ]; - - function ticks(start, stop, count) { - const reverse = stop < start; - if (reverse) [start, stop] = [stop, start]; - const interval = count && typeof count.range === "function" ? count : tickInterval(start, stop, count); - const ticks = interval ? interval.range(start, +stop + 1) : []; // inclusive stop - return reverse ? ticks.reverse() : ticks; - } - - function tickInterval(start, stop, count) { - const target = Math.abs(stop - start) / count; - const i = d3Array.exports.bisector(([,, step]) => step).right(tickIntervals, target); - if (i === tickIntervals.length) return year.every(d3Array.exports.tickStep(start / durationYear, stop / durationYear, count)); - if (i === 0) return millisecond.every(Math.max(d3Array.exports.tickStep(start, stop, count), 1)); - const [t, step] = tickIntervals[target / tickIntervals[i - 1][2] < tickIntervals[i][2] / target ? i - 1 : i]; - return t.every(step); - } - - return [ticks, tickInterval]; -} - -const [utcTicks, utcTickInterval] = ticker(utcYear, utcMonth, utcSunday, utcDay, utcHour, utcMinute); -const [timeTicks, timeTickInterval] = ticker(year, month, sunday, day, hour, minute); - -var src = /*#__PURE__*/Object.freeze({ - __proto__: null, - timeInterval: newInterval, - timeMillisecond: millisecond, - timeMilliseconds: milliseconds, - utcMillisecond: millisecond, - utcMilliseconds: milliseconds, - timeSecond: second, - timeSeconds: seconds, - utcSecond: second, - utcSeconds: seconds, - timeMinute: minute, - timeMinutes: minutes, - timeHour: hour, - timeHours: hours, - timeDay: day, - timeDays: days, - timeWeek: sunday, - timeWeeks: sundays, - timeSunday: sunday, - timeSundays: sundays, - timeMonday: monday, - timeMondays: mondays, - timeTuesday: tuesday, - timeTuesdays: tuesdays, - timeWednesday: wednesday, - timeWednesdays: wednesdays, - timeThursday: thursday, - timeThursdays: thursdays, - timeFriday: friday, - timeFridays: fridays, - timeSaturday: saturday, - timeSaturdays: saturdays, - timeMonth: month, - timeMonths: months, - timeYear: year, - timeYears: years, - utcMinute: utcMinute, - utcMinutes: utcMinutes, - utcHour: utcHour, - utcHours: utcHours, - utcDay: utcDay, - utcDays: utcDays, - utcWeek: utcSunday, - utcWeeks: utcSundays, - utcSunday: utcSunday, - utcSundays: utcSundays, - utcMonday: utcMonday, - utcMondays: utcMondays, - utcTuesday: utcTuesday, - utcTuesdays: utcTuesdays, - utcWednesday: utcWednesday, - utcWednesdays: utcWednesdays, - utcThursday: utcThursday, - utcThursdays: utcThursdays, - utcFriday: utcFriday, - utcFridays: utcFridays, - utcSaturday: utcSaturday, - utcSaturdays: utcSaturdays, - utcMonth: utcMonth, - utcMonths: utcMonths, - utcYear: utcYear, - utcYears: utcYears, - utcTicks: utcTicks, - utcTickInterval: utcTickInterval, - timeTicks: timeTicks, - timeTickInterval: timeTickInterval -}); - -var d3TimeFormat = {exports: {}}; - -var require$$0 = /*@__PURE__*/getAugmentedNamespace(src); - -(function (module, exports) { - // https://d3js.org/d3-time-format/ v3.0.0 Copyright 2020 Mike Bostock - (function (global, factory) { - factory(exports, require$$0) ; - }(commonjsGlobal$1, function (exports, d3Time) { - function localDate(d) { - if (0 <= d.y && d.y < 100) { - var date = new Date(-1, d.m, d.d, d.H, d.M, d.S, d.L); - date.setFullYear(d.y); - return date; - } - return new Date(d.y, d.m, d.d, d.H, d.M, d.S, d.L); - } - - function utcDate(d) { - if (0 <= d.y && d.y < 100) { - var date = new Date(Date.UTC(-1, d.m, d.d, d.H, d.M, d.S, d.L)); - date.setUTCFullYear(d.y); - return date; - } - return new Date(Date.UTC(d.y, d.m, d.d, d.H, d.M, d.S, d.L)); - } - - function newDate(y, m, d) { - return {y: y, m: m, d: d, H: 0, M: 0, S: 0, L: 0}; - } - - function formatLocale(locale) { - var locale_dateTime = locale.dateTime, - locale_date = locale.date, - locale_time = locale.time, - locale_periods = locale.periods, - locale_weekdays = locale.days, - locale_shortWeekdays = locale.shortDays, - locale_months = locale.months, - locale_shortMonths = locale.shortMonths; - - var periodRe = formatRe(locale_periods), - periodLookup = formatLookup(locale_periods), - weekdayRe = formatRe(locale_weekdays), - weekdayLookup = formatLookup(locale_weekdays), - shortWeekdayRe = formatRe(locale_shortWeekdays), - shortWeekdayLookup = formatLookup(locale_shortWeekdays), - monthRe = formatRe(locale_months), - monthLookup = formatLookup(locale_months), - shortMonthRe = formatRe(locale_shortMonths), - shortMonthLookup = formatLookup(locale_shortMonths); - - var formats = { - "a": formatShortWeekday, - "A": formatWeekday, - "b": formatShortMonth, - "B": formatMonth, - "c": null, - "d": formatDayOfMonth, - "e": formatDayOfMonth, - "f": formatMicroseconds, - "g": formatYearISO, - "G": formatFullYearISO, - "H": formatHour24, - "I": formatHour12, - "j": formatDayOfYear, - "L": formatMilliseconds, - "m": formatMonthNumber, - "M": formatMinutes, - "p": formatPeriod, - "q": formatQuarter, - "Q": formatUnixTimestamp, - "s": formatUnixTimestampSeconds, - "S": formatSeconds, - "u": formatWeekdayNumberMonday, - "U": formatWeekNumberSunday, - "V": formatWeekNumberISO, - "w": formatWeekdayNumberSunday, - "W": formatWeekNumberMonday, - "x": null, - "X": null, - "y": formatYear, - "Y": formatFullYear, - "Z": formatZone, - "%": formatLiteralPercent - }; - - var utcFormats = { - "a": formatUTCShortWeekday, - "A": formatUTCWeekday, - "b": formatUTCShortMonth, - "B": formatUTCMonth, - "c": null, - "d": formatUTCDayOfMonth, - "e": formatUTCDayOfMonth, - "f": formatUTCMicroseconds, - "g": formatUTCYearISO, - "G": formatUTCFullYearISO, - "H": formatUTCHour24, - "I": formatUTCHour12, - "j": formatUTCDayOfYear, - "L": formatUTCMilliseconds, - "m": formatUTCMonthNumber, - "M": formatUTCMinutes, - "p": formatUTCPeriod, - "q": formatUTCQuarter, - "Q": formatUnixTimestamp, - "s": formatUnixTimestampSeconds, - "S": formatUTCSeconds, - "u": formatUTCWeekdayNumberMonday, - "U": formatUTCWeekNumberSunday, - "V": formatUTCWeekNumberISO, - "w": formatUTCWeekdayNumberSunday, - "W": formatUTCWeekNumberMonday, - "x": null, - "X": null, - "y": formatUTCYear, - "Y": formatUTCFullYear, - "Z": formatUTCZone, - "%": formatLiteralPercent - }; - - var parses = { - "a": parseShortWeekday, - "A": parseWeekday, - "b": parseShortMonth, - "B": parseMonth, - "c": parseLocaleDateTime, - "d": parseDayOfMonth, - "e": parseDayOfMonth, - "f": parseMicroseconds, - "g": parseYear, - "G": parseFullYear, - "H": parseHour24, - "I": parseHour24, - "j": parseDayOfYear, - "L": parseMilliseconds, - "m": parseMonthNumber, - "M": parseMinutes, - "p": parsePeriod, - "q": parseQuarter, - "Q": parseUnixTimestamp, - "s": parseUnixTimestampSeconds, - "S": parseSeconds, - "u": parseWeekdayNumberMonday, - "U": parseWeekNumberSunday, - "V": parseWeekNumberISO, - "w": parseWeekdayNumberSunday, - "W": parseWeekNumberMonday, - "x": parseLocaleDate, - "X": parseLocaleTime, - "y": parseYear, - "Y": parseFullYear, - "Z": parseZone, - "%": parseLiteralPercent - }; - - // These recursive directive definitions must be deferred. - formats.x = newFormat(locale_date, formats); - formats.X = newFormat(locale_time, formats); - formats.c = newFormat(locale_dateTime, formats); - utcFormats.x = newFormat(locale_date, utcFormats); - utcFormats.X = newFormat(locale_time, utcFormats); - utcFormats.c = newFormat(locale_dateTime, utcFormats); - - function newFormat(specifier, formats) { - return function(date) { - var string = [], - i = -1, - j = 0, - n = specifier.length, - c, - pad, - format; - - if (!(date instanceof Date)) date = new Date(+date); - - while (++i < n) { - if (specifier.charCodeAt(i) === 37) { - string.push(specifier.slice(j, i)); - if ((pad = pads[c = specifier.charAt(++i)]) != null) c = specifier.charAt(++i); - else pad = c === "e" ? " " : "0"; - if (format = formats[c]) c = format(date, pad); - string.push(c); - j = i + 1; - } - } - - string.push(specifier.slice(j, i)); - return string.join(""); - }; - } - - function newParse(specifier, Z) { - return function(string) { - var d = newDate(1900, undefined, 1), - i = parseSpecifier(d, specifier, string += "", 0), - week, day; - if (i != string.length) return null; - - // If a UNIX timestamp is specified, return it. - if ("Q" in d) return new Date(d.Q); - if ("s" in d) return new Date(d.s * 1000 + ("L" in d ? d.L : 0)); - - // If this is utcParse, never use the local timezone. - if (Z && !("Z" in d)) d.Z = 0; - - // The am-pm flag is 0 for AM, and 1 for PM. - if ("p" in d) d.H = d.H % 12 + d.p * 12; - - // If the month was not specified, inherit from the quarter. - if (d.m === undefined) d.m = "q" in d ? d.q : 0; - - // Convert day-of-week and week-of-year to day-of-year. - if ("V" in d) { - if (d.V < 1 || d.V > 53) return null; - if (!("w" in d)) d.w = 1; - if ("Z" in d) { - week = utcDate(newDate(d.y, 0, 1)), day = week.getUTCDay(); - week = day > 4 || day === 0 ? d3Time.utcMonday.ceil(week) : d3Time.utcMonday(week); - week = d3Time.utcDay.offset(week, (d.V - 1) * 7); - d.y = week.getUTCFullYear(); - d.m = week.getUTCMonth(); - d.d = week.getUTCDate() + (d.w + 6) % 7; - } else { - week = localDate(newDate(d.y, 0, 1)), day = week.getDay(); - week = day > 4 || day === 0 ? d3Time.timeMonday.ceil(week) : d3Time.timeMonday(week); - week = d3Time.timeDay.offset(week, (d.V - 1) * 7); - d.y = week.getFullYear(); - d.m = week.getMonth(); - d.d = week.getDate() + (d.w + 6) % 7; - } - } else if ("W" in d || "U" in d) { - if (!("w" in d)) d.w = "u" in d ? d.u % 7 : "W" in d ? 1 : 0; - day = "Z" in d ? utcDate(newDate(d.y, 0, 1)).getUTCDay() : localDate(newDate(d.y, 0, 1)).getDay(); - d.m = 0; - d.d = "W" in d ? (d.w + 6) % 7 + d.W * 7 - (day + 5) % 7 : d.w + d.U * 7 - (day + 6) % 7; - } - - // If a time zone is specified, all fields are interpreted as UTC and then - // offset according to the specified time zone. - if ("Z" in d) { - d.H += d.Z / 100 | 0; - d.M += d.Z % 100; - return utcDate(d); - } - - // Otherwise, all fields are in local time. - return localDate(d); - }; - } - - function parseSpecifier(d, specifier, string, j) { - var i = 0, - n = specifier.length, - m = string.length, - c, - parse; - - while (i < n) { - if (j >= m) return -1; - c = specifier.charCodeAt(i++); - if (c === 37) { - c = specifier.charAt(i++); - parse = parses[c in pads ? specifier.charAt(i++) : c]; - if (!parse || ((j = parse(d, string, j)) < 0)) return -1; - } else if (c != string.charCodeAt(j++)) { - return -1; - } - } - - return j; - } - - function parsePeriod(d, string, i) { - var n = periodRe.exec(string.slice(i)); - return n ? (d.p = periodLookup.get(n[0].toLowerCase()), i + n[0].length) : -1; - } - - function parseShortWeekday(d, string, i) { - var n = shortWeekdayRe.exec(string.slice(i)); - return n ? (d.w = shortWeekdayLookup.get(n[0].toLowerCase()), i + n[0].length) : -1; - } - - function parseWeekday(d, string, i) { - var n = weekdayRe.exec(string.slice(i)); - return n ? (d.w = weekdayLookup.get(n[0].toLowerCase()), i + n[0].length) : -1; - } - - function parseShortMonth(d, string, i) { - var n = shortMonthRe.exec(string.slice(i)); - return n ? (d.m = shortMonthLookup.get(n[0].toLowerCase()), i + n[0].length) : -1; - } - - function parseMonth(d, string, i) { - var n = monthRe.exec(string.slice(i)); - return n ? (d.m = monthLookup.get(n[0].toLowerCase()), i + n[0].length) : -1; - } - - function parseLocaleDateTime(d, string, i) { - return parseSpecifier(d, locale_dateTime, string, i); - } - - function parseLocaleDate(d, string, i) { - return parseSpecifier(d, locale_date, string, i); - } - - function parseLocaleTime(d, string, i) { - return parseSpecifier(d, locale_time, string, i); - } - - function formatShortWeekday(d) { - return locale_shortWeekdays[d.getDay()]; - } - - function formatWeekday(d) { - return locale_weekdays[d.getDay()]; - } - - function formatShortMonth(d) { - return locale_shortMonths[d.getMonth()]; - } - - function formatMonth(d) { - return locale_months[d.getMonth()]; - } - - function formatPeriod(d) { - return locale_periods[+(d.getHours() >= 12)]; - } - - function formatQuarter(d) { - return 1 + ~~(d.getMonth() / 3); - } - - function formatUTCShortWeekday(d) { - return locale_shortWeekdays[d.getUTCDay()]; - } - - function formatUTCWeekday(d) { - return locale_weekdays[d.getUTCDay()]; - } - - function formatUTCShortMonth(d) { - return locale_shortMonths[d.getUTCMonth()]; - } - - function formatUTCMonth(d) { - return locale_months[d.getUTCMonth()]; - } - - function formatUTCPeriod(d) { - return locale_periods[+(d.getUTCHours() >= 12)]; - } - - function formatUTCQuarter(d) { - return 1 + ~~(d.getUTCMonth() / 3); - } - - return { - format: function(specifier) { - var f = newFormat(specifier += "", formats); - f.toString = function() { return specifier; }; - return f; - }, - parse: function(specifier) { - var p = newParse(specifier += "", false); - p.toString = function() { return specifier; }; - return p; - }, - utcFormat: function(specifier) { - var f = newFormat(specifier += "", utcFormats); - f.toString = function() { return specifier; }; - return f; - }, - utcParse: function(specifier) { - var p = newParse(specifier += "", true); - p.toString = function() { return specifier; }; - return p; - } - }; - } - - var pads = {"-": "", "_": " ", "0": "0"}, - numberRe = /^\s*\d+/, // note: ignores next directive - percentRe = /^%/, - requoteRe = /[\\^$*+?|[\]().{}]/g; - - function pad(value, fill, width) { - var sign = value < 0 ? "-" : "", - string = (sign ? -value : value) + "", - length = string.length; - return sign + (length < width ? new Array(width - length + 1).join(fill) + string : string); - } - - function requote(s) { - return s.replace(requoteRe, "\\$&"); - } - - function formatRe(names) { - return new RegExp("^(?:" + names.map(requote).join("|") + ")", "i"); - } - - function formatLookup(names) { - return new Map(names.map((name, i) => [name.toLowerCase(), i])); - } - - function parseWeekdayNumberSunday(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 1)); - return n ? (d.w = +n[0], i + n[0].length) : -1; - } - - function parseWeekdayNumberMonday(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 1)); - return n ? (d.u = +n[0], i + n[0].length) : -1; - } - - function parseWeekNumberSunday(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.U = +n[0], i + n[0].length) : -1; - } - - function parseWeekNumberISO(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.V = +n[0], i + n[0].length) : -1; - } - - function parseWeekNumberMonday(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.W = +n[0], i + n[0].length) : -1; - } - - function parseFullYear(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 4)); - return n ? (d.y = +n[0], i + n[0].length) : -1; - } - - function parseYear(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.y = +n[0] + (+n[0] > 68 ? 1900 : 2000), i + n[0].length) : -1; - } - - function parseZone(d, string, i) { - var n = /^(Z)|([+-]\d\d)(?::?(\d\d))?/.exec(string.slice(i, i + 6)); - return n ? (d.Z = n[1] ? 0 : -(n[2] + (n[3] || "00")), i + n[0].length) : -1; - } - - function parseQuarter(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 1)); - return n ? (d.q = n[0] * 3 - 3, i + n[0].length) : -1; - } - - function parseMonthNumber(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.m = n[0] - 1, i + n[0].length) : -1; - } - - function parseDayOfMonth(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.d = +n[0], i + n[0].length) : -1; - } - - function parseDayOfYear(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 3)); - return n ? (d.m = 0, d.d = +n[0], i + n[0].length) : -1; - } - - function parseHour24(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.H = +n[0], i + n[0].length) : -1; - } - - function parseMinutes(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.M = +n[0], i + n[0].length) : -1; - } - - function parseSeconds(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 2)); - return n ? (d.S = +n[0], i + n[0].length) : -1; - } - - function parseMilliseconds(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 3)); - return n ? (d.L = +n[0], i + n[0].length) : -1; - } - - function parseMicroseconds(d, string, i) { - var n = numberRe.exec(string.slice(i, i + 6)); - return n ? (d.L = Math.floor(n[0] / 1000), i + n[0].length) : -1; - } - - function parseLiteralPercent(d, string, i) { - var n = percentRe.exec(string.slice(i, i + 1)); - return n ? i + n[0].length : -1; - } - - function parseUnixTimestamp(d, string, i) { - var n = numberRe.exec(string.slice(i)); - return n ? (d.Q = +n[0], i + n[0].length) : -1; - } - - function parseUnixTimestampSeconds(d, string, i) { - var n = numberRe.exec(string.slice(i)); - return n ? (d.s = +n[0], i + n[0].length) : -1; - } - - function formatDayOfMonth(d, p) { - return pad(d.getDate(), p, 2); - } - - function formatHour24(d, p) { - return pad(d.getHours(), p, 2); - } - - function formatHour12(d, p) { - return pad(d.getHours() % 12 || 12, p, 2); - } - - function formatDayOfYear(d, p) { - return pad(1 + d3Time.timeDay.count(d3Time.timeYear(d), d), p, 3); - } - - function formatMilliseconds(d, p) { - return pad(d.getMilliseconds(), p, 3); - } - - function formatMicroseconds(d, p) { - return formatMilliseconds(d, p) + "000"; - } - - function formatMonthNumber(d, p) { - return pad(d.getMonth() + 1, p, 2); - } - - function formatMinutes(d, p) { - return pad(d.getMinutes(), p, 2); - } - - function formatSeconds(d, p) { - return pad(d.getSeconds(), p, 2); - } - - function formatWeekdayNumberMonday(d) { - var day = d.getDay(); - return day === 0 ? 7 : day; - } - - function formatWeekNumberSunday(d, p) { - return pad(d3Time.timeSunday.count(d3Time.timeYear(d) - 1, d), p, 2); - } - - function dISO(d) { - var day = d.getDay(); - return (day >= 4 || day === 0) ? d3Time.timeThursday(d) : d3Time.timeThursday.ceil(d); - } - - function formatWeekNumberISO(d, p) { - d = dISO(d); - return pad(d3Time.timeThursday.count(d3Time.timeYear(d), d) + (d3Time.timeYear(d).getDay() === 4), p, 2); - } - - function formatWeekdayNumberSunday(d) { - return d.getDay(); - } - - function formatWeekNumberMonday(d, p) { - return pad(d3Time.timeMonday.count(d3Time.timeYear(d) - 1, d), p, 2); - } - - function formatYear(d, p) { - return pad(d.getFullYear() % 100, p, 2); - } - - function formatYearISO(d, p) { - d = dISO(d); - return pad(d.getFullYear() % 100, p, 2); - } - - function formatFullYear(d, p) { - return pad(d.getFullYear() % 10000, p, 4); - } - - function formatFullYearISO(d, p) { - var day = d.getDay(); - d = (day >= 4 || day === 0) ? d3Time.timeThursday(d) : d3Time.timeThursday.ceil(d); - return pad(d.getFullYear() % 10000, p, 4); - } - - function formatZone(d) { - var z = d.getTimezoneOffset(); - return (z > 0 ? "-" : (z *= -1, "+")) - + pad(z / 60 | 0, "0", 2) - + pad(z % 60, "0", 2); - } - - function formatUTCDayOfMonth(d, p) { - return pad(d.getUTCDate(), p, 2); - } - - function formatUTCHour24(d, p) { - return pad(d.getUTCHours(), p, 2); - } - - function formatUTCHour12(d, p) { - return pad(d.getUTCHours() % 12 || 12, p, 2); - } - - function formatUTCDayOfYear(d, p) { - return pad(1 + d3Time.utcDay.count(d3Time.utcYear(d), d), p, 3); - } - - function formatUTCMilliseconds(d, p) { - return pad(d.getUTCMilliseconds(), p, 3); - } - - function formatUTCMicroseconds(d, p) { - return formatUTCMilliseconds(d, p) + "000"; - } - - function formatUTCMonthNumber(d, p) { - return pad(d.getUTCMonth() + 1, p, 2); - } - - function formatUTCMinutes(d, p) { - return pad(d.getUTCMinutes(), p, 2); - } - - function formatUTCSeconds(d, p) { - return pad(d.getUTCSeconds(), p, 2); - } - - function formatUTCWeekdayNumberMonday(d) { - var dow = d.getUTCDay(); - return dow === 0 ? 7 : dow; - } - - function formatUTCWeekNumberSunday(d, p) { - return pad(d3Time.utcSunday.count(d3Time.utcYear(d) - 1, d), p, 2); - } - - function UTCdISO(d) { - var day = d.getUTCDay(); - return (day >= 4 || day === 0) ? d3Time.utcThursday(d) : d3Time.utcThursday.ceil(d); - } - - function formatUTCWeekNumberISO(d, p) { - d = UTCdISO(d); - return pad(d3Time.utcThursday.count(d3Time.utcYear(d), d) + (d3Time.utcYear(d).getUTCDay() === 4), p, 2); - } - - function formatUTCWeekdayNumberSunday(d) { - return d.getUTCDay(); - } - - function formatUTCWeekNumberMonday(d, p) { - return pad(d3Time.utcMonday.count(d3Time.utcYear(d) - 1, d), p, 2); - } - - function formatUTCYear(d, p) { - return pad(d.getUTCFullYear() % 100, p, 2); - } - - function formatUTCYearISO(d, p) { - d = UTCdISO(d); - return pad(d.getUTCFullYear() % 100, p, 2); - } - - function formatUTCFullYear(d, p) { - return pad(d.getUTCFullYear() % 10000, p, 4); - } - - function formatUTCFullYearISO(d, p) { - var day = d.getUTCDay(); - d = (day >= 4 || day === 0) ? d3Time.utcThursday(d) : d3Time.utcThursday.ceil(d); - return pad(d.getUTCFullYear() % 10000, p, 4); - } - - function formatUTCZone() { - return "+0000"; - } - - function formatLiteralPercent() { - return "%"; - } - - function formatUnixTimestamp(d) { - return +d; - } - - function formatUnixTimestampSeconds(d) { - return Math.floor(+d / 1000); - } - - var locale; - - defaultLocale({ - dateTime: "%x, %X", - date: "%-m/%-d/%Y", - time: "%-I:%M:%S %p", - periods: ["AM", "PM"], - days: ["Sunday", "Monday", "Tuesday", "Wednesday", "Thursday", "Friday", "Saturday"], - shortDays: ["Sun", "Mon", "Tue", "Wed", "Thu", "Fri", "Sat"], - months: ["January", "February", "March", "April", "May", "June", "July", "August", "September", "October", "November", "December"], - shortMonths: ["Jan", "Feb", "Mar", "Apr", "May", "Jun", "Jul", "Aug", "Sep", "Oct", "Nov", "Dec"] - }); - - function defaultLocale(definition) { - locale = formatLocale(definition); - exports.timeFormat = locale.format; - exports.timeParse = locale.parse; - exports.utcFormat = locale.utcFormat; - exports.utcParse = locale.utcParse; - return locale; - } - - var isoSpecifier = "%Y-%m-%dT%H:%M:%S.%LZ"; - - function formatIsoNative(date) { - return date.toISOString(); - } - - var formatIso = Date.prototype.toISOString - ? formatIsoNative - : exports.utcFormat(isoSpecifier); - - function parseIsoNative(string) { - var date = new Date(string); - return isNaN(date) ? null : date; - } - - var parseIso = +new Date("2000-01-01T00:00:00.000Z") - ? parseIsoNative - : exports.utcParse(isoSpecifier); - - exports.isoFormat = formatIso; - exports.isoParse = parseIso; - exports.timeFormatDefaultLocale = defaultLocale; - exports.timeFormatLocale = formatLocale; - - Object.defineProperty(exports, '__esModule', { value: true }); - - })); -} (d3TimeFormat, d3TimeFormat.exports)); - -function date(t) { - return new Date(t); -} - -function number(t) { - return t instanceof Date ? +t : +new Date(+t); -} - -function calendar(ticks, tickInterval, year, month, week, day, hour, minute, second, format) { - var scale = continuous(), - invert = scale.invert, - domain = scale.domain; - - var formatMillisecond = format(".%L"), - formatSecond = format(":%S"), - formatMinute = format("%I:%M"), - formatHour = format("%I %p"), - formatDay = format("%a %d"), - formatWeek = format("%b %d"), - formatMonth = format("%B"), - formatYear = format("%Y"); - - function tickFormat(date) { - return (second(date) < date ? formatMillisecond - : minute(date) < date ? formatSecond - : hour(date) < date ? formatMinute - : day(date) < date ? formatHour - : month(date) < date ? (week(date) < date ? formatDay : formatWeek) - : year(date) < date ? formatMonth - : formatYear)(date); - } - - scale.invert = function(y) { - return new Date(invert(y)); - }; - - scale.domain = function(_) { - return arguments.length ? domain(Array.from(_, number)) : domain().map(date); - }; - - scale.ticks = function(interval) { - var d = domain(); - return ticks(d[0], d[d.length - 1], interval == null ? 10 : interval); - }; - - scale.tickFormat = function(count, specifier) { - return specifier == null ? tickFormat : format(specifier); - }; - - scale.nice = function(interval) { - var d = domain(); - if (!interval || typeof interval.range !== "function") interval = tickInterval(d[0], d[d.length - 1], interval == null ? 10 : interval); - return interval ? domain(nice(d, interval)) : scale; - }; - - scale.copy = function() { - return copy$1(scale, calendar(ticks, tickInterval, year, month, week, day, hour, minute, second, format)); - }; - - return scale; -} - -function time() { - return initRange.apply(calendar(timeTicks, timeTickInterval, year, month, sunday, day, hour, minute, second, d3TimeFormat.exports.timeFormat).domain([new Date(2000, 0, 1), new Date(2000, 0, 2)]), arguments); -} - -function utcTime() { - return initRange.apply(calendar(utcTicks, utcTickInterval, utcYear, utcMonth, utcSunday, utcDay, utcHour, utcMinute, second, d3TimeFormat.exports.utcFormat).domain([Date.UTC(2000, 0, 1), Date.UTC(2000, 0, 2)]), arguments); -} - -function transformer$1() { - var x0 = 0, - x1 = 1, - t0, - t1, - k10, - transform, - interpolator = identity$1, - clamp = false, - unknown; - - function scale(x) { - return x == null || isNaN(x = +x) ? unknown : interpolator(k10 === 0 ? 0.5 : (x = (transform(x) - t0) * k10, clamp ? Math.max(0, Math.min(1, x)) : x)); - } - - scale.domain = function(_) { - return arguments.length ? ([x0, x1] = _, t0 = transform(x0 = +x0), t1 = transform(x1 = +x1), k10 = t0 === t1 ? 0 : 1 / (t1 - t0), scale) : [x0, x1]; - }; - - scale.clamp = function(_) { - return arguments.length ? (clamp = !!_, scale) : clamp; - }; - - scale.interpolator = function(_) { - return arguments.length ? (interpolator = _, scale) : interpolator; - }; - - function range(interpolate) { - return function(_) { - var r0, r1; - return arguments.length ? ([r0, r1] = _, interpolator = interpolate(r0, r1), scale) : [interpolator(0), interpolator(1)]; - }; - } - - scale.range = range(interpolate); - - scale.rangeRound = range(interpolateRound); - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : unknown; - }; - - return function(t) { - transform = t, t0 = t(x0), t1 = t(x1), k10 = t0 === t1 ? 0 : 1 / (t1 - t0); - return scale; - }; -} - -function copy(source, target) { - return target - .domain(source.domain()) - .interpolator(source.interpolator()) - .clamp(source.clamp()) - .unknown(source.unknown()); -} - -function sequential() { - var scale = linearish(transformer$1()(identity$1)); - - scale.copy = function() { - return copy(scale, sequential()); - }; - - return initInterpolator.apply(scale, arguments); -} - -function sequentialLog() { - var scale = loggish(transformer$1()).domain([1, 10]); - - scale.copy = function() { - return copy(scale, sequentialLog()).base(scale.base()); - }; - - return initInterpolator.apply(scale, arguments); -} - -function sequentialSymlog() { - var scale = symlogish(transformer$1()); - - scale.copy = function() { - return copy(scale, sequentialSymlog()).constant(scale.constant()); - }; - - return initInterpolator.apply(scale, arguments); -} - -function sequentialPow() { - var scale = powish(transformer$1()); - - scale.copy = function() { - return copy(scale, sequentialPow()).exponent(scale.exponent()); - }; - - return initInterpolator.apply(scale, arguments); -} - -function sequentialSqrt() { - return sequentialPow.apply(null, arguments).exponent(0.5); -} - -function sequentialQuantile() { - var domain = [], - interpolator = identity$1; - - function scale(x) { - if (x != null && !isNaN(x = +x)) return interpolator((d3Array.exports.bisect(domain, x, 1) - 1) / (domain.length - 1)); - } - - scale.domain = function(_) { - if (!arguments.length) return domain.slice(); - domain = []; - for (let d of _) if (d != null && !isNaN(d = +d)) domain.push(d); - domain.sort(d3Array.exports.ascending); - return scale; - }; - - scale.interpolator = function(_) { - return arguments.length ? (interpolator = _, scale) : interpolator; - }; - - scale.range = function() { - return domain.map((d, i) => interpolator(i / (domain.length - 1))); - }; - - scale.quantiles = function(n) { - return Array.from({length: n + 1}, (_, i) => d3Array.exports.quantile(domain, i / n)); - }; - - scale.copy = function() { - return sequentialQuantile(interpolator).domain(domain); - }; - - return initInterpolator.apply(scale, arguments); -} - -function transformer() { - var x0 = 0, - x1 = 0.5, - x2 = 1, - s = 1, - t0, - t1, - t2, - k10, - k21, - interpolator = identity$1, - transform, - clamp = false, - unknown; - - function scale(x) { - return isNaN(x = +x) ? unknown : (x = 0.5 + ((x = +transform(x)) - t1) * (s * x < s * t1 ? k10 : k21), interpolator(clamp ? Math.max(0, Math.min(1, x)) : x)); - } - - scale.domain = function(_) { - return arguments.length ? ([x0, x1, x2] = _, t0 = transform(x0 = +x0), t1 = transform(x1 = +x1), t2 = transform(x2 = +x2), k10 = t0 === t1 ? 0 : 0.5 / (t1 - t0), k21 = t1 === t2 ? 0 : 0.5 / (t2 - t1), s = t1 < t0 ? -1 : 1, scale) : [x0, x1, x2]; - }; - - scale.clamp = function(_) { - return arguments.length ? (clamp = !!_, scale) : clamp; - }; - - scale.interpolator = function(_) { - return arguments.length ? (interpolator = _, scale) : interpolator; - }; - - function range(interpolate) { - return function(_) { - var r0, r1, r2; - return arguments.length ? ([r0, r1, r2] = _, interpolator = piecewise(interpolate, [r0, r1, r2]), scale) : [interpolator(0), interpolator(0.5), interpolator(1)]; - }; - } - - scale.range = range(interpolate); - - scale.rangeRound = range(interpolateRound); - - scale.unknown = function(_) { - return arguments.length ? (unknown = _, scale) : unknown; - }; - - return function(t) { - transform = t, t0 = t(x0), t1 = t(x1), t2 = t(x2), k10 = t0 === t1 ? 0 : 0.5 / (t1 - t0), k21 = t1 === t2 ? 0 : 0.5 / (t2 - t1), s = t1 < t0 ? -1 : 1; - return scale; - }; -} - -function diverging() { - var scale = linearish(transformer()(identity$1)); - - scale.copy = function() { - return copy(scale, diverging()); - }; - - return initInterpolator.apply(scale, arguments); -} - -function divergingLog() { - var scale = loggish(transformer()).domain([0.1, 1, 10]); - - scale.copy = function() { - return copy(scale, divergingLog()).base(scale.base()); - }; - - return initInterpolator.apply(scale, arguments); -} - -function divergingSymlog() { - var scale = symlogish(transformer()); - - scale.copy = function() { - return copy(scale, divergingSymlog()).constant(scale.constant()); - }; - - return initInterpolator.apply(scale, arguments); -} - -function divergingPow() { - var scale = powish(transformer()); - - scale.copy = function() { - return copy(scale, divergingPow()).exponent(scale.exponent()); - }; - - return initInterpolator.apply(scale, arguments); -} - -function divergingSqrt() { - return divergingPow.apply(null, arguments).exponent(0.5); -} - -var d3Scales = /*#__PURE__*/Object.freeze({ - __proto__: null, - scaleBand: band, - scalePoint: point, - scaleIdentity: identity, - scaleLinear: linear, - scaleLog: log, - scaleSymlog: symlog, - scaleOrdinal: ordinal, - scaleImplicit: implicit, - scalePow: pow, - scaleSqrt: sqrt, - scaleRadial: radial, - scaleQuantile: quantile, - scaleQuantize: quantize, - scaleThreshold: threshold, - scaleTime: time, - scaleUtc: utcTime, - scaleSequential: sequential, - scaleSequentialLog: sequentialLog, - scaleSequentialPow: sequentialPow, - scaleSequentialSqrt: sequentialSqrt, - scaleSequentialSymlog: sequentialSymlog, - scaleSequentialQuantile: sequentialQuantile, - scaleDiverging: diverging, - scaleDivergingLog: divergingLog, - scaleDivergingPow: divergingPow, - scaleDivergingSqrt: divergingSqrt, - scaleDivergingSymlog: divergingSymlog, - tickFormat: tickFormat -}); - -function _toConsumableArray$5(arr) { return _arrayWithoutHoles$5(arr) || _iterableToArray$5(arr) || _unsupportedIterableToArray$9(arr) || _nonIterableSpread$5(); } - -function _nonIterableSpread$5() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$9(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$9(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$9(o, minLen); } - -function _iterableToArray$5(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$5(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$9(arr); } - -function _arrayLikeToArray$9(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function ownKeys$q(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$q(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$q(Object(source), true).forEach(function (key) { _defineProperty$q(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$q(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$q(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } -function getValueByDataKey(obj, dataKey, defaultValue) { - if (isNil_1(obj) || isNil_1(dataKey)) { - return defaultValue; - } - - if (isNumOrStr(dataKey)) { - return get_1(obj, dataKey, defaultValue); - } - - if (isFunction_1(dataKey)) { - return dataKey(obj); - } - - return defaultValue; -} -/** - * Get domain of data by key - * @param {Array} data The data displayed in the chart - * @param {String} key The unique key of a group of data - * @param {String} type The type of axis - * @param {Boolean} filterNil Whether or not filter nil values - * @return {Array} Domain of data - */ - -function getDomainOfDataByKey(data, key, type, filterNil) { - var flattenData = flatMap_1(data, function (entry) { - return getValueByDataKey(entry, key); - }); - - if (type === 'number') { - var domain = flattenData.filter(function (entry) { - return isNumber(entry) || parseFloat(entry); - }); - return domain.length ? [min_1(domain), max_1(domain)] : [Infinity, -Infinity]; - } - - var validateData = filterNil ? flattenData.filter(function (entry) { - return !isNil_1(entry); - }) : flattenData; // 支持Date类型的x轴 - - return validateData.map(function (entry) { - return isNumOrStr(entry) || entry instanceof Date ? entry : ''; - }); -} -var calculateActiveTickIndex = function calculateActiveTickIndex(coordinate) { - var _ticks$length; - - var ticks = arguments.length > 1 && arguments[1] !== undefined ? arguments[1] : []; - var unsortedTicks = arguments.length > 2 ? arguments[2] : undefined; - var axis = arguments.length > 3 ? arguments[3] : undefined; - var index = -1; - var len = (_ticks$length = ticks === null || ticks === void 0 ? void 0 : ticks.length) !== null && _ticks$length !== void 0 ? _ticks$length : 0; - - if (len > 1) { - if (axis && axis.axisType === 'angleAxis' && Math.abs(Math.abs(axis.range[1] - axis.range[0]) - 360) <= 1e-6) { - var range = axis.range; // ticks are distributed in a circle - - for (var i = 0; i < len; i++) { - var before = i > 0 ? unsortedTicks[i - 1].coordinate : unsortedTicks[len - 1].coordinate; - var cur = unsortedTicks[i].coordinate; - var after = i >= len - 1 ? unsortedTicks[0].coordinate : unsortedTicks[i + 1].coordinate; - var sameDirectionCoord = void 0; - - if (mathSign(cur - before) !== mathSign(after - cur)) { - var diffInterval = []; - - if (mathSign(after - cur) === mathSign(range[1] - range[0])) { - sameDirectionCoord = after; - var curInRange = cur + range[1] - range[0]; - diffInterval[0] = Math.min(curInRange, (curInRange + before) / 2); - diffInterval[1] = Math.max(curInRange, (curInRange + before) / 2); - } else { - sameDirectionCoord = before; - var afterInRange = after + range[1] - range[0]; - diffInterval[0] = Math.min(cur, (afterInRange + cur) / 2); - diffInterval[1] = Math.max(cur, (afterInRange + cur) / 2); - } - - var sameInterval = [Math.min(cur, (sameDirectionCoord + cur) / 2), Math.max(cur, (sameDirectionCoord + cur) / 2)]; - - if (coordinate > sameInterval[0] && coordinate <= sameInterval[1] || coordinate >= diffInterval[0] && coordinate <= diffInterval[1]) { - index = unsortedTicks[i].index; - break; - } - } else { - var min = Math.min(before, after); - var max = Math.max(before, after); - - if (coordinate > (min + cur) / 2 && coordinate <= (max + cur) / 2) { - index = unsortedTicks[i].index; - break; - } - } - } - } else { - // ticks are distributed in a single direction - for (var _i = 0; _i < len; _i++) { - if (_i === 0 && coordinate <= (ticks[_i].coordinate + ticks[_i + 1].coordinate) / 2 || _i > 0 && _i < len - 1 && coordinate > (ticks[_i].coordinate + ticks[_i - 1].coordinate) / 2 && coordinate <= (ticks[_i].coordinate + ticks[_i + 1].coordinate) / 2 || _i === len - 1 && coordinate > (ticks[_i].coordinate + ticks[_i - 1].coordinate) / 2) { - index = ticks[_i].index; - break; - } - } - } - } else { - index = 0; - } - - return index; -}; -/** - * Get the main color of each graphic item - * @param {ReactElement} item A graphic item - * @return {String} Color - */ - -var getMainColorOfGraphicItem = function getMainColorOfGraphicItem(item) { - var _ref = item, - displayName = _ref.type.displayName; // TODO: check if displayName is valid. - - var _item$props = item.props, - stroke = _item$props.stroke, - fill = _item$props.fill; - var result; - - switch (displayName) { - case 'Line': - result = stroke; - break; - - case 'Area': - case 'Radar': - result = stroke && stroke !== 'none' ? stroke : fill; - break; - - default: - result = fill; - break; - } - - return result; -}; -var getLegendProps = function getLegendProps(_ref2) { - var children = _ref2.children, - formattedGraphicalItems = _ref2.formattedGraphicalItems, - legendWidth = _ref2.legendWidth, - legendContent = _ref2.legendContent; - var legendItem = findChildByType(children, Legend.displayName); - - if (!legendItem) { - return null; - } - - var legendData; - - if (legendItem.props && legendItem.props.payload) { - legendData = legendItem.props && legendItem.props.payload; - } else if (legendContent === 'children') { - legendData = (formattedGraphicalItems || []).reduce(function (result, _ref3) { - var item = _ref3.item, - props = _ref3.props; - var data = props.sectors || props.data || []; - return result.concat(data.map(function (entry) { - return { - type: legendItem.props.iconType || item.props.legendType, - value: entry.name, - color: entry.fill, - payload: entry - }; - })); - }, []); - } else { - legendData = (formattedGraphicalItems || []).map(function (_ref4) { - var item = _ref4.item; - var _item$props2 = item.props, - dataKey = _item$props2.dataKey, - name = _item$props2.name, - legendType = _item$props2.legendType, - hide = _item$props2.hide; - return { - inactive: hide, - dataKey: dataKey, - type: legendItem.props.iconType || legendType || 'square', - color: getMainColorOfGraphicItem(item), - value: name || dataKey, - payload: item.props - }; - }); - } - - return _objectSpread$q(_objectSpread$q(_objectSpread$q({}, legendItem.props), Legend.getWithHeight(legendItem, legendWidth)), {}, { - payload: legendData, - item: legendItem - }); -}; -/** - * Calculate the size of all groups for stacked bar graph - * @param {Object} stackGroups The items grouped by axisId and stackId - * @return {Object} The size of all groups - */ - -var getBarSizeList = function getBarSizeList(_ref5) { - var globalSize = _ref5.barSize, - _ref5$stackGroups = _ref5.stackGroups, - stackGroups = _ref5$stackGroups === void 0 ? {} : _ref5$stackGroups; - - if (!stackGroups) { - return {}; - } - - var result = {}; - var numericAxisIds = Object.keys(stackGroups); - - for (var i = 0, len = numericAxisIds.length; i < len; i++) { - var sgs = stackGroups[numericAxisIds[i]].stackGroups; - var stackIds = Object.keys(sgs); - - for (var j = 0, sLen = stackIds.length; j < sLen; j++) { - var _sgs$stackIds$j = sgs[stackIds[j]], - items = _sgs$stackIds$j.items, - cateAxisId = _sgs$stackIds$j.cateAxisId; - var barItems = items.filter(function (item) { - return getDisplayName(item.type).indexOf('Bar') >= 0; - }); - - if (barItems && barItems.length) { - var selfSize = barItems[0].props.barSize; - var cateId = barItems[0].props[cateAxisId]; - - if (!result[cateId]) { - result[cateId] = []; - } - - result[cateId].push({ - item: barItems[0], - stackList: barItems.slice(1), - barSize: isNil_1(selfSize) ? globalSize : selfSize - }); - } - } - } - - return result; -}; -/** - * Calculate the size of each bar and the gap between two bars - * @param {Number} bandSize The size of each category - * @param {sizeList} sizeList The size of all groups - * @param {maxBarSize} maxBarSize The maximum size of bar - * @return {Number} The size of each bar and the gap between two bars - */ - -var getBarPosition = function getBarPosition(_ref6) { - var barGap = _ref6.barGap, - barCategoryGap = _ref6.barCategoryGap, - bandSize = _ref6.bandSize, - _ref6$sizeList = _ref6.sizeList, - sizeList = _ref6$sizeList === void 0 ? [] : _ref6$sizeList, - maxBarSize = _ref6.maxBarSize; - var len = sizeList.length; - if (len < 1) return null; - var realBarGap = getPercentValue(barGap, bandSize, 0, true); - var result; // whether or not is barSize setted by user - - if (sizeList[0].barSize === +sizeList[0].barSize) { - var useFull = false; - var fullBarSize = bandSize / len; - var sum = sizeList.reduce(function (res, entry) { - return res + entry.barSize || 0; - }, 0); - sum += (len - 1) * realBarGap; - - if (sum >= bandSize) { - sum -= (len - 1) * realBarGap; - realBarGap = 0; - } - - if (sum >= bandSize && fullBarSize > 0) { - useFull = true; - fullBarSize *= 0.9; - sum = len * fullBarSize; - } - - var offset = (bandSize - sum) / 2 >> 0; - var prev = { - offset: offset - realBarGap, - size: 0 - }; - result = sizeList.reduce(function (res, entry) { - var newRes = [].concat(_toConsumableArray$5(res), [{ - item: entry.item, - position: { - offset: prev.offset + prev.size + realBarGap, - size: useFull ? fullBarSize : entry.barSize - } - }]); - prev = newRes[newRes.length - 1].position; - - if (entry.stackList && entry.stackList.length) { - entry.stackList.forEach(function (item) { - newRes.push({ - item: item, - position: prev - }); - }); - } - - return newRes; - }, []); - } else { - var _offset = getPercentValue(barCategoryGap, bandSize, 0, true); - - if (bandSize - 2 * _offset - (len - 1) * realBarGap <= 0) { - realBarGap = 0; - } - - var originalSize = (bandSize - 2 * _offset - (len - 1) * realBarGap) / len; - - if (originalSize > 1) { - originalSize >>= 0; - } - - var size = maxBarSize === +maxBarSize ? Math.min(originalSize, maxBarSize) : originalSize; - result = sizeList.reduce(function (res, entry, i) { - var newRes = [].concat(_toConsumableArray$5(res), [{ - item: entry.item, - position: { - offset: _offset + (originalSize + realBarGap) * i + (originalSize - size) / 2, - size: size - } - }]); - - if (entry.stackList && entry.stackList.length) { - entry.stackList.forEach(function (item) { - newRes.push({ - item: item, - position: newRes[newRes.length - 1].position - }); - }); - } - - return newRes; - }, []); - } - - return result; -}; -var appendOffsetOfLegend = function appendOffsetOfLegend(offset, items, props, legendBox) { - var children = props.children, - width = props.width, - margin = props.margin; - var legendWidth = width - (margin.left || 0) - (margin.right || 0); // const legendHeight = height - (margin.top || 0) - (margin.bottom || 0); - - var legendProps = getLegendProps({ - children: children, - legendWidth: legendWidth - }); - var newOffset = offset; - - if (legendProps) { - var box = legendBox || {}; - var align = legendProps.align, - verticalAlign = legendProps.verticalAlign, - layout = legendProps.layout; - - if ((layout === 'vertical' || layout === 'horizontal' && verticalAlign === 'center') && isNumber(offset[align])) { - newOffset = _objectSpread$q(_objectSpread$q({}, offset), {}, _defineProperty$q({}, align, newOffset[align] + (box.width || 0))); - } - - if ((layout === 'horizontal' || layout === 'vertical' && align === 'center') && isNumber(offset[verticalAlign])) { - newOffset = _objectSpread$q(_objectSpread$q({}, offset), {}, _defineProperty$q({}, verticalAlign, newOffset[verticalAlign] + (box.height || 0))); - } - } - - return newOffset; -}; - -var isErrorBarRelevantForAxis = function isErrorBarRelevantForAxis(layout, axisType, direction) { - if (isNil_1(axisType)) { - return true; - } - - if (layout === 'horizontal') { - return axisType === 'yAxis'; - } - - if (layout === 'vertical') { - return axisType === 'xAxis'; - } - - if (direction === 'x') { - return axisType === 'xAxis'; - } - - if (direction === 'y') { - return axisType === 'yAxis'; - } - - return true; -}; - -var getDomainOfErrorBars = function getDomainOfErrorBars(data, item, dataKey, layout, axisType) { - var children = item.props.children; - var errorBars = findAllByType(children, 'ErrorBar').filter(function (errorBarChild) { - return isErrorBarRelevantForAxis(layout, axisType, errorBarChild.props.direction); - }); - - if (errorBars && errorBars.length) { - var keys = errorBars.map(function (errorBarChild) { - return errorBarChild.props.dataKey; - }); - return data.reduce(function (result, entry) { - var entryValue = getValueByDataKey(entry, dataKey, 0); - var mainValue = isArray_1(entryValue) ? [min_1(entryValue), max_1(entryValue)] : [entryValue, entryValue]; - var errorDomain = keys.reduce(function (prevErrorArr, k) { - var errorValue = getValueByDataKey(entry, k, 0); - var lowerValue = mainValue[0] - Math.abs(isArray_1(errorValue) ? errorValue[0] : errorValue); - var upperValue = mainValue[1] + Math.abs(isArray_1(errorValue) ? errorValue[1] : errorValue); - return [Math.min(lowerValue, prevErrorArr[0]), Math.max(upperValue, prevErrorArr[1])]; - }, [Infinity, -Infinity]); - return [Math.min(errorDomain[0], result[0]), Math.max(errorDomain[1], result[1])]; - }, [Infinity, -Infinity]); - } - - return null; -}; -var parseErrorBarsOfAxis = function parseErrorBarsOfAxis(data, items, dataKey, axisType, layout) { - var domains = items.map(function (item) { - return getDomainOfErrorBars(data, item, dataKey, layout, axisType); - }).filter(function (entry) { - return !isNil_1(entry); - }); - - if (domains && domains.length) { - return domains.reduce(function (result, entry) { - return [Math.min(result[0], entry[0]), Math.max(result[1], entry[1])]; - }, [Infinity, -Infinity]); - } - - return null; -}; -/** - * Get domain of data by the configuration of item element - * @param {Array} data The data displayed in the chart - * @param {Array} items The instances of item - * @param {String} type The type of axis, number - Number Axis, category - Category Axis - * @param {LayoutType} layout The type of layout - * @param {Boolean} filterNil Whether or not filter nil values - * @return {Array} Domain - */ - -var getDomainOfItemsWithSameAxis = function getDomainOfItemsWithSameAxis(data, items, type, layout, filterNil) { - var domains = items.map(function (item) { - var dataKey = item.props.dataKey; - - if (type === 'number' && dataKey) { - return getDomainOfErrorBars(data, item, dataKey, layout) || getDomainOfDataByKey(data, dataKey, type, filterNil); - } - - return getDomainOfDataByKey(data, dataKey, type, filterNil); - }); - - if (type === 'number') { - // Calculate the domain of number axis - return domains.reduce(function (result, entry) { - return [Math.min(result[0], entry[0]), Math.max(result[1], entry[1])]; - }, [Infinity, -Infinity]); - } - - var tag = {}; // Get the union set of category axis - - return domains.reduce(function (result, entry) { - for (var i = 0, len = entry.length; i < len; i++) { - if (!tag[entry[i]]) { - tag[entry[i]] = true; - result.push(entry[i]); - } - } - - return result; - }, []); -}; -var isCategoricalAxis = function isCategoricalAxis(layout, axisType) { - return layout === 'horizontal' && axisType === 'xAxis' || layout === 'vertical' && axisType === 'yAxis' || layout === 'centric' && axisType === 'angleAxis' || layout === 'radial' && axisType === 'radiusAxis'; -}; -/** - * Calculate the Coordinates of grid - * @param {Array} ticks The ticks in axis - * @param {Number} min The minimun value of axis - * @param {Number} max The maximun value of axis - * @return {Array} Coordinates - */ - -var getCoordinatesOfGrid = function getCoordinatesOfGrid(ticks, min, max) { - var hasMin, hasMax; - var values = ticks.map(function (entry) { - if (entry.coordinate === min) { - hasMin = true; - } - - if (entry.coordinate === max) { - hasMax = true; - } - - return entry.coordinate; - }); - - if (!hasMin) { - values.push(min); - } - - if (!hasMax) { - values.push(max); - } - - return values; -}; -/** - * Get the ticks of an axis - * @param {Object} axis The configuration of an axis - * @param {Boolean} isGrid Whether or not are the ticks in grid - * @param {Boolean} isAll Return the ticks of all the points or not - * @return {Array} Ticks - */ - -var getTicksOfAxis = function getTicksOfAxis(axis, isGrid, isAll) { - if (!axis) return null; - var scale = axis.scale; - var duplicateDomain = axis.duplicateDomain, - type = axis.type, - range = axis.range; - var offsetForBand = axis.realScaleType === 'scaleBand' ? scale.bandwidth() / 2 : 2; - var offset = (isGrid || isAll) && type === 'category' && scale.bandwidth ? scale.bandwidth() / offsetForBand : 0; - offset = axis.axisType === 'angleAxis' ? mathSign(range[0] - range[1]) * 2 * offset : offset; // The ticks setted by user should only affect the ticks adjacent to axis line - - if (isGrid && (axis.ticks || axis.niceTicks)) { - return (axis.ticks || axis.niceTicks).map(function (entry) { - var scaleContent = duplicateDomain ? duplicateDomain.indexOf(entry) : entry; - return { - coordinate: scale(scaleContent) + offset, - value: entry, - offset: offset - }; - }); - } // When axis is a categorial axis, but the type of axis is number or the scale of axis is not "auto" - - - if (axis.isCategorical && axis.categoricalDomain) { - return axis.categoricalDomain.map(function (entry, index) { - return { - coordinate: scale(entry) + offset, - value: entry, - index: index, - offset: offset - }; - }); - } - - if (scale.ticks && !isAll) { - return scale.ticks(axis.tickCount).map(function (entry) { - return { - coordinate: scale(entry) + offset, - value: entry, - offset: offset - }; - }); - } // When axis has duplicated text, serial numbers are used to generate scale - - - return scale.domain().map(function (entry, index) { - return { - coordinate: scale(entry) + offset, - value: duplicateDomain ? duplicateDomain[entry] : entry, - index: index, - offset: offset - }; - }); -}; -/** - * combine the handlers - * @param {Function} defaultHandler Internal private handler - * @param {Function} parentHandler Handler function specified in parent component - * @param {Function} childHandler Handler function specified in child component - * @return {Function} The combined handler - */ - -var combineEventHandlers = function combineEventHandlers(defaultHandler, parentHandler, childHandler) { - var customizedHandler; - - if (isFunction_1(childHandler)) { - customizedHandler = childHandler; - } else if (isFunction_1(parentHandler)) { - customizedHandler = parentHandler; - } - - if (isFunction_1(defaultHandler) || customizedHandler) { - return function (arg1, arg2, arg3, arg4) { - if (isFunction_1(defaultHandler)) { - defaultHandler(arg1, arg2, arg3, arg4); - } - - if (isFunction_1(customizedHandler)) { - customizedHandler(arg1, arg2, arg3, arg4); - } - }; - } - - return null; -}; -/** - * Parse the scale function of axis - * @param {Object} axis The option of axis - * @param {String} chartType The displayName of chart - * @param {Boolean} hasBar if it has a bar - * @return {Function} The scale function - */ - -var parseScale = function parseScale(axis, chartType, hasBar) { - var scale = axis.scale, - type = axis.type, - layout = axis.layout, - axisType = axis.axisType; - - if (scale === 'auto') { - if (layout === 'radial' && axisType === 'radiusAxis') { - return { - scale: band(), - realScaleType: 'band' - }; - } - - if (layout === 'radial' && axisType === 'angleAxis') { - return { - scale: linear(), - realScaleType: 'linear' - }; - } - - if (type === 'category' && chartType && (chartType.indexOf('LineChart') >= 0 || chartType.indexOf('AreaChart') >= 0 || chartType.indexOf('ComposedChart') >= 0 && !hasBar)) { - return { - scale: point(), - realScaleType: 'point' - }; - } - - if (type === 'category') { - return { - scale: band(), - realScaleType: 'band' - }; - } - - return { - scale: linear(), - realScaleType: 'linear' - }; - } - - if (isString_1(scale)) { - var name = "scale".concat(upperFirst_1(scale)); - return { - scale: (d3Scales[name] || point)(), - realScaleType: d3Scales[name] ? name : 'point' - }; - } - - return isFunction_1(scale) ? { - scale: scale - } : { - scale: point(), - realScaleType: 'point' - }; -}; -var EPS = 1e-4; -var checkDomainOfScale = function checkDomainOfScale(scale) { - var domain = scale.domain(); - - if (!domain || domain.length <= 2) { - return; - } - - var len = domain.length; - var range = scale.range(); - var min = Math.min(range[0], range[1]) - EPS; - var max = Math.max(range[0], range[1]) + EPS; - var first = scale(domain[0]); - var last = scale(domain[len - 1]); - - if (first < min || first > max || last < min || last > max) { - scale.domain([domain[0], domain[len - 1]]); - } -}; -var findPositionOfBar = function findPositionOfBar(barPosition, child) { - if (!barPosition) { - return null; - } - - for (var i = 0, len = barPosition.length; i < len; i++) { - if (barPosition[i].item === child) { - return barPosition[i].position; - } - } - - return null; -}; -var truncateByDomain = function truncateByDomain(value, domain) { - if (!domain || domain.length !== 2 || !isNumber(domain[0]) || !isNumber(domain[1])) { - return value; - } - - var min = Math.min(domain[0], domain[1]); - var max = Math.max(domain[0], domain[1]); - var result = [value[0], value[1]]; - - if (!isNumber(value[0]) || value[0] < min) { - result[0] = min; - } - - if (!isNumber(value[1]) || value[1] > max) { - result[1] = max; - } - - if (result[0] > max) { - result[0] = max; - } - - if (result[1] < min) { - result[1] = min; - } - - return result; -}; -/* eslint no-param-reassign: 0 */ - -var offsetSign = function offsetSign(series) { - var n = series.length; - - if (n <= 0) { - return; - } - - for (var j = 0, m = series[0].length; j < m; ++j) { - var positive = 0; - var negative = 0; - - for (var i = 0; i < n; ++i) { - var value = _isNaN(series[i][j][1]) ? series[i][j][0] : series[i][j][1]; - /* eslint-disable prefer-destructuring */ - - if (value >= 0) { - series[i][j][0] = positive; - series[i][j][1] = positive + value; - positive = series[i][j][1]; - } else { - series[i][j][0] = negative; - series[i][j][1] = negative + value; - negative = series[i][j][1]; - } - /* eslint-enable prefer-destructuring */ - - } - } -}; -/* eslint no-param-reassign: 0 */ - -var offsetPositive = function offsetPositive(series) { - var n = series.length; - - if (n <= 0) { - return; - } - - for (var j = 0, m = series[0].length; j < m; ++j) { - var positive = 0; - - for (var i = 0; i < n; ++i) { - var value = _isNaN(series[i][j][1]) ? series[i][j][0] : series[i][j][1]; - /* eslint-disable prefer-destructuring */ - - if (value >= 0) { - series[i][j][0] = positive; - series[i][j][1] = positive + value; - positive = series[i][j][1]; - } else { - series[i][j][0] = 0; - series[i][j][1] = 0; - } - /* eslint-enable prefer-destructuring */ - - } - } -}; -var STACK_OFFSET_MAP = { - sign: offsetSign, - expand: stackOffsetExpand, - none: stackOffsetNone, - silhouette: stackOffsetSilhouette, - wiggle: stackOffsetWiggle, - positive: offsetPositive -}; -var getStackedData = function getStackedData(data, stackItems, offsetType) { - var dataKeys = stackItems.map(function (item) { - return item.props.dataKey; - }); - var stack = shapeStack().keys(dataKeys).value(function (d, key) { - return +getValueByDataKey(d, key, 0); - }).order(stackOrderNone).offset(STACK_OFFSET_MAP[offsetType]); - return stack(data); -}; -var getStackGroupsByAxisId = function getStackGroupsByAxisId(data, _items, numericAxisId, cateAxisId, offsetType, reverseStackOrder) { - if (!data) { - return null; - } // reversing items to affect render order (for layering) - - - var items = reverseStackOrder ? _items.reverse() : _items; - var stackGroups = items.reduce(function (result, item) { - var _item$props3 = item.props, - stackId = _item$props3.stackId, - hide = _item$props3.hide; - - if (hide) { - return result; - } - - var axisId = item.props[numericAxisId]; - var parentGroup = result[axisId] || { - hasStack: false, - stackGroups: {} - }; - - if (isNumOrStr(stackId)) { - var childGroup = parentGroup.stackGroups[stackId] || { - numericAxisId: numericAxisId, - cateAxisId: cateAxisId, - items: [] - }; - childGroup.items.push(item); - parentGroup.hasStack = true; - parentGroup.stackGroups[stackId] = childGroup; - } else { - parentGroup.stackGroups[uniqueId('_stackId_')] = { - numericAxisId: numericAxisId, - cateAxisId: cateAxisId, - items: [item] - }; - } - - return _objectSpread$q(_objectSpread$q({}, result), {}, _defineProperty$q({}, axisId, parentGroup)); - }, {}); - return Object.keys(stackGroups).reduce(function (result, axisId) { - var group = stackGroups[axisId]; - - if (group.hasStack) { - group.stackGroups = Object.keys(group.stackGroups).reduce(function (res, stackId) { - var g = group.stackGroups[stackId]; - return _objectSpread$q(_objectSpread$q({}, res), {}, _defineProperty$q({}, stackId, { - numericAxisId: numericAxisId, - cateAxisId: cateAxisId, - items: g.items, - stackedData: getStackedData(data, g.items, offsetType) - })); - }, {}); - } - - return _objectSpread$q(_objectSpread$q({}, result), {}, _defineProperty$q({}, axisId, group)); - }, {}); -}; -/** - * get domain of ticks - * @param {Array} ticks Ticks of axis - * @param {String} type The type of axis - * @return {Array} domain - */ - -var calculateDomainOfTicks = function calculateDomainOfTicks(ticks, type) { - if (type === 'number') { - return [min_1(ticks), max_1(ticks)]; - } - - return ticks; -}; -/** - * Configure the scale function of axis - * @param {Object} scale The scale function - * @param {Object} opts The configuration of axis - * @return {Object} null - */ - -var getTicksOfScale = function getTicksOfScale(scale, opts) { - var realScaleType = opts.realScaleType, - type = opts.type, - tickCount = opts.tickCount, - originalDomain = opts.originalDomain, - allowDecimals = opts.allowDecimals; - var scaleType = realScaleType || opts.scale; - - if (scaleType !== 'auto' && scaleType !== 'linear') { - return null; - } - - if (tickCount && type === 'number' && originalDomain && (originalDomain[0] === 'auto' || originalDomain[1] === 'auto')) { - // Calculate the ticks by the number of grid when the axis is a number axis - var domain = scale.domain(); - - if (!domain.length) { - return null; - } - - var tickValues = lib.getNiceTickValues(domain, tickCount, allowDecimals); - scale.domain(calculateDomainOfTicks(tickValues, type)); - return { - niceTicks: tickValues - }; - } - - if (tickCount && type === 'number') { - var _domain = scale.domain(); - - var _tickValues = lib.getTickValuesFixedDomain(_domain, tickCount, allowDecimals); - - return { - niceTicks: _tickValues - }; - } - - return null; -}; -var getCateCoordinateOfLine = function getCateCoordinateOfLine(_ref7) { - var axis = _ref7.axis, - ticks = _ref7.ticks, - bandSize = _ref7.bandSize, - entry = _ref7.entry, - index = _ref7.index, - dataKey = _ref7.dataKey; - - if (axis.type === 'category') { - // find coordinate of category axis by the value of category - if (!axis.allowDuplicatedCategory && axis.dataKey && !isNil_1(entry[axis.dataKey])) { - var matchedTick = findEntryInArray(ticks, 'value', entry[axis.dataKey]); - - if (matchedTick) { - return matchedTick.coordinate + bandSize / 2; - } - } - - return ticks[index] ? ticks[index].coordinate + bandSize / 2 : null; - } - - var value = getValueByDataKey(entry, !isNil_1(dataKey) ? dataKey : axis.dataKey); - return !isNil_1(value) ? axis.scale(value) : null; -}; -var getCateCoordinateOfBar = function getCateCoordinateOfBar(_ref8) { - var axis = _ref8.axis, - ticks = _ref8.ticks, - offset = _ref8.offset, - bandSize = _ref8.bandSize, - entry = _ref8.entry, - index = _ref8.index; - - if (axis.type === 'category') { - return ticks[index] ? ticks[index].coordinate + offset : null; - } - - var value = getValueByDataKey(entry, axis.dataKey, axis.domain[index]); - return !isNil_1(value) ? axis.scale(value) - bandSize / 2 + offset : null; -}; -var getBaseValueOfBar = function getBaseValueOfBar(_ref9) { - var numericAxis = _ref9.numericAxis; - var domain = numericAxis.scale.domain(); - - if (numericAxis.type === 'number') { - var min = Math.min(domain[0], domain[1]); - var max = Math.max(domain[0], domain[1]); - - if (min <= 0 && max >= 0) { - return 0; - } - - if (max < 0) { - return max; - } - - return min; - } - - return domain[0]; -}; -var getStackedDataOfItem = function getStackedDataOfItem(item, stackGroups) { - var stackId = item.props.stackId; - - if (isNumOrStr(stackId)) { - var group = stackGroups[stackId]; - - if (group && group.items.length) { - var itemIndex = -1; - - for (var i = 0, len = group.items.length; i < len; i++) { - if (group.items[i] === item) { - itemIndex = i; - break; - } - } - - return itemIndex >= 0 ? group.stackedData[itemIndex] : null; - } - } - - return null; -}; - -var getDomainOfSingle = function getDomainOfSingle(data) { - return data.reduce(function (result, entry) { - return [min_1(entry.concat([result[0]]).filter(isNumber)), max_1(entry.concat([result[1]]).filter(isNumber))]; - }, [Infinity, -Infinity]); -}; - -var getDomainOfStackGroups = function getDomainOfStackGroups(stackGroups, startIndex, endIndex) { - return Object.keys(stackGroups).reduce(function (result, stackId) { - var group = stackGroups[stackId]; - var stackedData = group.stackedData; - var domain = stackedData.reduce(function (res, entry) { - var s = getDomainOfSingle(entry.slice(startIndex, endIndex + 1)); - return [Math.min(res[0], s[0]), Math.max(res[1], s[1])]; - }, [Infinity, -Infinity]); - return [Math.min(domain[0], result[0]), Math.max(domain[1], result[1])]; - }, [Infinity, -Infinity]).map(function (result) { - return result === Infinity || result === -Infinity ? 0 : result; - }); -}; -var MIN_VALUE_REG = /^dataMin[\s]*-[\s]*([0-9]+([.]{1}[0-9]+){0,1})$/; -var MAX_VALUE_REG = /^dataMax[\s]*\+[\s]*([0-9]+([.]{1}[0-9]+){0,1})$/; -var parseSpecifiedDomain = function parseSpecifiedDomain(specifiedDomain, dataDomain, allowDataOverflow) { - if (isFunction_1(specifiedDomain)) { - return specifiedDomain(dataDomain, allowDataOverflow); - } - - if (!isArray_1(specifiedDomain)) { - return dataDomain; - } - - var domain = []; - /* eslint-disable prefer-destructuring */ - - if (isNumber(specifiedDomain[0])) { - domain[0] = allowDataOverflow ? specifiedDomain[0] : Math.min(specifiedDomain[0], dataDomain[0]); - } else if (MIN_VALUE_REG.test(specifiedDomain[0])) { - var value = +MIN_VALUE_REG.exec(specifiedDomain[0])[1]; - domain[0] = dataDomain[0] - value; - } else if (isFunction_1(specifiedDomain[0])) { - domain[0] = specifiedDomain[0](dataDomain[0]); - } else { - domain[0] = dataDomain[0]; - } - - if (isNumber(specifiedDomain[1])) { - domain[1] = allowDataOverflow ? specifiedDomain[1] : Math.max(specifiedDomain[1], dataDomain[1]); - } else if (MAX_VALUE_REG.test(specifiedDomain[1])) { - var _value = +MAX_VALUE_REG.exec(specifiedDomain[1])[1]; - - domain[1] = dataDomain[1] + _value; - } else if (isFunction_1(specifiedDomain[1])) { - domain[1] = specifiedDomain[1](dataDomain[1]); - } else { - domain[1] = dataDomain[1]; - } - /* eslint-enable prefer-destructuring */ - - - return domain; -}; -/** - * Calculate the size between two category - * @param {Object} axis The options of axis - * @param {Array} ticks The ticks of axis - * @param {Boolean} isBar if items in axis are bars - * @return {Number} Size - */ - -var getBandSizeOfAxis = function getBandSizeOfAxis(axis, ticks, isBar) { - if (axis && axis.scale && axis.scale.bandwidth) { - var bandWidth = axis.scale.bandwidth(); - - if (!isBar || bandWidth > 0) { - return bandWidth; - } - } - - if (axis && ticks && ticks.length >= 2) { - var orderedTicks = sortBy_1(ticks, function (o) { - return o.coordinate; - }); - - var bandSize = Infinity; - - for (var i = 1, len = orderedTicks.length; i < len; i++) { - var cur = orderedTicks[i]; - var prev = orderedTicks[i - 1]; - bandSize = Math.min((cur.coordinate || 0) - (prev.coordinate || 0), bandSize); - } - - return bandSize === Infinity ? 0 : bandSize; - } - - return isBar ? undefined : 0; -}; -/** - * parse the domain of a category axis when a domain is specified - * @param {Array} specifiedDomain The domain specified by users - * @param {Array} calculatedDomain The domain calculated by dateKey - * @param {ReactElement} axisChild The axis element - * @returns {Array} domains - */ - -var parseDomainOfCategoryAxis = function parseDomainOfCategoryAxis(specifiedDomain, calculatedDomain, axisChild) { - if (!specifiedDomain || !specifiedDomain.length) { - return calculatedDomain; - } - - if (isEqual_1(specifiedDomain, get_1(axisChild, 'type.defaultProps.domain'))) { - return calculatedDomain; - } - - return specifiedDomain; -}; -var getTooltipItem = function getTooltipItem(graphicalItem, payload) { - var _graphicalItem$props = graphicalItem.props, - dataKey = _graphicalItem$props.dataKey, - name = _graphicalItem$props.name, - unit = _graphicalItem$props.unit, - formatter = _graphicalItem$props.formatter, - tooltipType = _graphicalItem$props.tooltipType, - chartType = _graphicalItem$props.chartType; - return _objectSpread$q(_objectSpread$q({}, filterProps(graphicalItem)), {}, { - dataKey: dataKey, - unit: unit, - formatter: formatter, - name: name || dataKey, - color: getMainColorOfGraphicItem(graphicalItem), - value: getValueByDataKey(payload, dataKey), - type: tooltipType, - payload: payload, - chartType: chartType - }); -}; - -function ownKeys$p(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$p(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$p(Object(source), true).forEach(function (key) { _defineProperty$p(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$p(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$p(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _slicedToArray$4(arr, i) { return _arrayWithHoles$4(arr) || _iterableToArrayLimit$4(arr, i) || _unsupportedIterableToArray$8(arr, i) || _nonIterableRest$4(); } - -function _nonIterableRest$4() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$8(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$8(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$8(o, minLen); } - -function _arrayLikeToArray$8(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit$4(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$4(arr) { if (Array.isArray(arr)) return arr; } -var RADIAN$1 = Math.PI / 180; -var radianToDegree = function radianToDegree(angleInRadian) { - return angleInRadian * 180 / Math.PI; -}; -var polarToCartesian = function polarToCartesian(cx, cy, radius, angle) { - return { - x: cx + Math.cos(-RADIAN$1 * angle) * radius, - y: cy + Math.sin(-RADIAN$1 * angle) * radius - }; -}; -var getMaxRadius = function getMaxRadius(width, height) { - var offset = arguments.length > 2 && arguments[2] !== undefined ? arguments[2] : { - top: 0, - right: 0, - bottom: 0, - left: 0 - }; - return Math.min(Math.abs(width - (offset.left || 0) - (offset.right || 0)), Math.abs(height - (offset.top || 0) - (offset.bottom || 0))) / 2; -}; -/** - * Calculate the scale function, position, width, height of axes - * @param {Object} props Latest props - * @param {Object} axisMap The configuration of axes - * @param {Object} offset The offset of main part in the svg element - * @param {Object} axisType The type of axes, radius-axis or angle-axis - * @param {String} chartName The name of chart - * @return {Object} Configuration - */ - -var formatAxisMap$1 = function formatAxisMap(props, axisMap, offset, axisType, chartName) { - var width = props.width, - height = props.height; - var startAngle = props.startAngle, - endAngle = props.endAngle; - var cx = getPercentValue(props.cx, width, width / 2); - var cy = getPercentValue(props.cy, height, height / 2); - var maxRadius = getMaxRadius(width, height, offset); - var innerRadius = getPercentValue(props.innerRadius, maxRadius, 0); - var outerRadius = getPercentValue(props.outerRadius, maxRadius, maxRadius * 0.8); - var ids = Object.keys(axisMap); - return ids.reduce(function (result, id) { - var axis = axisMap[id]; - var domain = axis.domain, - reversed = axis.reversed; - var range; - - if (isNil_1(axis.range)) { - if (axisType === 'angleAxis') { - range = [startAngle, endAngle]; - } else if (axisType === 'radiusAxis') { - range = [innerRadius, outerRadius]; - } - - if (reversed) { - range = [range[1], range[0]]; - } - } else { - range = axis.range; - var _range = range; - - var _range2 = _slicedToArray$4(_range, 2); - - startAngle = _range2[0]; - endAngle = _range2[1]; - } - - var _parseScale = parseScale(axis, chartName), - realScaleType = _parseScale.realScaleType, - scale = _parseScale.scale; - - scale.domain(domain).range(range); - checkDomainOfScale(scale); - var ticks = getTicksOfScale(scale, _objectSpread$p(_objectSpread$p({}, axis), {}, { - realScaleType: realScaleType - })); - - var finalAxis = _objectSpread$p(_objectSpread$p(_objectSpread$p({}, axis), ticks), {}, { - range: range, - radius: outerRadius, - realScaleType: realScaleType, - scale: scale, - cx: cx, - cy: cy, - innerRadius: innerRadius, - outerRadius: outerRadius, - startAngle: startAngle, - endAngle: endAngle - }); - - return _objectSpread$p(_objectSpread$p({}, result), {}, _defineProperty$p({}, id, finalAxis)); - }, {}); -}; -var distanceBetweenPoints = function distanceBetweenPoints(point, anotherPoint) { - var x1 = point.x, - y1 = point.y; - var x2 = anotherPoint.x, - y2 = anotherPoint.y; - return Math.sqrt(Math.pow(x1 - x2, 2) + Math.pow(y1 - y2, 2)); -}; -var getAngleOfPoint = function getAngleOfPoint(_ref, _ref2) { - var x = _ref.x, - y = _ref.y; - var cx = _ref2.cx, - cy = _ref2.cy; - var radius = distanceBetweenPoints({ - x: x, - y: y - }, { - x: cx, - y: cy - }); - - if (radius <= 0) { - return { - radius: radius - }; - } - - var cos = (x - cx) / radius; - var angleInRadian = Math.acos(cos); - - if (y > cy) { - angleInRadian = 2 * Math.PI - angleInRadian; - } - - return { - radius: radius, - angle: radianToDegree(angleInRadian), - angleInRadian: angleInRadian - }; -}; -var formatAngleOfSector = function formatAngleOfSector(_ref3) { - var startAngle = _ref3.startAngle, - endAngle = _ref3.endAngle; - var startCnt = Math.floor(startAngle / 360); - var endCnt = Math.floor(endAngle / 360); - var min = Math.min(startCnt, endCnt); - return { - startAngle: startAngle - min * 360, - endAngle: endAngle - min * 360 - }; -}; - -var reverseFormatAngleOfSetor = function reverseFormatAngleOfSetor(angle, _ref4) { - var startAngle = _ref4.startAngle, - endAngle = _ref4.endAngle; - var startCnt = Math.floor(startAngle / 360); - var endCnt = Math.floor(endAngle / 360); - var min = Math.min(startCnt, endCnt); - return angle + min * 360; -}; - -var inRangeOfSector = function inRangeOfSector(_ref5, sector) { - var x = _ref5.x, - y = _ref5.y; - - var _getAngleOfPoint = getAngleOfPoint({ - x: x, - y: y - }, sector), - radius = _getAngleOfPoint.radius, - angle = _getAngleOfPoint.angle; - - var innerRadius = sector.innerRadius, - outerRadius = sector.outerRadius; - - if (radius < innerRadius || radius > outerRadius) { - return false; - } - - if (radius === 0) { - return true; - } - - var _formatAngleOfSector = formatAngleOfSector(sector), - startAngle = _formatAngleOfSector.startAngle, - endAngle = _formatAngleOfSector.endAngle; - - var formatAngle = angle; - var inRange; - - if (startAngle <= endAngle) { - while (formatAngle > endAngle) { - formatAngle -= 360; - } - - while (formatAngle < startAngle) { - formatAngle += 360; - } - - inRange = formatAngle >= startAngle && formatAngle <= endAngle; - } else { - while (formatAngle > startAngle) { - formatAngle -= 360; - } - - while (formatAngle < endAngle) { - formatAngle += 360; - } - - inRange = formatAngle >= endAngle && formatAngle <= startAngle; - } - - if (inRange) { - return _objectSpread$p(_objectSpread$p({}, sector), {}, { - radius: radius, - angle: reverseFormatAngleOfSetor(formatAngle, sector) - }); - } - - return null; -}; - -function _toConsumableArray$4(arr) { return _arrayWithoutHoles$4(arr) || _iterableToArray$4(arr) || _unsupportedIterableToArray$7(arr) || _nonIterableSpread$4(); } - -function _nonIterableSpread$4() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$7(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$7(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$7(o, minLen); } - -function _iterableToArray$4(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$4(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$7(arr); } - -function _arrayLikeToArray$7(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function ownKeys$o(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$o(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$o(Object(source), true).forEach(function (key) { _defineProperty$o(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$o(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$o(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _extends$s() { _extends$s = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$s.apply(this, arguments); } - -var getLabel = function getLabel(props) { - var value = props.value, - formatter = props.formatter; - var label = isNil_1(props.children) ? value : props.children; - - if (isFunction_1(formatter)) { - return formatter(label); - } - - return label; -}; - -var getDeltaAngle$1 = function getDeltaAngle(startAngle, endAngle) { - var sign = mathSign(endAngle - startAngle); - var deltaAngle = Math.min(Math.abs(endAngle - startAngle), 360); - return sign * deltaAngle; -}; - -var renderRadialLabel = function renderRadialLabel(labelProps, label, attrs) { - var position = labelProps.position, - viewBox = labelProps.viewBox, - offset = labelProps.offset, - className = labelProps.className; - var _ref = viewBox, - cx = _ref.cx, - cy = _ref.cy, - innerRadius = _ref.innerRadius, - outerRadius = _ref.outerRadius, - startAngle = _ref.startAngle, - endAngle = _ref.endAngle, - clockWise = _ref.clockWise; - var radius = (innerRadius + outerRadius) / 2; - var deltaAngle = getDeltaAngle$1(startAngle, endAngle); - var sign = deltaAngle >= 0 ? 1 : -1; - var labelAngle, direction; - - if (position === 'insideStart') { - labelAngle = startAngle + sign * offset; - direction = clockWise; - } else if (position === 'insideEnd') { - labelAngle = endAngle - sign * offset; - direction = !clockWise; - } else if (position === 'end') { - labelAngle = endAngle + sign * offset; - direction = clockWise; - } - - direction = deltaAngle <= 0 ? direction : !direction; - var startPoint = polarToCartesian(cx, cy, radius, labelAngle); - var endPoint = polarToCartesian(cx, cy, radius, labelAngle + (direction ? 1 : -1) * 359); - var path = "M".concat(startPoint.x, ",").concat(startPoint.y, "\n A").concat(radius, ",").concat(radius, ",0,1,").concat(direction ? 0 : 1, ",\n ").concat(endPoint.x, ",").concat(endPoint.y); - var id = isNil_1(labelProps.id) ? uniqueId('recharts-radial-line-') : labelProps.id; - return /*#__PURE__*/React__default.createElement("text", _extends$s({}, attrs, { - dominantBaseline: "central", - className: classNames('recharts-radial-bar-label', className) - }), /*#__PURE__*/React__default.createElement("defs", null, /*#__PURE__*/React__default.createElement("path", { - id: id, - d: path - })), /*#__PURE__*/React__default.createElement("textPath", { - xlinkHref: "#".concat(id) - }, label)); -}; - -var getAttrsOfPolarLabel = function getAttrsOfPolarLabel(props) { - var viewBox = props.viewBox, - offset = props.offset, - position = props.position; - var _ref2 = viewBox, - cx = _ref2.cx, - cy = _ref2.cy, - innerRadius = _ref2.innerRadius, - outerRadius = _ref2.outerRadius, - startAngle = _ref2.startAngle, - endAngle = _ref2.endAngle; - var midAngle = (startAngle + endAngle) / 2; - - if (position === 'outside') { - var _polarToCartesian = polarToCartesian(cx, cy, outerRadius + offset, midAngle), - _x = _polarToCartesian.x, - _y = _polarToCartesian.y; - - return { - x: _x, - y: _y, - textAnchor: _x >= cx ? 'start' : 'end', - verticalAnchor: 'middle' - }; - } - - if (position === 'center') { - return { - x: cx, - y: cy, - textAnchor: 'middle', - verticalAnchor: 'middle' - }; - } - - if (position === 'centerTop') { - return { - x: cx, - y: cy, - textAnchor: 'middle', - verticalAnchor: 'start' - }; - } - - if (position === 'centerBottom') { - return { - x: cx, - y: cy, - textAnchor: 'middle', - verticalAnchor: 'end' - }; - } - - var r = (innerRadius + outerRadius) / 2; - - var _polarToCartesian2 = polarToCartesian(cx, cy, r, midAngle), - x = _polarToCartesian2.x, - y = _polarToCartesian2.y; - - return { - x: x, - y: y, - textAnchor: 'middle', - verticalAnchor: 'middle' - }; -}; - -var getAttrsOfCartesianLabel = function getAttrsOfCartesianLabel(props) { - var viewBox = props.viewBox, - parentViewBox = props.parentViewBox, - offset = props.offset, - position = props.position; - var _ref3 = viewBox, - x = _ref3.x, - y = _ref3.y, - width = _ref3.width, - height = _ref3.height; // Define vertical offsets and position inverts based on the value being positive or negative - - var verticalSign = height >= 0 ? 1 : -1; - var verticalOffset = verticalSign * offset; - var verticalEnd = verticalSign > 0 ? 'end' : 'start'; - var verticalStart = verticalSign > 0 ? 'start' : 'end'; // Define horizontal offsets and position inverts based on the value being positive or negative - - var horizontalSign = width >= 0 ? 1 : -1; - var horizontalOffset = horizontalSign * offset; - var horizontalEnd = horizontalSign > 0 ? 'end' : 'start'; - var horizontalStart = horizontalSign > 0 ? 'start' : 'end'; - - if (position === 'top') { - var attrs = { - x: x + width / 2, - y: y - verticalSign * offset, - textAnchor: 'middle', - verticalAnchor: verticalEnd - }; - return _objectSpread$o(_objectSpread$o({}, attrs), parentViewBox ? { - height: Math.max(y - parentViewBox.y, 0), - width: width - } : {}); - } - - if (position === 'bottom') { - var _attrs = { - x: x + width / 2, - y: y + height + verticalOffset, - textAnchor: 'middle', - verticalAnchor: verticalStart - }; - return _objectSpread$o(_objectSpread$o({}, _attrs), parentViewBox ? { - height: Math.max(parentViewBox.y + parentViewBox.height - (y + height), 0), - width: width - } : {}); - } - - if (position === 'left') { - var _attrs2 = { - x: x - horizontalOffset, - y: y + height / 2, - textAnchor: horizontalEnd, - verticalAnchor: 'middle' - }; - return _objectSpread$o(_objectSpread$o({}, _attrs2), parentViewBox ? { - width: Math.max(_attrs2.x - parentViewBox.x, 0), - height: height - } : {}); - } - - if (position === 'right') { - var _attrs3 = { - x: x + width + horizontalOffset, - y: y + height / 2, - textAnchor: horizontalStart, - verticalAnchor: 'middle' - }; - return _objectSpread$o(_objectSpread$o({}, _attrs3), parentViewBox ? { - width: Math.max(parentViewBox.x + parentViewBox.width - _attrs3.x, 0), - height: height - } : {}); - } - - var sizeAttrs = parentViewBox ? { - width: width, - height: height - } : {}; - - if (position === 'insideLeft') { - return _objectSpread$o({ - x: x + horizontalOffset, - y: y + height / 2, - textAnchor: horizontalStart, - verticalAnchor: 'middle' - }, sizeAttrs); - } - - if (position === 'insideRight') { - return _objectSpread$o({ - x: x + width - horizontalOffset, - y: y + height / 2, - textAnchor: horizontalEnd, - verticalAnchor: 'middle' - }, sizeAttrs); - } - - if (position === 'insideTop') { - return _objectSpread$o({ - x: x + width / 2, - y: y + verticalOffset, - textAnchor: 'middle', - verticalAnchor: verticalStart - }, sizeAttrs); - } - - if (position === 'insideBottom') { - return _objectSpread$o({ - x: x + width / 2, - y: y + height - verticalOffset, - textAnchor: 'middle', - verticalAnchor: verticalEnd - }, sizeAttrs); - } - - if (position === 'insideTopLeft') { - return _objectSpread$o({ - x: x + horizontalOffset, - y: y + verticalOffset, - textAnchor: horizontalStart, - verticalAnchor: verticalStart - }, sizeAttrs); - } - - if (position === 'insideTopRight') { - return _objectSpread$o({ - x: x + width - horizontalOffset, - y: y + verticalOffset, - textAnchor: horizontalEnd, - verticalAnchor: verticalStart - }, sizeAttrs); - } - - if (position === 'insideBottomLeft') { - return _objectSpread$o({ - x: x + horizontalOffset, - y: y + height - verticalOffset, - textAnchor: horizontalStart, - verticalAnchor: verticalEnd - }, sizeAttrs); - } - - if (position === 'insideBottomRight') { - return _objectSpread$o({ - x: x + width - horizontalOffset, - y: y + height - verticalOffset, - textAnchor: horizontalEnd, - verticalAnchor: verticalEnd - }, sizeAttrs); - } - - if (isObject_1$1(position) && (isNumber(position.x) || isPercent(position.x)) && (isNumber(position.y) || isPercent(position.y))) { - return _objectSpread$o({ - x: x + getPercentValue(position.x, width), - y: y + getPercentValue(position.y, height), - textAnchor: 'end', - verticalAnchor: 'end' - }, sizeAttrs); - } - - return _objectSpread$o({ - x: x + width / 2, - y: y + height / 2, - textAnchor: 'middle', - verticalAnchor: 'middle' - }, sizeAttrs); -}; - -var isPolar = function isPolar(viewBox) { - return isNumber(viewBox.cx); -}; - -function Label(props) { - var viewBox = props.viewBox, - position = props.position, - value = props.value, - children = props.children, - content = props.content, - _props$className = props.className, - className = _props$className === void 0 ? '' : _props$className, - textBreakAll = props.textBreakAll; - - if (!viewBox || isNil_1(value) && isNil_1(children) && ! /*#__PURE__*/isValidElement(content) && !isFunction_1(content)) { - return null; - } - - if ( /*#__PURE__*/isValidElement(content)) { - return /*#__PURE__*/cloneElement(content, props); - } - - var label; - - if (isFunction_1(content)) { - label = /*#__PURE__*/createElement(content, props); - - if ( /*#__PURE__*/isValidElement(label)) { - return label; - } - } else { - label = getLabel(props); - } - - var isPolarLabel = isPolar(viewBox); - var attrs = filterProps(props, true); - - if (isPolarLabel && (position === 'insideStart' || position === 'insideEnd' || position === 'end')) { - return renderRadialLabel(props, label, attrs); - } - - var positionAttrs = isPolarLabel ? getAttrsOfPolarLabel(props) : getAttrsOfCartesianLabel(props); - return /*#__PURE__*/React__default.createElement(Text, _extends$s({ - className: classNames('recharts-label', className) - }, attrs, positionAttrs, { - breakAll: textBreakAll - }), label); -} -Label.displayName = 'Label'; -Label.defaultProps = { - offset: 5 -}; - -var parseViewBox = function parseViewBox(props) { - var cx = props.cx, - cy = props.cy, - angle = props.angle, - startAngle = props.startAngle, - endAngle = props.endAngle, - r = props.r, - radius = props.radius, - innerRadius = props.innerRadius, - outerRadius = props.outerRadius, - x = props.x, - y = props.y, - top = props.top, - left = props.left, - width = props.width, - height = props.height, - clockWise = props.clockWise, - labelViewBox = props.labelViewBox; - - if (labelViewBox) { - return labelViewBox; - } - - if (isNumber(width) && isNumber(height)) { - if (isNumber(x) && isNumber(y)) { - return { - x: x, - y: y, - width: width, - height: height - }; - } - - if (isNumber(top) && isNumber(left)) { - return { - x: top, - y: left, - width: width, - height: height - }; - } - } - - if (isNumber(x) && isNumber(y)) { - return { - x: x, - y: y, - width: 0, - height: 0 - }; - } - - if (isNumber(cx) && isNumber(cy)) { - return { - cx: cx, - cy: cy, - startAngle: startAngle || angle || 0, - endAngle: endAngle || angle || 0, - innerRadius: innerRadius || 0, - outerRadius: outerRadius || radius || r || 0, - clockWise: clockWise - }; - } - - if (props.viewBox) { - return props.viewBox; - } - - return {}; -}; - -var parseLabel = function parseLabel(label, viewBox) { - if (!label) { - return null; - } - - if (label === true) { - return /*#__PURE__*/React__default.createElement(Label, { - key: "label-implicit", - viewBox: viewBox - }); - } - - if (isNumOrStr(label)) { - return /*#__PURE__*/React__default.createElement(Label, { - key: "label-implicit", - viewBox: viewBox, - value: label - }); - } - - if ( /*#__PURE__*/isValidElement(label)) { - if (label.type === Label) { - return /*#__PURE__*/cloneElement(label, { - key: 'label-implicit', - viewBox: viewBox - }); - } - - return /*#__PURE__*/React__default.createElement(Label, { - key: "label-implicit", - content: label, - viewBox: viewBox - }); - } - - if (isFunction_1(label)) { - return /*#__PURE__*/React__default.createElement(Label, { - key: "label-implicit", - content: label, - viewBox: viewBox - }); - } - - if (isObject_1$1(label)) { - return /*#__PURE__*/React__default.createElement(Label, _extends$s({ - viewBox: viewBox - }, label, { - key: "label-implicit" - })); - } - - return null; -}; - -var renderCallByParent$1 = function renderCallByParent(parentProps, viewBox) { - var checkPropsLabel = arguments.length > 2 && arguments[2] !== undefined ? arguments[2] : true; - - if (!parentProps || !parentProps.children && checkPropsLabel && !parentProps.label) { - return null; - } - - var children = parentProps.children; - var parentViewBox = parseViewBox(parentProps); - var explicitChildren = findAllByType(children, Label.displayName).map(function (child, index) { - return /*#__PURE__*/cloneElement(child, { - viewBox: viewBox || parentViewBox, - key: "label-".concat(index) - }); - }); - - if (!checkPropsLabel) { - return explicitChildren; - } - - var implicitLabel = parseLabel(parentProps.label, viewBox || parentViewBox); - return [implicitLabel].concat(_toConsumableArray$4(explicitChildren)); -}; - -Label.parseViewBox = parseViewBox; -Label.renderCallByParent = renderCallByParent$1; - -/** - * Gets the last element of `array`. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Array - * @param {Array} array The array to query. - * @returns {*} Returns the last element of `array`. - * @example - * - * _.last([1, 2, 3]); - * // => 3 - */ - -function last$1(array) { - var length = array == null ? 0 : array.length; - return length ? array[length - 1] : undefined; -} - -var last_1 = last$1; - -function _toConsumableArray$3(arr) { return _arrayWithoutHoles$3(arr) || _iterableToArray$3(arr) || _unsupportedIterableToArray$6(arr) || _nonIterableSpread$3(); } - -function _nonIterableSpread$3() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$6(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$6(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$6(o, minLen); } - -function _iterableToArray$3(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$3(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$6(arr); } - -function _arrayLikeToArray$6(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _extends$r() { _extends$r = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$r.apply(this, arguments); } - -function ownKeys$n(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$n(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$n(Object(source), true).forEach(function (key) { _defineProperty$n(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$n(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$n(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _objectWithoutProperties$d(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$d(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$d(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } -var defaultProps = { - valueAccessor: function valueAccessor(entry) { - return isArray_1(entry.value) ? last_1(entry.value) : entry.value; - } -}; -function LabelList(props) { - var data = props.data, - valueAccessor = props.valueAccessor, - dataKey = props.dataKey, - clockWise = props.clockWise, - id = props.id, - textBreakAll = props.textBreakAll, - others = _objectWithoutProperties$d(props, ["data", "valueAccessor", "dataKey", "clockWise", "id", "textBreakAll"]); - - if (!data || !data.length) { - return null; - } - - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-label-list" - }, data.map(function (entry, index) { - var value = isNil_1(dataKey) ? valueAccessor(entry, index) : getValueByDataKey(entry && entry.payload, dataKey); - var idProps = isNil_1(id) ? {} : { - id: "".concat(id, "-").concat(index) - }; - return /*#__PURE__*/React__default.createElement(Label, _extends$r({}, filterProps(entry, true), others, idProps, { - parentViewBox: entry.parentViewBox, - index: index, - value: value, - textBreakAll: textBreakAll, - viewBox: Label.parseViewBox(isNil_1(clockWise) ? entry : _objectSpread$n(_objectSpread$n({}, entry), {}, { - clockWise: clockWise - })), - key: "label-".concat(index) // eslint-disable-line react/no-array-index-key - - })); - })); -} -LabelList.displayName = 'LabelList'; - -function parseLabelList(label, data) { - if (!label) { - return null; - } - - if (label === true) { - return /*#__PURE__*/React__default.createElement(LabelList, { - key: "labelList-implicit", - data: data - }); - } - - if ( /*#__PURE__*/React__default.isValidElement(label) || isFunction_1(label)) { - return /*#__PURE__*/React__default.createElement(LabelList, { - key: "labelList-implicit", - data: data, - content: label - }); - } - - if (isObject_1$1(label)) { - return /*#__PURE__*/React__default.createElement(LabelList, _extends$r({ - data: data - }, label, { - key: "labelList-implicit" - })); - } - - return null; -} - -function renderCallByParent(parentProps, data) { - var ckeckPropsLabel = arguments.length > 2 && arguments[2] !== undefined ? arguments[2] : true; - - if (!parentProps || !parentProps.children && ckeckPropsLabel && !parentProps.label) { - return null; - } - - var children = parentProps.children; - var explicitChilren = findAllByType(children, LabelList.displayName).map(function (child, index) { - return /*#__PURE__*/cloneElement(child, { - data: data, - key: "labelList-".concat(index) - }); - }); - - if (!ckeckPropsLabel) { - return explicitChilren; - } - - var implicitLabelList = parseLabelList(parentProps.label, data); - return [implicitLabelList].concat(_toConsumableArray$3(explicitChilren)); -} - -LabelList.renderCallByParent = renderCallByParent; -LabelList.defaultProps = defaultProps; - -function _typeof$o(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$o = function _typeof(obj) { return typeof obj; }; } else { _typeof$o = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$o(obj); } - -function _objectWithoutProperties$c(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$c(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$c(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -/** - * custom svg elements by rechart instance props and state. - * @returns {Object} svg elements - */ -function Customized(_ref) { - var component = _ref.component, - props = _objectWithoutProperties$c(_ref, ["component"]); - - var child; - - if ( /*#__PURE__*/isValidElement(component)) { - child = /*#__PURE__*/cloneElement(component, props); - } else if (isFunction_1(component)) { - child = /*#__PURE__*/createElement(component, props); - } else { - warn(false, "Customized's props `component` must be React.element or Function, but got %s.", _typeof$o(component)); - } - - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-customized-wrapper" - }, child); -} -Customized.displayName = 'Customized'; - -function _typeof$n(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$n = function _typeof(obj) { return typeof obj; }; } else { _typeof$n = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$n(obj); } - -function _extends$q() { _extends$q = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$q.apply(this, arguments); } - -function _classCallCheck$o(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$o(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$o(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$o(Constructor.prototype, protoProps); if (staticProps) _defineProperties$o(Constructor, staticProps); return Constructor; } - -function _inherits$n(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$n(subClass, superClass); } - -function _setPrototypeOf$n(o, p) { _setPrototypeOf$n = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$n(o, p); } - -function _createSuper$n(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$n(); return function _createSuperInternal() { var Super = _getPrototypeOf$n(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$n(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$n(this, result); }; } - -function _possibleConstructorReturn$n(self, call) { if (call && (_typeof$n(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$n(self); } - -function _assertThisInitialized$n(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$n() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$n(o) { _getPrototypeOf$n = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$n(o); } - -var getDeltaAngle = function getDeltaAngle(startAngle, endAngle) { - var sign = mathSign(endAngle - startAngle); - var deltaAngle = Math.min(Math.abs(endAngle - startAngle), 359.999); - return sign * deltaAngle; -}; - -var getTangentCircle = function getTangentCircle(_ref) { - var cx = _ref.cx, - cy = _ref.cy, - radius = _ref.radius, - angle = _ref.angle, - sign = _ref.sign, - isExternal = _ref.isExternal, - cornerRadius = _ref.cornerRadius, - cornerIsExternal = _ref.cornerIsExternal; - var centerRadius = cornerRadius * (isExternal ? 1 : -1) + radius; - var theta = Math.asin(cornerRadius / centerRadius) / RADIAN$1; - var centerAngle = cornerIsExternal ? angle : angle + sign * theta; - var center = polarToCartesian(cx, cy, centerRadius, centerAngle); // The coordinate of point which is tangent to the circle - - var circleTangency = polarToCartesian(cx, cy, radius, centerAngle); // The coordinate of point which is tangent to the radius line - - var lineTangencyAngle = cornerIsExternal ? angle - sign * theta : angle; - var lineTangency = polarToCartesian(cx, cy, centerRadius * Math.cos(theta * RADIAN$1), lineTangencyAngle); - return { - center: center, - circleTangency: circleTangency, - lineTangency: lineTangency, - theta: theta - }; -}; - -var getSectorPath = function getSectorPath(_ref2) { - var cx = _ref2.cx, - cy = _ref2.cy, - innerRadius = _ref2.innerRadius, - outerRadius = _ref2.outerRadius, - startAngle = _ref2.startAngle, - endAngle = _ref2.endAngle; - var angle = getDeltaAngle(startAngle, endAngle); // When the angle of sector equals to 360, star point and end point coincide - - var tempEndAngle = startAngle + angle; - var outerStartPoint = polarToCartesian(cx, cy, outerRadius, startAngle); - var outerEndPoint = polarToCartesian(cx, cy, outerRadius, tempEndAngle); - var path = "M ".concat(outerStartPoint.x, ",").concat(outerStartPoint.y, "\n A ").concat(outerRadius, ",").concat(outerRadius, ",0,\n ").concat(+(Math.abs(angle) > 180), ",").concat(+(startAngle > tempEndAngle), ",\n ").concat(outerEndPoint.x, ",").concat(outerEndPoint.y, "\n "); - - if (innerRadius > 0) { - var innerStartPoint = polarToCartesian(cx, cy, innerRadius, startAngle); - var innerEndPoint = polarToCartesian(cx, cy, innerRadius, tempEndAngle); - path += "L ".concat(innerEndPoint.x, ",").concat(innerEndPoint.y, "\n A ").concat(innerRadius, ",").concat(innerRadius, ",0,\n ").concat(+(Math.abs(angle) > 180), ",").concat(+(startAngle <= tempEndAngle), ",\n ").concat(innerStartPoint.x, ",").concat(innerStartPoint.y, " Z"); - } else { - path += "L ".concat(cx, ",").concat(cy, " Z"); - } - - return path; -}; - -var getSectorWithCorner = function getSectorWithCorner(_ref3) { - var cx = _ref3.cx, - cy = _ref3.cy, - innerRadius = _ref3.innerRadius, - outerRadius = _ref3.outerRadius, - cornerRadius = _ref3.cornerRadius, - forceCornerRadius = _ref3.forceCornerRadius, - cornerIsExternal = _ref3.cornerIsExternal, - startAngle = _ref3.startAngle, - endAngle = _ref3.endAngle; - var sign = mathSign(endAngle - startAngle); - - var _getTangentCircle = getTangentCircle({ - cx: cx, - cy: cy, - radius: outerRadius, - angle: startAngle, - sign: sign, - cornerRadius: cornerRadius, - cornerIsExternal: cornerIsExternal - }), - soct = _getTangentCircle.circleTangency, - solt = _getTangentCircle.lineTangency, - sot = _getTangentCircle.theta; - - var _getTangentCircle2 = getTangentCircle({ - cx: cx, - cy: cy, - radius: outerRadius, - angle: endAngle, - sign: -sign, - cornerRadius: cornerRadius, - cornerIsExternal: cornerIsExternal - }), - eoct = _getTangentCircle2.circleTangency, - eolt = _getTangentCircle2.lineTangency, - eot = _getTangentCircle2.theta; - - var outerArcAngle = cornerIsExternal ? Math.abs(startAngle - endAngle) : Math.abs(startAngle - endAngle) - sot - eot; - - if (outerArcAngle < 0) { - if (forceCornerRadius) { - return "M ".concat(solt.x, ",").concat(solt.y, "\n a").concat(cornerRadius, ",").concat(cornerRadius, ",0,0,1,").concat(cornerRadius * 2, ",0\n a").concat(cornerRadius, ",").concat(cornerRadius, ",0,0,1,").concat(-cornerRadius * 2, ",0\n "); - } - - return getSectorPath({ - cx: cx, - cy: cy, - innerRadius: innerRadius, - outerRadius: outerRadius, - startAngle: startAngle, - endAngle: endAngle - }); - } - - var path = "M ".concat(solt.x, ",").concat(solt.y, "\n A").concat(cornerRadius, ",").concat(cornerRadius, ",0,0,").concat(+(sign < 0), ",").concat(soct.x, ",").concat(soct.y, "\n A").concat(outerRadius, ",").concat(outerRadius, ",0,").concat(+(outerArcAngle > 180), ",").concat(+(sign < 0), ",").concat(eoct.x, ",").concat(eoct.y, "\n A").concat(cornerRadius, ",").concat(cornerRadius, ",0,0,").concat(+(sign < 0), ",").concat(eolt.x, ",").concat(eolt.y, "\n "); - - if (innerRadius > 0) { - var _getTangentCircle3 = getTangentCircle({ - cx: cx, - cy: cy, - radius: innerRadius, - angle: startAngle, - sign: sign, - isExternal: true, - cornerRadius: cornerRadius, - cornerIsExternal: cornerIsExternal - }), - sict = _getTangentCircle3.circleTangency, - silt = _getTangentCircle3.lineTangency, - sit = _getTangentCircle3.theta; - - var _getTangentCircle4 = getTangentCircle({ - cx: cx, - cy: cy, - radius: innerRadius, - angle: endAngle, - sign: -sign, - isExternal: true, - cornerRadius: cornerRadius, - cornerIsExternal: cornerIsExternal - }), - eict = _getTangentCircle4.circleTangency, - eilt = _getTangentCircle4.lineTangency, - eit = _getTangentCircle4.theta; - - var innerArcAngle = cornerIsExternal ? Math.abs(startAngle - endAngle) : Math.abs(startAngle - endAngle) - sit - eit; - - if (innerArcAngle < 0 && cornerRadius === 0) { - return "".concat(path, "L").concat(cx, ",").concat(cy, "Z"); - } - - path += "L".concat(eilt.x, ",").concat(eilt.y, "\n A").concat(cornerRadius, ",").concat(cornerRadius, ",0,0,").concat(+(sign < 0), ",").concat(eict.x, ",").concat(eict.y, "\n A").concat(innerRadius, ",").concat(innerRadius, ",0,").concat(+(innerArcAngle > 180), ",").concat(+(sign > 0), ",").concat(sict.x, ",").concat(sict.y, "\n A").concat(cornerRadius, ",").concat(cornerRadius, ",0,0,").concat(+(sign < 0), ",").concat(silt.x, ",").concat(silt.y, "Z"); - } else { - path += "L".concat(cx, ",").concat(cy, "Z"); - } - - return path; -}; - -var Sector = /*#__PURE__*/function (_PureComponent) { - _inherits$n(Sector, _PureComponent); - - var _super = _createSuper$n(Sector); - - function Sector() { - _classCallCheck$o(this, Sector); - - return _super.apply(this, arguments); - } - - _createClass$o(Sector, [{ - key: "render", - value: function render() { - var _this$props = this.props, - cx = _this$props.cx, - cy = _this$props.cy, - innerRadius = _this$props.innerRadius, - outerRadius = _this$props.outerRadius, - cornerRadius = _this$props.cornerRadius, - forceCornerRadius = _this$props.forceCornerRadius, - cornerIsExternal = _this$props.cornerIsExternal, - startAngle = _this$props.startAngle, - endAngle = _this$props.endAngle, - className = _this$props.className; - - if (outerRadius < innerRadius || startAngle === endAngle) { - return null; - } - - var layerClass = classNames('recharts-sector', className); - var deltaRadius = outerRadius - innerRadius; - var cr = getPercentValue(cornerRadius, deltaRadius, 0, true); - var path; - - if (cr > 0 && Math.abs(startAngle - endAngle) < 360) { - path = getSectorWithCorner({ - cx: cx, - cy: cy, - innerRadius: innerRadius, - outerRadius: outerRadius, - cornerRadius: Math.min(cr, deltaRadius / 2), - forceCornerRadius: forceCornerRadius, - cornerIsExternal: cornerIsExternal, - startAngle: startAngle, - endAngle: endAngle - }); - } else { - path = getSectorPath({ - cx: cx, - cy: cy, - innerRadius: innerRadius, - outerRadius: outerRadius, - startAngle: startAngle, - endAngle: endAngle - }); - } - - return /*#__PURE__*/React__default.createElement("path", _extends$q({}, filterProps(this.props, true), { - className: layerClass, - d: path - })); - } - }]); - - return Sector; -}(PureComponent); -Sector.defaultProps = { - cx: 0, - cy: 0, - innerRadius: 0, - outerRadius: 0, - startAngle: 0, - endAngle: 0, - cornerRadius: 0, - forceCornerRadius: false, - cornerIsExternal: false -}; - -function _typeof$m(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$m = function _typeof(obj) { return typeof obj; }; } else { _typeof$m = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$m(obj); } - -function _extends$p() { _extends$p = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$p.apply(this, arguments); } - -function ownKeys$m(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$m(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$m(Object(source), true).forEach(function (key) { _defineProperty$m(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$m(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$m(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$n(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$n(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$n(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$n(Constructor.prototype, protoProps); if (staticProps) _defineProperties$n(Constructor, staticProps); return Constructor; } - -function _inherits$m(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$m(subClass, superClass); } - -function _setPrototypeOf$m(o, p) { _setPrototypeOf$m = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$m(o, p); } - -function _createSuper$m(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$m(); return function _createSuperInternal() { var Super = _getPrototypeOf$m(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$m(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$m(this, result); }; } - -function _possibleConstructorReturn$m(self, call) { if (call && (_typeof$m(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$m(self); } - -function _assertThisInitialized$m(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$m() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$m(o) { _getPrototypeOf$m = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$m(o); } -var CURVE_FACTORIES = { - curveBasisClosed: curveBasisClosed, - curveBasisOpen: curveBasisOpen, - curveBasis: curveBasis, - curveLinearClosed: curveLinearClosed, - curveLinear: curveLinear, - curveMonotoneX: monotoneX, - curveMonotoneY: monotoneY, - curveNatural: curveNatural, - curveStep: curveStep, - curveStepAfter: stepAfter, - curveStepBefore: stepBefore -}; - -var defined = function defined(p) { - return p.x === +p.x && p.y === +p.y; -}; - -var getX = function getX(p) { - return p.x; -}; - -var getY = function getY(p) { - return p.y; -}; - -var getCurveFactory = function getCurveFactory(type, layout) { - if (isFunction_1(type)) { - return type; - } - - var name = "curve".concat(upperFirst_1(type)); - - if (name === 'curveMonotone' && layout) { - return CURVE_FACTORIES["".concat(name).concat(layout === 'vertical' ? 'Y' : 'X')]; - } - - return CURVE_FACTORIES[name] || curveLinear; -}; - -var Curve = /*#__PURE__*/function (_PureComponent) { - _inherits$m(Curve, _PureComponent); - - var _super = _createSuper$m(Curve); - - function Curve() { - _classCallCheck$n(this, Curve); - - return _super.apply(this, arguments); - } - - _createClass$n(Curve, [{ - key: "getPath", - value: - /** - * Calculate the path of curve - * @return {String} path - */ - function getPath() { - var _this$props = this.props, - type = _this$props.type, - points = _this$props.points, - baseLine = _this$props.baseLine, - layout = _this$props.layout, - connectNulls = _this$props.connectNulls; - var curveFactory = getCurveFactory(type, layout); - var formatPoints = connectNulls ? points.filter(function (entry) { - return defined(entry); - }) : points; - var lineFunction; - - if (isArray_1(baseLine)) { - var formatBaseLine = connectNulls ? baseLine.filter(function (base) { - return defined(base); - }) : baseLine; - var areaPoints = formatPoints.map(function (entry, index) { - return _objectSpread$m(_objectSpread$m({}, entry), {}, { - base: formatBaseLine[index] - }); - }); - - if (layout === 'vertical') { - lineFunction = shapeArea().y(getY).x1(getX).x0(function (d) { - return d.base.x; - }); - } else { - lineFunction = shapeArea().x(getX).y1(getY).y0(function (d) { - return d.base.y; - }); - } - - lineFunction.defined(defined).curve(curveFactory); - return lineFunction(areaPoints); - } - - if (layout === 'vertical' && isNumber(baseLine)) { - lineFunction = shapeArea().y(getY).x1(getX).x0(baseLine); - } else if (isNumber(baseLine)) { - lineFunction = shapeArea().x(getX).y1(getY).y0(baseLine); - } else { - lineFunction = shapeLine().x(getX).y(getY); - } - - lineFunction.defined(defined).curve(curveFactory); - return lineFunction(formatPoints); - } - }, { - key: "render", - value: function render() { - var _this$props2 = this.props, - className = _this$props2.className, - points = _this$props2.points, - path = _this$props2.path, - pathRef = _this$props2.pathRef; - - if ((!points || !points.length) && !path) { - return null; - } - - var realPath = points && points.length ? this.getPath() : path; - return /*#__PURE__*/React__default.createElement("path", _extends$p({}, filterProps(this.props), adaptEventHandlers(this.props), { - className: classNames('recharts-curve', className), - d: realPath, - ref: pathRef - })); - } - }]); - - return Curve; -}(PureComponent); -Curve.defaultProps = { - type: 'linear', - points: [], - connectNulls: false -}; - -function _typeof$l(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$l = function _typeof(obj) { return typeof obj; }; } else { _typeof$l = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$l(obj); } - -function _extends$o() { _extends$o = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$o.apply(this, arguments); } - -function _classCallCheck$m(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$m(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$m(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$m(Constructor.prototype, protoProps); if (staticProps) _defineProperties$m(Constructor, staticProps); return Constructor; } - -function _inherits$l(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$l(subClass, superClass); } - -function _setPrototypeOf$l(o, p) { _setPrototypeOf$l = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$l(o, p); } - -function _createSuper$l(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$l(); return function _createSuperInternal() { var Super = _getPrototypeOf$l(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$l(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$l(this, result); }; } - -function _possibleConstructorReturn$l(self, call) { if (call && (_typeof$l(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$l(self); } - -function _assertThisInitialized$l(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$l() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$l(o) { _getPrototypeOf$l = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$l(o); } - -var getRectanglePath = function getRectanglePath(x, y, width, height, radius) { - var maxRadius = Math.min(Math.abs(width) / 2, Math.abs(height) / 2); - var ySign = height >= 0 ? 1 : -1; - var xSign = width >= 0 ? 1 : -1; - var clockWise = height >= 0 && width >= 0 || height < 0 && width < 0 ? 1 : 0; - var path; - - if (maxRadius > 0 && radius instanceof Array) { - var newRadius = [0, 0, 0, 0]; - - for (var i = 0, len = 4; i < len; i++) { - newRadius[i] = radius[i] > maxRadius ? maxRadius : radius[i]; - } - - path = "M".concat(x, ",").concat(y + ySign * newRadius[0]); - - if (newRadius[0] > 0) { - path += "A ".concat(newRadius[0], ",").concat(newRadius[0], ",0,0,").concat(clockWise, ",").concat(x + xSign * newRadius[0], ",").concat(y); - } - - path += "L ".concat(x + width - xSign * newRadius[1], ",").concat(y); - - if (newRadius[1] > 0) { - path += "A ".concat(newRadius[1], ",").concat(newRadius[1], ",0,0,").concat(clockWise, ",\n ").concat(x + width, ",").concat(y + ySign * newRadius[1]); - } - - path += "L ".concat(x + width, ",").concat(y + height - ySign * newRadius[2]); - - if (newRadius[2] > 0) { - path += "A ".concat(newRadius[2], ",").concat(newRadius[2], ",0,0,").concat(clockWise, ",\n ").concat(x + width - xSign * newRadius[2], ",").concat(y + height); - } - - path += "L ".concat(x + xSign * newRadius[3], ",").concat(y + height); - - if (newRadius[3] > 0) { - path += "A ".concat(newRadius[3], ",").concat(newRadius[3], ",0,0,").concat(clockWise, ",\n ").concat(x, ",").concat(y + height - ySign * newRadius[3]); - } - - path += 'Z'; - } else if (maxRadius > 0 && radius === +radius && radius > 0) { - var _newRadius = Math.min(maxRadius, radius); - - path = "M ".concat(x, ",").concat(y + ySign * _newRadius, "\n A ").concat(_newRadius, ",").concat(_newRadius, ",0,0,").concat(clockWise, ",").concat(x + xSign * _newRadius, ",").concat(y, "\n L ").concat(x + width - xSign * _newRadius, ",").concat(y, "\n A ").concat(_newRadius, ",").concat(_newRadius, ",0,0,").concat(clockWise, ",").concat(x + width, ",").concat(y + ySign * _newRadius, "\n L ").concat(x + width, ",").concat(y + height - ySign * _newRadius, "\n A ").concat(_newRadius, ",").concat(_newRadius, ",0,0,").concat(clockWise, ",").concat(x + width - xSign * _newRadius, ",").concat(y + height, "\n L ").concat(x + xSign * _newRadius, ",").concat(y + height, "\n A ").concat(_newRadius, ",").concat(_newRadius, ",0,0,").concat(clockWise, ",").concat(x, ",").concat(y + height - ySign * _newRadius, " Z"); - } else { - path = "M ".concat(x, ",").concat(y, " h ").concat(width, " v ").concat(height, " h ").concat(-width, " Z"); - } - - return path; -}; - -var isInRectangle = function isInRectangle(point, rect) { - if (!point || !rect) { - return false; - } - - var px = point.x, - py = point.y; - var x = rect.x, - y = rect.y, - width = rect.width, - height = rect.height; - - if (Math.abs(width) > 0 && Math.abs(height) > 0) { - var minX = Math.min(x, x + width); - var maxX = Math.max(x, x + width); - var minY = Math.min(y, y + height); - var maxY = Math.max(y, y + height); - return px >= minX && px <= maxX && py >= minY && py <= maxY; - } - - return false; -}; -var Rectangle = /*#__PURE__*/function (_PureComponent) { - _inherits$l(Rectangle, _PureComponent); - - var _super = _createSuper$l(Rectangle); - - function Rectangle() { - var _this; - - _classCallCheck$m(this, Rectangle); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - totalLength: -1 - }; - _this.node = void 0; - return _this; - } - - _createClass$m(Rectangle, [{ - key: "componentDidMount", - value: - /* eslint-disable react/no-did-mount-set-state */ - function componentDidMount() { - if (this.node && this.node.getTotalLength) { - try { - var totalLength = this.node.getTotalLength(); - - if (totalLength) { - this.setState({ - totalLength: totalLength - }); - } - } catch (err) {// calculate total length error - } - } - } - }, { - key: "render", - value: function render() { - var _this2 = this; - - var _this$props = this.props, - x = _this$props.x, - y = _this$props.y, - width = _this$props.width, - height = _this$props.height, - radius = _this$props.radius, - className = _this$props.className; - var totalLength = this.state.totalLength; - var _this$props2 = this.props, - animationEasing = _this$props2.animationEasing, - animationDuration = _this$props2.animationDuration, - animationBegin = _this$props2.animationBegin, - isAnimationActive = _this$props2.isAnimationActive, - isUpdateAnimationActive = _this$props2.isUpdateAnimationActive; - - if (x !== +x || y !== +y || width !== +width || height !== +height || width === 0 || height === 0) { - return null; - } - - var layerClass = classNames('recharts-rectangle', className); - - if (!isUpdateAnimationActive) { - return /*#__PURE__*/React__default.createElement("path", _extends$o({}, filterProps(this.props, true), { - className: layerClass, - d: getRectanglePath(x, y, width, height, radius) - })); - } - - return /*#__PURE__*/React__default.createElement(Animate, { - canBegin: totalLength > 0, - from: { - width: width, - height: height, - x: x, - y: y - }, - to: { - width: width, - height: height, - x: x, - y: y - }, - duration: animationDuration, - animationEasing: animationEasing, - isActive: isUpdateAnimationActive - }, function (_ref) { - var currWidth = _ref.width, - currHeight = _ref.height, - currX = _ref.x, - currY = _ref.y; - return /*#__PURE__*/React__default.createElement(Animate, { - canBegin: totalLength > 0, - from: "0px ".concat(totalLength === -1 ? 1 : totalLength, "px"), - to: "".concat(totalLength, "px 0px"), - attributeName: "strokeDasharray", - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing - }, /*#__PURE__*/React__default.createElement("path", _extends$o({}, filterProps(_this2.props, true), { - className: layerClass, - d: getRectanglePath(currX, currY, currWidth, currHeight, radius), - ref: function ref(node) { - _this2.node = node; - } - }))); - }); - } - }]); - - return Rectangle; -}(PureComponent); -Rectangle.defaultProps = { - x: 0, - y: 0, - width: 0, - height: 0, - // The radius of border - // The radius of four corners when radius is a number - // The radius of left-top, right-top, right-bottom, left-bottom when radius is an array - radius: 0, - isAnimationActive: false, - isUpdateAnimationActive: false, - animationBegin: 0, - animationDuration: 1500, - animationEasing: 'ease' -}; - -function _typeof$k(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$k = function _typeof(obj) { return typeof obj; }; } else { _typeof$k = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$k(obj); } - -function _extends$n() { _extends$n = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$n.apply(this, arguments); } - -function _objectWithoutProperties$b(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$b(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$b(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _classCallCheck$l(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$l(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$l(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$l(Constructor.prototype, protoProps); if (staticProps) _defineProperties$l(Constructor, staticProps); return Constructor; } - -function _inherits$k(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$k(subClass, superClass); } - -function _setPrototypeOf$k(o, p) { _setPrototypeOf$k = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$k(o, p); } - -function _createSuper$k(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$k(); return function _createSuperInternal() { var Super = _getPrototypeOf$k(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$k(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$k(this, result); }; } - -function _possibleConstructorReturn$k(self, call) { if (call && (_typeof$k(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$k(self); } - -function _assertThisInitialized$k(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$k() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$k(o) { _getPrototypeOf$k = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$k(o); } - -function _toConsumableArray$2(arr) { return _arrayWithoutHoles$2(arr) || _iterableToArray$2(arr) || _unsupportedIterableToArray$5(arr) || _nonIterableSpread$2(); } - -function _nonIterableSpread$2() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$5(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$5(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$5(o, minLen); } - -function _iterableToArray$2(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$2(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$5(arr); } - -function _arrayLikeToArray$5(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -var isValidatePoint = function isValidatePoint(point) { - return point && point.x === +point.x && point.y === +point.y; -}; - -var getParsedPoints = function getParsedPoints() { - var points = arguments.length > 0 && arguments[0] !== undefined ? arguments[0] : []; - var segmentPoints = [[]]; - points.forEach(function (entry) { - if (isValidatePoint(entry)) { - segmentPoints[segmentPoints.length - 1].push(entry); - } else if (segmentPoints[segmentPoints.length - 1].length > 0) { - // add another path - segmentPoints.push([]); - } - }); - - if (isValidatePoint(points[0])) { - segmentPoints[segmentPoints.length - 1].push(points[0]); - } - - if (segmentPoints[segmentPoints.length - 1].length <= 0) { - segmentPoints = segmentPoints.slice(0, -1); - } - - return segmentPoints; -}; - -var getSinglePolygonPath = function getSinglePolygonPath(points, connectNulls) { - var segmentPoints = getParsedPoints(points); - - if (connectNulls) { - segmentPoints = [segmentPoints.reduce(function (res, segPoints) { - return [].concat(_toConsumableArray$2(res), _toConsumableArray$2(segPoints)); - }, [])]; - } - - var polygonPath = segmentPoints.map(function (segPoints) { - return segPoints.reduce(function (path, point, index) { - return "".concat(path).concat(index === 0 ? 'M' : 'L').concat(point.x, ",").concat(point.y); - }, ''); - }).join(''); - return segmentPoints.length === 1 ? "".concat(polygonPath, "Z") : polygonPath; -}; - -var getRanglePath = function getRanglePath(points, baseLinePoints, connectNulls) { - var outerPath = getSinglePolygonPath(points, connectNulls); - return "".concat(outerPath.slice(-1) === 'Z' ? outerPath.slice(0, -1) : outerPath, "L").concat(getSinglePolygonPath(baseLinePoints.reverse(), connectNulls).slice(1)); -}; - -var Polygon = /*#__PURE__*/function (_PureComponent) { - _inherits$k(Polygon, _PureComponent); - - var _super = _createSuper$k(Polygon); - - function Polygon() { - _classCallCheck$l(this, Polygon); - - return _super.apply(this, arguments); - } - - _createClass$l(Polygon, [{ - key: "render", - value: function render() { - var _this$props = this.props, - points = _this$props.points, - className = _this$props.className, - baseLinePoints = _this$props.baseLinePoints, - connectNulls = _this$props.connectNulls, - others = _objectWithoutProperties$b(_this$props, ["points", "className", "baseLinePoints", "connectNulls"]); - - if (!points || !points.length) { - return null; - } - - var layerClass = classNames('recharts-polygon', className); - - if (baseLinePoints && baseLinePoints.length) { - var hasStroke = others.stroke && others.stroke !== 'none'; - var rangePath = getRanglePath(points, baseLinePoints, connectNulls); - return /*#__PURE__*/React__default.createElement("g", { - className: layerClass - }, /*#__PURE__*/React__default.createElement("path", _extends$n({}, filterProps(others, true), { - fill: rangePath.slice(-1) === 'Z' ? others.fill : 'none', - stroke: "none", - d: rangePath - })), hasStroke ? /*#__PURE__*/React__default.createElement("path", _extends$n({}, filterProps(others, true), { - fill: "none", - d: getSinglePolygonPath(points, connectNulls) - })) : null, hasStroke ? /*#__PURE__*/React__default.createElement("path", _extends$n({}, filterProps(others, true), { - fill: "none", - d: getSinglePolygonPath(baseLinePoints, connectNulls) - })) : null); - } - - var singlePath = getSinglePolygonPath(points, connectNulls); - return /*#__PURE__*/React__default.createElement("path", _extends$n({}, filterProps(others, true), { - fill: singlePath.slice(-1) === 'Z' ? others.fill : 'none', - className: layerClass, - d: singlePath - })); - } - }]); - - return Polygon; -}(PureComponent); - -function _typeof$j(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$j = function _typeof(obj) { return typeof obj; }; } else { _typeof$j = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$j(obj); } - -function _extends$m() { _extends$m = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$m.apply(this, arguments); } - -function _classCallCheck$k(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$k(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$k(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$k(Constructor.prototype, protoProps); if (staticProps) _defineProperties$k(Constructor, staticProps); return Constructor; } - -function _inherits$j(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$j(subClass, superClass); } - -function _setPrototypeOf$j(o, p) { _setPrototypeOf$j = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$j(o, p); } - -function _createSuper$j(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$j(); return function _createSuperInternal() { var Super = _getPrototypeOf$j(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$j(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$j(this, result); }; } - -function _possibleConstructorReturn$j(self, call) { if (call && (_typeof$j(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$j(self); } - -function _assertThisInitialized$j(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$j() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$j(o) { _getPrototypeOf$j = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$j(o); } -var Dot = /*#__PURE__*/function (_PureComponent) { - _inherits$j(Dot, _PureComponent); - - var _super = _createSuper$j(Dot); - - function Dot() { - _classCallCheck$k(this, Dot); - - return _super.apply(this, arguments); - } - - _createClass$k(Dot, [{ - key: "render", - value: function render() { - var _this$props = this.props, - cx = _this$props.cx, - cy = _this$props.cy, - r = _this$props.r, - className = _this$props.className; - var layerClass = classNames('recharts-dot', className); - - if (cx === +cx && cy === +cy && r === +r) { - return /*#__PURE__*/React__default.createElement("circle", _extends$m({}, filterProps(this.props), adaptEventHandlers(this.props), { - className: layerClass, - cx: cx, - cy: cy, - r: r - })); - } - - return null; - } - }]); - - return Dot; -}(PureComponent); - -function _typeof$i(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$i = function _typeof(obj) { return typeof obj; }; } else { _typeof$i = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$i(obj); } - -function _extends$l() { _extends$l = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$l.apply(this, arguments); } - -function _classCallCheck$j(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$j(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$j(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$j(Constructor.prototype, protoProps); if (staticProps) _defineProperties$j(Constructor, staticProps); return Constructor; } - -function _inherits$i(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$i(subClass, superClass); } - -function _setPrototypeOf$i(o, p) { _setPrototypeOf$i = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$i(o, p); } - -function _createSuper$i(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$i(); return function _createSuperInternal() { var Super = _getPrototypeOf$i(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$i(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$i(this, result); }; } - -function _possibleConstructorReturn$i(self, call) { if (call && (_typeof$i(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$i(self); } - -function _assertThisInitialized$i(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$i() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$i(o) { _getPrototypeOf$i = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$i(o); } -var Cross = /*#__PURE__*/function (_PureComponent) { - _inherits$i(Cross, _PureComponent); - - var _super = _createSuper$i(Cross); - - function Cross() { - _classCallCheck$j(this, Cross); - - return _super.apply(this, arguments); - } - - _createClass$j(Cross, [{ - key: "render", - value: function render() { - var _this$props = this.props, - x = _this$props.x, - y = _this$props.y, - width = _this$props.width, - height = _this$props.height, - top = _this$props.top, - left = _this$props.left, - className = _this$props.className; - - if (!isNumber(x) || !isNumber(y) || !isNumber(width) || !isNumber(height) || !isNumber(top) || !isNumber(left)) { - return null; - } - - return /*#__PURE__*/React__default.createElement("path", _extends$l({}, filterProps(this.props, true), { - className: classNames('recharts-cross', className), - d: Cross.getPath(x, y, width, height, top, left) - })); - } - }], [{ - key: "getPath", - value: function getPath(x, y, width, height, top, left) { - return "M".concat(x, ",").concat(top, "v").concat(height, "M").concat(left, ",").concat(y, "h").concat(width); - } - }]); - - return Cross; -}(PureComponent); -Cross.defaultProps = { - x: 0, - y: 0, - top: 0, - left: 0, - width: 0, - height: 0 -}; - -function _typeof$h(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$h = function _typeof(obj) { return typeof obj; }; } else { _typeof$h = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$h(obj); } - -function _extends$k() { _extends$k = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$k.apply(this, arguments); } - -function ownKeys$l(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$l(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$l(Object(source), true).forEach(function (key) { _defineProperty$l(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$l(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$l(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$i(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$i(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$i(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$i(Constructor.prototype, protoProps); if (staticProps) _defineProperties$i(Constructor, staticProps); return Constructor; } - -function _inherits$h(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$h(subClass, superClass); } - -function _setPrototypeOf$h(o, p) { _setPrototypeOf$h = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$h(o, p); } - -function _createSuper$h(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$h(); return function _createSuperInternal() { var Super = _getPrototypeOf$h(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$h(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$h(this, result); }; } - -function _possibleConstructorReturn$h(self, call) { if (call && (_typeof$h(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$h(self); } - -function _assertThisInitialized$h(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$h() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$h(o) { _getPrototypeOf$h = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$h(o); } -var PolarGrid = /*#__PURE__*/function (_PureComponent) { - _inherits$h(PolarGrid, _PureComponent); - - var _super = _createSuper$h(PolarGrid); - - function PolarGrid() { - _classCallCheck$i(this, PolarGrid); - - return _super.apply(this, arguments); - } - - _createClass$i(PolarGrid, [{ - key: "getPolygonPath", - value: function getPolygonPath(radius) { - var _this$props = this.props, - cx = _this$props.cx, - cy = _this$props.cy, - polarAngles = _this$props.polarAngles; - var path = ''; - polarAngles.forEach(function (angle, i) { - var point = polarToCartesian(cx, cy, radius, angle); - - if (i) { - path += "L ".concat(point.x, ",").concat(point.y); - } else { - path += "M ".concat(point.x, ",").concat(point.y); - } - }); - path += 'Z'; - return path; - } - /** - * Draw axis of radial line - * @return {[type]} The lines - */ - - }, { - key: "renderPolarAngles", - value: function renderPolarAngles() { - var _this$props2 = this.props, - cx = _this$props2.cx, - cy = _this$props2.cy, - innerRadius = _this$props2.innerRadius, - outerRadius = _this$props2.outerRadius, - polarAngles = _this$props2.polarAngles, - radialLines = _this$props2.radialLines; - - if (!polarAngles || !polarAngles.length || !radialLines) { - return null; - } - - var props = _objectSpread$l({ - stroke: '#ccc' - }, filterProps(this.props)); - - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-polar-grid-angle" - }, polarAngles.map(function (entry, i) { - var start = polarToCartesian(cx, cy, innerRadius, entry); - var end = polarToCartesian(cx, cy, outerRadius, entry); - return /*#__PURE__*/React__default.createElement("line", _extends$k({}, props, { - key: "line-".concat(i) // eslint-disable-line react/no-array-index-key - , - x1: start.x, - y1: start.y, - x2: end.x, - y2: end.y - })); - })); - } - /** - * Draw concentric circles - * @param {Number} radius The radius of circle - * @param {Number} index The index of circle - * @param {Object} extraProps Extra props - * @return {ReactElement} circle - */ - - }, { - key: "renderConcentricCircle", - value: function renderConcentricCircle(radius, index, extraProps) { - var _this$props3 = this.props, - cx = _this$props3.cx, - cy = _this$props3.cy; - - var props = _objectSpread$l(_objectSpread$l({ - stroke: '#ccc' - }, filterProps(this.props)), {}, { - fill: 'none' - }, extraProps); - - return /*#__PURE__*/React__default.createElement("circle", _extends$k({}, props, { - className: "recharts-polar-grid-concentric-circle", - key: "circle-".concat(index), - cx: cx, - cy: cy, - r: radius - })); - } - /** - * Draw concentric polygons - * @param {Number} radius The radius of polygon - * @param {Number} index The index of polygon - * @param {Object} extraProps Extra props - * @return {ReactElement} polygon - */ - - }, { - key: "renderConcentricPolygon", - value: function renderConcentricPolygon(radius, index, extraProps) { - var props = _objectSpread$l(_objectSpread$l({ - stroke: '#ccc' - }, filterProps(this.props)), {}, { - fill: 'none' - }, extraProps); - - return /*#__PURE__*/React__default.createElement("path", _extends$k({}, props, { - className: "recharts-polar-grid-concentric-polygon", - key: "path-".concat(index), - d: this.getPolygonPath(radius) - })); - } - /** - * Draw concentric axis - * @return {ReactElement} Concentric axis - * @todo Optimize the name - */ - - }, { - key: "renderConcentricPath", - value: function renderConcentricPath() { - var _this = this; - - var _this$props4 = this.props, - polarRadius = _this$props4.polarRadius, - gridType = _this$props4.gridType; - - if (!polarRadius || !polarRadius.length) { - return null; - } - - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-polar-grid-concentric" - }, polarRadius.map(function (entry, i) { - return gridType === 'circle' ? _this.renderConcentricCircle(entry, i) : _this.renderConcentricPolygon(entry, i); - })); - } - }, { - key: "render", - value: function render() { - var outerRadius = this.props.outerRadius; - - if (outerRadius <= 0) { - return null; - } - - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-polar-grid" - }, this.renderPolarAngles(), this.renderConcentricPath()); - } - }]); - - return PolarGrid; -}(PureComponent); -PolarGrid.displayName = 'PolarGrid'; -PolarGrid.defaultProps = { - cx: 0, - cy: 0, - innerRadius: 0, - outerRadius: 0, - gridType: 'polygon', - radialLines: true -}; - -var baseExtremum$1 = _baseExtremum, - baseIteratee$7 = _baseIteratee, - baseLt = _baseLt; - -/** - * This method is like `_.min` except that it accepts `iteratee` which is - * invoked for each element in `array` to generate the criterion by which - * the value is ranked. The iteratee is invoked with one argument: (value). - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Math - * @param {Array} array The array to iterate over. - * @param {Function} [iteratee=_.identity] The iteratee invoked per element. - * @returns {*} Returns the minimum value. - * @example - * - * var objects = [{ 'n': 1 }, { 'n': 2 }]; - * - * _.minBy(objects, function(o) { return o.n; }); - * // => { 'n': 1 } - * - * // The `_.property` iteratee shorthand. - * _.minBy(objects, 'n'); - * // => { 'n': 1 } - */ -function minBy(array, iteratee) { - return (array && array.length) - ? baseExtremum$1(array, baseIteratee$7(iteratee), baseLt) - : undefined; -} - -var minBy_1 = minBy; - -var baseExtremum = _baseExtremum, - baseGt = _baseGt, - baseIteratee$6 = _baseIteratee; - -/** - * This method is like `_.max` except that it accepts `iteratee` which is - * invoked for each element in `array` to generate the criterion by which - * the value is ranked. The iteratee is invoked with one argument: (value). - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Math - * @param {Array} array The array to iterate over. - * @param {Function} [iteratee=_.identity] The iteratee invoked per element. - * @returns {*} Returns the maximum value. - * @example - * - * var objects = [{ 'n': 1 }, { 'n': 2 }]; - * - * _.maxBy(objects, function(o) { return o.n; }); - * // => { 'n': 2 } - * - * // The `_.property` iteratee shorthand. - * _.maxBy(objects, 'n'); - * // => { 'n': 2 } - */ -function maxBy(array, iteratee) { - return (array && array.length) - ? baseExtremum(array, baseIteratee$6(iteratee), baseGt) - : undefined; -} - -var maxBy_1 = maxBy; - -function _typeof$g(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$g = function _typeof(obj) { return typeof obj; }; } else { _typeof$g = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$g(obj); } - -function _extends$j() { _extends$j = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$j.apply(this, arguments); } - -function ownKeys$k(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$k(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$k(Object(source), true).forEach(function (key) { _defineProperty$k(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$k(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$k(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _objectWithoutProperties$a(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$a(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$a(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _classCallCheck$h(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$h(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$h(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$h(Constructor.prototype, protoProps); if (staticProps) _defineProperties$h(Constructor, staticProps); return Constructor; } - -function _inherits$g(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$g(subClass, superClass); } - -function _setPrototypeOf$g(o, p) { _setPrototypeOf$g = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$g(o, p); } - -function _createSuper$g(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$g(); return function _createSuperInternal() { var Super = _getPrototypeOf$g(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$g(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$g(this, result); }; } - -function _possibleConstructorReturn$g(self, call) { if (call && (_typeof$g(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$g(self); } - -function _assertThisInitialized$g(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$g() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$g(o) { _getPrototypeOf$g = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$g(o); } -var PolarRadiusAxis = /*#__PURE__*/function (_PureComponent) { - _inherits$g(PolarRadiusAxis, _PureComponent); - - var _super = _createSuper$g(PolarRadiusAxis); - - function PolarRadiusAxis() { - _classCallCheck$h(this, PolarRadiusAxis); - - return _super.apply(this, arguments); - } - - _createClass$h(PolarRadiusAxis, [{ - key: "getTickValueCoord", - value: - /** - * Calculate the coordinate of tick - * @param {Number} coordinate The radius of tick - * @return {Object} (x, y) - */ - function getTickValueCoord(_ref) { - var coordinate = _ref.coordinate; - var _this$props = this.props, - angle = _this$props.angle, - cx = _this$props.cx, - cy = _this$props.cy; - return polarToCartesian(cx, cy, coordinate, angle); - } - }, { - key: "getTickTextAnchor", - value: function getTickTextAnchor() { - var orientation = this.props.orientation; - var textAnchor; - - switch (orientation) { - case 'left': - textAnchor = 'end'; - break; - - case 'right': - textAnchor = 'start'; - break; - - default: - textAnchor = 'middle'; - break; - } - - return textAnchor; - } - }, { - key: "getViewBox", - value: function getViewBox() { - var _this$props2 = this.props, - cx = _this$props2.cx, - cy = _this$props2.cy, - angle = _this$props2.angle, - ticks = _this$props2.ticks; - - var maxRadiusTick = maxBy_1(ticks, function (entry) { - return entry.coordinate || 0; - }); - - var minRadiusTick = minBy_1(ticks, function (entry) { - return entry.coordinate || 0; - }); - - return { - cx: cx, - cy: cy, - startAngle: angle, - endAngle: angle, - innerRadius: minRadiusTick.coordinate || 0, - outerRadius: maxRadiusTick.coordinate || 0 - }; - } - }, { - key: "renderAxisLine", - value: function renderAxisLine() { - var _this$props3 = this.props, - cx = _this$props3.cx, - cy = _this$props3.cy, - angle = _this$props3.angle, - ticks = _this$props3.ticks, - axisLine = _this$props3.axisLine, - others = _objectWithoutProperties$a(_this$props3, ["cx", "cy", "angle", "ticks", "axisLine"]); - - var extent = ticks.reduce(function (result, entry) { - return [Math.min(result[0], entry.coordinate), Math.max(result[1], entry.coordinate)]; - }, [Infinity, -Infinity]); - var point0 = polarToCartesian(cx, cy, extent[0], angle); - var point1 = polarToCartesian(cx, cy, extent[1], angle); - - var props = _objectSpread$k(_objectSpread$k(_objectSpread$k({}, filterProps(others)), {}, { - fill: 'none' - }, filterProps(axisLine)), {}, { - x1: point0.x, - y1: point0.y, - x2: point1.x, - y2: point1.y - }); - - return /*#__PURE__*/React__default.createElement("line", _extends$j({ - className: "recharts-polar-radius-axis-line" - }, props)); - } - }, { - key: "renderTicks", - value: function renderTicks() { - var _this = this; - - var _this$props4 = this.props, - ticks = _this$props4.ticks, - tick = _this$props4.tick, - angle = _this$props4.angle, - tickFormatter = _this$props4.tickFormatter, - stroke = _this$props4.stroke, - others = _objectWithoutProperties$a(_this$props4, ["ticks", "tick", "angle", "tickFormatter", "stroke"]); - - var textAnchor = this.getTickTextAnchor(); - var axisProps = filterProps(others); - var customTickProps = filterProps(tick); - var items = ticks.map(function (entry, i) { - var coord = _this.getTickValueCoord(entry); - - var tickProps = _objectSpread$k(_objectSpread$k(_objectSpread$k(_objectSpread$k({ - textAnchor: textAnchor, - transform: "rotate(".concat(90 - angle, ", ").concat(coord.x, ", ").concat(coord.y, ")") - }, axisProps), {}, { - stroke: 'none', - fill: stroke - }, customTickProps), {}, { - index: i - }, coord), {}, { - payload: entry - }); - - return /*#__PURE__*/React__default.createElement(Layer, _extends$j({ - className: "recharts-polar-radius-axis-tick", - key: "tick-".concat(i) // eslint-disable-line react/no-array-index-key - - }, adaptEventsOfChild(_this.props, entry, i)), PolarRadiusAxis.renderTickItem(tick, tickProps, tickFormatter ? tickFormatter(entry.value, i) : entry.value)); - }); - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-polar-radius-axis-ticks" - }, items); - } - }, { - key: "render", - value: function render() { - var _this$props5 = this.props, - ticks = _this$props5.ticks, - axisLine = _this$props5.axisLine, - tick = _this$props5.tick; - - if (!ticks || !ticks.length) { - return null; - } - - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-polar-radius-axis" - }, axisLine && this.renderAxisLine(), tick && this.renderTicks(), Label.renderCallByParent(this.props, this.getViewBox())); - } - }], [{ - key: "renderTickItem", - value: function renderTickItem(option, props, value) { - var tickItem; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - tickItem = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - tickItem = option(props); - } else { - tickItem = /*#__PURE__*/React__default.createElement(Text, _extends$j({}, props, { - className: "recharts-polar-radius-axis-tick-value" - }), value); - } - - return tickItem; - } - }]); - - return PolarRadiusAxis; -}(PureComponent); -PolarRadiusAxis.displayName = 'PolarRadiusAxis'; -PolarRadiusAxis.axisType = 'radiusAxis'; -PolarRadiusAxis.defaultProps = { - type: 'number', - radiusAxisId: 0, - cx: 0, - cy: 0, - angle: 0, - orientation: 'right', - stroke: '#ccc', - axisLine: true, - tick: true, - tickCount: 5, - domain: [0, 'auto'], - allowDataOverflow: false, - scale: 'auto', - allowDuplicatedCategory: true -}; - -function _typeof$f(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$f = function _typeof(obj) { return typeof obj; }; } else { _typeof$f = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$f(obj); } - -function _extends$i() { _extends$i = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$i.apply(this, arguments); } - -function ownKeys$j(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$j(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$j(Object(source), true).forEach(function (key) { _defineProperty$j(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$j(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$j(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$g(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$g(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$g(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$g(Constructor.prototype, protoProps); if (staticProps) _defineProperties$g(Constructor, staticProps); return Constructor; } - -function _inherits$f(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$f(subClass, superClass); } - -function _setPrototypeOf$f(o, p) { _setPrototypeOf$f = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$f(o, p); } - -function _createSuper$f(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$f(); return function _createSuperInternal() { var Super = _getPrototypeOf$f(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$f(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$f(this, result); }; } - -function _possibleConstructorReturn$f(self, call) { if (call && (_typeof$f(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$f(self); } - -function _assertThisInitialized$f(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$f() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$f(o) { _getPrototypeOf$f = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$f(o); } -var RADIAN = Math.PI / 180; -var eps = 1e-5; -var PolarAngleAxis = /*#__PURE__*/function (_PureComponent) { - _inherits$f(PolarAngleAxis, _PureComponent); - - var _super = _createSuper$f(PolarAngleAxis); - - function PolarAngleAxis() { - _classCallCheck$g(this, PolarAngleAxis); - - return _super.apply(this, arguments); - } - - _createClass$g(PolarAngleAxis, [{ - key: "getTickLineCoord", - value: - /** - * Calculate the coordinate of line endpoint - * @param {Object} data The Data if ticks - * @return {Object} (x0, y0): The start point of text, - * (x1, y1): The end point close to text, - * (x2, y2): The end point close to axis - */ - function getTickLineCoord(data) { - var _this$props = this.props, - cx = _this$props.cx, - cy = _this$props.cy, - radius = _this$props.radius, - orientation = _this$props.orientation, - tickSize = _this$props.tickSize; - var tickLineSize = tickSize || 8; - var p1 = polarToCartesian(cx, cy, radius, data.coordinate); - var p2 = polarToCartesian(cx, cy, radius + (orientation === 'inner' ? -1 : 1) * tickLineSize, data.coordinate); - return { - x1: p1.x, - y1: p1.y, - x2: p2.x, - y2: p2.y - }; - } - /** - * Get the text-anchor of each tick - * @param {Object} data Data of ticks - * @return {String} text-anchor - */ - - }, { - key: "getTickTextAnchor", - value: function getTickTextAnchor(data) { - var orientation = this.props.orientation; - var cos = Math.cos(-data.coordinate * RADIAN); - var textAnchor; - - if (cos > eps) { - textAnchor = orientation === 'outer' ? 'start' : 'end'; - } else if (cos < -eps) { - textAnchor = orientation === 'outer' ? 'end' : 'start'; - } else { - textAnchor = 'middle'; - } - - return textAnchor; - } - }, { - key: "renderAxisLine", - value: function renderAxisLine() { - var _this$props2 = this.props, - cx = _this$props2.cx, - cy = _this$props2.cy, - radius = _this$props2.radius, - axisLine = _this$props2.axisLine, - axisLineType = _this$props2.axisLineType; - - var props = _objectSpread$j(_objectSpread$j({}, filterProps(this.props)), {}, { - fill: 'none' - }, filterProps(axisLine)); - - if (axisLineType === 'circle') { - return /*#__PURE__*/React__default.createElement(Dot, _extends$i({ - className: "recharts-polar-angle-axis-line" - }, props, { - cx: cx, - cy: cy, - r: radius - })); - } - - var ticks = this.props.ticks; - var points = ticks.map(function (entry) { - return polarToCartesian(cx, cy, radius, entry.coordinate); - }); - return /*#__PURE__*/React__default.createElement(Polygon, _extends$i({ - className: "recharts-polar-angle-axis-line" - }, props, { - points: points - })); - } - }, { - key: "renderTicks", - value: function renderTicks() { - var _this = this; - - var _this$props3 = this.props, - ticks = _this$props3.ticks, - tick = _this$props3.tick, - tickLine = _this$props3.tickLine, - tickFormatter = _this$props3.tickFormatter, - stroke = _this$props3.stroke; - var axisProps = filterProps(this.props); - var customTickProps = filterProps(tick); - - var tickLineProps = _objectSpread$j(_objectSpread$j({}, axisProps), {}, { - fill: 'none' - }, filterProps(tickLine)); - - var items = ticks.map(function (entry, i) { - var lineCoord = _this.getTickLineCoord(entry); - - var textAnchor = _this.getTickTextAnchor(entry); - - var tickProps = _objectSpread$j(_objectSpread$j(_objectSpread$j({ - textAnchor: textAnchor - }, axisProps), {}, { - stroke: 'none', - fill: stroke - }, customTickProps), {}, { - index: i, - payload: entry, - x: lineCoord.x2, - y: lineCoord.y2 - }); - - return /*#__PURE__*/React__default.createElement(Layer, _extends$i({ - className: "recharts-polar-angle-axis-tick", - key: "tick-".concat(i) // eslint-disable-line react/no-array-index-key - - }, adaptEventsOfChild(_this.props, entry, i)), tickLine && /*#__PURE__*/React__default.createElement("line", _extends$i({ - className: "recharts-polar-angle-axis-tick-line" - }, tickLineProps, lineCoord)), tick && PolarAngleAxis.renderTickItem(tick, tickProps, tickFormatter ? tickFormatter(entry.value, i) : entry.value)); - }); - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-polar-angle-axis-ticks" - }, items); - } - }, { - key: "render", - value: function render() { - var _this$props4 = this.props, - ticks = _this$props4.ticks, - radius = _this$props4.radius, - axisLine = _this$props4.axisLine; - - if (radius <= 0 || !ticks || !ticks.length) { - return null; - } - - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-polar-angle-axis" - }, axisLine && this.renderAxisLine(), this.renderTicks()); - } - }], [{ - key: "renderTickItem", - value: function renderTickItem(option, props, value) { - var tickItem; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - tickItem = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - tickItem = option(props); - } else { - tickItem = /*#__PURE__*/React__default.createElement(Text, _extends$i({}, props, { - className: "recharts-polar-angle-axis-tick-value" - }), value); - } - - return tickItem; - } - }]); - - return PolarAngleAxis; -}(PureComponent); -PolarAngleAxis.displayName = 'PolarAngleAxis'; -PolarAngleAxis.axisType = 'angleAxis'; -PolarAngleAxis.defaultProps = { - type: 'category', - angleAxisId: 0, - scale: 'auto', - cx: 0, - cy: 0, - domain: [0, 'auto'], - orientation: 'outer', - axisLine: true, - tickLine: true, - tickSize: 8, - tick: true, - hide: false, - allowDuplicatedCategory: true -}; - -var overArg = _overArg; - -/** Built-in value references. */ -var getPrototype$3 = overArg(Object.getPrototypeOf, Object); - -var _getPrototype = getPrototype$3; - -var baseGetTag$1 = _baseGetTag$1, - getPrototype$2 = _getPrototype, - isObjectLike$3 = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var objectTag$1 = '[object Object]'; - -/** Used for built-in method references. */ -var funcProto = Function.prototype, - objectProto$3 = Object.prototype; - -/** Used to resolve the decompiled source of functions. */ -var funcToString = funcProto.toString; - -/** Used to check objects for own properties. */ -var hasOwnProperty$3 = objectProto$3.hasOwnProperty; - -/** Used to infer the `Object` constructor. */ -var objectCtorString = funcToString.call(Object); - -/** - * Checks if `value` is a plain object, that is, an object created by the - * `Object` constructor or one with a `[[Prototype]]` of `null`. - * - * @static - * @memberOf _ - * @since 0.8.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a plain object, else `false`. - * @example - * - * function Foo() { - * this.a = 1; - * } - * - * _.isPlainObject(new Foo); - * // => false - * - * _.isPlainObject([1, 2, 3]); - * // => false - * - * _.isPlainObject({ 'x': 0, 'y': 0 }); - * // => true - * - * _.isPlainObject(Object.create(null)); - * // => true - */ -function isPlainObject$1(value) { - if (!isObjectLike$3(value) || baseGetTag$1(value) != objectTag$1) { - return false; - } - var proto = getPrototype$2(value); - if (proto === null) { - return true; - } - var Ctor = hasOwnProperty$3.call(proto, 'constructor') && proto.constructor; - return typeof Ctor == 'function' && Ctor instanceof Ctor && - funcToString.call(Ctor) == objectCtorString; -} - -var isPlainObject_1 = isPlainObject$1; - -function _typeof$e(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$e = function _typeof(obj) { return typeof obj; }; } else { _typeof$e = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$e(obj); } - -function _extends$h() { _extends$h = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$h.apply(this, arguments); } - -function ownKeys$i(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$i(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$i(Object(source), true).forEach(function (key) { _defineProperty$i(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$i(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$i(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$f(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$f(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$f(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$f(Constructor.prototype, protoProps); if (staticProps) _defineProperties$f(Constructor, staticProps); return Constructor; } - -function _inherits$e(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$e(subClass, superClass); } - -function _setPrototypeOf$e(o, p) { _setPrototypeOf$e = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$e(o, p); } - -function _createSuper$e(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$e(); return function _createSuperInternal() { var Super = _getPrototypeOf$e(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$e(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$e(this, result); }; } - -function _possibleConstructorReturn$e(self, call) { if (call && (_typeof$e(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$e(self); } - -function _assertThisInitialized$e(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$e() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$e(o) { _getPrototypeOf$e = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$e(o); } -var Pie = /*#__PURE__*/function (_PureComponent) { - _inherits$e(Pie, _PureComponent); - - var _super = _createSuper$e(Pie); - - function Pie(props) { - var _this; - - _classCallCheck$f(this, Pie); - - _this = _super.call(this, props); - _this.state = void 0; - _this.id = uniqueId('recharts-pie-'); - - _this.handleAnimationEnd = function () { - var onAnimationEnd = _this.props.onAnimationEnd; - - _this.setState({ - isAnimationFinished: true - }); - - if (isFunction_1(onAnimationEnd)) { - onAnimationEnd(); - } - }; - - _this.handleAnimationStart = function () { - var onAnimationStart = _this.props.onAnimationStart; - - _this.setState({ - isAnimationFinished: false - }); - - if (isFunction_1(onAnimationStart)) { - onAnimationStart(); - } - }; - - _this.state = { - isAnimationFinished: !props.isAnimationActive, - prevIsAnimationActive: props.isAnimationActive, - prevAnimationId: props.animationId - }; - return _this; - } - - _createClass$f(Pie, [{ - key: "isActiveIndex", - value: function isActiveIndex(i) { - var activeIndex = this.props.activeIndex; - - if (Array.isArray(activeIndex)) { - return activeIndex.indexOf(i) !== -1; - } - - return i === activeIndex; - } - }, { - key: "hasActiveIndex", - value: function hasActiveIndex() { - var activeIndex = this.props.activeIndex; - return Array.isArray(activeIndex) ? activeIndex.length !== 0 : activeIndex || activeIndex === 0; - } - }, { - key: "renderLabels", - value: function renderLabels(sectors) { - var isAnimationActive = this.props.isAnimationActive; - - if (isAnimationActive && !this.state.isAnimationFinished) { - return null; - } - - var _this$props = this.props, - label = _this$props.label, - labelLine = _this$props.labelLine, - dataKey = _this$props.dataKey, - valueKey = _this$props.valueKey; - var pieProps = filterProps(this.props); - var customLabelProps = filterProps(label); - var customLabelLineProps = filterProps(labelLine); - var offsetRadius = label && label.offsetRadius || 20; - var labels = sectors.map(function (entry, i) { - var midAngle = (entry.startAngle + entry.endAngle) / 2; - var endPoint = polarToCartesian(entry.cx, entry.cy, entry.outerRadius + offsetRadius, midAngle); - - var labelProps = _objectSpread$i(_objectSpread$i(_objectSpread$i(_objectSpread$i({}, pieProps), entry), {}, { - stroke: 'none' - }, customLabelProps), {}, { - index: i, - textAnchor: Pie.getTextAnchor(endPoint.x, entry.cx) - }, endPoint); - - var lineProps = _objectSpread$i(_objectSpread$i(_objectSpread$i(_objectSpread$i({}, pieProps), entry), {}, { - fill: 'none', - stroke: entry.fill - }, customLabelLineProps), {}, { - index: i, - points: [polarToCartesian(entry.cx, entry.cy, entry.outerRadius, midAngle), endPoint], - key: 'line' - }); - - var realDataKey = dataKey; // TODO: compatible to lower versions - - if (isNil_1(dataKey) && isNil_1(valueKey)) { - realDataKey = 'value'; - } else if (isNil_1(dataKey)) { - realDataKey = valueKey; - } - - return ( - /*#__PURE__*/ - // eslint-disable-next-line react/no-array-index-key - React__default.createElement(Layer, { - key: "label-".concat(i) - }, labelLine && Pie.renderLabelLineItem(labelLine, lineProps), Pie.renderLabelItem(label, labelProps, getValueByDataKey(entry, realDataKey))) - ); - }); - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-pie-labels" - }, labels); - } - }, { - key: "renderSectorsStatically", - value: function renderSectorsStatically(sectors) { - var _this2 = this; - - var _this$props2 = this.props, - activeShape = _this$props2.activeShape, - blendStroke = _this$props2.blendStroke, - inactiveShapeProp = _this$props2.inactiveShape; - return sectors.map(function (entry, i) { - var inactiveShape = inactiveShapeProp && _this2.hasActiveIndex() ? inactiveShapeProp : null; - var sectorOptions = _this2.isActiveIndex(i) ? activeShape : inactiveShape; - - var sectorProps = _objectSpread$i(_objectSpread$i({}, entry), {}, { - stroke: blendStroke ? entry.fill : entry.stroke - }); - - return /*#__PURE__*/React__default.createElement(Layer, _extends$h({ - className: "recharts-pie-sector" - }, adaptEventsOfChild(_this2.props, entry, i), { - key: "sector-".concat(i) // eslint-disable-line react/no-array-index-key - - }), Pie.renderSectorItem(sectorOptions, sectorProps)); - }); - } - }, { - key: "renderSectorsWithAnimation", - value: function renderSectorsWithAnimation() { - var _this3 = this; - - var _this$props3 = this.props, - sectors = _this$props3.sectors, - isAnimationActive = _this$props3.isAnimationActive, - animationBegin = _this$props3.animationBegin, - animationDuration = _this$props3.animationDuration, - animationEasing = _this$props3.animationEasing, - animationId = _this$props3.animationId; - var _this$state = this.state, - prevSectors = _this$state.prevSectors, - prevIsAnimationActive = _this$state.prevIsAnimationActive; - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - from: { - t: 0 - }, - to: { - t: 1 - }, - key: "pie-".concat(animationId, "-").concat(prevIsAnimationActive), - onAnimationStart: this.handleAnimationStart, - onAnimationEnd: this.handleAnimationEnd - }, function (_ref) { - var t = _ref.t; - var stepData = []; - var first = sectors && sectors[0]; - var curAngle = first.startAngle; - sectors.forEach(function (entry, index) { - var prev = prevSectors && prevSectors[index]; - var paddingAngle = index > 0 ? get_1(entry, 'paddingAngle', 0) : 0; - - if (prev) { - var angleIp = interpolateNumber$2(prev.endAngle - prev.startAngle, entry.endAngle - entry.startAngle); - - var latest = _objectSpread$i(_objectSpread$i({}, entry), {}, { - startAngle: curAngle + paddingAngle, - endAngle: curAngle + angleIp(t) + paddingAngle - }); - - stepData.push(latest); - curAngle = latest.endAngle; - } else { - var endAngle = entry.endAngle, - startAngle = entry.startAngle; - var interpolatorAngle = interpolateNumber$2(0, endAngle - startAngle); - var deltaAngle = interpolatorAngle(t); - - var _latest = _objectSpread$i(_objectSpread$i({}, entry), {}, { - startAngle: curAngle + paddingAngle, - endAngle: curAngle + deltaAngle + paddingAngle - }); - - stepData.push(_latest); - curAngle = _latest.endAngle; - } - }); - return /*#__PURE__*/React__default.createElement(Layer, null, _this3.renderSectorsStatically(stepData)); - }); - } - }, { - key: "renderSectors", - value: function renderSectors() { - var _this$props4 = this.props, - sectors = _this$props4.sectors, - isAnimationActive = _this$props4.isAnimationActive; - var prevSectors = this.state.prevSectors; - - if (isAnimationActive && sectors && sectors.length && (!prevSectors || !isEqual_1(prevSectors, sectors))) { - return this.renderSectorsWithAnimation(); - } - - return this.renderSectorsStatically(sectors); - } - }, { - key: "render", - value: function render() { - var _this$props5 = this.props, - hide = _this$props5.hide, - sectors = _this$props5.sectors, - className = _this$props5.className, - label = _this$props5.label, - cx = _this$props5.cx, - cy = _this$props5.cy, - innerRadius = _this$props5.innerRadius, - outerRadius = _this$props5.outerRadius, - isAnimationActive = _this$props5.isAnimationActive; - var isAnimationFinished = this.state.isAnimationFinished; - - if (hide || !sectors || !sectors.length || !isNumber(cx) || !isNumber(cy) || !isNumber(innerRadius) || !isNumber(outerRadius)) { - return null; - } - - var layerClass = classNames('recharts-pie', className); - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass - }, this.renderSectors(), label && this.renderLabels(sectors), Label.renderCallByParent(this.props, null, false), (!isAnimationActive || isAnimationFinished) && LabelList.renderCallByParent(this.props, sectors, false)); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (prevState.prevIsAnimationActive !== nextProps.isAnimationActive) { - return { - prevIsAnimationActive: nextProps.isAnimationActive, - prevAnimationId: nextProps.animationId, - curSectors: nextProps.sectors, - prevSectors: [], - isAnimationFinished: true - }; - } - - if (nextProps.isAnimationActive && nextProps.animationId !== prevState.prevAnimationId) { - return { - prevAnimationId: nextProps.animationId, - curSectors: nextProps.sectors, - prevSectors: prevState.curSectors, - isAnimationFinished: true - }; - } - - if (nextProps.sectors !== prevState.curSectors) { - return { - curSectors: nextProps.sectors, - isAnimationFinished: true - }; - } - - return null; - } - }, { - key: "getTextAnchor", - value: function getTextAnchor(x, cx) { - if (x > cx) { - return 'start'; - } - - if (x < cx) { - return 'end'; - } - - return 'middle'; - } - }, { - key: "renderLabelLineItem", - value: function renderLabelLineItem(option, props) { - if ( /*#__PURE__*/React__default.isValidElement(option)) { - return /*#__PURE__*/React__default.cloneElement(option, props); - } - - if (isFunction_1(option)) { - return option(props); - } - - return /*#__PURE__*/React__default.createElement(Curve, _extends$h({}, props, { - type: "linear", - className: "recharts-pie-label-line" - })); - } - }, { - key: "renderLabelItem", - value: function renderLabelItem(option, props, value) { - if ( /*#__PURE__*/React__default.isValidElement(option)) { - return /*#__PURE__*/React__default.cloneElement(option, props); - } - - var label = value; - - if (isFunction_1(option)) { - label = option(props); - - if ( /*#__PURE__*/React__default.isValidElement(label)) { - return label; - } - } - - return /*#__PURE__*/React__default.createElement(Text, _extends$h({}, props, { - alignmentBaseline: "middle", - className: "recharts-pie-label-text" - }), label); - } - }, { - key: "renderSectorItem", - value: function renderSectorItem(option, props) { - if ( /*#__PURE__*/React__default.isValidElement(option)) { - return /*#__PURE__*/React__default.cloneElement(option, props); - } - - if (isFunction_1(option)) { - return option(props); - } - - if (isPlainObject_1(option)) { - return /*#__PURE__*/React__default.createElement(Sector, _extends$h({}, props, option)); - } - - return /*#__PURE__*/React__default.createElement(Sector, props); - } - }]); - - return Pie; -}(PureComponent); -Pie.displayName = 'Pie'; -Pie.defaultProps = { - stroke: '#fff', - fill: '#808080', - legendType: 'rect', - cx: '50%', - cy: '50%', - startAngle: 0, - endAngle: 360, - innerRadius: 0, - outerRadius: '80%', - paddingAngle: 0, - labelLine: true, - hide: false, - minAngle: 0, - isAnimationActive: !Global.isSsr, - animationBegin: 400, - animationDuration: 1500, - animationEasing: 'ease', - nameKey: 'name', - blendStroke: false -}; - -Pie.parseDeltaAngle = function (startAngle, endAngle) { - var sign = mathSign(endAngle - startAngle); - var deltaAngle = Math.min(Math.abs(endAngle - startAngle), 360); - return sign * deltaAngle; -}; - -Pie.getRealPieData = function (item) { - var _item$props = item.props, - data = _item$props.data, - children = _item$props.children; - var presentationProps = filterProps(item.props); - var cells = findAllByType(children, Cell.displayName); - - if (data && data.length) { - return data.map(function (entry, index) { - return _objectSpread$i(_objectSpread$i(_objectSpread$i({ - payload: entry - }, presentationProps), entry), cells && cells[index] && cells[index].props); - }); - } - - if (cells && cells.length) { - return cells.map(function (cell) { - return _objectSpread$i(_objectSpread$i({}, presentationProps), cell.props); - }); - } - - return []; -}; - -Pie.parseCoordinateOfPie = function (item, offset) { - var top = offset.top, - left = offset.left, - width = offset.width, - height = offset.height; - var maxPieRadius = getMaxRadius(width, height); - var cx = left + getPercentValue(item.props.cx, width, width / 2); - var cy = top + getPercentValue(item.props.cy, height, height / 2); - var innerRadius = getPercentValue(item.props.innerRadius, maxPieRadius, 0); - var outerRadius = getPercentValue(item.props.outerRadius, maxPieRadius, maxPieRadius * 0.8); - var maxRadius = item.props.maxRadius || Math.sqrt(width * width + height * height) / 2; - return { - cx: cx, - cy: cy, - innerRadius: innerRadius, - outerRadius: outerRadius, - maxRadius: maxRadius - }; -}; - -Pie.getComposedData = function (_ref2) { - var item = _ref2.item, - offset = _ref2.offset; - var pieData = Pie.getRealPieData(item); - - if (!pieData || !pieData.length) { - return null; - } - - var _item$props2 = item.props, - cornerRadius = _item$props2.cornerRadius, - startAngle = _item$props2.startAngle, - endAngle = _item$props2.endAngle, - paddingAngle = _item$props2.paddingAngle, - dataKey = _item$props2.dataKey, - nameKey = _item$props2.nameKey, - valueKey = _item$props2.valueKey, - tooltipType = _item$props2.tooltipType; - var minAngle = Math.abs(item.props.minAngle); - var coordinate = Pie.parseCoordinateOfPie(item, offset); - var deltaAngle = Pie.parseDeltaAngle(startAngle, endAngle); - var absDeltaAngle = Math.abs(deltaAngle); - var realDataKey = dataKey; - - if (isNil_1(dataKey) && isNil_1(valueKey)) { - warn(false, "Use \"dataKey\" to specify the value of pie,\n the props \"valueKey\" will be deprecated in 1.1.0"); - realDataKey = 'value'; - } else if (isNil_1(dataKey)) { - warn(false, "Use \"dataKey\" to specify the value of pie,\n the props \"valueKey\" will be deprecated in 1.1.0"); - realDataKey = valueKey; - } - - var notZeroItemCount = pieData.filter(function (entry) { - return getValueByDataKey(entry, realDataKey, 0) !== 0; - }).length; - var totalPadingAngle = (absDeltaAngle >= 360 ? notZeroItemCount : notZeroItemCount - 1) * paddingAngle; - var realTotalAngle = absDeltaAngle - notZeroItemCount * minAngle - totalPadingAngle; - var sum = pieData.reduce(function (result, entry) { - var val = getValueByDataKey(entry, realDataKey, 0); - return result + (isNumber(val) ? val : 0); - }, 0); - var sectors; - - if (sum > 0) { - var prev; - sectors = pieData.map(function (entry, i) { - var val = getValueByDataKey(entry, realDataKey, 0); - var name = getValueByDataKey(entry, nameKey, i); - var percent = (isNumber(val) ? val : 0) / sum; - var tempStartAngle; - - if (i) { - tempStartAngle = prev.endAngle + mathSign(deltaAngle) * paddingAngle * (val !== 0 ? 1 : 0); - } else { - tempStartAngle = startAngle; - } - - var tempEndAngle = tempStartAngle + mathSign(deltaAngle) * ((val !== 0 ? minAngle : 0) + percent * realTotalAngle); - var midAngle = (tempStartAngle + tempEndAngle) / 2; - var middleRadius = (coordinate.innerRadius + coordinate.outerRadius) / 2; - var tooltipPayload = [{ - name: name, - value: val, - payload: entry, - dataKey: realDataKey, - type: tooltipType - }]; - var tooltipPosition = polarToCartesian(coordinate.cx, coordinate.cy, middleRadius, midAngle); - prev = _objectSpread$i(_objectSpread$i(_objectSpread$i({ - percent: percent, - cornerRadius: cornerRadius, - name: name, - tooltipPayload: tooltipPayload, - midAngle: midAngle, - middleRadius: middleRadius, - tooltipPosition: tooltipPosition - }, entry), coordinate), {}, { - value: getValueByDataKey(entry, realDataKey), - startAngle: tempStartAngle, - endAngle: tempEndAngle, - payload: entry, - paddingAngle: mathSign(deltaAngle) * paddingAngle - }); - return prev; - }); - } - - return _objectSpread$i(_objectSpread$i({}, coordinate), {}, { - sectors: sectors, - data: pieData - }); -}; - -var first = {exports: {}}; - -/** - * Gets the first element of `array`. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @alias first - * @category Array - * @param {Array} array The array to query. - * @returns {*} Returns the first element of `array`. - * @example - * - * _.head([1, 2, 3]); - * // => 1 - * - * _.head([]); - * // => undefined - */ - -function head(array) { - return (array && array.length) ? array[0] : undefined; -} - -var head_1 = head; - -(function (module) { - module.exports = head_1; -} (first)); - -var _first = /*@__PURE__*/getDefaultExportFromCjs(first.exports); - -function _typeof$d(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$d = function _typeof(obj) { return typeof obj; }; } else { _typeof$d = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$d(obj); } - -function _extends$g() { _extends$g = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$g.apply(this, arguments); } - -function ownKeys$h(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$h(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$h(Object(source), true).forEach(function (key) { _defineProperty$h(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$h(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$h(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$e(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$e(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$e(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$e(Constructor.prototype, protoProps); if (staticProps) _defineProperties$e(Constructor, staticProps); return Constructor; } - -function _inherits$d(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$d(subClass, superClass); } - -function _setPrototypeOf$d(o, p) { _setPrototypeOf$d = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$d(o, p); } - -function _createSuper$d(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$d(); return function _createSuperInternal() { var Super = _getPrototypeOf$d(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$d(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$d(this, result); }; } - -function _possibleConstructorReturn$d(self, call) { if (call && (_typeof$d(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$d(self); } - -function _assertThisInitialized$d(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$d() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$d(o) { _getPrototypeOf$d = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$d(o); } -var Radar = /*#__PURE__*/function (_PureComponent) { - _inherits$d(Radar, _PureComponent); - - var _super = _createSuper$d(Radar); - - function Radar() { - var _this; - - _classCallCheck$e(this, Radar); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - isAnimationFinished: false - }; - - _this.handleAnimationEnd = function () { - var onAnimationEnd = _this.props.onAnimationEnd; - - _this.setState({ - isAnimationFinished: true - }); - - if (isFunction_1(onAnimationEnd)) { - onAnimationEnd(); - } - }; - - _this.handleAnimationStart = function () { - var onAnimationStart = _this.props.onAnimationStart; - - _this.setState({ - isAnimationFinished: false - }); - - if (isFunction_1(onAnimationStart)) { - onAnimationStart(); - } - }; - - _this.handleMouseEnter = function (e) { - var onMouseEnter = _this.props.onMouseEnter; - - if (onMouseEnter) { - onMouseEnter(_this.props, e); - } - }; - - _this.handleMouseLeave = function (e) { - var onMouseLeave = _this.props.onMouseLeave; - - if (onMouseLeave) { - onMouseLeave(_this.props, e); - } - }; - - return _this; - } - - _createClass$e(Radar, [{ - key: "renderDots", - value: function renderDots(points) { - var _this$props = this.props, - dot = _this$props.dot, - dataKey = _this$props.dataKey; - var baseProps = filterProps(this.props); - var customDotProps = filterProps(dot); - var dots = points.map(function (entry, i) { - var dotProps = _objectSpread$h(_objectSpread$h(_objectSpread$h({ - key: "dot-".concat(i), - r: 3 - }, baseProps), customDotProps), {}, { - dataKey: dataKey, - cx: entry.x, - cy: entry.y, - index: i, - payload: entry - }); - - return Radar.renderDotItem(dot, dotProps); - }); - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-radar-dots" - }, dots); - } - }, { - key: "renderPolygonStatically", - value: function renderPolygonStatically(points) { - var _this$props2 = this.props, - shape = _this$props2.shape, - dot = _this$props2.dot, - isRange = _this$props2.isRange, - baseLinePoints = _this$props2.baseLinePoints, - connectNulls = _this$props2.connectNulls; - var radar; - - if ( /*#__PURE__*/React__default.isValidElement(shape)) { - radar = /*#__PURE__*/React__default.cloneElement(shape, _objectSpread$h(_objectSpread$h({}, this.props), {}, { - points: points - })); - } else if (isFunction_1(shape)) { - radar = shape(_objectSpread$h(_objectSpread$h({}, this.props), {}, { - points: points - })); - } else { - radar = /*#__PURE__*/React__default.createElement(Polygon, _extends$g({}, filterProps(this.props, true), { - onMouseEnter: this.handleMouseEnter, - onMouseLeave: this.handleMouseLeave, - points: points, - baseLinePoints: isRange ? baseLinePoints : null, - connectNulls: connectNulls - })); - } - - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-radar-polygon" - }, radar, dot ? this.renderDots(points) : null); - } - }, { - key: "renderPolygonWithAnimation", - value: function renderPolygonWithAnimation() { - var _this2 = this; - - var _this$props3 = this.props, - points = _this$props3.points, - isAnimationActive = _this$props3.isAnimationActive, - animationBegin = _this$props3.animationBegin, - animationDuration = _this$props3.animationDuration, - animationEasing = _this$props3.animationEasing, - animationId = _this$props3.animationId; - var prevPoints = this.state.prevPoints; - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - from: { - t: 0 - }, - to: { - t: 1 - }, - key: "radar-".concat(animationId), - onAnimationEnd: this.handleAnimationEnd, - onAnimationStart: this.handleAnimationStart - }, function (_ref) { - var t = _ref.t; - var prevPointsDiffFactor = prevPoints && prevPoints.length / points.length; - var stepData = points.map(function (entry, index) { - var prev = prevPoints && prevPoints[Math.floor(index * prevPointsDiffFactor)]; - - if (prev) { - var _interpolatorX = interpolateNumber$2(prev.x, entry.x); - - var _interpolatorY = interpolateNumber$2(prev.y, entry.y); - - return _objectSpread$h(_objectSpread$h({}, entry), {}, { - x: _interpolatorX(t), - y: _interpolatorY(t) - }); - } - - var interpolatorX = interpolateNumber$2(entry.cx, entry.x); - var interpolatorY = interpolateNumber$2(entry.cy, entry.y); - return _objectSpread$h(_objectSpread$h({}, entry), {}, { - x: interpolatorX(t), - y: interpolatorY(t) - }); - }); - return _this2.renderPolygonStatically(stepData); - }); - } - }, { - key: "renderPolygon", - value: function renderPolygon() { - var _this$props4 = this.props, - points = _this$props4.points, - isAnimationActive = _this$props4.isAnimationActive, - isRange = _this$props4.isRange; - var prevPoints = this.state.prevPoints; - - if (isAnimationActive && points && points.length && !isRange && (!prevPoints || !isEqual_1(prevPoints, points))) { - return this.renderPolygonWithAnimation(); - } - - return this.renderPolygonStatically(points); - } - }, { - key: "render", - value: function render() { - var _this$props5 = this.props, - hide = _this$props5.hide, - className = _this$props5.className, - points = _this$props5.points, - isAnimationActive = _this$props5.isAnimationActive; - - if (hide || !points || !points.length) { - return null; - } - - var isAnimationFinished = this.state.isAnimationFinished; - var layerClass = classNames('recharts-radar', className); - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass - }, this.renderPolygon(), (!isAnimationActive || isAnimationFinished) && LabelList.renderCallByParent(this.props, points)); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.animationId !== prevState.prevAnimationId) { - return { - prevAnimationId: nextProps.animationId, - curPoints: nextProps.points, - prevPoints: prevState.curPoints - }; - } - - if (nextProps.points !== prevState.curPoints) { - return { - curPoints: nextProps.points - }; - } - - return null; - } - }, { - key: "renderDotItem", - value: function renderDotItem(option, props) { - var dotItem; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - dotItem = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - dotItem = option(props); - } else { - dotItem = /*#__PURE__*/React__default.createElement(Dot, _extends$g({}, props, { - className: "recharts-radar-dot" - })); - } - - return dotItem; - } - }]); - - return Radar; -}(PureComponent); -Radar.displayName = 'Radar'; -Radar.defaultProps = { - angleAxisId: 0, - radiusAxisId: 0, - hide: false, - activeDot: true, - dot: false, - legendType: 'rect', - isAnimationActive: !Global.isSsr, - animationBegin: 0, - animationDuration: 1500, - animationEasing: 'ease' -}; - -Radar.getComposedData = function (_ref2) { - var radiusAxis = _ref2.radiusAxis, - angleAxis = _ref2.angleAxis, - displayedData = _ref2.displayedData, - dataKey = _ref2.dataKey, - bandSize = _ref2.bandSize; - var cx = angleAxis.cx, - cy = angleAxis.cy; - var isRange = false; - var points = []; - displayedData.forEach(function (entry, i) { - var name = getValueByDataKey(entry, angleAxis.dataKey, i); - var value = getValueByDataKey(entry, dataKey); - var angle = angleAxis.scale(name) + (bandSize || 0); - var pointValue = isArray_1(value) ? last_1(value) : value; - var radius = isNil_1(pointValue) ? undefined : radiusAxis.scale(pointValue); - - if (isArray_1(value) && value.length >= 2) { - isRange = true; - } - - points.push(_objectSpread$h(_objectSpread$h({}, polarToCartesian(cx, cy, radius, angle)), {}, { - name: name, - value: value, - cx: cx, - cy: cy, - radius: radius, - angle: angle, - payload: entry - })); - }); - var baseLinePoints = []; - - if (isRange) { - points.forEach(function (point) { - if (isArray_1(point.value)) { - var baseValue = _first(point.value); - - var radius = isNil_1(baseValue) ? undefined : radiusAxis.scale(baseValue); - baseLinePoints.push(_objectSpread$h(_objectSpread$h({}, point), {}, { - radius: radius - }, polarToCartesian(cx, cy, radius, point.angle))); - } else { - baseLinePoints.push(point); - } - }); - } - - return { - points: points, - isRange: isRange, - baseLinePoints: baseLinePoints - }; -}; - -function _typeof$c(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$c = function _typeof(obj) { return typeof obj; }; } else { _typeof$c = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$c(obj); } - -function ownKeys$g(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$g(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$g(Object(source), true).forEach(function (key) { _defineProperty$g(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$g(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$g(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _objectWithoutProperties$9(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$9(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$9(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _classCallCheck$d(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$d(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$d(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$d(Constructor.prototype, protoProps); if (staticProps) _defineProperties$d(Constructor, staticProps); return Constructor; } - -function _inherits$c(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$c(subClass, superClass); } - -function _setPrototypeOf$c(o, p) { _setPrototypeOf$c = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$c(o, p); } - -function _createSuper$c(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$c(); return function _createSuperInternal() { var Super = _getPrototypeOf$c(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$c(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$c(this, result); }; } - -function _possibleConstructorReturn$c(self, call) { if (call && (_typeof$c(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$c(self); } - -function _assertThisInitialized$c(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$c() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$c(o) { _getPrototypeOf$c = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$c(o); } -// import { AngleAxisProps, RadiusAxisProps } from './types'; - -var RadialBar = /*#__PURE__*/function (_PureComponent) { - _inherits$c(RadialBar, _PureComponent); - - var _super = _createSuper$c(RadialBar); - - function RadialBar() { - var _this; - - _classCallCheck$d(this, RadialBar); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - isAnimationFinished: false - }; - - _this.handleAnimationEnd = function () { - var onAnimationEnd = _this.props.onAnimationEnd; - - _this.setState({ - isAnimationFinished: true - }); - - if (isFunction_1(onAnimationEnd)) { - onAnimationEnd(); - } - }; - - _this.handleAnimationStart = function () { - var onAnimationStart = _this.props.onAnimationStart; - - _this.setState({ - isAnimationFinished: false - }); - - if (isFunction_1(onAnimationStart)) { - onAnimationStart(); - } - }; - - return _this; - } - - _createClass$d(RadialBar, [{ - key: "getDeltaAngle", - value: function getDeltaAngle() { - var _this$props = this.props, - startAngle = _this$props.startAngle, - endAngle = _this$props.endAngle; - var sign = mathSign(endAngle - startAngle); - var deltaAngle = Math.min(Math.abs(endAngle - startAngle), 360); - return sign * deltaAngle; - } - }, { - key: "renderSectorsStatically", - value: function renderSectorsStatically(sectors) { - var _this2 = this; - - var _this$props2 = this.props, - shape = _this$props2.shape, - activeShape = _this$props2.activeShape, - activeIndex = _this$props2.activeIndex, - cornerRadius = _this$props2.cornerRadius, - others = _objectWithoutProperties$9(_this$props2, ["shape", "activeShape", "activeIndex", "cornerRadius"]); - - var baseProps = filterProps(others); - return sectors.map(function (entry, i) { - var props = _objectSpread$g(_objectSpread$g(_objectSpread$g(_objectSpread$g({}, baseProps), {}, { - cornerRadius: cornerRadius - }, entry), adaptEventsOfChild(_this2.props, entry, i)), {}, { - key: "sector-".concat(i), - className: 'recharts-radial-bar-sector', - forceCornerRadius: others.forceCornerRadius, - cornerIsExternal: others.cornerIsExternal - }); - - return RadialBar.renderSectorShape(i === activeIndex ? activeShape : shape, props); - }); - } - }, { - key: "renderSectorsWithAnimation", - value: function renderSectorsWithAnimation() { - var _this3 = this; - - var _this$props3 = this.props, - data = _this$props3.data, - isAnimationActive = _this$props3.isAnimationActive, - animationBegin = _this$props3.animationBegin, - animationDuration = _this$props3.animationDuration, - animationEasing = _this$props3.animationEasing, - animationId = _this$props3.animationId; - var prevData = this.state.prevData; - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - from: { - t: 0 - }, - to: { - t: 1 - }, - key: "radialBar-".concat(animationId), - onAnimationStart: this.handleAnimationStart, - onAnimationEnd: this.handleAnimationEnd - }, function (_ref) { - var t = _ref.t; - var stepData = data.map(function (entry, index) { - var prev = prevData && prevData[index]; - - if (prev) { - var interpolatorStartAngle = interpolateNumber$2(prev.startAngle, entry.startAngle); - var interpolatorEndAngle = interpolateNumber$2(prev.endAngle, entry.endAngle); - return _objectSpread$g(_objectSpread$g({}, entry), {}, { - startAngle: interpolatorStartAngle(t), - endAngle: interpolatorEndAngle(t) - }); - } - - var endAngle = entry.endAngle, - startAngle = entry.startAngle; - var interpolator = interpolateNumber$2(startAngle, endAngle); - return _objectSpread$g(_objectSpread$g({}, entry), {}, { - endAngle: interpolator(t) - }); - }); - return /*#__PURE__*/React__default.createElement(Layer, null, _this3.renderSectorsStatically(stepData)); - }); - } - }, { - key: "renderSectors", - value: function renderSectors() { - var _this$props4 = this.props, - data = _this$props4.data, - isAnimationActive = _this$props4.isAnimationActive; - var prevData = this.state.prevData; - - if (isAnimationActive && data && data.length && (!prevData || !isEqual_1(prevData, data))) { - return this.renderSectorsWithAnimation(); - } - - return this.renderSectorsStatically(data); - } - }, { - key: "renderBackground", - value: function renderBackground(sectors) { - var _this4 = this; - - var cornerRadius = this.props.cornerRadius; - var backgroundProps = filterProps(this.props.background); - return sectors.map(function (entry, i) { - // eslint-disable-next-line @typescript-eslint/no-unused-vars - entry.value; - var background = entry.background, - rest = _objectWithoutProperties$9(entry, ["value", "background"]); - - if (!background) { - return null; - } - - var props = _objectSpread$g(_objectSpread$g(_objectSpread$g(_objectSpread$g(_objectSpread$g({ - cornerRadius: cornerRadius - }, rest), {}, { - fill: '#eee' - }, background), backgroundProps), adaptEventsOfChild(_this4.props, entry, i)), {}, { - index: i, - key: "sector-".concat(i), - className: 'recharts-radial-bar-background-sector' - }); - - return RadialBar.renderSectorShape(background, props); - }); - } - }, { - key: "render", - value: function render() { - var _this$props5 = this.props, - hide = _this$props5.hide, - data = _this$props5.data, - className = _this$props5.className, - background = _this$props5.background, - isAnimationActive = _this$props5.isAnimationActive; - - if (hide || !data || !data.length) { - return null; - } - - var isAnimationFinished = this.state.isAnimationFinished; - var layerClass = classNames('recharts-area', className); - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass - }, background && /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-radial-bar-background" - }, this.renderBackground(data)), /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-radial-bar-sectors" - }, this.renderSectors()), (!isAnimationActive || isAnimationFinished) && LabelList.renderCallByParent(_objectSpread$g(_objectSpread$g({}, this.props), {}, { - clockWise: this.getDeltaAngle() < 0 - }), data)); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.animationId !== prevState.prevAnimationId) { - return { - prevAnimationId: nextProps.animationId, - curData: nextProps.data, - prevData: prevState.curData - }; - } - - if (nextProps.data !== prevState.curData) { - return { - curData: nextProps.data - }; - } - - return null; - } - }, { - key: "renderSectorShape", - value: function renderSectorShape(shape, props) { - var sectorShape; - - if ( /*#__PURE__*/React__default.isValidElement(shape)) { - sectorShape = /*#__PURE__*/React__default.cloneElement(shape, props); - } else if (isFunction_1(shape)) { - sectorShape = shape(props); - } else { - sectorShape = /*#__PURE__*/React__default.createElement(Sector, props); - } - - return sectorShape; - } - }]); - - return RadialBar; -}(PureComponent); -RadialBar.displayName = 'RadialBar'; -RadialBar.defaultProps = { - angleAxisId: 0, - radiusAxisId: 0, - minPointSize: 0, - hide: false, - legendType: 'rect', - data: [], - isAnimationActive: !Global.isSsr, - animationBegin: 0, - animationDuration: 1500, - animationEasing: 'ease', - forceCornerRadius: false, - cornerIsExternal: false -}; - -RadialBar.getComposedData = function (_ref2) { - var item = _ref2.item, - props = _ref2.props, - radiusAxis = _ref2.radiusAxis, - radiusAxisTicks = _ref2.radiusAxisTicks, - angleAxis = _ref2.angleAxis, - angleAxisTicks = _ref2.angleAxisTicks, - displayedData = _ref2.displayedData, - dataKey = _ref2.dataKey, - stackedData = _ref2.stackedData, - barPosition = _ref2.barPosition, - bandSize = _ref2.bandSize, - dataStartIndex = _ref2.dataStartIndex; - var pos = findPositionOfBar(barPosition, item); - - if (!pos) { - return null; - } - - var cx = angleAxis.cx, - cy = angleAxis.cy; - var layout = props.layout; - var _item$props = item.props, - children = _item$props.children, - minPointSize = _item$props.minPointSize; - var numericAxis = layout === 'radial' ? angleAxis : radiusAxis; - var stackedDomain = stackedData ? numericAxis.scale.domain() : null; - var baseValue = getBaseValueOfBar({ - numericAxis: numericAxis - }); - var cells = findAllByType(children, Cell.displayName); - var sectors = displayedData.map(function (entry, index) { - var value, innerRadius, outerRadius, startAngle, endAngle, backgroundSector; - - if (stackedData) { - value = truncateByDomain(stackedData[dataStartIndex + index], stackedDomain); - } else { - value = getValueByDataKey(entry, dataKey); - - if (!isArray_1(value)) { - value = [baseValue, value]; - } - } - - if (layout === 'radial') { - innerRadius = getCateCoordinateOfBar({ - axis: radiusAxis, - ticks: radiusAxisTicks, - bandSize: bandSize, - offset: pos.offset, - entry: entry, - index: index - }); - endAngle = angleAxis.scale(value[1]); - startAngle = angleAxis.scale(value[0]); - outerRadius = innerRadius + pos.size; - var deltaAngle = endAngle - startAngle; - - if (Math.abs(minPointSize) > 0 && Math.abs(deltaAngle) < Math.abs(minPointSize)) { - var delta = mathSign(deltaAngle || minPointSize) * (Math.abs(minPointSize) - Math.abs(deltaAngle)); - endAngle += delta; - } - - backgroundSector = { - background: { - cx: cx, - cy: cy, - innerRadius: innerRadius, - outerRadius: outerRadius, - startAngle: props.startAngle, - endAngle: props.endAngle - } - }; - } else { - innerRadius = radiusAxis.scale(value[0]); - outerRadius = radiusAxis.scale(value[1]); - startAngle = getCateCoordinateOfBar({ - axis: angleAxis, - ticks: angleAxisTicks, - bandSize: bandSize, - offset: pos.offset, - entry: entry, - index: index - }); - endAngle = startAngle + pos.size; - var deltaRadius = outerRadius - innerRadius; - - if (Math.abs(minPointSize) > 0 && Math.abs(deltaRadius) < Math.abs(minPointSize)) { - var _delta = mathSign(deltaRadius || minPointSize) * (Math.abs(minPointSize) - Math.abs(deltaRadius)); - - outerRadius += _delta; - } - } - - return _objectSpread$g(_objectSpread$g(_objectSpread$g(_objectSpread$g({}, entry), backgroundSector), {}, { - payload: entry, - value: stackedData ? value : value[1], - cx: cx, - cy: cy, - innerRadius: innerRadius, - outerRadius: outerRadius, - startAngle: startAngle, - endAngle: endAngle - }, cells && cells[index] && cells[index].props), {}, { - tooltipPayload: [getTooltipItem(item, entry)], - tooltipPosition: polarToCartesian(cx, cy, (innerRadius + outerRadius) / 2, (startAngle + endAngle) / 2) - }); - }); - return { - data: sectors, - layout: layout - }; -}; - -/* Built-in method references for those with the same name as other `lodash` methods. */ - -var nativeCeil = Math.ceil, - nativeMax$1 = Math.max; - -/** - * The base implementation of `_.range` and `_.rangeRight` which doesn't - * coerce arguments. - * - * @private - * @param {number} start The start of the range. - * @param {number} end The end of the range. - * @param {number} step The value to increment or decrement by. - * @param {boolean} [fromRight] Specify iterating from right to left. - * @returns {Array} Returns the range of numbers. - */ -function baseRange$1(start, end, step, fromRight) { - var index = -1, - length = nativeMax$1(nativeCeil((end - start) / (step || 1)), 0), - result = Array(length); - - while (length--) { - result[fromRight ? length : ++index] = start; - start += step; - } - return result; -} - -var _baseRange = baseRange$1; - -var toNumber = toNumber_1$1; - -/** Used as references for various `Number` constants. */ -var INFINITY = 1 / 0, - MAX_INTEGER = 1.7976931348623157e+308; - -/** - * Converts `value` to a finite number. - * - * @static - * @memberOf _ - * @since 4.12.0 - * @category Lang - * @param {*} value The value to convert. - * @returns {number} Returns the converted number. - * @example - * - * _.toFinite(3.2); - * // => 3.2 - * - * _.toFinite(Number.MIN_VALUE); - * // => 5e-324 - * - * _.toFinite(Infinity); - * // => 1.7976931348623157e+308 - * - * _.toFinite('3.2'); - * // => 3.2 - */ -function toFinite$2(value) { - if (!value) { - return value === 0 ? value : 0; - } - value = toNumber(value); - if (value === INFINITY || value === -INFINITY) { - var sign = (value < 0 ? -1 : 1); - return sign * MAX_INTEGER; - } - return value === value ? value : 0; -} - -var toFinite_1 = toFinite$2; - -var baseRange = _baseRange, - isIterateeCall$2 = _isIterateeCall, - toFinite$1 = toFinite_1; - -/** - * Creates a `_.range` or `_.rangeRight` function. - * - * @private - * @param {boolean} [fromRight] Specify iterating from right to left. - * @returns {Function} Returns the new range function. - */ -function createRange$1(fromRight) { - return function(start, end, step) { - if (step && typeof step != 'number' && isIterateeCall$2(start, end, step)) { - end = step = undefined; - } - // Ensure the sign of `-0` is preserved. - start = toFinite$1(start); - if (end === undefined) { - end = start; - start = 0; - } else { - end = toFinite$1(end); - } - step = step === undefined ? (start < end ? 1 : -1) : toFinite$1(step); - return baseRange(start, end, step, fromRight); - }; -} - -var _createRange = createRange$1; - -var createRange = _createRange; - -/** - * Creates an array of numbers (positive and/or negative) progressing from - * `start` up to, but not including, `end`. A step of `-1` is used if a negative - * `start` is specified without an `end` or `step`. If `end` is not specified, - * it's set to `start` with `start` then set to `0`. - * - * **Note:** JavaScript follows the IEEE-754 standard for resolving - * floating-point values which can produce unexpected results. - * - * @static - * @since 0.1.0 - * @memberOf _ - * @category Util - * @param {number} [start=0] The start of the range. - * @param {number} end The end of the range. - * @param {number} [step=1] The value to increment or decrement by. - * @returns {Array} Returns the range of numbers. - * @see _.inRange, _.rangeRight - * @example - * - * _.range(4); - * // => [0, 1, 2, 3] - * - * _.range(-4); - * // => [0, -1, -2, -3] - * - * _.range(1, 5); - * // => [1, 2, 3, 4] - * - * _.range(0, 20, 5); - * // => [0, 5, 10, 15] - * - * _.range(0, -4, -1); - * // => [0, -1, -2, -3] - * - * _.range(1, 4, 0); - * // => [1, 1, 1] - * - * _.range(0); - * // => [] - */ -var range = createRange(); - -var range_1 = range; - -function ownKeys$f(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$f(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$f(Object(source), true).forEach(function (key) { _defineProperty$f(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$f(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$f(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -var PREFIX_LIST = ['Webkit', 'Moz', 'O', 'ms']; -var generatePrefixStyle = function generatePrefixStyle(name, value) { - if (!name) { - return null; - } - - var camelName = name.replace(/(\w)/, function (v) { - return v.toUpperCase(); - }); - var result = PREFIX_LIST.reduce(function (res, entry) { - return _objectSpread$f(_objectSpread$f({}, res), {}, _defineProperty$f({}, entry + camelName, value)); - }, {}); - result[name] = value; - return result; -}; - -function _typeof$b(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$b = function _typeof(obj) { return typeof obj; }; } else { _typeof$b = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$b(obj); } - -function _extends$f() { _extends$f = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$f.apply(this, arguments); } - -function ownKeys$e(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$e(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$e(Object(source), true).forEach(function (key) { _defineProperty$e(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$e(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$e(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$c(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$c(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$c(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$c(Constructor.prototype, protoProps); if (staticProps) _defineProperties$c(Constructor, staticProps); return Constructor; } - -function _inherits$b(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$b(subClass, superClass); } - -function _setPrototypeOf$b(o, p) { _setPrototypeOf$b = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$b(o, p); } - -function _createSuper$b(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$b(); return function _createSuperInternal() { var Super = _getPrototypeOf$b(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$b(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$b(this, result); }; } - -function _possibleConstructorReturn$b(self, call) { if (call && (_typeof$b(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$b(self); } - -function _assertThisInitialized$b(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$b() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$b(o) { _getPrototypeOf$b = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$b(o); } - -var createScale = function createScale(_ref) { - var data = _ref.data, - startIndex = _ref.startIndex, - endIndex = _ref.endIndex, - x = _ref.x, - width = _ref.width, - travellerWidth = _ref.travellerWidth; - - if (!data || !data.length) { - return {}; - } - - var len = data.length; - var scale = point().domain(range_1(0, len)).range([x, x + width - travellerWidth]); - var scaleValues = scale.domain().map(function (entry) { - return scale(entry); - }); - return { - isTextActive: false, - isSlideMoving: false, - isTravellerMoving: false, - startX: scale(startIndex), - endX: scale(endIndex), - scale: scale, - scaleValues: scaleValues - }; -}; - -var isTouch = function isTouch(e) { - return e.changedTouches && !!e.changedTouches.length; -}; - -var Brush = /*#__PURE__*/function (_PureComponent) { - _inherits$b(Brush, _PureComponent); - - var _super = _createSuper$b(Brush); - - function Brush(props) { - var _this; - - _classCallCheck$c(this, Brush); - - _this = _super.call(this, props); - _this.leaveTimer = void 0; - _this.travellerDragStartHandlers = void 0; - - _this.handleDrag = function (e) { - if (_this.leaveTimer) { - clearTimeout(_this.leaveTimer); - _this.leaveTimer = null; - } - - if (_this.state.isTravellerMoving) { - _this.handleTravellerMove(e); - } else if (_this.state.isSlideMoving) { - _this.handleSlideDrag(e); - } - }; - - _this.handleTouchMove = function (e) { - if (e.changedTouches != null && e.changedTouches.length > 0) { - _this.handleDrag(e.changedTouches[0]); - } - }; - - _this.handleDragEnd = function () { - _this.setState({ - isTravellerMoving: false, - isSlideMoving: false - }); - - _this.detachDragEndListener(); - }; - - _this.handleLeaveWrapper = function () { - if (_this.state.isTravellerMoving || _this.state.isSlideMoving) { - _this.leaveTimer = window.setTimeout(_this.handleDragEnd, _this.props.leaveTimeOut); - } - }; - - _this.handleEnterSlideOrTraveller = function () { - _this.setState({ - isTextActive: true - }); - }; - - _this.handleLeaveSlideOrTraveller = function () { - _this.setState({ - isTextActive: false - }); - }; - - _this.handleSlideDragStart = function (e) { - var event = isTouch(e) ? e.changedTouches[0] : e; - - _this.setState({ - isTravellerMoving: false, - isSlideMoving: true, - slideMoveStartX: event.pageX - }); - - _this.attachDragEndListener(); - }; - - _this.travellerDragStartHandlers = { - startX: _this.handleTravellerDragStart.bind(_assertThisInitialized$b(_this), 'startX'), - endX: _this.handleTravellerDragStart.bind(_assertThisInitialized$b(_this), 'endX') - }; - _this.state = {}; - return _this; - } - - _createClass$c(Brush, [{ - key: "componentWillUnmount", - value: function componentWillUnmount() { - if (this.leaveTimer) { - clearTimeout(this.leaveTimer); - this.leaveTimer = null; - } - - this.detachDragEndListener(); - } - }, { - key: "getIndex", - value: function getIndex(_ref2) { - var startX = _ref2.startX, - endX = _ref2.endX; - var scaleValues = this.state.scaleValues; - var _this$props = this.props, - gap = _this$props.gap, - data = _this$props.data; - var lastIndex = data.length - 1; - var min = Math.min(startX, endX); - var max = Math.max(startX, endX); - var minIndex = Brush.getIndexInRange(scaleValues, min); - var maxIndex = Brush.getIndexInRange(scaleValues, max); - return { - startIndex: minIndex - minIndex % gap, - endIndex: maxIndex === lastIndex ? lastIndex : maxIndex - maxIndex % gap - }; - } - }, { - key: "getTextOfTick", - value: function getTextOfTick(index) { - var _this$props2 = this.props, - data = _this$props2.data, - tickFormatter = _this$props2.tickFormatter, - dataKey = _this$props2.dataKey; - var text = getValueByDataKey(data[index], dataKey, index); - return isFunction_1(tickFormatter) ? tickFormatter(text, index) : text; - } - }, { - key: "attachDragEndListener", - value: function attachDragEndListener() { - window.addEventListener('mouseup', this.handleDragEnd, true); - window.addEventListener('touchend', this.handleDragEnd, true); - } - }, { - key: "detachDragEndListener", - value: function detachDragEndListener() { - window.removeEventListener('mouseup', this.handleDragEnd, true); - window.removeEventListener('touchend', this.handleDragEnd, true); - } - }, { - key: "handleSlideDrag", - value: function handleSlideDrag(e) { - var _this$state = this.state, - slideMoveStartX = _this$state.slideMoveStartX, - startX = _this$state.startX, - endX = _this$state.endX; - var _this$props3 = this.props, - x = _this$props3.x, - width = _this$props3.width, - travellerWidth = _this$props3.travellerWidth, - startIndex = _this$props3.startIndex, - endIndex = _this$props3.endIndex, - onChange = _this$props3.onChange; - var delta = e.pageX - slideMoveStartX; - - if (delta > 0) { - delta = Math.min(delta, x + width - travellerWidth - endX, x + width - travellerWidth - startX); - } else if (delta < 0) { - delta = Math.max(delta, x - startX, x - endX); - } - - var newIndex = this.getIndex({ - startX: startX + delta, - endX: endX + delta - }); - - if ((newIndex.startIndex !== startIndex || newIndex.endIndex !== endIndex) && onChange) { - onChange(newIndex); - } - - this.setState({ - startX: startX + delta, - endX: endX + delta, - slideMoveStartX: e.pageX - }); - } - }, { - key: "handleTravellerDragStart", - value: function handleTravellerDragStart(id, e) { - var event = isTouch(e) ? e.changedTouches[0] : e; - this.setState({ - isSlideMoving: false, - isTravellerMoving: true, - movingTravellerId: id, - brushMoveStartX: event.pageX - }); - this.attachDragEndListener(); - } - }, { - key: "handleTravellerMove", - value: function handleTravellerMove(e) { - var _this$setState; - - var _this$state2 = this.state, - brushMoveStartX = _this$state2.brushMoveStartX, - movingTravellerId = _this$state2.movingTravellerId, - endX = _this$state2.endX, - startX = _this$state2.startX; - var prevValue = this.state[movingTravellerId]; - var _this$props4 = this.props, - x = _this$props4.x, - width = _this$props4.width, - travellerWidth = _this$props4.travellerWidth, - onChange = _this$props4.onChange, - gap = _this$props4.gap, - data = _this$props4.data; - var params = { - startX: this.state.startX, - endX: this.state.endX - }; - var delta = e.pageX - brushMoveStartX; - - if (delta > 0) { - delta = Math.min(delta, x + width - travellerWidth - prevValue); - } else if (delta < 0) { - delta = Math.max(delta, x - prevValue); - } - - params[movingTravellerId] = prevValue + delta; - var newIndex = this.getIndex(params); - var startIndex = newIndex.startIndex, - endIndex = newIndex.endIndex; - - var isFullGap = function isFullGap() { - var lastIndex = data.length - 1; - - if (movingTravellerId === 'startX' && (endX > startX ? startIndex % gap === 0 : endIndex % gap === 0) || endX < startX && endIndex === lastIndex || movingTravellerId === 'endX' && (endX > startX ? endIndex % gap === 0 : startIndex % gap === 0) || endX > startX && endIndex === lastIndex) { - return true; - } - - return false; - }; - - this.setState((_this$setState = {}, _defineProperty$e(_this$setState, movingTravellerId, prevValue + delta), _defineProperty$e(_this$setState, "brushMoveStartX", e.pageX), _this$setState), function () { - if (onChange) { - if (isFullGap()) { - onChange(newIndex); - } - } - }); - } - }, { - key: "renderBackground", - value: function renderBackground() { - var _this$props5 = this.props, - x = _this$props5.x, - y = _this$props5.y, - width = _this$props5.width, - height = _this$props5.height, - fill = _this$props5.fill, - stroke = _this$props5.stroke; - return /*#__PURE__*/React__default.createElement("rect", { - stroke: stroke, - fill: fill, - x: x, - y: y, - width: width, - height: height - }); - } - }, { - key: "renderPanorama", - value: function renderPanorama() { - var _this$props6 = this.props, - x = _this$props6.x, - y = _this$props6.y, - width = _this$props6.width, - height = _this$props6.height, - data = _this$props6.data, - children = _this$props6.children, - padding = _this$props6.padding; - var chartElement = Children.only(children); - - if (!chartElement) { - return null; - } - - return /*#__PURE__*/React__default.cloneElement(chartElement, { - x: x, - y: y, - width: width, - height: height, - margin: padding, - compact: true, - data: data - }); - } - }, { - key: "renderTravellerLayer", - value: function renderTravellerLayer(travellerX, id) { - var _this$props7 = this.props, - y = _this$props7.y, - travellerWidth = _this$props7.travellerWidth, - height = _this$props7.height, - traveller = _this$props7.traveller; - var x = Math.max(travellerX, this.props.x); - - var travellerProps = _objectSpread$e(_objectSpread$e({}, filterProps(this.props)), {}, { - x: x, - y: y, - width: travellerWidth, - height: height - }); - - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-brush-traveller", - onMouseEnter: this.handleEnterSlideOrTraveller, - onMouseLeave: this.handleLeaveSlideOrTraveller, - onMouseDown: this.travellerDragStartHandlers[id], - onTouchStart: this.travellerDragStartHandlers[id], - style: { - cursor: 'col-resize' - } - }, Brush.renderTraveller(traveller, travellerProps)); - } - }, { - key: "renderSlide", - value: function renderSlide(startX, endX) { - var _this$props8 = this.props, - y = _this$props8.y, - height = _this$props8.height, - stroke = _this$props8.stroke, - travellerWidth = _this$props8.travellerWidth; - var x = Math.min(startX, endX) + travellerWidth; - var width = Math.max(Math.abs(endX - startX) - travellerWidth, 0); - return /*#__PURE__*/React__default.createElement("rect", { - className: "recharts-brush-slide", - onMouseEnter: this.handleEnterSlideOrTraveller, - onMouseLeave: this.handleLeaveSlideOrTraveller, - onMouseDown: this.handleSlideDragStart, - onTouchStart: this.handleSlideDragStart, - style: { - cursor: 'move' - }, - stroke: "none", - fill: stroke, - fillOpacity: 0.2, - x: x, - y: y, - width: width, - height: height - }); - } - }, { - key: "renderText", - value: function renderText() { - var _this$props9 = this.props, - startIndex = _this$props9.startIndex, - endIndex = _this$props9.endIndex, - y = _this$props9.y, - height = _this$props9.height, - travellerWidth = _this$props9.travellerWidth, - stroke = _this$props9.stroke; - var _this$state3 = this.state, - startX = _this$state3.startX, - endX = _this$state3.endX; - var offset = 5; - var attrs = { - pointerEvents: 'none', - fill: stroke - }; - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-brush-texts" - }, /*#__PURE__*/React__default.createElement(Text, _extends$f({ - textAnchor: "end", - verticalAnchor: "middle", - x: Math.min(startX, endX) - offset, - y: y + height / 2 - }, attrs), this.getTextOfTick(startIndex)), /*#__PURE__*/React__default.createElement(Text, _extends$f({ - textAnchor: "start", - verticalAnchor: "middle", - x: Math.max(startX, endX) + travellerWidth + offset, - y: y + height / 2 - }, attrs), this.getTextOfTick(endIndex))); - } - }, { - key: "render", - value: function render() { - var _this$props10 = this.props, - data = _this$props10.data, - className = _this$props10.className, - children = _this$props10.children, - x = _this$props10.x, - y = _this$props10.y, - width = _this$props10.width, - height = _this$props10.height, - alwaysShowText = _this$props10.alwaysShowText; - var _this$state4 = this.state, - startX = _this$state4.startX, - endX = _this$state4.endX, - isTextActive = _this$state4.isTextActive, - isSlideMoving = _this$state4.isSlideMoving, - isTravellerMoving = _this$state4.isTravellerMoving; - - if (!data || !data.length || !isNumber(x) || !isNumber(y) || !isNumber(width) || !isNumber(height) || width <= 0 || height <= 0) { - return null; - } - - var layerClass = classNames('recharts-brush', className); - var isPanoramic = React__default.Children.count(children) === 1; - var style = generatePrefixStyle('userSelect', 'none'); - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass, - onMouseMove: this.handleDrag, - onMouseLeave: this.handleLeaveWrapper, - onTouchMove: this.handleTouchMove, - style: style - }, this.renderBackground(), isPanoramic && this.renderPanorama(), this.renderSlide(startX, endX), this.renderTravellerLayer(startX, 'startX'), this.renderTravellerLayer(endX, 'endX'), (isTextActive || isSlideMoving || isTravellerMoving || alwaysShowText) && this.renderText()); - } - }], [{ - key: "renderDefaultTraveller", - value: function renderDefaultTraveller(props) { - var x = props.x, - y = props.y, - width = props.width, - height = props.height, - stroke = props.stroke; - var lineY = Math.floor(y + height / 2) - 1; - return /*#__PURE__*/React__default.createElement(React__default.Fragment, null, /*#__PURE__*/React__default.createElement("rect", { - x: x, - y: y, - width: width, - height: height, - fill: stroke, - stroke: "none" - }), /*#__PURE__*/React__default.createElement("line", { - x1: x + 1, - y1: lineY, - x2: x + width - 1, - y2: lineY, - fill: "none", - stroke: "#fff" - }), /*#__PURE__*/React__default.createElement("line", { - x1: x + 1, - y1: lineY + 2, - x2: x + width - 1, - y2: lineY + 2, - fill: "none", - stroke: "#fff" - })); - } - }, { - key: "renderTraveller", - value: function renderTraveller(option, props) { - var rectangle; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - rectangle = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - rectangle = option(props); - } else { - rectangle = Brush.renderDefaultTraveller(props); - } - - return rectangle; - } - }, { - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - var data = nextProps.data, - width = nextProps.width, - x = nextProps.x, - travellerWidth = nextProps.travellerWidth, - updateId = nextProps.updateId, - startIndex = nextProps.startIndex, - endIndex = nextProps.endIndex; - - if (data !== prevState.prevData || updateId !== prevState.prevUpdateId) { - return _objectSpread$e({ - prevData: data, - prevTravellerWidth: travellerWidth, - prevUpdateId: updateId, - prevX: x, - prevWidth: width - }, data && data.length ? createScale({ - data: data, - width: width, - x: x, - travellerWidth: travellerWidth, - startIndex: startIndex, - endIndex: endIndex - }) : { - scale: null, - scaleValues: null - }); - } - - if (prevState.scale && (width !== prevState.prevWidth || x !== prevState.prevX || travellerWidth !== prevState.prevTravellerWidth)) { - prevState.scale.range([x, x + width - travellerWidth]); - var scaleValues = prevState.scale.domain().map(function (entry) { - return prevState.scale(entry); - }); - return { - prevData: data, - prevTravellerWidth: travellerWidth, - prevUpdateId: updateId, - prevX: x, - prevWidth: width, - startX: prevState.scale(nextProps.startIndex), - endX: prevState.scale(nextProps.endIndex), - scaleValues: scaleValues - }; - } - - return null; - } - }, { - key: "getIndexInRange", - value: function getIndexInRange(range, x) { - var len = range.length; - var start = 0; - var end = len - 1; - - while (end - start > 1) { - var middle = Math.floor((start + end) / 2); - - if (range[middle] > x) { - end = middle; - } else { - start = middle; - } - } - - return x >= range[end] ? end : start; - } - }]); - - return Brush; -}(PureComponent); -Brush.displayName = 'Brush'; -Brush.defaultProps = { - height: 40, - travellerWidth: 5, - gap: 1, - fill: '#fff', - stroke: '#666', - padding: { - top: 1, - right: 1, - bottom: 1, - left: 1 - }, - leaveTimeOut: 1000, - alwaysShowText: false -}; - -var baseEach$1 = _baseEach; - -/** - * The base implementation of `_.some` without support for iteratee shorthands. - * - * @private - * @param {Array|Object} collection The collection to iterate over. - * @param {Function} predicate The function invoked per iteration. - * @returns {boolean} Returns `true` if any element passes the predicate check, - * else `false`. - */ -function baseSome$1(collection, predicate) { - var result; - - baseEach$1(collection, function(value, index, collection) { - result = predicate(value, index, collection); - return !result; - }); - return !!result; -} - -var _baseSome = baseSome$1; - -var arraySome = _arraySome, - baseIteratee$5 = _baseIteratee, - baseSome = _baseSome, - isArray$2 = isArray_1, - isIterateeCall$1 = _isIterateeCall; - -/** - * Checks if `predicate` returns truthy for **any** element of `collection`. - * Iteration is stopped once `predicate` returns truthy. The predicate is - * invoked with three arguments: (value, index|key, collection). - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Collection - * @param {Array|Object} collection The collection to iterate over. - * @param {Function} [predicate=_.identity] The function invoked per iteration. - * @param- {Object} [guard] Enables use as an iteratee for methods like `_.map`. - * @returns {boolean} Returns `true` if any element passes the predicate check, - * else `false`. - * @example - * - * _.some([null, 0, 'yes', false], Boolean); - * // => true - * - * var users = [ - * { 'user': 'barney', 'active': true }, - * { 'user': 'fred', 'active': false } - * ]; - * - * // The `_.matches` iteratee shorthand. - * _.some(users, { 'user': 'barney', 'active': false }); - * // => false - * - * // The `_.matchesProperty` iteratee shorthand. - * _.some(users, ['active', false]); - * // => true - * - * // The `_.property` iteratee shorthand. - * _.some(users, 'active'); - * // => true - */ -function some(collection, predicate, guard) { - var func = isArray$2(collection) ? arraySome : baseSome; - if (guard && isIterateeCall$1(collection, predicate, guard)) { - predicate = undefined; - } - return func(collection, baseIteratee$5(predicate)); -} - -var some_1 = some; - -var ifOverflowMatches = function ifOverflowMatches(props, value) { - var alwaysShow = props.alwaysShow; - var ifOverflow = props.ifOverflow; - - if (alwaysShow) { - ifOverflow = 'extendDomain'; - } - - return ifOverflow === value; -}; - -/** - * A specialized version of `_.every` for arrays without support for - * iteratee shorthands. - * - * @private - * @param {Array} [array] The array to iterate over. - * @param {Function} predicate The function invoked per iteration. - * @returns {boolean} Returns `true` if all elements pass the predicate check, - * else `false`. - */ - -function arrayEvery$1(array, predicate) { - var index = -1, - length = array == null ? 0 : array.length; - - while (++index < length) { - if (!predicate(array[index], index, array)) { - return false; - } - } - return true; -} - -var _arrayEvery = arrayEvery$1; - -var baseEach = _baseEach; - -/** - * The base implementation of `_.every` without support for iteratee shorthands. - * - * @private - * @param {Array|Object} collection The collection to iterate over. - * @param {Function} predicate The function invoked per iteration. - * @returns {boolean} Returns `true` if all elements pass the predicate check, - * else `false` - */ -function baseEvery$1(collection, predicate) { - var result = true; - baseEach(collection, function(value, index, collection) { - result = !!predicate(value, index, collection); - return result; - }); - return result; -} - -var _baseEvery = baseEvery$1; - -var arrayEvery = _arrayEvery, - baseEvery = _baseEvery, - baseIteratee$4 = _baseIteratee, - isArray$1 = isArray_1, - isIterateeCall = _isIterateeCall; - -/** - * Checks if `predicate` returns truthy for **all** elements of `collection`. - * Iteration is stopped once `predicate` returns falsey. The predicate is - * invoked with three arguments: (value, index|key, collection). - * - * **Note:** This method returns `true` for - * [empty collections](https://en.wikipedia.org/wiki/Empty_set) because - * [everything is true](https://en.wikipedia.org/wiki/Vacuous_truth) of - * elements of empty collections. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Collection - * @param {Array|Object} collection The collection to iterate over. - * @param {Function} [predicate=_.identity] The function invoked per iteration. - * @param- {Object} [guard] Enables use as an iteratee for methods like `_.map`. - * @returns {boolean} Returns `true` if all elements pass the predicate check, - * else `false`. - * @example - * - * _.every([true, 1, null, 'yes'], Boolean); - * // => false - * - * var users = [ - * { 'user': 'barney', 'age': 36, 'active': false }, - * { 'user': 'fred', 'age': 40, 'active': false } - * ]; - * - * // The `_.matches` iteratee shorthand. - * _.every(users, { 'user': 'barney', 'active': false }); - * // => false - * - * // The `_.matchesProperty` iteratee shorthand. - * _.every(users, ['active', false]); - * // => true - * - * // The `_.property` iteratee shorthand. - * _.every(users, 'active'); - * // => false - */ -function every(collection, predicate, guard) { - var func = isArray$1(collection) ? arrayEvery : baseEvery; - if (guard && isIterateeCall(collection, predicate, guard)) { - predicate = undefined; - } - return func(collection, baseIteratee$4(predicate)); -} - -var every_1 = every; - -var defineProperty = _defineProperty$v; - -/** - * The base implementation of `assignValue` and `assignMergeValue` without - * value checks. - * - * @private - * @param {Object} object The object to modify. - * @param {string} key The key of the property to assign. - * @param {*} value The value to assign. - */ -function baseAssignValue$3(object, key, value) { - if (key == '__proto__' && defineProperty) { - defineProperty(object, key, { - 'configurable': true, - 'enumerable': true, - 'value': value, - 'writable': true - }); - } else { - object[key] = value; - } -} - -var _baseAssignValue = baseAssignValue$3; - -var baseAssignValue$2 = _baseAssignValue, - baseForOwn = _baseForOwn, - baseIteratee$3 = _baseIteratee; - -/** - * Creates an object with the same keys as `object` and values generated - * by running each own enumerable string keyed property of `object` thru - * `iteratee`. The iteratee is invoked with three arguments: - * (value, key, object). - * - * @static - * @memberOf _ - * @since 2.4.0 - * @category Object - * @param {Object} object The object to iterate over. - * @param {Function} [iteratee=_.identity] The function invoked per iteration. - * @returns {Object} Returns the new mapped object. - * @see _.mapKeys - * @example - * - * var users = { - * 'fred': { 'user': 'fred', 'age': 40 }, - * 'pebbles': { 'user': 'pebbles', 'age': 1 } - * }; - * - * _.mapValues(users, function(o) { return o.age; }); - * // => { 'fred': 40, 'pebbles': 1 } (iteration order is not guaranteed) - * - * // The `_.property` iteratee shorthand. - * _.mapValues(users, 'age'); - * // => { 'fred': 40, 'pebbles': 1 } (iteration order is not guaranteed) - */ -function mapValues(object, iteratee) { - var result = {}; - iteratee = baseIteratee$3(iteratee); - - baseForOwn(object, function(value, key, object) { - baseAssignValue$2(result, key, iteratee(value, key, object)); - }); - return result; -} - -var mapValues_1 = mapValues; - -function _classCallCheck$b(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$b(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$b(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$b(Constructor.prototype, protoProps); if (staticProps) _defineProperties$b(Constructor, staticProps); return Constructor; } - -function ownKeys$d(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$d(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$d(Object(source), true).forEach(function (key) { _defineProperty$d(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$d(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$d(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } -/** - * Calculate the scale function, position, width, height of axes - * @param {Object} props Latest props - * @param {Object} axisMap The configuration of axes - * @param {Object} offset The offset of main part in the svg element - * @param {String} axisType The type of axes, x-axis or y-axis - * @param {String} chartName The name of chart - * @return {Object} Configuration - */ - -var formatAxisMap = function formatAxisMap(props, axisMap, offset, axisType, chartName) { - var width = props.width, - height = props.height, - layout = props.layout, - children = props.children; - var ids = Object.keys(axisMap); - var steps = { - left: offset.left, - leftMirror: offset.left, - right: width - offset.right, - rightMirror: width - offset.right, - top: offset.top, - topMirror: offset.top, - bottom: height - offset.bottom, - bottomMirror: height - offset.bottom - }; - var hasBar = !!findChildByType(children, 'Bar'); - return ids.reduce(function (result, id) { - var axis = axisMap[id]; - var orientation = axis.orientation, - domain = axis.domain, - _axis$padding = axis.padding, - padding = _axis$padding === void 0 ? {} : _axis$padding, - mirror = axis.mirror, - reversed = axis.reversed; - var offsetKey = "".concat(orientation).concat(mirror ? 'Mirror' : ''); - var calculatedPadding, range, x, y, needSpace; - - if (axis.type === 'number' && (axis.padding === 'gap' || axis.padding === 'no-gap')) { - var diff = domain[1] - domain[0]; - var smallestDistanceBetweenValues = Infinity; - var sortedValues = axis.categoricalDomain.sort(); - sortedValues.forEach(function (value, index) { - if (index > 0) { - smallestDistanceBetweenValues = Math.min((value || 0) - (sortedValues[index - 1] || 0), smallestDistanceBetweenValues); - } - }); - var smallestDistanceInPercent = smallestDistanceBetweenValues / diff; - var rangeWidth = axis.layout === 'vertical' ? offset.height : offset.width; - - if (axis.padding === 'gap') { - calculatedPadding = smallestDistanceInPercent * rangeWidth / 2; - } - - if (axis.padding === 'no-gap') { - var gap = getPercentValue(props.barCategoryGap, smallestDistanceInPercent * rangeWidth); - var halfBand = smallestDistanceInPercent * rangeWidth / 2; - calculatedPadding = halfBand - gap - (halfBand - gap) / rangeWidth * gap; - } - } - - if (axisType === 'xAxis') { - range = [offset.left + (padding.left || 0) + (calculatedPadding || 0), offset.left + offset.width - (padding.right || 0) - (calculatedPadding || 0)]; - } else if (axisType === 'yAxis') { - range = layout === 'horizontal' ? [offset.top + offset.height - (padding.bottom || 0), offset.top + (padding.top || 0)] : [offset.top + (padding.top || 0) + (calculatedPadding || 0), offset.top + offset.height - (padding.bottom || 0) - (calculatedPadding || 0)]; - } else { - range = axis.range; - } - - if (reversed) { - range = [range[1], range[0]]; - } - - var _parseScale = parseScale(axis, chartName, hasBar), - scale = _parseScale.scale, - realScaleType = _parseScale.realScaleType; - - scale.domain(domain).range(range); - checkDomainOfScale(scale); - var ticks = getTicksOfScale(scale, _objectSpread$d(_objectSpread$d({}, axis), {}, { - realScaleType: realScaleType - })); - - if (axisType === 'xAxis') { - needSpace = orientation === 'top' && !mirror || orientation === 'bottom' && mirror; - x = offset.left; - y = steps[offsetKey] - needSpace * axis.height; - } else if (axisType === 'yAxis') { - needSpace = orientation === 'left' && !mirror || orientation === 'right' && mirror; - x = steps[offsetKey] - needSpace * axis.width; - y = offset.top; - } - - var finalAxis = _objectSpread$d(_objectSpread$d(_objectSpread$d({}, axis), ticks), {}, { - realScaleType: realScaleType, - x: x, - y: y, - scale: scale, - width: axisType === 'xAxis' ? offset.width : axis.width, - height: axisType === 'yAxis' ? offset.height : axis.height - }); - - finalAxis.bandSize = getBandSizeOfAxis(finalAxis, ticks); - - if (!axis.hide && axisType === 'xAxis') { - steps[offsetKey] += (needSpace ? -1 : 1) * finalAxis.height; - } else if (!axis.hide) { - steps[offsetKey] += (needSpace ? -1 : 1) * finalAxis.width; - } - - return _objectSpread$d(_objectSpread$d({}, result), {}, _defineProperty$d({}, id, finalAxis)); - }, {}); -}; -var rectWithPoints = function rectWithPoints(_ref, _ref2) { - var x1 = _ref.x, - y1 = _ref.y; - var x2 = _ref2.x, - y2 = _ref2.y; - return { - x: Math.min(x1, x2), - y: Math.min(y1, y2), - width: Math.abs(x2 - x1), - height: Math.abs(y2 - y1) - }; -}; -/** - * Compute the x, y, width, and height of a box from two reference points. - * @param {Object} coords x1, x2, y1, and y2 - * @return {Object} object - */ - -var rectWithCoords = function rectWithCoords(_ref3) { - var x1 = _ref3.x1, - y1 = _ref3.y1, - x2 = _ref3.x2, - y2 = _ref3.y2; - return rectWithPoints({ - x: x1, - y: y1 - }, { - x: x2, - y: y2 - }); -}; -var ScaleHelper = /*#__PURE__*/function () { - function ScaleHelper(scale) { - _classCallCheck$b(this, ScaleHelper); - - this.scale = void 0; - this.scale = scale; - } - - _createClass$b(ScaleHelper, [{ - key: "domain", - get: function get() { - return this.scale.domain; - } - }, { - key: "range", - get: function get() { - return this.scale.range; - } - }, { - key: "rangeMin", - get: function get() { - return this.range()[0]; - } - }, { - key: "rangeMax", - get: function get() { - return this.range()[1]; - } - }, { - key: "bandwidth", - get: function get() { - return this.scale.bandwidth; - } - }, { - key: "apply", - value: function apply(value) { - var _ref4 = arguments.length > 1 && arguments[1] !== undefined ? arguments[1] : {}, - bandAware = _ref4.bandAware, - position = _ref4.position; - - if (value === undefined) { - return undefined; - } - - if (position) { - switch (position) { - case 'start': - { - return this.scale(value); - } - - case 'middle': - { - var offset = this.bandwidth ? this.bandwidth() / 2 : 0; - return this.scale(value) + offset; - } - - case 'end': - { - var _offset = this.bandwidth ? this.bandwidth() : 0; - - return this.scale(value) + _offset; - } - - default: - { - return this.scale(value); - } - } - } - - if (bandAware) { - var _offset2 = this.bandwidth ? this.bandwidth() / 2 : 0; - - return this.scale(value) + _offset2; - } - - return this.scale(value); - } - }, { - key: "isInRange", - value: function isInRange(value) { - var range = this.range(); - var first = range[0]; - var last = range[range.length - 1]; - return first <= last ? value >= first && value <= last : value >= last && value <= first; - } - }], [{ - key: "create", - value: function create(obj) { - return new ScaleHelper(obj); - } - }]); - - return ScaleHelper; -}(); -ScaleHelper.EPS = 1e-4; -var createLabeledScales = function createLabeledScales(options) { - var scales = Object.keys(options).reduce(function (res, key) { - return _objectSpread$d(_objectSpread$d({}, res), {}, _defineProperty$d({}, key, ScaleHelper.create(options[key]))); - }, {}); - return _objectSpread$d(_objectSpread$d({}, scales), {}, { - apply: function apply(coord) { - var _ref5 = arguments.length > 1 && arguments[1] !== undefined ? arguments[1] : {}, - bandAware = _ref5.bandAware, - position = _ref5.position; - - return mapValues_1(coord, function (value, label) { - return scales[label].apply(value, { - bandAware: bandAware, - position: position - }); - }); - }, - isInRange: function isInRange(coord) { - return every_1(coord, function (value, label) { - return scales[label].isInRange(value); - }); - } - }); -}; - -function ownKeys$c(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$c(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$c(Object(source), true).forEach(function (key) { _defineProperty$c(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$c(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$c(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _slicedToArray$3(arr, i) { return _arrayWithHoles$3(arr) || _iterableToArrayLimit$3(arr, i) || _unsupportedIterableToArray$4(arr, i) || _nonIterableRest$3(); } - -function _nonIterableRest$3() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$4(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$4(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$4(o, minLen); } - -function _arrayLikeToArray$4(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit$3(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$3(arr) { if (Array.isArray(arr)) return arr; } - -function _extends$e() { _extends$e = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$e.apply(this, arguments); } - -var renderLine = function renderLine(option, props) { - var line; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - line = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - line = option(props); - } else { - line = /*#__PURE__*/React__default.createElement("line", _extends$e({}, props, { - className: "recharts-reference-line-line" - })); - } - - return line; -}; // TODO: ScaleHelper - - -var getEndPoints = function getEndPoints(scales, isFixedX, isFixedY, isSegment, props) { - var _props$viewBox = props.viewBox, - x = _props$viewBox.x, - y = _props$viewBox.y, - width = _props$viewBox.width, - height = _props$viewBox.height, - position = props.position; - - if (isFixedY) { - var yCoord = props.y, - orientation = props.yAxis.orientation; - var coord = scales.y.apply(yCoord, { - position: position - }); - - if (ifOverflowMatches(props, 'discard') && !scales.y.isInRange(coord)) { - return null; - } - - var points = [{ - x: x + width, - y: coord - }, { - x: x, - y: coord - }]; - return orientation === 'left' ? points.reverse() : points; - } - - if (isFixedX) { - var xCoord = props.x, - _orientation = props.xAxis.orientation; - - var _coord = scales.x.apply(xCoord, { - position: position - }); - - if (ifOverflowMatches(props, 'discard') && !scales.x.isInRange(_coord)) { - return null; - } - - var _points = [{ - x: _coord, - y: y + height - }, { - x: _coord, - y: y - }]; - return _orientation === 'top' ? _points.reverse() : _points; - } - - if (isSegment) { - var segment = props.segment; - - var _points2 = segment.map(function (p) { - return scales.apply(p, { - position: position - }); - }); - - if (ifOverflowMatches(props, 'discard') && some_1(_points2, function (p) { - return !scales.isInRange(p); - })) { - return null; - } - - return _points2; - } - - return null; -}; - -function ReferenceLine(props) { - var fixedX = props.x, - fixedY = props.y, - segment = props.segment, - xAxis = props.xAxis, - yAxis = props.yAxis, - shape = props.shape, - className = props.className, - alwaysShow = props.alwaysShow, - clipPathId = props.clipPathId; - warn(alwaysShow === undefined, 'The alwaysShow prop is deprecated. Please use ifOverflow="extendDomain" instead.'); - var scales = createLabeledScales({ - x: xAxis.scale, - y: yAxis.scale - }); - var isX = isNumOrStr(fixedX); - var isY = isNumOrStr(fixedY); - var isSegment = segment && segment.length === 2; - var endPoints = getEndPoints(scales, isX, isY, isSegment, props); - - if (!endPoints) { - return null; - } - - var _endPoints = _slicedToArray$3(endPoints, 2), - _endPoints$ = _endPoints[0], - x1 = _endPoints$.x, - y1 = _endPoints$.y, - _endPoints$2 = _endPoints[1], - x2 = _endPoints$2.x, - y2 = _endPoints$2.y; - - var clipPath = ifOverflowMatches(props, 'hidden') ? "url(#".concat(clipPathId, ")") : undefined; - - var lineProps = _objectSpread$c(_objectSpread$c({ - clipPath: clipPath - }, filterProps(props, true)), {}, { - x1: x1, - y1: y1, - x2: x2, - y2: y2 - }); - - return /*#__PURE__*/React__default.createElement(Layer, { - className: classNames('recharts-reference-line', className) - }, renderLine(shape, lineProps), Label.renderCallByParent(props, rectWithCoords({ - x1: x1, - y1: y1, - x2: x2, - y2: y2 - }))); -} -ReferenceLine.displayName = 'ReferenceLine'; -ReferenceLine.defaultProps = { - isFront: false, - ifOverflow: 'discard', - xAxisId: 0, - yAxisId: 0, - fill: 'none', - stroke: '#ccc', - fillOpacity: 1, - strokeWidth: 1, - position: 'middle' -}; - -function _extends$d() { _extends$d = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$d.apply(this, arguments); } - -function ownKeys$b(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$b(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$b(Object(source), true).forEach(function (key) { _defineProperty$b(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$b(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$b(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -var getCoordinate = function getCoordinate(props) { - var x = props.x, - y = props.y, - xAxis = props.xAxis, - yAxis = props.yAxis; - var scales = createLabeledScales({ - x: xAxis.scale, - y: yAxis.scale - }); - var result = scales.apply({ - x: x, - y: y - }, { - bandAware: true - }); - - if (ifOverflowMatches(props, 'discard') && !scales.isInRange(result)) { - return null; - } - - return result; -}; - -function ReferenceDot(props) { - var x = props.x, - y = props.y, - r = props.r, - alwaysShow = props.alwaysShow, - clipPathId = props.clipPathId; - var isX = isNumOrStr(x); - var isY = isNumOrStr(y); - warn(alwaysShow === undefined, 'The alwaysShow prop is deprecated. Please use ifOverflow="extendDomain" instead.'); - - if (!isX || !isY) { - return null; - } - - var coordinate = getCoordinate(props); - - if (!coordinate) { - return null; - } - - var cx = coordinate.x, - cy = coordinate.y; - var shape = props.shape, - className = props.className; - var clipPath = ifOverflowMatches(props, 'hidden') ? "url(#".concat(clipPathId, ")") : undefined; - - var dotProps = _objectSpread$b(_objectSpread$b({ - clipPath: clipPath - }, filterProps(props, true)), {}, { - cx: cx, - cy: cy - }); - - return /*#__PURE__*/React__default.createElement(Layer, { - className: classNames('recharts-reference-dot', className) - }, ReferenceDot.renderDot(shape, dotProps), Label.renderCallByParent(props, { - x: cx - r, - y: cy - r, - width: 2 * r, - height: 2 * r - })); -} -ReferenceDot.displayName = 'ReferenceDot'; -ReferenceDot.defaultProps = { - isFront: false, - ifOverflow: 'discard', - xAxisId: 0, - yAxisId: 0, - r: 10, - fill: '#fff', - stroke: '#ccc', - fillOpacity: 1, - strokeWidth: 1 -}; - -ReferenceDot.renderDot = function (option, props) { - var dot; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - dot = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - dot = option(props); - } else { - dot = /*#__PURE__*/React__default.createElement(Dot, _extends$d({}, props, { - cx: props.cx, - cy: props.cy, - className: "recharts-reference-dot-dot" - })); - } - - return dot; -}; - -function _extends$c() { _extends$c = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$c.apply(this, arguments); } - -function ownKeys$a(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$a(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$a(Object(source), true).forEach(function (key) { _defineProperty$a(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$a(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$a(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -var getRect = function getRect(hasX1, hasX2, hasY1, hasY2, props) { - var xValue1 = props.x1, - xValue2 = props.x2, - yValue1 = props.y1, - yValue2 = props.y2, - xAxis = props.xAxis, - yAxis = props.yAxis; - if (!xAxis || !yAxis) return null; - var scales = createLabeledScales({ - x: xAxis.scale, - y: yAxis.scale - }); - var p1 = { - x: hasX1 ? scales.x.apply(xValue1, { - position: 'start' - }) : scales.x.rangeMin, - y: hasY1 ? scales.y.apply(yValue1, { - position: 'start' - }) : scales.y.rangeMin - }; - var p2 = { - x: hasX2 ? scales.x.apply(xValue2, { - position: 'end' - }) : scales.x.rangeMax, - y: hasY2 ? scales.y.apply(yValue2, { - position: 'end' - }) : scales.y.rangeMax - }; - - if (ifOverflowMatches(props, 'discard') && (!scales.isInRange(p1) || !scales.isInRange(p2))) { - return null; - } - - return rectWithPoints(p1, p2); -}; - -function ReferenceArea(props) { - var x1 = props.x1, - x2 = props.x2, - y1 = props.y1, - y2 = props.y2, - className = props.className, - alwaysShow = props.alwaysShow, - clipPathId = props.clipPathId; - warn(alwaysShow === undefined, 'The alwaysShow prop is deprecated. Please use ifOverflow="extendDomain" instead.'); - var hasX1 = isNumOrStr(x1); - var hasX2 = isNumOrStr(x2); - var hasY1 = isNumOrStr(y1); - var hasY2 = isNumOrStr(y2); - var shape = props.shape; - - if (!hasX1 && !hasX2 && !hasY1 && !hasY2 && !shape) { - return null; - } - - var rect = getRect(hasX1, hasX2, hasY1, hasY2, props); - - if (!rect && !shape) { - return null; - } - - var clipPath = ifOverflowMatches(props, 'hidden') ? "url(#".concat(clipPathId, ")") : undefined; - return /*#__PURE__*/React__default.createElement(Layer, { - className: classNames('recharts-reference-area', className) - }, ReferenceArea.renderRect(shape, _objectSpread$a(_objectSpread$a({ - clipPath: clipPath - }, filterProps(props, true)), rect)), Label.renderCallByParent(props, rect)); -} -ReferenceArea.displayName = 'ReferenceArea'; -ReferenceArea.defaultProps = { - isFront: false, - ifOverflow: 'discard', - xAxisId: 0, - yAxisId: 0, - r: 10, - fill: '#ccc', - fillOpacity: 0.5, - stroke: 'none', - strokeWidth: 1 -}; - -ReferenceArea.renderRect = function (option, props) { - var rect; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - rect = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - rect = option(props); - } else { - rect = /*#__PURE__*/React__default.createElement(Rectangle, _extends$c({}, props, { - className: "recharts-reference-area-rect" - })); - } - - return rect; -}; - -function _typeof$a(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$a = function _typeof(obj) { return typeof obj; }; } else { _typeof$a = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$a(obj); } - -function _extends$b() { _extends$b = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$b.apply(this, arguments); } - -function ownKeys$9(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$9(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$9(Object(source), true).forEach(function (key) { _defineProperty$9(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$9(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$9(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _objectWithoutProperties$8(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$8(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$8(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _classCallCheck$a(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$a(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$a(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$a(Constructor.prototype, protoProps); if (staticProps) _defineProperties$a(Constructor, staticProps); return Constructor; } - -function _inherits$a(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$a(subClass, superClass); } - -function _setPrototypeOf$a(o, p) { _setPrototypeOf$a = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$a(o, p); } - -function _createSuper$a(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$a(); return function _createSuperInternal() { var Super = _getPrototypeOf$a(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$a(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$a(this, result); }; } - -function _possibleConstructorReturn$a(self, call) { if (call && (_typeof$a(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$a(self); } - -function _assertThisInitialized$a(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$a() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$a(o) { _getPrototypeOf$a = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$a(o); } -var CartesianAxis = /*#__PURE__*/function (_Component) { - _inherits$a(CartesianAxis, _Component); - - var _super = _createSuper$a(CartesianAxis); - - function CartesianAxis(props) { - var _this; - - _classCallCheck$a(this, CartesianAxis); - - _this = _super.call(this, props); - _this.layerReference = void 0; - _this.state = { - fontSize: '', - letterSpacing: '' - }; - return _this; - } // todo Array - - - _createClass$a(CartesianAxis, [{ - key: "shouldComponentUpdate", - value: function shouldComponentUpdate(_ref, nextState) { - var viewBox = _ref.viewBox, - restProps = _objectWithoutProperties$8(_ref, ["viewBox"]); - - // props.viewBox is sometimes generated every time - - // check that specially as object equality is likely to fail - var _this$props = this.props, - viewBoxOld = _this$props.viewBox, - restPropsOld = _objectWithoutProperties$8(_this$props, ["viewBox"]); - - return !shallowEqual(viewBox, viewBoxOld) || !shallowEqual(restProps, restPropsOld) || !shallowEqual(nextState, this.state); - } - }, { - key: "componentDidMount", - value: function componentDidMount() { - var htmlLayer = this.layerReference; - if (!htmlLayer) return; - var tick = htmlLayer.getElementsByClassName('recharts-cartesian-axis-tick-value')[0]; - - if (tick) { - this.setState({ - fontSize: window.getComputedStyle(tick).fontSize, - letterSpacing: window.getComputedStyle(tick).letterSpacing - }); - } - } - /** - * Calculate the coordinates of endpoints in ticks - * @param {Object} data The data of a simple tick - * @return {Object} (x1, y1): The coordinate of endpoint close to tick text - * (x2, y2): The coordinate of endpoint close to axis - */ - - }, { - key: "getTickLineCoord", - value: function getTickLineCoord(data) { - var _this$props2 = this.props, - x = _this$props2.x, - y = _this$props2.y, - width = _this$props2.width, - height = _this$props2.height, - orientation = _this$props2.orientation, - tickSize = _this$props2.tickSize, - mirror = _this$props2.mirror, - tickMargin = _this$props2.tickMargin; - var x1, x2, y1, y2, tx, ty; - var sign = mirror ? -1 : 1; - var finalTickSize = data.tickSize || tickSize; - var tickCoord = isNumber(data.tickCoord) ? data.tickCoord : data.coordinate; - - switch (orientation) { - case 'top': - x1 = x2 = data.coordinate; - y2 = y + +!mirror * height; - y1 = y2 - sign * finalTickSize; - ty = y1 - sign * tickMargin; - tx = tickCoord; - break; - - case 'left': - y1 = y2 = data.coordinate; - x2 = x + +!mirror * width; - x1 = x2 - sign * finalTickSize; - tx = x1 - sign * tickMargin; - ty = tickCoord; - break; - - case 'right': - y1 = y2 = data.coordinate; - x2 = x + +mirror * width; - x1 = x2 + sign * finalTickSize; - tx = x1 + sign * tickMargin; - ty = tickCoord; - break; - - default: - x1 = x2 = data.coordinate; - y2 = y + +mirror * height; - y1 = y2 + sign * finalTickSize; - ty = y1 + sign * tickMargin; - tx = tickCoord; - break; - } - - return { - line: { - x1: x1, - y1: y1, - x2: x2, - y2: y2 - }, - tick: { - x: tx, - y: ty - } - }; - } - }, { - key: "getTickTextAnchor", - value: function getTickTextAnchor() { - var _this$props3 = this.props, - orientation = _this$props3.orientation, - mirror = _this$props3.mirror; - var textAnchor; - - switch (orientation) { - case 'left': - textAnchor = mirror ? 'start' : 'end'; - break; - - case 'right': - textAnchor = mirror ? 'end' : 'start'; - break; - - default: - textAnchor = 'middle'; - break; - } - - return textAnchor; - } - }, { - key: "getTickVerticalAnchor", - value: function getTickVerticalAnchor() { - var _this$props4 = this.props, - orientation = _this$props4.orientation, - mirror = _this$props4.mirror; - var verticalAnchor = 'end'; - - switch (orientation) { - case 'left': - case 'right': - verticalAnchor = 'middle'; - break; - - case 'top': - verticalAnchor = mirror ? 'start' : 'end'; - break; - - default: - verticalAnchor = mirror ? 'end' : 'start'; - break; - } - - return verticalAnchor; - } - }, { - key: "renderAxisLine", - value: function renderAxisLine() { - var _this$props5 = this.props, - x = _this$props5.x, - y = _this$props5.y, - width = _this$props5.width, - height = _this$props5.height, - orientation = _this$props5.orientation, - mirror = _this$props5.mirror, - axisLine = _this$props5.axisLine; - - var props = _objectSpread$9(_objectSpread$9(_objectSpread$9({}, filterProps(this.props)), filterProps(axisLine)), {}, { - fill: 'none' - }); - - if (orientation === 'top' || orientation === 'bottom') { - var needHeight = +(orientation === 'top' && !mirror || orientation === 'bottom' && mirror); - props = _objectSpread$9(_objectSpread$9({}, props), {}, { - x1: x, - y1: y + needHeight * height, - x2: x + width, - y2: y + needHeight * height - }); - } else { - var needWidth = +(orientation === 'left' && !mirror || orientation === 'right' && mirror); - props = _objectSpread$9(_objectSpread$9({}, props), {}, { - x1: x + needWidth * width, - y1: y, - x2: x + needWidth * width, - y2: y + height - }); - } - - return /*#__PURE__*/React__default.createElement("line", _extends$b({}, props, { - className: classNames('recharts-cartesian-axis-line', get_1(axisLine, 'className')) - })); - } - }, { - key: "renderTicks", - value: - /** - * render the ticks - * @param {Array} ticks The ticks to actually render (overrides what was passed in props) - * @param {string} fontSize Fontsize to consider for tick spacing - * @param {string} letterSpacing Letterspacing to consider for tick spacing - * @return {ReactComponent} renderedTicks - */ - function renderTicks(ticks, fontSize, letterSpacing) { - var _this2 = this; - - var _this$props6 = this.props, - tickLine = _this$props6.tickLine, - stroke = _this$props6.stroke, - tick = _this$props6.tick, - tickFormatter = _this$props6.tickFormatter, - unit = _this$props6.unit; - var finalTicks = CartesianAxis.getTicks(_objectSpread$9(_objectSpread$9({}, this.props), {}, { - ticks: ticks - }), fontSize, letterSpacing); - var textAnchor = this.getTickTextAnchor(); - var verticalAnchor = this.getTickVerticalAnchor(); - var axisProps = filterProps(this.props); - var customTickProps = filterProps(tick); - - var tickLineProps = _objectSpread$9(_objectSpread$9({}, axisProps), {}, { - fill: 'none' - }, filterProps(tickLine)); - - var items = finalTicks.map(function (entry, i) { - var _this2$getTickLineCoo = _this2.getTickLineCoord(entry), - lineCoord = _this2$getTickLineCoo.line, - tickCoord = _this2$getTickLineCoo.tick; - - var tickProps = _objectSpread$9(_objectSpread$9(_objectSpread$9(_objectSpread$9({ - textAnchor: textAnchor, - verticalAnchor: verticalAnchor - }, axisProps), {}, { - stroke: 'none', - fill: stroke - }, customTickProps), tickCoord), {}, { - index: i, - payload: entry, - visibleTicksCount: finalTicks.length, - tickFormatter: tickFormatter - }); - - return /*#__PURE__*/React__default.createElement(Layer, _extends$b({ - className: "recharts-cartesian-axis-tick", - key: "tick-".concat(i) // eslint-disable-line react/no-array-index-key - - }, adaptEventsOfChild(_this2.props, entry, i)), tickLine && /*#__PURE__*/React__default.createElement("line", _extends$b({}, tickLineProps, lineCoord, { - className: classNames('recharts-cartesian-axis-tick-line', get_1(tickLine, 'className')) - })), tick && CartesianAxis.renderTickItem(tick, tickProps, "".concat(isFunction_1(tickFormatter) ? tickFormatter(entry.value, i) : entry.value).concat(unit || ''))); - }); - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-cartesian-axis-ticks" - }, items); - } - }, { - key: "render", - value: function render() { - var _this3 = this; - - var _this$props7 = this.props, - axisLine = _this$props7.axisLine, - width = _this$props7.width, - height = _this$props7.height, - ticksGenerator = _this$props7.ticksGenerator, - className = _this$props7.className, - hide = _this$props7.hide; - - if (hide) { - return null; - } - - var _this$props8 = this.props, - ticks = _this$props8.ticks, - noTicksProps = _objectWithoutProperties$8(_this$props8, ["ticks"]); - - var finalTicks = ticks; - - if (isFunction_1(ticksGenerator)) { - finalTicks = ticks && ticks.length > 0 ? ticksGenerator(this.props) : ticksGenerator(noTicksProps); - } - - if (width <= 0 || height <= 0 || !finalTicks || !finalTicks.length) { - return null; - } - - return /*#__PURE__*/React__default.createElement(Layer, { - className: classNames('recharts-cartesian-axis', className), - ref: function ref(_ref2) { - _this3.layerReference = _ref2; - } - }, axisLine && this.renderAxisLine(), this.renderTicks(finalTicks, this.state.fontSize, this.state.letterSpacing), Label.renderCallByParent(this.props)); - } - }], [{ - key: "getTicks", - value: function getTicks(props, fontSize, letterSpacing) { - var tick = props.tick, - ticks = props.ticks, - viewBox = props.viewBox, - minTickGap = props.minTickGap, - orientation = props.orientation, - interval = props.interval, - tickFormatter = props.tickFormatter, - unit = props.unit; - - if (!ticks || !ticks.length || !tick) { - return []; - } - - if (isNumber(interval) || Global.isSsr) { - return CartesianAxis.getNumberIntervalTicks(ticks, typeof interval === 'number' && isNumber(interval) ? interval : 0); - } - - if (interval === 'preserveStartEnd') { - return CartesianAxis.getTicksStart({ - ticks: ticks, - tickFormatter: tickFormatter, - viewBox: viewBox, - orientation: orientation, - minTickGap: minTickGap, - unit: unit, - fontSize: fontSize, - letterSpacing: letterSpacing - }, true); - } - - if (interval === 'preserveStart') { - return CartesianAxis.getTicksStart({ - ticks: ticks, - tickFormatter: tickFormatter, - viewBox: viewBox, - orientation: orientation, - minTickGap: minTickGap, - unit: unit, - fontSize: fontSize, - letterSpacing: letterSpacing - }); - } - - return CartesianAxis.getTicksEnd({ - ticks: ticks, - tickFormatter: tickFormatter, - viewBox: viewBox, - orientation: orientation, - minTickGap: minTickGap, - unit: unit, - fontSize: fontSize, - letterSpacing: letterSpacing - }); - } - }, { - key: "getNumberIntervalTicks", - value: function getNumberIntervalTicks(ticks, interval) { - return ticks.filter(function (entry, i) { - return i % (interval + 1) === 0; - }); - } - }, { - key: "getTicksStart", - value: function getTicksStart(_ref3, preserveEnd) { - var ticks = _ref3.ticks, - tickFormatter = _ref3.tickFormatter, - viewBox = _ref3.viewBox, - orientation = _ref3.orientation, - minTickGap = _ref3.minTickGap, - unit = _ref3.unit, - fontSize = _ref3.fontSize, - letterSpacing = _ref3.letterSpacing; - var x = viewBox.x, - y = viewBox.y, - width = viewBox.width, - height = viewBox.height; - var sizeKey = orientation === 'top' || orientation === 'bottom' ? 'width' : 'height'; - var result = (ticks || []).slice(); // we need add the width of 'unit' only when sizeKey === 'width' - - var unitSize = unit && sizeKey === 'width' ? getStringSize(unit, { - fontSize: fontSize, - letterSpacing: letterSpacing - })[sizeKey] : 0; - var len = result.length; - var sign = len >= 2 ? mathSign(result[1].coordinate - result[0].coordinate) : 1; - var start, end; - - if (sign === 1) { - start = sizeKey === 'width' ? x : y; - end = sizeKey === 'width' ? x + width : y + height; - } else { - start = sizeKey === 'width' ? x + width : y + height; - end = sizeKey === 'width' ? x : y; - } - - if (preserveEnd) { - // Try to guarantee the tail to be displayed - var tail = ticks[len - 1]; - var tailContent = isFunction_1(tickFormatter) ? tickFormatter(tail.value, len - 1) : tail.value; - var tailSize = getStringSize(tailContent, { - fontSize: fontSize, - letterSpacing: letterSpacing - })[sizeKey] + unitSize; - var tailGap = sign * (tail.coordinate + sign * tailSize / 2 - end); - result[len - 1] = tail = _objectSpread$9(_objectSpread$9({}, tail), {}, { - tickCoord: tailGap > 0 ? tail.coordinate - tailGap * sign : tail.coordinate - }); - var isTailShow = sign * (tail.tickCoord - sign * tailSize / 2 - start) >= 0 && sign * (tail.tickCoord + sign * tailSize / 2 - end) <= 0; - - if (isTailShow) { - end = tail.tickCoord - sign * (tailSize / 2 + minTickGap); - result[len - 1] = _objectSpread$9(_objectSpread$9({}, tail), {}, { - isShow: true - }); - } - } - - var count = preserveEnd ? len - 1 : len; - - for (var i = 0; i < count; i++) { - var entry = result[i]; - var content = isFunction_1(tickFormatter) ? tickFormatter(entry.value, i) : entry.value; - var size = getStringSize(content, { - fontSize: fontSize, - letterSpacing: letterSpacing - })[sizeKey] + unitSize; - - if (i === 0) { - var gap = sign * (entry.coordinate - sign * size / 2 - start); - result[i] = entry = _objectSpread$9(_objectSpread$9({}, entry), {}, { - tickCoord: gap < 0 ? entry.coordinate - gap * sign : entry.coordinate - }); - } else { - result[i] = entry = _objectSpread$9(_objectSpread$9({}, entry), {}, { - tickCoord: entry.coordinate - }); - } - - var isShow = sign * (entry.tickCoord - sign * size / 2 - start) >= 0 && sign * (entry.tickCoord + sign * size / 2 - end) <= 0; - - if (isShow) { - start = entry.tickCoord + sign * (size / 2 + minTickGap); - result[i] = _objectSpread$9(_objectSpread$9({}, entry), {}, { - isShow: true - }); - } - } - - return result.filter(function (entry) { - return entry.isShow; - }); - } - }, { - key: "getTicksEnd", - value: function getTicksEnd(_ref4) { - var ticks = _ref4.ticks, - tickFormatter = _ref4.tickFormatter, - viewBox = _ref4.viewBox, - orientation = _ref4.orientation, - minTickGap = _ref4.minTickGap, - unit = _ref4.unit, - fontSize = _ref4.fontSize, - letterSpacing = _ref4.letterSpacing; - var x = viewBox.x, - y = viewBox.y, - width = viewBox.width, - height = viewBox.height; - var sizeKey = orientation === 'top' || orientation === 'bottom' ? 'width' : 'height'; // we need add the width of 'unit' only when sizeKey === 'width' - - var unitSize = unit && sizeKey === 'width' ? getStringSize(unit, { - fontSize: fontSize, - letterSpacing: letterSpacing - })[sizeKey] : 0; - var result = (ticks || []).slice(); - var len = result.length; - var sign = len >= 2 ? mathSign(result[1].coordinate - result[0].coordinate) : 1; - var start, end; - - if (sign === 1) { - start = sizeKey === 'width' ? x : y; - end = sizeKey === 'width' ? x + width : y + height; - } else { - start = sizeKey === 'width' ? x + width : y + height; - end = sizeKey === 'width' ? x : y; - } - - for (var i = len - 1; i >= 0; i--) { - var entry = result[i]; - var content = isFunction_1(tickFormatter) ? tickFormatter(entry.value, len - i - 1) : entry.value; - var size = getStringSize(content, { - fontSize: fontSize, - letterSpacing: letterSpacing - })[sizeKey] + unitSize; - - if (i === len - 1) { - var gap = sign * (entry.coordinate + sign * size / 2 - end); - result[i] = entry = _objectSpread$9(_objectSpread$9({}, entry), {}, { - tickCoord: gap > 0 ? entry.coordinate - gap * sign : entry.coordinate - }); - } else { - result[i] = entry = _objectSpread$9(_objectSpread$9({}, entry), {}, { - tickCoord: entry.coordinate - }); - } - - var isShow = sign * (entry.tickCoord - sign * size / 2 - start) >= 0 && sign * (entry.tickCoord + sign * size / 2 - end) <= 0; - - if (isShow) { - end = entry.tickCoord - sign * (size / 2 + minTickGap); - result[i] = _objectSpread$9(_objectSpread$9({}, entry), {}, { - isShow: true - }); - } - } - - return result.filter(function (entry) { - return entry.isShow; - }); - } - }, { - key: "renderTickItem", - value: function renderTickItem(option, props, value) { - var tickItem; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - tickItem = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - tickItem = option(props); - } else { - tickItem = /*#__PURE__*/React__default.createElement(Text, _extends$b({}, props, { - className: "recharts-cartesian-axis-tick-value" - }), value); - } - - return tickItem; - } - }]); - - return CartesianAxis; -}(Component); -CartesianAxis.displayName = 'CartesianAxis'; -CartesianAxis.defaultProps = { - x: 0, - y: 0, - width: 0, - height: 0, - viewBox: { - x: 0, - y: 0, - width: 0, - height: 0 - }, - // The orientation of axis - orientation: 'bottom', - // The ticks - ticks: [], - stroke: '#666', - tickLine: true, - axisLine: true, - tick: true, - mirror: false, - minTickGap: 5, - // The width or height of tick - tickSize: 6, - tickMargin: 2, - interval: 'preserveEnd' -}; - -function _typeof$9(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$9 = function _typeof(obj) { return typeof obj; }; } else { _typeof$9 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$9(obj); } - -function _extends$a() { _extends$a = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$a.apply(this, arguments); } - -function _objectWithoutProperties$7(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$7(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$7(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function ownKeys$8(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$8(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$8(Object(source), true).forEach(function (key) { _defineProperty$8(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$8(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$8(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$9(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$9(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$9(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$9(Constructor.prototype, protoProps); if (staticProps) _defineProperties$9(Constructor, staticProps); return Constructor; } - -function _inherits$9(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$9(subClass, superClass); } - -function _setPrototypeOf$9(o, p) { _setPrototypeOf$9 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$9(o, p); } - -function _createSuper$9(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$9(); return function _createSuperInternal() { var Super = _getPrototypeOf$9(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$9(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$9(this, result); }; } - -function _possibleConstructorReturn$9(self, call) { if (call && (_typeof$9(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$9(self); } - -function _assertThisInitialized$9(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$9() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$9(o) { _getPrototypeOf$9 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$9(o); } -var CartesianGrid = /*#__PURE__*/function (_PureComponent) { - _inherits$9(CartesianGrid, _PureComponent); - - var _super = _createSuper$9(CartesianGrid); - - function CartesianGrid() { - _classCallCheck$9(this, CartesianGrid); - - return _super.apply(this, arguments); - } - - _createClass$9(CartesianGrid, [{ - key: "renderHorizontal", - value: - /** - * Draw the horizontal grid lines - * @param {Array} horizontalPoints either passed in as props or generated from function - * @return {Group} Horizontal lines - */ - function renderHorizontal(horizontalPoints) { - var _this = this; - - var _this$props = this.props, - x = _this$props.x, - width = _this$props.width, - horizontal = _this$props.horizontal; - - if (!horizontalPoints || !horizontalPoints.length) { - return null; - } - - var items = horizontalPoints.map(function (entry, i) { - var props = _objectSpread$8(_objectSpread$8({}, _this.props), {}, { - x1: x, - y1: entry, - x2: x + width, - y2: entry, - key: "line-".concat(i), - index: i - }); - - return CartesianGrid.renderLineItem(horizontal, props); - }); - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-cartesian-grid-horizontal" - }, items); - } - /** - * Draw vertical grid lines - * @param {Array} verticalPoints either passed in as props or generated from function - * @return {Group} Vertical lines - */ - - }, { - key: "renderVertical", - value: function renderVertical(verticalPoints) { - var _this2 = this; - - var _this$props2 = this.props, - y = _this$props2.y, - height = _this$props2.height, - vertical = _this$props2.vertical; - - if (!verticalPoints || !verticalPoints.length) { - return null; - } - - var items = verticalPoints.map(function (entry, i) { - var props = _objectSpread$8(_objectSpread$8({}, _this2.props), {}, { - x1: entry, - y1: y, - x2: entry, - y2: y + height, - key: "line-".concat(i), - index: i - }); - - return CartesianGrid.renderLineItem(vertical, props); - }); - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-cartesian-grid-vertical" - }, items); - } - /** - * Draw vertical grid stripes filled by colors - * @param {Array} verticalPoints either passed in as props or generated from function - * @return {Group} Vertical stripes - */ - - }, { - key: "renderVerticalStripes", - value: function renderVerticalStripes(verticalPoints) { - var verticalFill = this.props.verticalFill; - - if (!verticalFill || !verticalFill.length) { - return null; - } - - var _this$props3 = this.props, - fillOpacity = _this$props3.fillOpacity, - x = _this$props3.x, - y = _this$props3.y, - width = _this$props3.width, - height = _this$props3.height; - var verticalPointsUpdated = verticalPoints.slice().sort(function (a, b) { - return a - b; - }); - - if (x !== verticalPointsUpdated[0]) { - verticalPointsUpdated.unshift(0); - } - - var items = verticalPointsUpdated.map(function (entry, i) { - var lineWidth = verticalPointsUpdated[i + 1] ? verticalPointsUpdated[i + 1] - entry : x + width - entry; - - if (lineWidth <= 0) { - return null; - } - - var colorIndex = i % verticalFill.length; - return /*#__PURE__*/React__default.createElement("rect", { - key: "react-".concat(i) // eslint-disable-line react/no-array-index-key - , - x: Math.round(entry + x - x), - y: y, - width: lineWidth, - height: height, - stroke: "none", - fill: verticalFill[colorIndex], - fillOpacity: fillOpacity, - className: "recharts-cartesian-grid-bg" - }); - }); - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-cartesian-gridstripes-vertical" - }, items); - } - /** - * Draw horizontal grid stripes filled by colors - * @param {Array} horizontalPoints either passed in as props or generated from function - * @return {Group} Horizontal stripes - */ - - }, { - key: "renderHorizontalStripes", - value: function renderHorizontalStripes(horizontalPoints) { - var horizontalFill = this.props.horizontalFill; - - if (!horizontalFill || !horizontalFill.length) { - return null; - } - - var _this$props4 = this.props, - fillOpacity = _this$props4.fillOpacity, - x = _this$props4.x, - y = _this$props4.y, - width = _this$props4.width, - height = _this$props4.height; - var horizontalPointsUpdated = horizontalPoints.slice().sort(function (a, b) { - return a - b; - }); - - if (y !== horizontalPointsUpdated[0]) { - horizontalPointsUpdated.unshift(0); - } - - var items = horizontalPointsUpdated.map(function (entry, i) { - var lineHeight = horizontalPointsUpdated[i + 1] ? horizontalPointsUpdated[i + 1] - entry : y + height - entry; - - if (lineHeight <= 0) { - return null; - } - - var colorIndex = i % horizontalFill.length; - return /*#__PURE__*/React__default.createElement("rect", { - key: "react-".concat(i) // eslint-disable-line react/no-array-index-key - , - y: Math.round(entry + y - y), - x: x, - height: lineHeight, - width: width, - stroke: "none", - fill: horizontalFill[colorIndex], - fillOpacity: fillOpacity, - className: "recharts-cartesian-grid-bg" - }); - }); - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-cartesian-gridstripes-horizontal" - }, items); - } - }, { - key: "renderBackground", - value: function renderBackground() { - var fill = this.props.fill; - - if (!fill || fill === 'none') { - return null; - } - - var _this$props5 = this.props, - fillOpacity = _this$props5.fillOpacity, - x = _this$props5.x, - y = _this$props5.y, - width = _this$props5.width, - height = _this$props5.height; - return /*#__PURE__*/React__default.createElement("rect", { - x: x, - y: y, - width: width, - height: height, - stroke: "none", - fill: fill, - fillOpacity: fillOpacity, - className: "recharts-cartesian-grid-bg" - }); - } - }, { - key: "render", - value: function render() { - var _this$props6 = this.props, - x = _this$props6.x, - y = _this$props6.y, - width = _this$props6.width, - height = _this$props6.height, - horizontal = _this$props6.horizontal, - vertical = _this$props6.vertical, - horizontalCoordinatesGenerator = _this$props6.horizontalCoordinatesGenerator, - verticalCoordinatesGenerator = _this$props6.verticalCoordinatesGenerator, - xAxis = _this$props6.xAxis, - yAxis = _this$props6.yAxis, - offset = _this$props6.offset, - chartWidth = _this$props6.chartWidth, - chartHeight = _this$props6.chartHeight; - - if (!isNumber(width) || width <= 0 || !isNumber(height) || height <= 0 || !isNumber(x) || x !== +x || !isNumber(y) || y !== +y) { - return null; - } - - var _this$props7 = this.props, - horizontalPoints = _this$props7.horizontalPoints, - verticalPoints = _this$props7.verticalPoints; // No horizontal points are specified - - if ((!horizontalPoints || !horizontalPoints.length) && isFunction_1(horizontalCoordinatesGenerator)) { - horizontalPoints = horizontalCoordinatesGenerator({ - yAxis: yAxis, - width: chartWidth, - height: chartHeight, - offset: offset - }); - } // No vertical points are specified - - - if ((!verticalPoints || !verticalPoints.length) && isFunction_1(verticalCoordinatesGenerator)) { - verticalPoints = verticalCoordinatesGenerator({ - xAxis: xAxis, - width: chartWidth, - height: chartHeight, - offset: offset - }); - } - - return /*#__PURE__*/React__default.createElement("g", { - className: "recharts-cartesian-grid" - }, this.renderBackground(), horizontal && this.renderHorizontal(horizontalPoints), vertical && this.renderVertical(verticalPoints), horizontal && this.renderHorizontalStripes(horizontalPoints), vertical && this.renderVerticalStripes(verticalPoints)); - } - }], [{ - key: "renderLineItem", - value: function renderLineItem(option, props) { - var lineItem; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - lineItem = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - lineItem = option(props); - } else { - var x1 = props.x1, - y1 = props.y1, - x2 = props.x2, - y2 = props.y2, - key = props.key, - others = _objectWithoutProperties$7(props, ["x1", "y1", "x2", "y2", "key"]); - - lineItem = /*#__PURE__*/React__default.createElement("line", _extends$a({}, filterProps(others), { - x1: x1, - y1: y1, - x2: x2, - y2: y2, - fill: "none", - key: key - })); - } - - return lineItem; - } - }]); - - return CartesianGrid; -}(PureComponent); -CartesianGrid.displayName = 'CartesianGrid'; -CartesianGrid.defaultProps = { - horizontal: true, - vertical: true, - // The ordinates of horizontal grid lines - horizontalPoints: [], - // The abscissas of vertical grid lines - verticalPoints: [], - stroke: '#ccc', - fill: 'none', - // The fill of colors of grid lines - verticalFill: [], - horizontalFill: [] -}; - -function _extends$9() { _extends$9 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$9.apply(this, arguments); } - -function _slicedToArray$2(arr, i) { return _arrayWithHoles$2(arr) || _iterableToArrayLimit$2(arr, i) || _unsupportedIterableToArray$3(arr, i) || _nonIterableRest$2(); } - -function _nonIterableRest$2() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$3(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$3(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$3(o, minLen); } - -function _arrayLikeToArray$3(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit$2(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$2(arr) { if (Array.isArray(arr)) return arr; } - -function _objectWithoutProperties$6(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$6(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$6(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } -function ErrorBar(props) { - var offset = props.offset, - layout = props.layout, - width = props.width, - dataKey = props.dataKey, - data = props.data, - dataPointFormatter = props.dataPointFormatter, - xAxis = props.xAxis, - yAxis = props.yAxis, - others = _objectWithoutProperties$6(props, ["offset", "layout", "width", "dataKey", "data", "dataPointFormatter", "xAxis", "yAxis"]); - - var svgProps = filterProps(others); - var errorBars = data.map(function (entry, i) { - var _dataPointFormatter = dataPointFormatter(entry, dataKey), - x = _dataPointFormatter.x, - y = _dataPointFormatter.y, - value = _dataPointFormatter.value, - errorVal = _dataPointFormatter.errorVal; - - if (!errorVal) { - return null; - } - - var lineCoordinates = []; - var lowBound, highBound; - - if (Array.isArray(errorVal)) { - var _errorVal = _slicedToArray$2(errorVal, 2); - - lowBound = _errorVal[0]; - highBound = _errorVal[1]; - } else { - lowBound = highBound = errorVal; - } - - if (layout === 'vertical') { - // error bar for horizontal charts, the y is fixed, x is a range value - var scale = xAxis.scale; - var yMid = y + offset; - var yMin = yMid + width; - var yMax = yMid - width; - var xMin = scale(value - lowBound); - var xMax = scale(value + highBound); // the right line of |--| - - lineCoordinates.push({ - x1: xMax, - y1: yMin, - x2: xMax, - y2: yMax - }); // the middle line of |--| - - lineCoordinates.push({ - x1: xMin, - y1: yMid, - x2: xMax, - y2: yMid - }); // the left line of |--| - - lineCoordinates.push({ - x1: xMin, - y1: yMin, - x2: xMin, - y2: yMax - }); - } else if (layout === 'horizontal') { - // error bar for horizontal charts, the x is fixed, y is a range value - var _scale = yAxis.scale; - var xMid = x + offset; - - var _xMin = xMid - width; - - var _xMax = xMid + width; - - var _yMin = _scale(value - lowBound); - - var _yMax = _scale(value + highBound); // the top line - - - lineCoordinates.push({ - x1: _xMin, - y1: _yMax, - x2: _xMax, - y2: _yMax - }); // the middle line - - lineCoordinates.push({ - x1: xMid, - y1: _yMin, - x2: xMid, - y2: _yMax - }); // the bottom line - - lineCoordinates.push({ - x1: _xMin, - y1: _yMin, - x2: _xMax, - y2: _yMin - }); - } - - return ( - /*#__PURE__*/ - // eslint-disable-next-line react/no-array-index-key - React__default.createElement(Layer, _extends$9({ - className: "recharts-errorBar", - key: "bar-".concat(i) - }, svgProps), lineCoordinates.map(function (coordinates, index) { - return ( - /*#__PURE__*/ - // eslint-disable-next-line react/no-array-index-key - React__default.createElement("line", _extends$9({}, coordinates, { - key: "line-".concat(index) - })) - ); - })) - ); - }); - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-errorBars" - }, errorBars); -} -ErrorBar.defaultProps = { - stroke: 'black', - strokeWidth: 1.5, - width: 5, - offset: 0, - layout: 'horizontal' -}; -ErrorBar.displayName = 'ErrorBar'; - -function _typeof$8(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$8 = function _typeof(obj) { return typeof obj; }; } else { _typeof$8 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$8(obj); } - -function _objectWithoutProperties$5(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$5(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$5(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _extends$8() { _extends$8 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$8.apply(this, arguments); } - -function ownKeys$7(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$7(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$7(Object(source), true).forEach(function (key) { _defineProperty$7(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$7(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$7(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _toConsumableArray$1(arr) { return _arrayWithoutHoles$1(arr) || _iterableToArray$1(arr) || _unsupportedIterableToArray$2(arr) || _nonIterableSpread$1(); } - -function _nonIterableSpread$1() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$2(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$2(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$2(o, minLen); } - -function _iterableToArray$1(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles$1(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$2(arr); } - -function _arrayLikeToArray$2(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _classCallCheck$8(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$8(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$8(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$8(Constructor.prototype, protoProps); if (staticProps) _defineProperties$8(Constructor, staticProps); return Constructor; } - -function _inherits$8(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$8(subClass, superClass); } - -function _setPrototypeOf$8(o, p) { _setPrototypeOf$8 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$8(o, p); } - -function _createSuper$8(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$8(); return function _createSuperInternal() { var Super = _getPrototypeOf$8(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$8(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$8(this, result); }; } - -function _possibleConstructorReturn$8(self, call) { if (call && (_typeof$8(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$8(self); } - -function _assertThisInitialized$8(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$8() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$8(o) { _getPrototypeOf$8 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$8(o); } -var Line = /*#__PURE__*/function (_PureComponent) { - _inherits$8(Line, _PureComponent); - - var _super = _createSuper$8(Line); - - function Line() { - var _this; - - _classCallCheck$8(this, Line); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.mainCurve = void 0; - _this.state = { - isAnimationFinished: true, - totalLength: 0 - }; - - _this.getStrokeDasharray = function (length, totalLength, lines) { - var lineLength = lines.reduce(function (pre, next) { - return pre + next; - }); - var count = Math.floor(length / lineLength); - var remainLength = length % lineLength; - var restLength = totalLength - length; - var remainLines = []; - - for (var i = 0, sum = 0;; sum += lines[i], ++i) { - if (sum + lines[i] > remainLength) { - remainLines = [].concat(_toConsumableArray$1(lines.slice(0, i)), [remainLength - sum]); - break; - } - } - - var emptyLines = remainLines.length % 2 === 0 ? [0, restLength] : [restLength]; - return [].concat(_toConsumableArray$1(Line.repeat(lines, count)), _toConsumableArray$1(remainLines), emptyLines).map(function (line) { - return "".concat(line, "px"); - }).join(', '); - }; - - _this.id = uniqueId('recharts-line-'); - - _this.pathRef = function (node) { - _this.mainCurve = node; - }; - - _this.handleAnimationEnd = function () { - _this.setState({ - isAnimationFinished: true - }); - - if (_this.props.onAnimationEnd) { - _this.props.onAnimationEnd(); - } - }; - - _this.handleAnimationStart = function () { - _this.setState({ - isAnimationFinished: false - }); - - if (_this.props.onAnimationStart) { - _this.props.onAnimationStart(); - } - }; - - return _this; - } - - _createClass$8(Line, [{ - key: "componentDidMount", - value: - /* eslint-disable react/no-did-mount-set-state */ - function componentDidMount() { - if (!this.props.isAnimationActive) { - return; - } - - var totalLength = this.getTotalLength(); - this.setState({ - totalLength: totalLength - }); - } - }, { - key: "getTotalLength", - value: function getTotalLength() { - var curveDom = this.mainCurve; - - try { - return curveDom && curveDom.getTotalLength && curveDom.getTotalLength() || 0; - } catch (err) { - return 0; - } - } - }, { - key: "renderErrorBar", - value: function renderErrorBar() { - if (this.props.isAnimationActive && !this.state.isAnimationFinished) { - return null; - } - - var _this$props = this.props, - points = _this$props.points, - xAxis = _this$props.xAxis, - yAxis = _this$props.yAxis, - layout = _this$props.layout, - children = _this$props.children; - var errorBarItems = findAllByType(children, ErrorBar.displayName); - - if (!errorBarItems) { - return null; - } - - function dataPointFormatter(dataPoint, dataKey) { - return { - x: dataPoint.x, - y: dataPoint.y, - value: dataPoint.value, - errorVal: getValueByDataKey(dataPoint.payload, dataKey) - }; - } - - return errorBarItems.map(function (item, i) { - return /*#__PURE__*/React__default.cloneElement(item, { - // eslint-disable-next-line react/no-array-index-key - key: "bar-".concat(i), - data: points, - xAxis: xAxis, - yAxis: yAxis, - layout: layout, - dataPointFormatter: dataPointFormatter - }); - }); - } - }, { - key: "renderDots", - value: function renderDots(needClip, clipPathId) { - var isAnimationActive = this.props.isAnimationActive; - - if (isAnimationActive && !this.state.isAnimationFinished) { - return null; - } - - var _this$props2 = this.props, - dot = _this$props2.dot, - points = _this$props2.points, - dataKey = _this$props2.dataKey; - var lineProps = filterProps(this.props); - var customDotProps = filterProps(dot, true); - var dots = points.map(function (entry, i) { - var dotProps = _objectSpread$7(_objectSpread$7(_objectSpread$7({ - key: "dot-".concat(i), - r: 3 - }, lineProps), customDotProps), {}, { - value: entry.value, - dataKey: dataKey, - cx: entry.x, - cy: entry.y, - index: i, - payload: entry.payload - }); - - return Line.renderDotItem(dot, dotProps); - }); - var dotsProps = { - clipPath: needClip ? "url(#clipPath-".concat(clipPathId, ")") : null - }; - return /*#__PURE__*/React__default.createElement(Layer, _extends$8({ - className: "recharts-line-dots", - key: "dots" - }, dotsProps), dots); - } - }, { - key: "renderCurveStatically", - value: function renderCurveStatically(points, needClip, clipPathId, props) { - // eslint-disable-next-line @typescript-eslint/no-unused-vars - var _this$props3 = this.props, - type = _this$props3.type, - layout = _this$props3.layout, - connectNulls = _this$props3.connectNulls; - _this$props3.ref; - var others = _objectWithoutProperties$5(_this$props3, ["type", "layout", "connectNulls", "ref"]); - - var curveProps = _objectSpread$7(_objectSpread$7(_objectSpread$7({}, filterProps(others, true)), {}, { - fill: 'none', - className: 'recharts-line-curve', - clipPath: needClip ? "url(#clipPath-".concat(clipPathId, ")") : null, - points: points - }, props), {}, { - type: type, - layout: layout, - connectNulls: connectNulls - }); - - return /*#__PURE__*/React__default.createElement(Curve, _extends$8({}, curveProps, { - pathRef: this.pathRef - })); - } - }, { - key: "renderCurveWithAnimation", - value: function renderCurveWithAnimation(needClip, clipPathId) { - var _this2 = this; - - var _this$props4 = this.props, - points = _this$props4.points, - strokeDasharray = _this$props4.strokeDasharray, - isAnimationActive = _this$props4.isAnimationActive, - animationBegin = _this$props4.animationBegin, - animationDuration = _this$props4.animationDuration, - animationEasing = _this$props4.animationEasing, - animationId = _this$props4.animationId, - animateNewValues = _this$props4.animateNewValues, - width = _this$props4.width, - height = _this$props4.height; - var _this$state = this.state, - prevPoints = _this$state.prevPoints, - totalLength = _this$state.totalLength; - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - from: { - t: 0 - }, - to: { - t: 1 - }, - key: "line-".concat(animationId), - onAnimationEnd: this.handleAnimationEnd, - onAnimationStart: this.handleAnimationStart - }, function (_ref) { - var t = _ref.t; - - if (prevPoints) { - var prevPointsDiffFactor = prevPoints.length / points.length; - var stepData = points.map(function (entry, index) { - var prevPointIndex = Math.floor(index * prevPointsDiffFactor); - - if (prevPoints[prevPointIndex]) { - var prev = prevPoints[prevPointIndex]; - var interpolatorX = interpolateNumber$2(prev.x, entry.x); - var interpolatorY = interpolateNumber$2(prev.y, entry.y); - return _objectSpread$7(_objectSpread$7({}, entry), {}, { - x: interpolatorX(t), - y: interpolatorY(t) - }); - } // magic number of faking previous x and y location - - - if (animateNewValues) { - var _interpolatorX = interpolateNumber$2(width * 2, entry.x); - - var _interpolatorY = interpolateNumber$2(height / 2, entry.y); - - return _objectSpread$7(_objectSpread$7({}, entry), {}, { - x: _interpolatorX(t), - y: _interpolatorY(t) - }); - } - - return _objectSpread$7(_objectSpread$7({}, entry), {}, { - x: entry.x, - y: entry.y - }); - }); - return _this2.renderCurveStatically(stepData, needClip, clipPathId); - } - - var interpolator = interpolateNumber$2(0, totalLength); - var curLength = interpolator(t); - var currentStrokeDasharray; - - if (strokeDasharray) { - var lines = "".concat(strokeDasharray).split(/[,\s]+/gim).map(function (num) { - return parseFloat(num); - }); - currentStrokeDasharray = _this2.getStrokeDasharray(curLength, totalLength, lines); - } else { - currentStrokeDasharray = "".concat(curLength, "px ").concat(totalLength - curLength, "px"); - } - - return _this2.renderCurveStatically(points, needClip, clipPathId, { - strokeDasharray: currentStrokeDasharray - }); - }); - } - }, { - key: "renderCurve", - value: function renderCurve(needClip, clipPathId) { - var _this$props5 = this.props, - points = _this$props5.points, - isAnimationActive = _this$props5.isAnimationActive; - var _this$state2 = this.state, - prevPoints = _this$state2.prevPoints, - totalLength = _this$state2.totalLength; - - if (isAnimationActive && points && points.length && (!prevPoints && totalLength > 0 || !isEqual_1(prevPoints, points))) { - return this.renderCurveWithAnimation(needClip, clipPathId); - } - - return this.renderCurveStatically(points, needClip, clipPathId); - } - }, { - key: "render", - value: function render() { - var _this$props6 = this.props, - hide = _this$props6.hide, - dot = _this$props6.dot, - points = _this$props6.points, - className = _this$props6.className, - xAxis = _this$props6.xAxis, - yAxis = _this$props6.yAxis, - top = _this$props6.top, - left = _this$props6.left, - width = _this$props6.width, - height = _this$props6.height, - isAnimationActive = _this$props6.isAnimationActive, - id = _this$props6.id; - - if (hide || !points || !points.length) { - return null; - } - - var isAnimationFinished = this.state.isAnimationFinished; - var hasSinglePoint = points.length === 1; - var layerClass = classNames('recharts-line', className); - var needClip = xAxis && xAxis.allowDataOverflow || yAxis && yAxis.allowDataOverflow; - var clipPathId = isNil_1(id) ? this.id : id; - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass - }, needClip ? /*#__PURE__*/React__default.createElement("defs", null, /*#__PURE__*/React__default.createElement("clipPath", { - id: "clipPath-".concat(clipPathId) - }, /*#__PURE__*/React__default.createElement("rect", { - x: left, - y: top, - width: width, - height: height - }))) : null, !hasSinglePoint && this.renderCurve(needClip, clipPathId), this.renderErrorBar(), (hasSinglePoint || dot) && this.renderDots(needClip, clipPathId), (!isAnimationActive || isAnimationFinished) && LabelList.renderCallByParent(this.props, points)); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.animationId !== prevState.prevAnimationId) { - return { - prevAnimationId: nextProps.animationId, - curPoints: nextProps.points, - prevPoints: prevState.curPoints - }; - } - - if (nextProps.points !== prevState.curPoints) { - return { - curPoints: nextProps.points - }; - } - - return null; - } - }, { - key: "repeat", - value: function repeat(lines, count) { - var linesUnit = lines.length % 2 !== 0 ? [].concat(_toConsumableArray$1(lines), [0]) : lines; - var result = []; - - for (var i = 0; i < count; ++i) { - result = [].concat(_toConsumableArray$1(result), _toConsumableArray$1(linesUnit)); - } - - return result; - } - }, { - key: "renderDotItem", - value: function renderDotItem(option, props) { - var dotItem; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - dotItem = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - dotItem = option(props); - } else { - var className = classNames('recharts-line-dot', option ? option.className : ''); - dotItem = /*#__PURE__*/React__default.createElement(Dot, _extends$8({}, props, { - className: className - })); - } - - return dotItem; - } - }]); - - return Line; -}(PureComponent); -Line.displayName = 'Line'; -Line.defaultProps = { - xAxisId: 0, - yAxisId: 0, - connectNulls: false, - activeDot: true, - dot: true, - legendType: 'line', - stroke: '#3182bd', - strokeWidth: 1, - fill: '#fff', - points: [], - isAnimationActive: !Global.isSsr, - animateNewValues: true, - animationBegin: 0, - animationDuration: 1500, - animationEasing: 'ease', - hide: false -}; - -Line.getComposedData = function (_ref2) { - var props = _ref2.props, - xAxis = _ref2.xAxis, - yAxis = _ref2.yAxis, - xAxisTicks = _ref2.xAxisTicks, - yAxisTicks = _ref2.yAxisTicks, - dataKey = _ref2.dataKey, - bandSize = _ref2.bandSize, - displayedData = _ref2.displayedData, - offset = _ref2.offset; - var layout = props.layout; - var points = displayedData.map(function (entry, index) { - var value = getValueByDataKey(entry, dataKey); - - if (layout === 'horizontal') { - return { - x: getCateCoordinateOfLine({ - axis: xAxis, - ticks: xAxisTicks, - bandSize: bandSize, - entry: entry, - index: index - }), - y: isNil_1(value) ? null : yAxis.scale(value), - value: value, - payload: entry - }; - } - - return { - x: isNil_1(value) ? null : xAxis.scale(value), - y: getCateCoordinateOfLine({ - axis: yAxis, - ticks: yAxisTicks, - bandSize: bandSize, - entry: entry, - index: index - }), - value: value, - payload: entry - }; - }); - return _objectSpread$7({ - points: points, - layout: layout - }, offset); -}; - -function _typeof$7(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$7 = function _typeof(obj) { return typeof obj; }; } else { _typeof$7 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$7(obj); } - -function _objectWithoutProperties$4(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$4(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$4(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _extends$7() { _extends$7 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$7.apply(this, arguments); } - -function ownKeys$6(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$6(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$6(Object(source), true).forEach(function (key) { _defineProperty$6(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$6(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$6(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$7(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$7(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$7(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$7(Constructor.prototype, protoProps); if (staticProps) _defineProperties$7(Constructor, staticProps); return Constructor; } - -function _inherits$7(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$7(subClass, superClass); } - -function _setPrototypeOf$7(o, p) { _setPrototypeOf$7 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$7(o, p); } - -function _createSuper$7(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$7(); return function _createSuperInternal() { var Super = _getPrototypeOf$7(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$7(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$7(this, result); }; } - -function _possibleConstructorReturn$7(self, call) { if (call && (_typeof$7(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$7(self); } - -function _assertThisInitialized$7(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$7() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$7(o) { _getPrototypeOf$7 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$7(o); } -var Area = /*#__PURE__*/function (_PureComponent) { - _inherits$7(Area, _PureComponent); - - var _super = _createSuper$7(Area); - - function Area() { - var _this; - - _classCallCheck$7(this, Area); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - isAnimationFinished: true - }; - _this.id = uniqueId('recharts-area-'); - - _this.handleAnimationEnd = function () { - var onAnimationEnd = _this.props.onAnimationEnd; - - _this.setState({ - isAnimationFinished: true - }); - - if (isFunction_1(onAnimationEnd)) { - onAnimationEnd(); - } - }; - - _this.handleAnimationStart = function () { - var onAnimationStart = _this.props.onAnimationStart; - - _this.setState({ - isAnimationFinished: false - }); - - if (isFunction_1(onAnimationStart)) { - onAnimationStart(); - } - }; - - return _this; - } - - _createClass$7(Area, [{ - key: "renderDots", - value: function renderDots(needClip, clipPathId) { - var isAnimationActive = this.props.isAnimationActive; - var isAnimationFinished = this.state.isAnimationFinished; - - if (isAnimationActive && !isAnimationFinished) { - return null; - } - - var _this$props = this.props, - dot = _this$props.dot, - points = _this$props.points, - dataKey = _this$props.dataKey; - var areaProps = filterProps(this.props); - var customDotProps = filterProps(dot, true); - var dots = points.map(function (entry, i) { - var dotProps = _objectSpread$6(_objectSpread$6(_objectSpread$6({ - key: "dot-".concat(i), - r: 3 - }, areaProps), customDotProps), {}, { - dataKey: dataKey, - cx: entry.x, - cy: entry.y, - index: i, - value: entry.value, - payload: entry.payload - }); - - return Area.renderDotItem(dot, dotProps); - }); - var dotsProps = { - clipPath: needClip ? "url(#clipPath-".concat(clipPathId, ")") : null - }; - return /*#__PURE__*/React__default.createElement(Layer, _extends$7({ - className: "recharts-area-dots" - }, dotsProps), dots); - } - }, { - key: "renderHorizontalRect", - value: function renderHorizontalRect(alpha) { - var _this$props2 = this.props, - baseLine = _this$props2.baseLine, - points = _this$props2.points, - strokeWidth = _this$props2.strokeWidth; - var startX = points[0].x; - var endX = points[points.length - 1].x; - var width = alpha * Math.abs(startX - endX); - - var maxY = max_1(points.map(function (entry) { - return entry.y || 0; - })); - - if (isNumber(baseLine) && typeof baseLine === 'number') { - maxY = Math.max(baseLine, maxY); - } else if (baseLine && isArray_1(baseLine) && baseLine.length) { - maxY = Math.max(max_1(baseLine.map(function (entry) { - return entry.y || 0; - })), maxY); - } - - if (isNumber(maxY)) { - return /*#__PURE__*/React__default.createElement("rect", { - x: startX < endX ? startX : startX - width, - y: 0, - width: width, - height: Math.floor(maxY + (strokeWidth ? parseInt("".concat(strokeWidth), 10) : 1)) - }); - } - - return null; - } - }, { - key: "renderVerticalRect", - value: function renderVerticalRect(alpha) { - var _this$props3 = this.props, - baseLine = _this$props3.baseLine, - points = _this$props3.points, - strokeWidth = _this$props3.strokeWidth; - var startY = points[0].y; - var endY = points[points.length - 1].y; - var height = alpha * Math.abs(startY - endY); - - var maxX = max_1(points.map(function (entry) { - return entry.x || 0; - })); - - if (isNumber(baseLine) && typeof baseLine === 'number') { - maxX = Math.max(baseLine, maxX); - } else if (baseLine && isArray_1(baseLine) && baseLine.length) { - maxX = Math.max(max_1(baseLine.map(function (entry) { - return entry.x || 0; - })), maxX); - } - - if (isNumber(maxX)) { - return /*#__PURE__*/React__default.createElement("rect", { - x: 0, - y: startY < endY ? startY : startY - height, - width: maxX + (strokeWidth ? parseInt("".concat(strokeWidth), 10) : 1), - height: Math.floor(height) - }); - } - - return null; - } - }, { - key: "renderClipRect", - value: function renderClipRect(alpha) { - var layout = this.props.layout; - - if (layout === 'vertical') { - return this.renderVerticalRect(alpha); - } - - return this.renderHorizontalRect(alpha); - } - }, { - key: "renderAreaStatically", - value: function renderAreaStatically(points, baseLine, needClip, clipPathId) { - // eslint-disable-next-line @typescript-eslint/no-unused-vars - var _this$props4 = this.props, - layout = _this$props4.layout, - type = _this$props4.type, - stroke = _this$props4.stroke, - connectNulls = _this$props4.connectNulls, - isRange = _this$props4.isRange; - _this$props4.ref; - var others = _objectWithoutProperties$4(_this$props4, ["layout", "type", "stroke", "connectNulls", "isRange", "ref"]); - - return /*#__PURE__*/React__default.createElement(Layer, { - clipPath: needClip ? "url(#clipPath-".concat(clipPathId, ")") : null - }, /*#__PURE__*/React__default.createElement(Curve, _extends$7({}, filterProps(others, true), { - points: points, - connectNulls: connectNulls, - type: type, - baseLine: baseLine, - layout: layout, - stroke: "none", - className: "recharts-area-area" - })), stroke !== 'none' && /*#__PURE__*/React__default.createElement(Curve, _extends$7({}, filterProps(this.props), { - className: "recharts-area-curve", - layout: layout, - type: type, - connectNulls: connectNulls, - fill: "none", - points: points - })), stroke !== 'none' && isRange && /*#__PURE__*/React__default.createElement(Curve, _extends$7({}, filterProps(this.props), { - className: "recharts-area-curve", - layout: layout, - type: type, - connectNulls: connectNulls, - fill: "none", - points: baseLine - }))); - } - }, { - key: "renderAreaWithAnimation", - value: function renderAreaWithAnimation(needClip, clipPathId) { - var _this2 = this; - - var _this$props5 = this.props, - points = _this$props5.points, - baseLine = _this$props5.baseLine, - isAnimationActive = _this$props5.isAnimationActive, - animationBegin = _this$props5.animationBegin, - animationDuration = _this$props5.animationDuration, - animationEasing = _this$props5.animationEasing, - animationId = _this$props5.animationId; - var _this$state = this.state, - prevPoints = _this$state.prevPoints, - prevBaseLine = _this$state.prevBaseLine; // const clipPathId = _.isNil(id) ? this.id : id; - - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - from: { - t: 0 - }, - to: { - t: 1 - }, - key: "area-".concat(animationId), - onAnimationEnd: this.handleAnimationEnd, - onAnimationStart: this.handleAnimationStart - }, function (_ref) { - var t = _ref.t; - - if (prevPoints) { - var prevPointsDiffFactor = prevPoints.length / points.length; // update animtaion - - var stepPoints = points.map(function (entry, index) { - var prevPointIndex = Math.floor(index * prevPointsDiffFactor); - - if (prevPoints[prevPointIndex]) { - var prev = prevPoints[prevPointIndex]; - var interpolatorX = interpolateNumber$2(prev.x, entry.x); - var interpolatorY = interpolateNumber$2(prev.y, entry.y); - return _objectSpread$6(_objectSpread$6({}, entry), {}, { - x: interpolatorX(t), - y: interpolatorY(t) - }); - } - - return entry; - }); - var stepBaseLine; - - if (isNumber(baseLine) && typeof baseLine === 'number') { - var interpolator = interpolateNumber$2(prevBaseLine, baseLine); - stepBaseLine = interpolator(t); - } else if (isNil_1(baseLine) || _isNaN(baseLine)) { - var _interpolator = interpolateNumber$2(prevBaseLine, 0); - - stepBaseLine = _interpolator(t); - } else { - stepBaseLine = baseLine.map(function (entry, index) { - var prevPointIndex = Math.floor(index * prevPointsDiffFactor); - - if (prevBaseLine[prevPointIndex]) { - var prev = prevBaseLine[prevPointIndex]; - var interpolatorX = interpolateNumber$2(prev.x, entry.x); - var interpolatorY = interpolateNumber$2(prev.y, entry.y); - return _objectSpread$6(_objectSpread$6({}, entry), {}, { - x: interpolatorX(t), - y: interpolatorY(t) - }); - } - - return entry; - }); - } - - return _this2.renderAreaStatically(stepPoints, stepBaseLine, needClip, clipPathId); - } - - return /*#__PURE__*/React__default.createElement(Layer, null, /*#__PURE__*/React__default.createElement("defs", null, /*#__PURE__*/React__default.createElement("clipPath", { - id: "animationClipPath-".concat(clipPathId) - }, _this2.renderClipRect(t))), /*#__PURE__*/React__default.createElement(Layer, { - clipPath: "url(#animationClipPath-".concat(clipPathId, ")") - }, _this2.renderAreaStatically(points, baseLine, needClip, clipPathId))); - }); - } - }, { - key: "renderArea", - value: function renderArea(needClip, clipPathId) { - var _this$props6 = this.props, - points = _this$props6.points, - baseLine = _this$props6.baseLine, - isAnimationActive = _this$props6.isAnimationActive; - var _this$state2 = this.state, - prevPoints = _this$state2.prevPoints, - prevBaseLine = _this$state2.prevBaseLine, - totalLength = _this$state2.totalLength; - - if (isAnimationActive && points && points.length && (!prevPoints && totalLength > 0 || !isEqual_1(prevPoints, points) || !isEqual_1(prevBaseLine, baseLine))) { - return this.renderAreaWithAnimation(needClip, clipPathId); - } - - return this.renderAreaStatically(points, baseLine, needClip, clipPathId); - } - }, { - key: "render", - value: function render() { - var _this$props7 = this.props, - hide = _this$props7.hide, - dot = _this$props7.dot, - points = _this$props7.points, - className = _this$props7.className, - top = _this$props7.top, - left = _this$props7.left, - xAxis = _this$props7.xAxis, - yAxis = _this$props7.yAxis, - width = _this$props7.width, - height = _this$props7.height, - isAnimationActive = _this$props7.isAnimationActive, - id = _this$props7.id; - - if (hide || !points || !points.length) { - return null; - } - - var isAnimationFinished = this.state.isAnimationFinished; - var hasSinglePoint = points.length === 1; - var layerClass = classNames('recharts-area', className); - var needClip = xAxis && xAxis.allowDataOverflow || yAxis && yAxis.allowDataOverflow; - var clipPathId = isNil_1(id) ? this.id : id; - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass - }, needClip ? /*#__PURE__*/React__default.createElement("defs", null, /*#__PURE__*/React__default.createElement("clipPath", { - id: "clipPath-".concat(clipPathId) - }, /*#__PURE__*/React__default.createElement("rect", { - x: left, - y: top, - width: width, - height: Math.floor(height) - }))) : null, !hasSinglePoint ? this.renderArea(needClip, clipPathId) : null, (dot || hasSinglePoint) && this.renderDots(needClip, clipPathId), (!isAnimationActive || isAnimationFinished) && LabelList.renderCallByParent(this.props, points)); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.animationId !== prevState.prevAnimationId) { - return { - prevAnimationId: nextProps.animationId, - curPoints: nextProps.points, - curBaseLine: nextProps.baseLine, - prevPoints: prevState.curPoints, - prevBaseLine: prevState.curBaseLine - }; - } - - if (nextProps.points !== prevState.curPoints || nextProps.baseLine !== prevState.curBaseLine) { - return { - curPoints: nextProps.points, - curBaseLine: nextProps.baseLine - }; - } - - return null; - } - }]); - - return Area; -}(PureComponent); -Area.displayName = 'Area'; -Area.defaultProps = { - stroke: '#3182bd', - fill: '#3182bd', - fillOpacity: 0.6, - xAxisId: 0, - yAxisId: 0, - legendType: 'line', - connectNulls: false, - // points of area - points: [], - dot: false, - activeDot: true, - hide: false, - isAnimationActive: !Global.isSsr, - animationBegin: 0, - animationDuration: 1500, - animationEasing: 'ease' -}; - -Area.getBaseValue = function (props, item, xAxis, yAxis) { - var layout = props.layout; - var baseValue = item.props.baseValue; - - if (isNumber(baseValue) && typeof baseValue === 'number') { - return baseValue; - } - - var numericAxis = layout === 'horizontal' ? yAxis : xAxis; - var domain = numericAxis.scale.domain(); - - if (numericAxis.type === 'number') { - var max = Math.max(domain[0], domain[1]); - var min = Math.min(domain[0], domain[1]); - - if (baseValue === 'dataMin') { - return min; - } - - if (baseValue === 'dataMax') { - return max; - } - - return max < 0 ? max : Math.max(Math.min(domain[0], domain[1]), 0); - } - - if (baseValue === 'dataMin') { - return domain[0]; - } - - if (baseValue === 'dataMax') { - return domain[1]; - } - - return domain[0]; -}; - -Area.getComposedData = function (_ref2) { - var props = _ref2.props, - item = _ref2.item, - xAxis = _ref2.xAxis, - yAxis = _ref2.yAxis, - xAxisTicks = _ref2.xAxisTicks, - yAxisTicks = _ref2.yAxisTicks, - bandSize = _ref2.bandSize, - dataKey = _ref2.dataKey, - stackedData = _ref2.stackedData, - dataStartIndex = _ref2.dataStartIndex, - displayedData = _ref2.displayedData, - offset = _ref2.offset; - var layout = props.layout; - var hasStack = stackedData && stackedData.length; - var baseValue = Area.getBaseValue(props, item, xAxis, yAxis); - var isRange = false; - var points = displayedData.map(function (entry, index) { - var originalValue = getValueByDataKey(entry, dataKey); - var value; - - if (hasStack) { - value = stackedData[dataStartIndex + index]; - } else { - value = originalValue; - - if (!isArray_1(value)) { - value = [baseValue, value]; - } else { - isRange = true; - } - } - - var isBreakPoint = isNil_1(value[1]) || hasStack && isNil_1(originalValue); - - if (layout === 'horizontal') { - return { - x: getCateCoordinateOfLine({ - axis: xAxis, - ticks: xAxisTicks, - bandSize: bandSize, - entry: entry, - index: index - }), - y: isBreakPoint ? null : yAxis.scale(value[1]), - value: value, - payload: entry - }; - } - - return { - x: isBreakPoint ? null : xAxis.scale(value[1]), - y: getCateCoordinateOfLine({ - axis: yAxis, - ticks: yAxisTicks, - bandSize: bandSize, - entry: entry, - index: index - }), - value: value, - payload: entry - }; - }); - var baseLine; - - if (hasStack || isRange) { - baseLine = points.map(function (entry) { - if (layout === 'horizontal') { - return { - x: entry.x, - y: !isNil_1(get_1(entry, 'value[0]')) && !isNil_1(get_1(entry, 'y')) ? yAxis.scale(get_1(entry, 'value[0]')) : null - }; - } - - return { - x: !isNil_1(get_1(entry, 'value[0]')) ? xAxis.scale(get_1(entry, 'value[0]')) : null, - y: entry.y - }; - }); - } else if (layout === 'horizontal') { - baseLine = yAxis.scale(baseValue); - } else { - baseLine = xAxis.scale(baseValue); - } - - return _objectSpread$6({ - points: points, - baseLine: baseLine, - layout: layout, - isRange: isRange - }, offset); -}; - -Area.renderDotItem = function (option, props) { - var dotItem; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - dotItem = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - dotItem = option(props); - } else { - dotItem = /*#__PURE__*/React__default.createElement(Dot, _extends$7({}, props, { - className: "recharts-area-dot" - })); - } - - return dotItem; -}; - -function _typeof$6(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$6 = function _typeof(obj) { return typeof obj; }; } else { _typeof$6 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$6(obj); } - -function _objectWithoutProperties$3(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$3(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$3(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _extends$6() { _extends$6 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$6.apply(this, arguments); } - -function ownKeys$5(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$5(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$5(Object(source), true).forEach(function (key) { _defineProperty$5(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$5(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$5(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$6(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$6(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$6(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$6(Constructor.prototype, protoProps); if (staticProps) _defineProperties$6(Constructor, staticProps); return Constructor; } - -function _inherits$6(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$6(subClass, superClass); } - -function _setPrototypeOf$6(o, p) { _setPrototypeOf$6 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$6(o, p); } - -function _createSuper$6(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$6(); return function _createSuperInternal() { var Super = _getPrototypeOf$6(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$6(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$6(this, result); }; } - -function _possibleConstructorReturn$6(self, call) { if (call && (_typeof$6(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$6(self); } - -function _assertThisInitialized$6(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$6() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$6(o) { _getPrototypeOf$6 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$6(o); } -var Bar = /*#__PURE__*/function (_PureComponent) { - _inherits$6(Bar, _PureComponent); - - var _super = _createSuper$6(Bar); - - function Bar() { - var _this; - - _classCallCheck$6(this, Bar); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - isAnimationFinished: false - }; - _this.id = uniqueId('recharts-bar-'); - - _this.handleAnimationEnd = function () { - var onAnimationEnd = _this.props.onAnimationEnd; - - _this.setState({ - isAnimationFinished: true - }); - - if (onAnimationEnd) { - onAnimationEnd(); - } - }; - - _this.handleAnimationStart = function () { - var onAnimationStart = _this.props.onAnimationStart; - - _this.setState({ - isAnimationFinished: false - }); - - if (onAnimationStart) { - onAnimationStart(); - } - }; - - return _this; - } - - _createClass$6(Bar, [{ - key: "renderRectanglesStatically", - value: function renderRectanglesStatically(data) { - var _this2 = this; - - var shape = this.props.shape; - var baseProps = filterProps(this.props); - return data && data.map(function (entry, i) { - var props = _objectSpread$5(_objectSpread$5(_objectSpread$5({}, baseProps), entry), {}, { - index: i - }); - - return /*#__PURE__*/React__default.createElement(Layer, _extends$6({ - className: "recharts-bar-rectangle" - }, adaptEventsOfChild(_this2.props, entry, i), { - key: "rectangle-".concat(i) // eslint-disable-line react/no-array-index-key - - }), Bar.renderRectangle(shape, props)); - }); - } - }, { - key: "renderRectanglesWithAnimation", - value: function renderRectanglesWithAnimation() { - var _this3 = this; - - var _this$props = this.props, - data = _this$props.data, - layout = _this$props.layout, - isAnimationActive = _this$props.isAnimationActive, - animationBegin = _this$props.animationBegin, - animationDuration = _this$props.animationDuration, - animationEasing = _this$props.animationEasing, - animationId = _this$props.animationId; - var prevData = this.state.prevData; - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - from: { - t: 0 - }, - to: { - t: 1 - }, - key: "bar-".concat(animationId), - onAnimationEnd: this.handleAnimationEnd, - onAnimationStart: this.handleAnimationStart - }, function (_ref) { - var t = _ref.t; - var stepData = data.map(function (entry, index) { - var prev = prevData && prevData[index]; - - if (prev) { - var interpolatorX = interpolateNumber$2(prev.x, entry.x); - var interpolatorY = interpolateNumber$2(prev.y, entry.y); - var interpolatorWidth = interpolateNumber$2(prev.width, entry.width); - var interpolatorHeight = interpolateNumber$2(prev.height, entry.height); - return _objectSpread$5(_objectSpread$5({}, entry), {}, { - x: interpolatorX(t), - y: interpolatorY(t), - width: interpolatorWidth(t), - height: interpolatorHeight(t) - }); - } - - if (layout === 'horizontal') { - var _interpolatorHeight = interpolateNumber$2(0, entry.height); - - var h = _interpolatorHeight(t); - - return _objectSpread$5(_objectSpread$5({}, entry), {}, { - y: entry.y + entry.height - h, - height: h - }); - } - - var interpolator = interpolateNumber$2(0, entry.width); - var w = interpolator(t); - return _objectSpread$5(_objectSpread$5({}, entry), {}, { - width: w - }); - }); - return /*#__PURE__*/React__default.createElement(Layer, null, _this3.renderRectanglesStatically(stepData)); - }); - } - }, { - key: "renderRectangles", - value: function renderRectangles() { - var _this$props2 = this.props, - data = _this$props2.data, - isAnimationActive = _this$props2.isAnimationActive; - var prevData = this.state.prevData; - - if (isAnimationActive && data && data.length && (!prevData || !isEqual_1(prevData, data))) { - return this.renderRectanglesWithAnimation(); - } - - return this.renderRectanglesStatically(data); - } - }, { - key: "renderBackground", - value: function renderBackground() { - var _this4 = this; - - var data = this.props.data; - var backgroundProps = filterProps(this.props.background); - return data.map(function (entry, i) { - // eslint-disable-next-line @typescript-eslint/no-unused-vars - entry.value; - var background = entry.background, - rest = _objectWithoutProperties$3(entry, ["value", "background"]); - - if (!background) { - return null; - } - - var props = _objectSpread$5(_objectSpread$5(_objectSpread$5(_objectSpread$5(_objectSpread$5({}, rest), {}, { - fill: '#eee' - }, background), backgroundProps), adaptEventsOfChild(_this4.props, entry, i)), {}, { - index: i, - key: "background-bar-".concat(i), - className: 'recharts-bar-background-rectangle' - }); - - return Bar.renderRectangle(_this4.props.background, props); - }); - } - }, { - key: "renderErrorBar", - value: function renderErrorBar() { - if (this.props.isAnimationActive && !this.state.isAnimationFinished) { - return null; - } - - var _this$props3 = this.props, - data = _this$props3.data, - xAxis = _this$props3.xAxis, - yAxis = _this$props3.yAxis, - layout = _this$props3.layout, - children = _this$props3.children; - var errorBarItems = findAllByType(children, ErrorBar.displayName); - - if (!errorBarItems) { - return null; - } - - var offset = layout === 'vertical' ? data[0].height / 2 : data[0].width / 2; - - function dataPointFormatter(dataPoint, dataKey) { - return { - x: dataPoint.x, - y: dataPoint.y, - value: dataPoint.value, - errorVal: getValueByDataKey(dataPoint, dataKey) - }; - } - - return errorBarItems.map(function (item, i) { - return /*#__PURE__*/React__default.cloneElement(item, { - key: "error-bar-".concat(i), - // eslint-disable-line react/no-array-index-key - data: data, - xAxis: xAxis, - yAxis: yAxis, - layout: layout, - offset: offset, - dataPointFormatter: dataPointFormatter - }); - }); - } - }, { - key: "render", - value: function render() { - var _this$props4 = this.props, - hide = _this$props4.hide, - data = _this$props4.data, - className = _this$props4.className, - xAxis = _this$props4.xAxis, - yAxis = _this$props4.yAxis, - left = _this$props4.left, - top = _this$props4.top, - width = _this$props4.width, - height = _this$props4.height, - isAnimationActive = _this$props4.isAnimationActive, - background = _this$props4.background, - id = _this$props4.id; - - if (hide || !data || !data.length) { - return null; - } - - var isAnimationFinished = this.state.isAnimationFinished; - var layerClass = classNames('recharts-bar', className); - var needClip = xAxis && xAxis.allowDataOverflow || yAxis && yAxis.allowDataOverflow; - var clipPathId = isNil_1(id) ? this.id : id; - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass - }, needClip ? /*#__PURE__*/React__default.createElement("defs", null, /*#__PURE__*/React__default.createElement("clipPath", { - id: "clipPath-".concat(clipPathId) - }, /*#__PURE__*/React__default.createElement("rect", { - x: left, - y: top, - width: width, - height: height - }))) : null, /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-bar-rectangles", - clipPath: needClip ? "url(#clipPath-".concat(clipPathId, ")") : null - }, background ? this.renderBackground() : null, this.renderRectangles()), this.renderErrorBar(), (!isAnimationActive || isAnimationFinished) && LabelList.renderCallByParent(this.props, data)); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.animationId !== prevState.prevAnimationId) { - return { - prevAnimationId: nextProps.animationId, - curData: nextProps.data, - prevData: prevState.curData - }; - } - - if (nextProps.data !== prevState.curData) { - return { - curData: nextProps.data - }; - } - - return null; - } - }, { - key: "renderRectangle", - value: function renderRectangle(option, props) { - var rectangle; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - rectangle = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - rectangle = option(props); - } else { - rectangle = /*#__PURE__*/React__default.createElement(Rectangle, props); - } - - return rectangle; - } - }]); - - return Bar; -}(PureComponent); -Bar.displayName = 'Bar'; -Bar.defaultProps = { - xAxisId: 0, - yAxisId: 0, - legendType: 'rect', - minPointSize: 0, - hide: false, - // data of bar - data: [], - layout: 'vertical', - isAnimationActive: !Global.isSsr, - animationBegin: 0, - animationDuration: 400, - animationEasing: 'ease' -}; - -Bar.getComposedData = function (_ref2) { - var props = _ref2.props, - item = _ref2.item, - barPosition = _ref2.barPosition, - bandSize = _ref2.bandSize, - xAxis = _ref2.xAxis, - yAxis = _ref2.yAxis, - xAxisTicks = _ref2.xAxisTicks, - yAxisTicks = _ref2.yAxisTicks, - stackedData = _ref2.stackedData, - dataStartIndex = _ref2.dataStartIndex, - displayedData = _ref2.displayedData, - offset = _ref2.offset; - var pos = findPositionOfBar(barPosition, item); - - if (!pos) { - return null; - } - - var layout = props.layout; - var _item$props = item.props, - dataKey = _item$props.dataKey, - children = _item$props.children, - minPointSize = _item$props.minPointSize; - var numericAxis = layout === 'horizontal' ? yAxis : xAxis; - var stackedDomain = stackedData ? numericAxis.scale.domain() : null; - var baseValue = getBaseValueOfBar({ - numericAxis: numericAxis - }); - var cells = findAllByType(children, Cell.displayName); - var rects = displayedData.map(function (entry, index) { - var value, x, y, width, height, background; - - if (stackedData) { - value = truncateByDomain(stackedData[dataStartIndex + index], stackedDomain); - } else { - value = getValueByDataKey(entry, dataKey); - - if (!isArray_1(value)) { - value = [baseValue, value]; - } - } - - if (layout === 'horizontal') { - x = getCateCoordinateOfBar({ - axis: xAxis, - ticks: xAxisTicks, - bandSize: bandSize, - offset: pos.offset, - entry: entry, - index: index - }); - y = yAxis.scale(value[1]); - width = pos.size; - height = yAxis.scale(value[0]) - yAxis.scale(value[1]); - background = { - x: x, - y: yAxis.y, - width: width, - height: yAxis.height - }; - - if (Math.abs(minPointSize) > 0 && Math.abs(height) < Math.abs(minPointSize)) { - var delta = mathSign(height || minPointSize) * (Math.abs(minPointSize) - Math.abs(height)); - y -= delta; - height += delta; - } - } else { - x = xAxis.scale(value[0]); - y = getCateCoordinateOfBar({ - axis: yAxis, - ticks: yAxisTicks, - bandSize: bandSize, - offset: pos.offset, - entry: entry, - index: index - }); - width = xAxis.scale(value[1]) - xAxis.scale(value[0]); - height = pos.size; - background = { - x: xAxis.x, - y: y, - width: xAxis.width, - height: height - }; - - if (Math.abs(minPointSize) > 0 && Math.abs(width) < Math.abs(minPointSize)) { - var _delta = mathSign(width || minPointSize) * (Math.abs(minPointSize) - Math.abs(width)); - - width += _delta; - } - } - - return _objectSpread$5(_objectSpread$5(_objectSpread$5({}, entry), {}, { - x: x, - y: y, - width: width, - height: height, - value: stackedData ? value : value[1], - payload: entry, - background: background - }, cells && cells[index] && cells[index].props), {}, { - tooltipPayload: [getTooltipItem(item, entry)], - tooltipPosition: { - x: x + width / 2, - y: y + height / 2 - } - }); - }); - return _objectSpread$5({ - data: rects, - layout: layout - }, offset); -}; - -/** - * @fileOverview Z Axis - */ -var ZAxis = function ZAxis() { - return null; -}; -ZAxis.displayName = 'ZAxis'; -ZAxis.defaultProps = { - zAxisId: 0, - range: [64, 64], - scale: 'auto', - type: 'number' -}; - -function _typeof$5(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$5 = function _typeof(obj) { return typeof obj; }; } else { _typeof$5 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$5(obj); } - -function _extends$5() { _extends$5 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$5.apply(this, arguments); } - -function ownKeys$4(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$4(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$4(Object(source), true).forEach(function (key) { _defineProperty$4(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$4(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$4(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck$5(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$5(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$5(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$5(Constructor.prototype, protoProps); if (staticProps) _defineProperties$5(Constructor, staticProps); return Constructor; } - -function _inherits$5(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$5(subClass, superClass); } - -function _setPrototypeOf$5(o, p) { _setPrototypeOf$5 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$5(o, p); } - -function _createSuper$5(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$5(); return function _createSuperInternal() { var Super = _getPrototypeOf$5(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$5(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$5(this, result); }; } - -function _possibleConstructorReturn$5(self, call) { if (call && (_typeof$5(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$5(self); } - -function _assertThisInitialized$5(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$5() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$5(o) { _getPrototypeOf$5 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$5(o); } -var Scatter = /*#__PURE__*/function (_PureComponent) { - _inherits$5(Scatter, _PureComponent); - - var _super = _createSuper$5(Scatter); - - function Scatter() { - var _this; - - _classCallCheck$5(this, Scatter); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - isAnimationFinished: false - }; - - _this.handleAnimationEnd = function () { - _this.setState({ - isAnimationFinished: true - }); - }; - - _this.handleAnimationStart = function () { - _this.setState({ - isAnimationFinished: false - }); - }; - - _this.id = uniqueId('recharts-scatter-'); - return _this; - } - - _createClass$5(Scatter, [{ - key: "renderSymbolsStatically", - value: function renderSymbolsStatically(points) { - var _this2 = this; - - var _this$props = this.props, - shape = _this$props.shape, - activeShape = _this$props.activeShape, - activeIndex = _this$props.activeIndex; - var baseProps = filterProps(this.props); - return points.map(function (entry, i) { - var props = _objectSpread$4(_objectSpread$4({ - key: "symbol-".concat(i) - }, baseProps), entry); - - return /*#__PURE__*/React__default.createElement(Layer, _extends$5({ - className: "recharts-scatter-symbol" - }, adaptEventsOfChild(_this2.props, entry, i), { - key: "symbol-".concat(i) // eslint-disable-line react/no-array-index-key - - }), Scatter.renderSymbolItem(activeIndex === i ? activeShape : shape, props)); - }); - } - }, { - key: "renderSymbolsWithAnimation", - value: function renderSymbolsWithAnimation() { - var _this3 = this; - - var _this$props2 = this.props, - points = _this$props2.points, - isAnimationActive = _this$props2.isAnimationActive, - animationBegin = _this$props2.animationBegin, - animationDuration = _this$props2.animationDuration, - animationEasing = _this$props2.animationEasing, - animationId = _this$props2.animationId; - var prevPoints = this.state.prevPoints; - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - from: { - t: 0 - }, - to: { - t: 1 - }, - key: "pie-".concat(animationId), - onAnimationEnd: this.handleAnimationEnd, - onAnimationStart: this.handleAnimationStart - }, function (_ref) { - var t = _ref.t; - var stepData = points.map(function (entry, index) { - var prev = prevPoints && prevPoints[index]; - - if (prev) { - var interpolatorCx = interpolateNumber$2(prev.cx, entry.cx); - var interpolatorCy = interpolateNumber$2(prev.cy, entry.cy); - var interpolatorSize = interpolateNumber$2(prev.size, entry.size); - return _objectSpread$4(_objectSpread$4({}, entry), {}, { - cx: interpolatorCx(t), - cy: interpolatorCy(t), - size: interpolatorSize(t) - }); - } - - var interpolator = interpolateNumber$2(0, entry.size); - return _objectSpread$4(_objectSpread$4({}, entry), {}, { - size: interpolator(t) - }); - }); - return /*#__PURE__*/React__default.createElement(Layer, null, _this3.renderSymbolsStatically(stepData)); - }); - } - }, { - key: "renderSymbols", - value: function renderSymbols() { - var _this$props3 = this.props, - points = _this$props3.points, - isAnimationActive = _this$props3.isAnimationActive; - var prevPoints = this.state.prevPoints; - - if (isAnimationActive && points && points.length && (!prevPoints || !isEqual_1(prevPoints, points))) { - return this.renderSymbolsWithAnimation(); - } - - return this.renderSymbolsStatically(points); - } - }, { - key: "renderErrorBar", - value: function renderErrorBar() { - var isAnimationActive = this.props.isAnimationActive; - - if (isAnimationActive && !this.state.isAnimationFinished) { - return null; - } - - var _this$props4 = this.props, - points = _this$props4.points, - xAxis = _this$props4.xAxis, - yAxis = _this$props4.yAxis, - children = _this$props4.children; - var errorBarItems = findAllByType(children, ErrorBar.displayName); - - if (!errorBarItems) { - return null; - } - - function dataPointFormatterY(dataPoint, dataKey) { - return { - x: dataPoint.cx, - y: dataPoint.cy, - value: +dataPoint.node.y, - errorVal: getValueByDataKey(dataPoint, dataKey) - }; - } - - function dataPointFormatterX(dataPoint, dataKey) { - return { - x: dataPoint.cx, - y: dataPoint.cy, - value: +dataPoint.node.x, - errorVal: getValueByDataKey(dataPoint, dataKey) - }; - } - - return errorBarItems.map(function (item, i) { - var direction = item.props.direction; - return /*#__PURE__*/React__default.cloneElement(item, { - key: i, - // eslint-disable-line react/no-array-index-key - data: points, - xAxis: xAxis, - yAxis: yAxis, - layout: direction === 'x' ? 'vertical' : 'horizontal', - dataPointFormatter: direction === 'x' ? dataPointFormatterX : dataPointFormatterY - }); - }); - } - }, { - key: "renderLine", - value: function renderLine() { - var _this$props5 = this.props, - points = _this$props5.points, - line = _this$props5.line, - lineType = _this$props5.lineType, - lineJointType = _this$props5.lineJointType; - var scatterProps = filterProps(this.props); - var customLineProps = filterProps(line); - var linePoints, lineItem; - - if (lineType === 'joint') { - linePoints = points.map(function (entry) { - return { - x: entry.cx, - y: entry.cy - }; - }); - } else if (lineType === 'fitting') { - var _getLinearRegression = getLinearRegression(points), - xmin = _getLinearRegression.xmin, - xmax = _getLinearRegression.xmax, - a = _getLinearRegression.a, - b = _getLinearRegression.b; - - var linearExp = function linearExp(x) { - return a * x + b; - }; - - linePoints = [{ - x: xmin, - y: linearExp(xmin) - }, { - x: xmax, - y: linearExp(xmax) - }]; - } - - var lineProps = _objectSpread$4(_objectSpread$4(_objectSpread$4({}, scatterProps), {}, { - fill: 'none', - stroke: scatterProps && scatterProps.fill - }, customLineProps), {}, { - points: linePoints - }); - - if ( /*#__PURE__*/React__default.isValidElement(line)) { - lineItem = /*#__PURE__*/React__default.cloneElement(line, lineProps); - } else if (isFunction_1(line)) { - lineItem = line(lineProps); - } else { - lineItem = /*#__PURE__*/React__default.createElement(Curve, _extends$5({}, lineProps, { - type: lineJointType - })); - } - - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-scatter-line", - key: "recharts-scatter-line" - }, lineItem); - } - }, { - key: "render", - value: function render() { - var _this$props6 = this.props, - hide = _this$props6.hide, - points = _this$props6.points, - line = _this$props6.line, - className = _this$props6.className, - xAxis = _this$props6.xAxis, - yAxis = _this$props6.yAxis, - left = _this$props6.left, - top = _this$props6.top, - width = _this$props6.width, - height = _this$props6.height, - id = _this$props6.id, - isAnimationActive = _this$props6.isAnimationActive; - - if (hide || !points || !points.length) { - return null; - } - - var isAnimationFinished = this.state.isAnimationFinished; - var layerClass = classNames('recharts-scatter', className); - var needClip = xAxis && xAxis.allowDataOverflow || yAxis && yAxis.allowDataOverflow; - var clipPathId = isNil_1(id) ? this.id : id; - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass, - clipPath: needClip ? "url(#clipPath-".concat(clipPathId, ")") : null - }, needClip ? /*#__PURE__*/React__default.createElement("defs", null, /*#__PURE__*/React__default.createElement("clipPath", { - id: "clipPath-".concat(clipPathId) - }, /*#__PURE__*/React__default.createElement("rect", { - x: left, - y: top, - width: width, - height: height - }))) : null, line && this.renderLine(), this.renderErrorBar(), /*#__PURE__*/React__default.createElement(Layer, { - key: "recharts-scatter-symbols" - }, this.renderSymbols()), (!isAnimationActive || isAnimationFinished) && LabelList.renderCallByParent(this.props, points)); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.animationId !== prevState.prevAnimationId) { - return { - prevAnimationId: nextProps.animationId, - curPoints: nextProps.points, - prevPoints: prevState.curPoints - }; - } - - if (nextProps.points !== prevState.curPoints) { - return { - curPoints: nextProps.points - }; - } - - return null; - } - }, { - key: "renderSymbolItem", - value: function renderSymbolItem(option, props) { - var symbol; - - if ( /*#__PURE__*/React__default.isValidElement(option)) { - symbol = /*#__PURE__*/React__default.cloneElement(option, props); - } else if (isFunction_1(option)) { - symbol = option(props); - } else if (typeof option === 'string') { - symbol = /*#__PURE__*/React__default.createElement(Symbols, _extends$5({}, props, { - type: option - })); - } - - return symbol; - } - }]); - - return Scatter; -}(PureComponent); -Scatter.displayName = 'Scatter'; -Scatter.defaultProps = { - xAxisId: 0, - yAxisId: 0, - zAxisId: 0, - legendType: 'circle', - lineType: 'joint', - lineJointType: 'linear', - data: [], - shape: 'circle', - hide: false, - isAnimationActive: !Global.isSsr, - animationBegin: 0, - animationDuration: 400, - animationEasing: 'linear' -}; - -Scatter.getComposedData = function (_ref2) { - var xAxis = _ref2.xAxis, - yAxis = _ref2.yAxis, - zAxis = _ref2.zAxis, - item = _ref2.item, - displayedData = _ref2.displayedData, - xAxisTicks = _ref2.xAxisTicks, - yAxisTicks = _ref2.yAxisTicks, - offset = _ref2.offset; - var tooltipType = item.props.tooltipType; - var cells = findAllByType(item.props.children, Cell.displayName); - var xAxisDataKey = isNil_1(xAxis.dataKey) ? item.props.dataKey : xAxis.dataKey; - var yAxisDataKey = isNil_1(yAxis.dataKey) ? item.props.dataKey : yAxis.dataKey; - var zAxisDataKey = zAxis && zAxis.dataKey; - var defaultRangeZ = zAxis ? zAxis.range : ZAxis.defaultProps.range; - var defaultZ = defaultRangeZ && defaultRangeZ[0]; - var xBandSize = xAxis.scale.bandwidth ? xAxis.scale.bandwidth() : 0; - var yBandSize = yAxis.scale.bandwidth ? yAxis.scale.bandwidth() : 0; - var points = displayedData.map(function (entry, index) { - var x = getValueByDataKey(entry, xAxisDataKey); - var y = getValueByDataKey(entry, yAxisDataKey); - var z = !isNil_1(zAxisDataKey) && getValueByDataKey(entry, zAxisDataKey) || '-'; - var tooltipPayload = [{ - name: isNil_1(xAxis.dataKey) ? item.props.name : xAxis.name || xAxis.dataKey, - unit: xAxis.unit || '', - value: x, - payload: entry, - dataKey: xAxisDataKey, - type: tooltipType - }, { - name: isNil_1(yAxis.dataKey) ? item.props.name : yAxis.name || yAxis.dataKey, - unit: yAxis.unit || '', - value: y, - payload: entry, - dataKey: yAxisDataKey, - type: tooltipType - }]; - - if (z !== '-') { - tooltipPayload.push({ - name: zAxis.name || zAxis.dataKey, - unit: zAxis.unit || '', - value: z, - payload: entry, - dataKey: zAxisDataKey, - type: tooltipType - }); - } - - var cx = getCateCoordinateOfLine({ - axis: xAxis, - ticks: xAxisTicks, - bandSize: xBandSize, - entry: entry, - index: index, - dataKey: xAxisDataKey - }); - var cy = getCateCoordinateOfLine({ - axis: yAxis, - ticks: yAxisTicks, - bandSize: yBandSize, - entry: entry, - index: index, - dataKey: yAxisDataKey - }); - var size = z !== '-' ? zAxis.scale(z) : defaultZ; - var radius = Math.sqrt(Math.max(size, 0) / Math.PI); - return _objectSpread$4(_objectSpread$4({}, entry), {}, { - cx: cx, - cy: cy, - x: cx - radius, - y: cy - radius, - xAxis: xAxis, - yAxis: yAxis, - zAxis: zAxis, - width: 2 * radius, - height: 2 * radius, - size: size, - node: { - x: x, - y: y, - z: z - }, - tooltipPayload: tooltipPayload, - tooltipPosition: { - x: cx, - y: cy - }, - payload: entry - }, cells && cells[index] && cells[index].props); - }); - return _objectSpread$4({ - points: points - }, offset); -}; - -/** - * @fileOverview X Axis - */ - -/** Define of XAxis props */ -var XAxis = function XAxis() { - return null; -}; -XAxis.displayName = 'XAxis'; -XAxis.defaultProps = { - allowDecimals: true, - hide: false, - orientation: 'bottom', - width: 0, - height: 30, - mirror: false, - xAxisId: 0, - tickCount: 5, - type: 'category', - domain: [0, 'auto'], - padding: { - left: 0, - right: 0 - }, - allowDataOverflow: false, - scale: 'auto', - reversed: false, - allowDuplicatedCategory: true -}; - -/** - * @fileOverview Y Axis - */ -var YAxis = function YAxis() { - return null; -}; -YAxis.displayName = 'YAxis'; -YAxis.defaultProps = { - allowDuplicatedCategory: true, - allowDecimals: true, - hide: false, - orientation: 'left', - width: 60, - height: 0, - mirror: false, - yAxisId: 0, - tickCount: 5, - type: 'number', - domain: [0, 'auto'], - padding: { - top: 0, - bottom: 0 - }, - allowDataOverflow: false, - scale: 'auto', - reversed: false -}; - -var baseIteratee$2 = _baseIteratee, - isArrayLike$1 = isArrayLike_1, - keys$2 = keys_1; - -/** - * Creates a `_.find` or `_.findLast` function. - * - * @private - * @param {Function} findIndexFunc The function to find the collection index. - * @returns {Function} Returns the new find function. - */ -function createFind$1(findIndexFunc) { - return function(collection, predicate, fromIndex) { - var iterable = Object(collection); - if (!isArrayLike$1(collection)) { - var iteratee = baseIteratee$2(predicate); - collection = keys$2(collection); - predicate = function(key) { return iteratee(iterable[key], key, iterable); }; - } - var index = findIndexFunc(collection, predicate, fromIndex); - return index > -1 ? iterable[iteratee ? collection[index] : index] : undefined; - }; -} - -var _createFind = createFind$1; - -var toFinite = toFinite_1; - -/** - * Converts `value` to an integer. - * - * **Note:** This method is loosely based on - * [`ToInteger`](http://www.ecma-international.org/ecma-262/7.0/#sec-tointeger). - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Lang - * @param {*} value The value to convert. - * @returns {number} Returns the converted integer. - * @example - * - * _.toInteger(3.2); - * // => 3 - * - * _.toInteger(Number.MIN_VALUE); - * // => 0 - * - * _.toInteger(Infinity); - * // => 1.7976931348623157e+308 - * - * _.toInteger('3.2'); - * // => 3 - */ -function toInteger$1(value) { - var result = toFinite(value), - remainder = result % 1; - - return result === result ? (remainder ? result - remainder : result) : 0; -} - -var toInteger_1 = toInteger$1; - -var baseFindIndex = _baseFindIndex, - baseIteratee$1 = _baseIteratee, - toInteger = toInteger_1; - -/* Built-in method references for those with the same name as other `lodash` methods. */ -var nativeMax = Math.max; - -/** - * This method is like `_.find` except that it returns the index of the first - * element `predicate` returns truthy for instead of the element itself. - * - * @static - * @memberOf _ - * @since 1.1.0 - * @category Array - * @param {Array} array The array to inspect. - * @param {Function} [predicate=_.identity] The function invoked per iteration. - * @param {number} [fromIndex=0] The index to search from. - * @returns {number} Returns the index of the found element, else `-1`. - * @example - * - * var users = [ - * { 'user': 'barney', 'active': false }, - * { 'user': 'fred', 'active': false }, - * { 'user': 'pebbles', 'active': true } - * ]; - * - * _.findIndex(users, function(o) { return o.user == 'barney'; }); - * // => 0 - * - * // The `_.matches` iteratee shorthand. - * _.findIndex(users, { 'user': 'fred', 'active': false }); - * // => 1 - * - * // The `_.matchesProperty` iteratee shorthand. - * _.findIndex(users, ['active', false]); - * // => 0 - * - * // The `_.property` iteratee shorthand. - * _.findIndex(users, 'active'); - * // => 2 - */ -function findIndex$1(array, predicate, fromIndex) { - var length = array == null ? 0 : array.length; - if (!length) { - return -1; - } - var index = fromIndex == null ? 0 : toInteger(fromIndex); - if (index < 0) { - index = nativeMax(length + index, 0); - } - return baseFindIndex(array, baseIteratee$1(predicate), index); -} - -var findIndex_1 = findIndex$1; - -var createFind = _createFind, - findIndex = findIndex_1; - -/** - * Iterates over elements of `collection`, returning the first element - * `predicate` returns truthy for. The predicate is invoked with three - * arguments: (value, index|key, collection). - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Collection - * @param {Array|Object} collection The collection to inspect. - * @param {Function} [predicate=_.identity] The function invoked per iteration. - * @param {number} [fromIndex=0] The index to search from. - * @returns {*} Returns the matched element, else `undefined`. - * @example - * - * var users = [ - * { 'user': 'barney', 'age': 36, 'active': true }, - * { 'user': 'fred', 'age': 40, 'active': false }, - * { 'user': 'pebbles', 'age': 1, 'active': true } - * ]; - * - * _.find(users, function(o) { return o.age < 40; }); - * // => object for 'barney' - * - * // The `_.matches` iteratee shorthand. - * _.find(users, { 'age': 1, 'active': true }); - * // => object for 'pebbles' - * - * // The `_.matchesProperty` iteratee shorthand. - * _.find(users, ['active', false]); - * // => object for 'fred' - * - * // The `_.property` iteratee shorthand. - * _.find(users, 'active'); - * // => object for 'barney' - */ -var find = createFind(findIndex); - -var find_1 = find; - -var debounce = debounce_1$1, - isObject$3 = isObject_1$1; - -/** Error message constants. */ -var FUNC_ERROR_TEXT = 'Expected a function'; - -/** - * Creates a throttled function that only invokes `func` at most once per - * every `wait` milliseconds. The throttled function comes with a `cancel` - * method to cancel delayed `func` invocations and a `flush` method to - * immediately invoke them. Provide `options` to indicate whether `func` - * should be invoked on the leading and/or trailing edge of the `wait` - * timeout. The `func` is invoked with the last arguments provided to the - * throttled function. Subsequent calls to the throttled function return the - * result of the last `func` invocation. - * - * **Note:** If `leading` and `trailing` options are `true`, `func` is - * invoked on the trailing edge of the timeout only if the throttled function - * is invoked more than once during the `wait` timeout. - * - * If `wait` is `0` and `leading` is `false`, `func` invocation is deferred - * until to the next tick, similar to `setTimeout` with a timeout of `0`. - * - * See [David Corbacho's article](https://css-tricks.com/debouncing-throttling-explained-examples/) - * for details over the differences between `_.throttle` and `_.debounce`. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Function - * @param {Function} func The function to throttle. - * @param {number} [wait=0] The number of milliseconds to throttle invocations to. - * @param {Object} [options={}] The options object. - * @param {boolean} [options.leading=true] - * Specify invoking on the leading edge of the timeout. - * @param {boolean} [options.trailing=true] - * Specify invoking on the trailing edge of the timeout. - * @returns {Function} Returns the new throttled function. - * @example - * - * // Avoid excessively updating the position while scrolling. - * jQuery(window).on('scroll', _.throttle(updatePosition, 100)); - * - * // Invoke `renewToken` when the click event is fired, but not more than once every 5 minutes. - * var throttled = _.throttle(renewToken, 300000, { 'trailing': false }); - * jQuery(element).on('click', throttled); - * - * // Cancel the trailing throttled invocation. - * jQuery(window).on('popstate', throttled.cancel); - */ -function throttle(func, wait, options) { - var leading = true, - trailing = true; - - if (typeof func != 'function') { - throw new TypeError(FUNC_ERROR_TEXT); - } - if (isObject$3(options)) { - leading = 'leading' in options ? !!options.leading : leading; - trailing = 'trailing' in options ? !!options.trailing : trailing; - } - return debounce(func, wait, { - 'leading': leading, - 'maxWait': wait, - 'trailing': trailing - }); -} - -var throttle_1 = throttle; - -var baseGetTag = _baseGetTag$1, - isObjectLike$2 = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var boolTag$2 = '[object Boolean]'; - -/** - * Checks if `value` is classified as a boolean primitive or object. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a boolean, else `false`. - * @example - * - * _.isBoolean(false); - * // => true - * - * _.isBoolean(null); - * // => false - */ -function isBoolean(value) { - return value === true || value === false || - (isObjectLike$2(value) && baseGetTag(value) == boolTag$2); -} - -var isBoolean_1 = isBoolean; - -var detectReferenceElementsDomain = function detectReferenceElementsDomain(children, domain, axisId, axisType, specifiedTicks) { - var lines = findAllByType(children, ReferenceLine.displayName); - var dots = findAllByType(children, ReferenceDot.displayName); - var elements = lines.concat(dots); - var areas = findAllByType(children, ReferenceArea.displayName); - var idKey = "".concat(axisType, "Id"); - var valueKey = axisType[0]; - var finalDomain = domain; - - if (elements.length) { - finalDomain = elements.reduce(function (result, el) { - if (el.props[idKey] === axisId && ifOverflowMatches(el.props, 'extendDomain') && isNumber(el.props[valueKey])) { - var value = el.props[valueKey]; - return [Math.min(result[0], value), Math.max(result[1], value)]; - } - - return result; - }, finalDomain); - } - - if (areas.length) { - var key1 = "".concat(valueKey, "1"); - var key2 = "".concat(valueKey, "2"); - finalDomain = areas.reduce(function (result, el) { - if (el.props[idKey] === axisId && ifOverflowMatches(el.props, 'extendDomain') && isNumber(el.props[key1]) && isNumber(el.props[key2])) { - var value1 = el.props[key1]; - var value2 = el.props[key2]; - return [Math.min(result[0], value1, value2), Math.max(result[1], value1, value2)]; - } - - return result; - }, finalDomain); - } - - if (specifiedTicks && specifiedTicks.length) { - finalDomain = specifiedTicks.reduce(function (result, tick) { - if (isNumber(tick)) { - return [Math.min(result[0], tick), Math.max(result[1], tick)]; - } - - return result; - }, finalDomain); - } - - return finalDomain; -}; - -var eventemitter3 = {exports: {}}; - -(function (module) { - - var has = Object.prototype.hasOwnProperty - , prefix = '~'; - - /** - * Constructor to create a storage for our `EE` objects. - * An `Events` instance is a plain object whose properties are event names. - * - * @constructor - * @private - */ - function Events() {} - - // - // We try to not inherit from `Object.prototype`. In some engines creating an - // instance in this way is faster than calling `Object.create(null)` directly. - // If `Object.create(null)` is not supported we prefix the event names with a - // character to make sure that the built-in object properties are not - // overridden or used as an attack vector. - // - if (Object.create) { - Events.prototype = Object.create(null); - - // - // This hack is needed because the `__proto__` property is still inherited in - // some old browsers like Android 4, iPhone 5.1, Opera 11 and Safari 5. - // - if (!new Events().__proto__) prefix = false; - } - - /** - * Representation of a single event listener. - * - * @param {Function} fn The listener function. - * @param {*} context The context to invoke the listener with. - * @param {Boolean} [once=false] Specify if the listener is a one-time listener. - * @constructor - * @private - */ - function EE(fn, context, once) { - this.fn = fn; - this.context = context; - this.once = once || false; - } - - /** - * Add a listener for a given event. - * - * @param {EventEmitter} emitter Reference to the `EventEmitter` instance. - * @param {(String|Symbol)} event The event name. - * @param {Function} fn The listener function. - * @param {*} context The context to invoke the listener with. - * @param {Boolean} once Specify if the listener is a one-time listener. - * @returns {EventEmitter} - * @private - */ - function addListener(emitter, event, fn, context, once) { - if (typeof fn !== 'function') { - throw new TypeError('The listener must be a function'); - } - - var listener = new EE(fn, context || emitter, once) - , evt = prefix ? prefix + event : event; - - if (!emitter._events[evt]) emitter._events[evt] = listener, emitter._eventsCount++; - else if (!emitter._events[evt].fn) emitter._events[evt].push(listener); - else emitter._events[evt] = [emitter._events[evt], listener]; - - return emitter; - } - - /** - * Clear event by name. - * - * @param {EventEmitter} emitter Reference to the `EventEmitter` instance. - * @param {(String|Symbol)} evt The Event name. - * @private - */ - function clearEvent(emitter, evt) { - if (--emitter._eventsCount === 0) emitter._events = new Events(); - else delete emitter._events[evt]; - } - - /** - * Minimal `EventEmitter` interface that is molded against the Node.js - * `EventEmitter` interface. - * - * @constructor - * @public - */ - function EventEmitter() { - this._events = new Events(); - this._eventsCount = 0; - } - - /** - * Return an array listing the events for which the emitter has registered - * listeners. - * - * @returns {Array} - * @public - */ - EventEmitter.prototype.eventNames = function eventNames() { - var names = [] - , events - , name; - - if (this._eventsCount === 0) return names; - - for (name in (events = this._events)) { - if (has.call(events, name)) names.push(prefix ? name.slice(1) : name); - } - - if (Object.getOwnPropertySymbols) { - return names.concat(Object.getOwnPropertySymbols(events)); - } - - return names; - }; - - /** - * Return the listeners registered for a given event. - * - * @param {(String|Symbol)} event The event name. - * @returns {Array} The registered listeners. - * @public - */ - EventEmitter.prototype.listeners = function listeners(event) { - var evt = prefix ? prefix + event : event - , handlers = this._events[evt]; - - if (!handlers) return []; - if (handlers.fn) return [handlers.fn]; - - for (var i = 0, l = handlers.length, ee = new Array(l); i < l; i++) { - ee[i] = handlers[i].fn; - } - - return ee; - }; - - /** - * Return the number of listeners listening to a given event. - * - * @param {(String|Symbol)} event The event name. - * @returns {Number} The number of listeners. - * @public - */ - EventEmitter.prototype.listenerCount = function listenerCount(event) { - var evt = prefix ? prefix + event : event - , listeners = this._events[evt]; - - if (!listeners) return 0; - if (listeners.fn) return 1; - return listeners.length; - }; - - /** - * Calls each of the listeners registered for a given event. - * - * @param {(String|Symbol)} event The event name. - * @returns {Boolean} `true` if the event had listeners, else `false`. - * @public - */ - EventEmitter.prototype.emit = function emit(event, a1, a2, a3, a4, a5) { - var evt = prefix ? prefix + event : event; - - if (!this._events[evt]) return false; - - var listeners = this._events[evt] - , len = arguments.length - , args - , i; - - if (listeners.fn) { - if (listeners.once) this.removeListener(event, listeners.fn, undefined, true); - - switch (len) { - case 1: return listeners.fn.call(listeners.context), true; - case 2: return listeners.fn.call(listeners.context, a1), true; - case 3: return listeners.fn.call(listeners.context, a1, a2), true; - case 4: return listeners.fn.call(listeners.context, a1, a2, a3), true; - case 5: return listeners.fn.call(listeners.context, a1, a2, a3, a4), true; - case 6: return listeners.fn.call(listeners.context, a1, a2, a3, a4, a5), true; - } - - for (i = 1, args = new Array(len -1); i < len; i++) { - args[i - 1] = arguments[i]; - } - - listeners.fn.apply(listeners.context, args); - } else { - var length = listeners.length - , j; - - for (i = 0; i < length; i++) { - if (listeners[i].once) this.removeListener(event, listeners[i].fn, undefined, true); - - switch (len) { - case 1: listeners[i].fn.call(listeners[i].context); break; - case 2: listeners[i].fn.call(listeners[i].context, a1); break; - case 3: listeners[i].fn.call(listeners[i].context, a1, a2); break; - case 4: listeners[i].fn.call(listeners[i].context, a1, a2, a3); break; - default: - if (!args) for (j = 1, args = new Array(len -1); j < len; j++) { - args[j - 1] = arguments[j]; - } - - listeners[i].fn.apply(listeners[i].context, args); - } - } - } - - return true; - }; - - /** - * Add a listener for a given event. - * - * @param {(String|Symbol)} event The event name. - * @param {Function} fn The listener function. - * @param {*} [context=this] The context to invoke the listener with. - * @returns {EventEmitter} `this`. - * @public - */ - EventEmitter.prototype.on = function on(event, fn, context) { - return addListener(this, event, fn, context, false); - }; - - /** - * Add a one-time listener for a given event. - * - * @param {(String|Symbol)} event The event name. - * @param {Function} fn The listener function. - * @param {*} [context=this] The context to invoke the listener with. - * @returns {EventEmitter} `this`. - * @public - */ - EventEmitter.prototype.once = function once(event, fn, context) { - return addListener(this, event, fn, context, true); - }; - - /** - * Remove the listeners of a given event. - * - * @param {(String|Symbol)} event The event name. - * @param {Function} fn Only remove the listeners that match this function. - * @param {*} context Only remove the listeners that have this context. - * @param {Boolean} once Only remove one-time listeners. - * @returns {EventEmitter} `this`. - * @public - */ - EventEmitter.prototype.removeListener = function removeListener(event, fn, context, once) { - var evt = prefix ? prefix + event : event; - - if (!this._events[evt]) return this; - if (!fn) { - clearEvent(this, evt); - return this; - } - - var listeners = this._events[evt]; - - if (listeners.fn) { - if ( - listeners.fn === fn && - (!once || listeners.once) && - (!context || listeners.context === context) - ) { - clearEvent(this, evt); - } - } else { - for (var i = 0, events = [], length = listeners.length; i < length; i++) { - if ( - listeners[i].fn !== fn || - (once && !listeners[i].once) || - (context && listeners[i].context !== context) - ) { - events.push(listeners[i]); - } - } - - // - // Reset the array, or remove it completely if we have no more listeners. - // - if (events.length) this._events[evt] = events.length === 1 ? events[0] : events; - else clearEvent(this, evt); - } - - return this; - }; - - /** - * Remove all listeners, or those of the specified event. - * - * @param {(String|Symbol)} [event] The event name. - * @returns {EventEmitter} `this`. - * @public - */ - EventEmitter.prototype.removeAllListeners = function removeAllListeners(event) { - var evt; - - if (event) { - evt = prefix ? prefix + event : event; - if (this._events[evt]) clearEvent(this, evt); - } else { - this._events = new Events(); - this._eventsCount = 0; - } - - return this; - }; - - // - // Alias methods names because people roll like that. - // - EventEmitter.prototype.off = EventEmitter.prototype.removeListener; - EventEmitter.prototype.addListener = EventEmitter.prototype.on; - - // - // Expose the prefix. - // - EventEmitter.prefixed = prefix; - - // - // Allow `EventEmitter` to be imported as module namespace. - // - EventEmitter.EventEmitter = EventEmitter; - - // - // Expose the module. - // - { - module.exports = EventEmitter; - } -} (eventemitter3)); - -var EventEmitter = eventemitter3.exports; - -var eventCenter = new EventEmitter(); - -if (eventCenter.setMaxListeners) { - eventCenter.setMaxListeners(10); -} -var SYNC_EVENT = 'recharts.syncMouseEvents'; // eslint-disable-next-line no-redeclare - -function _typeof$4(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$4 = function _typeof(obj) { return typeof obj; }; } else { _typeof$4 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$4(obj); } - -function _slicedToArray$1(arr, i) { return _arrayWithHoles$1(arr) || _iterableToArrayLimit$1(arr, i) || _unsupportedIterableToArray$1(arr, i) || _nonIterableRest$1(); } - -function _nonIterableRest$1() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _iterableToArrayLimit$1(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles$1(arr) { if (Array.isArray(arr)) return arr; } - -function _extends$4() { _extends$4 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$4.apply(this, arguments); } - -function _objectWithoutProperties$2(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$2(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$2(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _classCallCheck$4(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$4(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$4(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$4(Constructor.prototype, protoProps); if (staticProps) _defineProperties$4(Constructor, staticProps); return Constructor; } - -function _inherits$4(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$4(subClass, superClass); } - -function _setPrototypeOf$4(o, p) { _setPrototypeOf$4 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$4(o, p); } - -function _createSuper$4(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$4(); return function _createSuperInternal() { var Super = _getPrototypeOf$4(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$4(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$4(this, result); }; } - -function _possibleConstructorReturn$4(self, call) { if (call && (_typeof$4(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$4(self); } - -function _assertThisInitialized$4(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$4() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$4(o) { _getPrototypeOf$4 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$4(o); } - -function _toConsumableArray(arr) { return _arrayWithoutHoles(arr) || _iterableToArray(arr) || _unsupportedIterableToArray$1(arr) || _nonIterableSpread(); } - -function _nonIterableSpread() { throw new TypeError("Invalid attempt to spread non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray$1(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray$1(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray$1(o, minLen); } - -function _iterableToArray(iter) { if (typeof Symbol !== "undefined" && Symbol.iterator in Object(iter)) return Array.from(iter); } - -function _arrayWithoutHoles(arr) { if (Array.isArray(arr)) return _arrayLikeToArray$1(arr); } - -function _arrayLikeToArray$1(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function ownKeys$3(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$3(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$3(Object(source), true).forEach(function (key) { _defineProperty$3(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$3(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$3(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } -var ORIENT_MAP = { - xAxis: ['bottom', 'top'], - yAxis: ['left', 'right'] -}; -var originCoordinate = { - x: 0, - y: 0 -}; // use legacy isFinite only if there is a problem (aka IE) -// eslint-disable-next-line no-restricted-globals - -var isFinit = Number.isFinite ? Number.isFinite : isFinite; -var defer = // eslint-disable-next-line no-nested-ternary -typeof requestAnimationFrame === 'function' ? requestAnimationFrame : typeof setImmediate === 'function' ? setImmediate : setTimeout; -var deferClear = // eslint-disable-next-line no-nested-ternary -typeof cancelAnimationFrame === 'function' ? cancelAnimationFrame : typeof clearImmediate === 'function' ? clearImmediate : clearTimeout; - -var calculateTooltipPos = function calculateTooltipPos(rangeObj, layout) { - if (layout === 'horizontal') { - return rangeObj.x; - } - - if (layout === 'vertical') { - return rangeObj.y; - } - - if (layout === 'centric') { - return rangeObj.angle; - } - - return rangeObj.radius; -}; - -var getActiveCoordinate = function getActiveCoordinate(layout, tooltipTicks, activeIndex, rangeObj) { - var entry = tooltipTicks.find(function (tick) { - return tick && tick.index === activeIndex; - }); - - if (entry) { - if (layout === 'horizontal') { - return { - x: entry.coordinate, - y: rangeObj.y - }; - } - - if (layout === 'vertical') { - return { - x: rangeObj.x, - y: entry.coordinate - }; - } - - if (layout === 'centric') { - var _angle = entry.coordinate; - var _radius = rangeObj.radius; - return _objectSpread$3(_objectSpread$3(_objectSpread$3({}, rangeObj), polarToCartesian(rangeObj.cx, rangeObj.cy, _radius, _angle)), {}, { - angle: _angle, - radius: _radius - }); - } - - var radius = entry.coordinate; - var angle = rangeObj.angle; - return _objectSpread$3(_objectSpread$3(_objectSpread$3({}, rangeObj), polarToCartesian(rangeObj.cx, rangeObj.cy, radius, angle)), {}, { - angle: angle, - radius: radius - }); - } - - return originCoordinate; -}; - -var getDisplayedData = function getDisplayedData(data, _ref, item) { - var graphicalItems = _ref.graphicalItems, - dataStartIndex = _ref.dataStartIndex, - dataEndIndex = _ref.dataEndIndex; - var itemsData = (graphicalItems || []).reduce(function (result, child) { - var itemData = child.props.data; - - if (itemData && itemData.length) { - return [].concat(_toConsumableArray(result), _toConsumableArray(itemData)); - } - - return result; - }, []); - - if (itemsData && itemsData.length > 0) { - return itemsData; - } - - if (item && item.props && item.props.data && item.props.data.length > 0) { - return item.props.data; - } - - if (data && data.length && isNumber(dataStartIndex) && isNumber(dataEndIndex)) { - return data.slice(dataStartIndex, dataEndIndex + 1); - } - - return []; -}; -/** - * Get the content to be displayed in the tooltip - * @param {Object} state Current state - * @param {Array} chartData The data defined in chart - * @param {Number} activeIndex Active index of data - * @param {String} activeLabel Active label of data - * @return {Array} The content of tooltip - */ - - -var getTooltipContent = function getTooltipContent(state, chartData, activeIndex, activeLabel) { - var graphicalItems = state.graphicalItems, - tooltipAxis = state.tooltipAxis; - var displayedData = getDisplayedData(chartData, state); - - if (activeIndex < 0 || !graphicalItems || !graphicalItems.length || activeIndex >= displayedData.length) { - return null; - } // get data by activeIndex when the axis don't allow duplicated category - - - return graphicalItems.reduce(function (result, child) { - var hide = child.props.hide; - - if (hide) { - return result; - } - - var data = child.props.data; - var payload; - - if (tooltipAxis.dataKey && !tooltipAxis.allowDuplicatedCategory) { - // graphic child has data props - var entries = data === undefined ? displayedData : data; - payload = findEntryInArray(entries, tooltipAxis.dataKey, activeLabel); - } else { - payload = data && data[activeIndex] || displayedData[activeIndex]; - } - - if (!payload) { - return result; - } - - return [].concat(_toConsumableArray(result), [getTooltipItem(child, payload)]); - }, []); -}; -/** - * Returns tooltip data based on a mouse position (as a parameter or in state) - * @param {Object} state current state - * @param {Array} chartData the data defined in chart - * @param {String} layout The layout type of chart - * @param {Object} rangeObj { x, y } coordinates - * @return {Object} Tooltip data data - */ - - -var getTooltipData = function getTooltipData(state, chartData, layout, rangeObj) { - var rangeData = rangeObj || { - x: state.chartX, - y: state.chartY - }; - var pos = calculateTooltipPos(rangeData, layout); - var ticks = state.orderedTooltipTicks, - axis = state.tooltipAxis, - tooltipTicks = state.tooltipTicks; - var activeIndex = calculateActiveTickIndex(pos, ticks, tooltipTicks, axis); - - if (activeIndex >= 0 && tooltipTicks) { - var activeLabel = tooltipTicks[activeIndex] && tooltipTicks[activeIndex].value; - var activePayload = getTooltipContent(state, chartData, activeIndex, activeLabel); - var activeCoordinate = getActiveCoordinate(layout, ticks, activeIndex, rangeData); - return { - activeTooltipIndex: activeIndex, - activeLabel: activeLabel, - activePayload: activePayload, - activeCoordinate: activeCoordinate - }; - } - - return null; -}; -/** - * Get the configuration of axis by the options of axis instance - * @param {Object} props Latest props - * @param {Array} axes The instance of axes - * @param {Array} graphicalItems The instances of item - * @param {String} axisType The type of axis, xAxis - x-axis, yAxis - y-axis - * @param {String} axisIdKey The unique id of an axis - * @param {Object} stackGroups The items grouped by axisId and stackId - * @param {Number} dataStartIndex The start index of the data series when a brush is applied - * @param {Number} dataEndIndex The end index of the data series when a brush is applied - * @return {Object} Configuration - */ - - -var getAxisMapByAxes = function getAxisMapByAxes(props, _ref2) { - var axes = _ref2.axes, - graphicalItems = _ref2.graphicalItems, - axisType = _ref2.axisType, - axisIdKey = _ref2.axisIdKey, - stackGroups = _ref2.stackGroups, - dataStartIndex = _ref2.dataStartIndex, - dataEndIndex = _ref2.dataEndIndex; - var layout = props.layout, - children = props.children, - stackOffset = props.stackOffset; - var isCategorical = isCategoricalAxis(layout, axisType); // Eliminate duplicated axes - - var axisMap = axes.reduce(function (result, child) { - var _child$props = child.props, - type = _child$props.type, - dataKey = _child$props.dataKey, - allowDataOverflow = _child$props.allowDataOverflow, - allowDuplicatedCategory = _child$props.allowDuplicatedCategory, - scale = _child$props.scale, - ticks = _child$props.ticks; - var axisId = child.props[axisIdKey]; - var displayedData = getDisplayedData(props.data, { - graphicalItems: graphicalItems.filter(function (item) { - return item.props[axisIdKey] === axisId; - }), - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex - }); - var len = displayedData.length; - - if (!result[axisId]) { - var domain, duplicateDomain, categoricalDomain; - - if (dataKey) { - // has dataKey in - domain = getDomainOfDataByKey(displayedData, dataKey, type); - - if (type === 'category' && isCategorical) { - // the field type is category data and this axis is catrgorical axis - var duplicate = hasDuplicate(domain); - - if (allowDuplicatedCategory && duplicate) { - duplicateDomain = domain; // When category axis has duplicated text, serial numbers are used to generate scale - - domain = range_1(0, len); - } else if (!allowDuplicatedCategory) { - // remove duplicated category - domain = parseDomainOfCategoryAxis(child.props.domain, domain, child).reduce(function (finalDomain, entry) { - return finalDomain.indexOf(entry) >= 0 ? finalDomain : [].concat(_toConsumableArray(finalDomain), [entry]); - }, []); - } - } else if (type === 'category') { - // the field type is category data and this axis is numerical axis - if (!allowDuplicatedCategory) { - domain = parseDomainOfCategoryAxis(child.props.domain, domain, child).reduce(function (finalDomain, entry) { - return finalDomain.indexOf(entry) >= 0 || entry === '' || isNil_1(entry) ? finalDomain : [].concat(_toConsumableArray(finalDomain), [entry]); - }, []); - } else { - // eliminate undefined or null or empty string - domain = domain.filter(function (entry) { - return entry !== '' && !isNil_1(entry); - }); - } - } else if (type === 'number') { - // the field type is numerical - var errorBarsDomain = parseErrorBarsOfAxis(displayedData, graphicalItems.filter(function (item) { - return item.props[axisIdKey] === axisId && !item.props.hide; - }), dataKey, axisType, layout); - - if (errorBarsDomain) { - domain = errorBarsDomain; - } - } - - if (isCategorical && (type === 'number' || scale !== 'auto')) { - categoricalDomain = getDomainOfDataByKey(displayedData, dataKey, 'category'); - } - } else if (isCategorical) { - // the axis is a categorical axis - domain = range_1(0, len); - } else if (stackGroups && stackGroups[axisId] && stackGroups[axisId].hasStack && type === 'number') { - // when stackOffset is 'expand', the domain may be calculated as [0, 1.000000000002] - domain = stackOffset === 'expand' ? [0, 1] : getDomainOfStackGroups(stackGroups[axisId].stackGroups, dataStartIndex, dataEndIndex); - } else { - domain = getDomainOfItemsWithSameAxis(displayedData, graphicalItems.filter(function (item) { - return item.props[axisIdKey] === axisId && !item.props.hide; - }), type, layout, true); - } - - if (type === 'number') { - // To detect wether there is any reference lines whose props alwaysShow is true - domain = detectReferenceElementsDomain(children, domain, axisId, axisType, ticks); - - if (child.props.domain) { - domain = parseSpecifiedDomain(child.props.domain, domain, allowDataOverflow); - } - } else if (type === 'category' && child.props.domain) { - var axisDomain = child.props.domain; - var isDomainValidate = domain.every(function (entry) { - return axisDomain.indexOf(entry) >= 0; - }); - - if (isDomainValidate) { - domain = axisDomain; - } - } - - return _objectSpread$3(_objectSpread$3({}, result), {}, _defineProperty$3({}, axisId, _objectSpread$3(_objectSpread$3({}, child.props), {}, { - axisType: axisType, - domain: domain, - categoricalDomain: categoricalDomain, - duplicateDomain: duplicateDomain, - originalDomain: child.props.domain, - isCategorical: isCategorical, - layout: layout - }))); - } - - return result; - }, {}); - return axisMap; -}; -/** - * Get the configuration of axis by the options of item, - * this kind of axis does not display in chart - * @param {Object} props Latest props - * @param {Array} graphicalItems The instances of item - * @param {ReactElement} Axis Axis Component - * @param {String} axisType The type of axis, xAxis - x-axis, yAxis - y-axis - * @param {String} axisIdKey The unique id of an axis - * @param {Object} stackGroups The items grouped by axisId and stackId - * @param {Number} dataStartIndex The start index of the data series when a brush is applied - * @param {Number} dataEndIndex The end index of the data series when a brush is applied - * @return {Object} Configuration - */ - - -var getAxisMapByItems = function getAxisMapByItems(props, _ref3) { - var graphicalItems = _ref3.graphicalItems, - Axis = _ref3.Axis, - axisType = _ref3.axisType, - axisIdKey = _ref3.axisIdKey, - stackGroups = _ref3.stackGroups, - dataStartIndex = _ref3.dataStartIndex, - dataEndIndex = _ref3.dataEndIndex; - var layout = props.layout, - children = props.children; - var displayedData = getDisplayedData(props.data, { - graphicalItems: graphicalItems, - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex - }); - var len = displayedData.length; - var isCategorical = isCategoricalAxis(layout, axisType); - var index = -1; // The default type of x-axis is category axis, - // The default contents of x-axis is the serial numbers of data - // The default type of y-axis is number axis - // The default contents of y-axis is the domain of data - - var axisMap = graphicalItems.reduce(function (result, child) { - var axisId = child.props[axisIdKey]; - - if (!result[axisId]) { - index++; - var domain; - - if (isCategorical) { - domain = range_1(0, len); - } else if (stackGroups && stackGroups[axisId] && stackGroups[axisId].hasStack) { - domain = getDomainOfStackGroups(stackGroups[axisId].stackGroups, dataStartIndex, dataEndIndex); - domain = detectReferenceElementsDomain(children, domain, axisId, axisType); - } else { - domain = parseSpecifiedDomain(Axis.defaultProps.domain, getDomainOfItemsWithSameAxis(displayedData, graphicalItems.filter(function (item) { - return item.props[axisIdKey] === axisId && !item.props.hide; - }), 'number', layout), Axis.defaultProps.allowDataOverflow); - domain = detectReferenceElementsDomain(children, domain, axisId, axisType); - } - - return _objectSpread$3(_objectSpread$3({}, result), {}, _defineProperty$3({}, axisId, _objectSpread$3(_objectSpread$3({ - axisType: axisType - }, Axis.defaultProps), {}, { - hide: true, - orientation: get_1(ORIENT_MAP, "".concat(axisType, ".").concat(index % 2), null), - domain: domain, - originalDomain: Axis.defaultProps.domain, - isCategorical: isCategorical, - layout: layout // specify scale when no Axis - // scale: isCategorical ? 'band' : 'linear', - - }))); - } - - return result; - }, {}); - return axisMap; -}; -/** - * Get the configuration of all x-axis or y-axis - * @param {Object} props Latest props - * @param {String} axisType The type of axis - * @param {Array} graphicalItems The instances of item - * @param {Object} stackGroups The items grouped by axisId and stackId - * @param {Number} dataStartIndex The start index of the data series when a brush is applied - * @param {Number} dataEndIndex The end index of the data series when a brush is applied - * @return {Object} Configuration - */ - - -var getAxisMap = function getAxisMap(props, _ref4) { - var _ref4$axisType = _ref4.axisType, - axisType = _ref4$axisType === void 0 ? 'xAxis' : _ref4$axisType, - AxisComp = _ref4.AxisComp, - graphicalItems = _ref4.graphicalItems, - stackGroups = _ref4.stackGroups, - dataStartIndex = _ref4.dataStartIndex, - dataEndIndex = _ref4.dataEndIndex; - var children = props.children; - var axisIdKey = "".concat(axisType, "Id"); // Get all the instance of Axis - - var axes = findAllByType(children, AxisComp); - var axisMap = {}; - - if (axes && axes.length) { - axisMap = getAxisMapByAxes(props, { - axes: axes, - graphicalItems: graphicalItems, - axisType: axisType, - axisIdKey: axisIdKey, - stackGroups: stackGroups, - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex - }); - } else if (graphicalItems && graphicalItems.length) { - axisMap = getAxisMapByItems(props, { - Axis: AxisComp, - graphicalItems: graphicalItems, - axisType: axisType, - axisIdKey: axisIdKey, - stackGroups: stackGroups, - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex - }); - } - - return axisMap; -}; - -var tooltipTicksGenerator = function tooltipTicksGenerator(axisMap) { - var axis = getAnyElementOfObject(axisMap); - var tooltipTicks = getTicksOfAxis(axis, false, true); - return { - tooltipTicks: tooltipTicks, - orderedTooltipTicks: sortBy_1(tooltipTicks, function (o) { - return o.coordinate; - }), - tooltipAxis: axis, - tooltipAxisBandSize: getBandSizeOfAxis(axis, tooltipTicks) - }; -}; -/** - * Returns default, reset state for the categorical chart. - * @param {Object} props Props object to use when creating the default state - * @return {Object} Whole new state - */ - - -var createDefaultState = function createDefaultState(props) { - var children = props.children, - defaultShowTooltip = props.defaultShowTooltip; - var brushItem = findChildByType(children, Brush.displayName); - var startIndex = brushItem && brushItem.props && brushItem.props.startIndex || 0; - var endIndex = brushItem && brushItem.props && brushItem.props.endIndex || props.data && props.data.length - 1 || 0; - return { - chartX: 0, - chartY: 0, - dataStartIndex: startIndex, - dataEndIndex: endIndex, - activeTooltipIndex: -1, - isTooltipActive: !isNil_1(defaultShowTooltip) ? defaultShowTooltip : false - }; -}; - -var hasGraphicalBarItem = function hasGraphicalBarItem(graphicalItems) { - if (!graphicalItems || !graphicalItems.length) { - return false; - } - - return graphicalItems.some(function (item) { - var name = getDisplayName(item && item.type); - return name && name.indexOf('Bar') >= 0; - }); -}; - -var getAxisNameByLayout = function getAxisNameByLayout(layout) { - if (layout === 'horizontal') { - return { - numericAxisName: 'yAxis', - cateAxisName: 'xAxis' - }; - } - - if (layout === 'vertical') { - return { - numericAxisName: 'xAxis', - cateAxisName: 'yAxis' - }; - } - - if (layout === 'centric') { - return { - numericAxisName: 'radiusAxis', - cateAxisName: 'angleAxis' - }; - } - - return { - numericAxisName: 'angleAxis', - cateAxisName: 'radiusAxis' - }; -}; -/** - * Calculate the offset of main part in the svg element - * @param {Object} props Latest props - * graphicalItems The instances of item - * xAxisMap The configuration of x-axis - * yAxisMap The configuration of y-axis - * @param {Object} prevLegendBBox the boundary box of legend - * @return {Object} The offset of main part in the svg element - */ - - -var calculateOffset = function calculateOffset(_ref5, prevLegendBBox) { - var props = _ref5.props, - graphicalItems = _ref5.graphicalItems, - _ref5$xAxisMap = _ref5.xAxisMap, - xAxisMap = _ref5$xAxisMap === void 0 ? {} : _ref5$xAxisMap, - _ref5$yAxisMap = _ref5.yAxisMap, - yAxisMap = _ref5$yAxisMap === void 0 ? {} : _ref5$yAxisMap; - var width = props.width, - height = props.height, - children = props.children; - var margin = props.margin || {}; - var brushItem = findChildByType(children, Brush.displayName); - var legendItem = findChildByType(children, Legend.displayName); - var offsetH = Object.keys(yAxisMap).reduce(function (result, id) { - var entry = yAxisMap[id]; - var orientation = entry.orientation; - - if (!entry.mirror && !entry.hide) { - return _objectSpread$3(_objectSpread$3({}, result), {}, _defineProperty$3({}, orientation, result[orientation] + entry.width)); - } - - return result; - }, { - left: margin.left || 0, - right: margin.right || 0 - }); - var offsetV = Object.keys(xAxisMap).reduce(function (result, id) { - var entry = xAxisMap[id]; - var orientation = entry.orientation; - - if (!entry.mirror && !entry.hide) { - return _objectSpread$3(_objectSpread$3({}, result), {}, _defineProperty$3({}, orientation, get_1(result, "".concat(orientation)) + entry.height)); - } - - return result; - }, { - top: margin.top || 0, - bottom: margin.bottom || 0 - }); - - var offset = _objectSpread$3(_objectSpread$3({}, offsetV), offsetH); - - var brushBottom = offset.bottom; - - if (brushItem) { - offset.bottom += brushItem.props.height || Brush.defaultProps.height; - } - - if (legendItem && prevLegendBBox) { - offset = appendOffsetOfLegend(offset, graphicalItems, props, prevLegendBBox); - } - - return _objectSpread$3(_objectSpread$3({ - brushBottom: brushBottom - }, offset), {}, { - width: width - offset.left - offset.right, - height: height - offset.top - offset.bottom - }); -}; - -var generateCategoricalChart = function generateCategoricalChart(_ref6) { - var _class, _temp; - - var chartName = _ref6.chartName, - GraphicalChild = _ref6.GraphicalChild, - _ref6$defaultTooltipE = _ref6.defaultTooltipEventType, - defaultTooltipEventType = _ref6$defaultTooltipE === void 0 ? 'axis' : _ref6$defaultTooltipE, - _ref6$validateTooltip = _ref6.validateTooltipEventTypes, - validateTooltipEventTypes = _ref6$validateTooltip === void 0 ? ['axis'] : _ref6$validateTooltip, - axisComponents = _ref6.axisComponents, - legendContent = _ref6.legendContent, - formatAxisMap = _ref6.formatAxisMap, - defaultProps = _ref6.defaultProps; - - var getFormatItems = function getFormatItems(props, currentState) { - var graphicalItems = currentState.graphicalItems, - stackGroups = currentState.stackGroups, - offset = currentState.offset, - updateId = currentState.updateId, - dataStartIndex = currentState.dataStartIndex, - dataEndIndex = currentState.dataEndIndex; - var barSize = props.barSize, - layout = props.layout, - barGap = props.barGap, - barCategoryGap = props.barCategoryGap, - globalMaxBarSize = props.maxBarSize; - - var _getAxisNameByLayout = getAxisNameByLayout(layout), - numericAxisName = _getAxisNameByLayout.numericAxisName, - cateAxisName = _getAxisNameByLayout.cateAxisName; - - var hasBar = hasGraphicalBarItem(graphicalItems); - var sizeList = hasBar && getBarSizeList({ - barSize: barSize, - stackGroups: stackGroups - }); - var formattedItems = []; - graphicalItems.forEach(function (item, index) { - var displayedData = getDisplayedData(props.data, { - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex - }, item); - var _item$props = item.props, - dataKey = _item$props.dataKey, - childMaxBarSize = _item$props.maxBarSize; - var numericAxisId = item.props["".concat(numericAxisName, "Id")]; - var cateAxisId = item.props["".concat(cateAxisName, "Id")]; - var axisObj = axisComponents.reduce(function (result, entry) { - var _objectSpread6; - - var axisMap = currentState["".concat(entry.axisType, "Map")]; - var id = item.props["".concat(entry.axisType, "Id")]; - var axis = axisMap && axisMap[id]; - return _objectSpread$3(_objectSpread$3({}, result), {}, (_objectSpread6 = {}, _defineProperty$3(_objectSpread6, entry.axisType, axis), _defineProperty$3(_objectSpread6, "".concat(entry.axisType, "Ticks"), getTicksOfAxis(axis)), _objectSpread6)); - }, {}); - var cateAxis = axisObj[cateAxisName]; - var cateTicks = axisObj["".concat(cateAxisName, "Ticks")]; - var stackedData = stackGroups && stackGroups[numericAxisId] && stackGroups[numericAxisId].hasStack && getStackedDataOfItem(item, stackGroups[numericAxisId].stackGroups); - var itemIsBar = getDisplayName(item.type).indexOf('Bar') >= 0; - var bandSize = getBandSizeOfAxis(cateAxis, cateTicks); - var barPosition = []; - - if (itemIsBar) { - var _ref7, _getBandSizeOfAxis; - - // 如果是bar,计算bar的位置 - var maxBarSize = isNil_1(childMaxBarSize) ? globalMaxBarSize : childMaxBarSize; - var barBandSize = (_ref7 = (_getBandSizeOfAxis = getBandSizeOfAxis(cateAxis, cateTicks, true)) !== null && _getBandSizeOfAxis !== void 0 ? _getBandSizeOfAxis : maxBarSize) !== null && _ref7 !== void 0 ? _ref7 : 0; - barPosition = getBarPosition({ - barGap: barGap, - barCategoryGap: barCategoryGap, - bandSize: barBandSize !== bandSize ? barBandSize : bandSize, - sizeList: sizeList[cateAxisId], - maxBarSize: maxBarSize - }); - - if (barBandSize !== bandSize) { - barPosition = barPosition.map(function (pos) { - return _objectSpread$3(_objectSpread$3({}, pos), {}, { - position: _objectSpread$3(_objectSpread$3({}, pos.position), {}, { - offset: pos.position.offset - barBandSize / 2 - }) - }); - }); - } - } - - var composedFn = item && item.type && item.type.getComposedData; - - if (composedFn) { - var _objectSpread7; - - formattedItems.push({ - props: _objectSpread$3(_objectSpread$3({}, composedFn(_objectSpread$3(_objectSpread$3({}, axisObj), {}, { - displayedData: displayedData, - props: props, - dataKey: dataKey, - item: item, - bandSize: bandSize, - barPosition: barPosition, - offset: offset, - stackedData: stackedData, - layout: layout, - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex - }))), {}, (_objectSpread7 = { - key: item.key || "item-".concat(index) - }, _defineProperty$3(_objectSpread7, numericAxisName, axisObj[numericAxisName]), _defineProperty$3(_objectSpread7, cateAxisName, axisObj[cateAxisName]), _defineProperty$3(_objectSpread7, "animationId", updateId), _objectSpread7)), - childIndex: parseChildIndex(item, props.children), - item: item - }); - } - }); - return formattedItems; - }; - /** - * The AxisMaps are expensive to render on large data sets - * so provide the ability to store them in state and only update them when necessary - * they are dependent upon the start and end index of - * the brush so it's important that this method is called _after_ - * the state is updated with any new start/end indices - * - * @param {Object} props The props object to be used for updating the axismaps - * dataStartIndex: The start index of the data series when a brush is applied - * dataEndIndex: The end index of the data series when a brush is applied - * updateId: The update id - * @param {Object} prevState Prev state - * @return {Object} state New state to set - */ - - - var updateStateOfAxisMapsOffsetAndStackGroups = function updateStateOfAxisMapsOffsetAndStackGroups(_ref8, prevState) { - var props = _ref8.props, - dataStartIndex = _ref8.dataStartIndex, - dataEndIndex = _ref8.dataEndIndex, - updateId = _ref8.updateId; - - if (!validateWidthHeight({ - props: props - })) { - return null; - } - - var children = props.children, - layout = props.layout, - stackOffset = props.stackOffset, - data = props.data, - reverseStackOrder = props.reverseStackOrder; - - var _getAxisNameByLayout2 = getAxisNameByLayout(layout), - numericAxisName = _getAxisNameByLayout2.numericAxisName, - cateAxisName = _getAxisNameByLayout2.cateAxisName; - - var graphicalItems = findAllByType(children, GraphicalChild); - var stackGroups = getStackGroupsByAxisId(data, graphicalItems, "".concat(numericAxisName, "Id"), "".concat(cateAxisName, "Id"), stackOffset, reverseStackOrder); - var axisObj = axisComponents.reduce(function (result, entry) { - var name = "".concat(entry.axisType, "Map"); - return _objectSpread$3(_objectSpread$3({}, result), {}, _defineProperty$3({}, name, getAxisMap(props, _objectSpread$3(_objectSpread$3({}, entry), {}, { - graphicalItems: graphicalItems, - stackGroups: entry.axisType === numericAxisName && stackGroups, - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex - })))); - }, {}); - var offset = calculateOffset(_objectSpread$3(_objectSpread$3({}, axisObj), {}, { - props: props, - graphicalItems: graphicalItems - }), prevState === null || prevState === void 0 ? void 0 : prevState.legendBBox); - Object.keys(axisObj).forEach(function (key) { - axisObj[key] = formatAxisMap(props, axisObj[key], offset, key.replace('Map', ''), chartName); - }); - var cateAxisMap = axisObj["".concat(cateAxisName, "Map")]; - var ticksObj = tooltipTicksGenerator(cateAxisMap); - var formattedGraphicalItems = getFormatItems(props, _objectSpread$3(_objectSpread$3({}, axisObj), {}, { - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex, - updateId: updateId, - graphicalItems: graphicalItems, - stackGroups: stackGroups, - offset: offset - })); - return _objectSpread$3(_objectSpread$3({ - formattedGraphicalItems: formattedGraphicalItems, - graphicalItems: graphicalItems, - offset: offset, - stackGroups: stackGroups - }, ticksObj), axisObj); - }; - - return _temp = _class = /*#__PURE__*/function (_Component) { - _inherits$4(CategoricalChartWrapper, _Component); - - var _super = _createSuper$4(CategoricalChartWrapper); - - // todo join specific chart propTypes - function CategoricalChartWrapper(_props) { - var _this; - - _classCallCheck$4(this, CategoricalChartWrapper); - - _this = _super.call(this, _props); - _this.uniqueChartId = void 0; - _this.clipPathId = void 0; - _this.legendInstance = void 0; - _this.deferId = void 0; - _this.container = void 0; - - _this.clearDeferId = function () { - if (!isNil_1(_this.deferId) && deferClear) { - deferClear(_this.deferId); - } - - _this.deferId = null; - }; - - _this.handleLegendBBoxUpdate = function (box) { - if (box) { - var _this$state = _this.state, - dataStartIndex = _this$state.dataStartIndex, - dataEndIndex = _this$state.dataEndIndex, - updateId = _this$state.updateId; - - _this.setState(_objectSpread$3({ - legendBBox: box - }, updateStateOfAxisMapsOffsetAndStackGroups({ - props: _this.props, - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex, - updateId: updateId - }, _objectSpread$3(_objectSpread$3({}, _this.state), {}, { - legendBBox: box - })))); - } - }; - - _this.handleReceiveSyncEvent = function (cId, chartId, data) { - var syncId = _this.props.syncId; - - if (syncId === cId && chartId !== _this.uniqueChartId) { - _this.clearDeferId(); - - _this.deferId = defer && defer(_this.applySyncEvent.bind(_assertThisInitialized$4(_this), data)); - } - }; - - _this.handleBrushChange = function (_ref9) { - var startIndex = _ref9.startIndex, - endIndex = _ref9.endIndex; - - // Only trigger changes if the extents of the brush have actually changed - if (startIndex !== _this.state.dataStartIndex || endIndex !== _this.state.dataEndIndex) { - var updateId = _this.state.updateId; - - _this.setState(function () { - return _objectSpread$3({ - dataStartIndex: startIndex, - dataEndIndex: endIndex - }, updateStateOfAxisMapsOffsetAndStackGroups({ - props: _this.props, - dataStartIndex: startIndex, - dataEndIndex: endIndex, - updateId: updateId - }, _this.state)); - }); - - _this.triggerSyncEvent({ - dataStartIndex: startIndex, - dataEndIndex: endIndex - }); - } - }; - - _this.handleMouseEnter = function (e) { - var onMouseEnter = _this.props.onMouseEnter; - - var mouse = _this.getMouseInfo(e); - - if (mouse) { - var _nextState = _objectSpread$3(_objectSpread$3({}, mouse), {}, { - isTooltipActive: true - }); - - _this.setState(_nextState); - - _this.triggerSyncEvent(_nextState); - - if (isFunction_1(onMouseEnter)) { - onMouseEnter(_nextState, e); - } - } - }; - - _this.triggeredAfterMouseMove = function (e) { - var onMouseMove = _this.props.onMouseMove; - - var mouse = _this.getMouseInfo(e); - - var nextState = mouse ? _objectSpread$3(_objectSpread$3({}, mouse), {}, { - isTooltipActive: true - }) : { - isTooltipActive: false - }; - - _this.setState(nextState); - - _this.triggerSyncEvent(nextState); - - if (isFunction_1(onMouseMove)) { - onMouseMove(nextState, e); - } - }; - - _this.handleItemMouseEnter = function (el) { - _this.setState(function () { - return { - isTooltipActive: true, - activeItem: el, - activePayload: el.tooltipPayload, - activeCoordinate: el.tooltipPosition || { - x: el.cx, - y: el.cy - } - }; - }); - }; - - _this.handleItemMouseLeave = function () { - _this.setState(function () { - return { - isTooltipActive: false - }; - }); - }; - - _this.handleMouseMove = function (e) { - if (e && isFunction_1(e.persist)) { - e.persist(); - } - - _this.triggeredAfterMouseMove(e); - }; - - _this.handleMouseLeave = function (e) { - var onMouseLeave = _this.props.onMouseLeave; - var nextState = { - isTooltipActive: false - }; - - _this.setState(nextState); - - _this.triggerSyncEvent(nextState); - - if (isFunction_1(onMouseLeave)) { - onMouseLeave(nextState, e); - } - - _this.cancelThrottledTriggerAfterMouseMove(); - }; - - _this.handleOuterEvent = function (e) { - var eventName = getReactEventByType(e); - - var event = get_1(_this.props, "".concat(eventName)); - - if (eventName && isFunction_1(event)) { - var mouse; - - if (/.*touch.*/i.test(eventName)) { - mouse = _this.getMouseInfo(e.changedTouches[0]); - } else { - mouse = _this.getMouseInfo(e); - } - - var handler = event; - handler(mouse, e); - } - }; - - _this.handleClick = function (e) { - var onClick = _this.props.onClick; - - var mouse = _this.getMouseInfo(e); - - if (mouse) { - var _nextState2 = _objectSpread$3(_objectSpread$3({}, mouse), {}, { - isTooltipActive: true - }); - - _this.setState(_nextState2); - - _this.triggerSyncEvent(_nextState2); - - if (isFunction_1(onClick)) { - onClick(_nextState2, e); - } - } - }; - - _this.handleMouseDown = function (e) { - var onMouseDown = _this.props.onMouseDown; - - if (isFunction_1(onMouseDown)) { - var _nextState3 = _this.getMouseInfo(e); - - onMouseDown(_nextState3, e); - } - }; - - _this.handleMouseUp = function (e) { - var onMouseUp = _this.props.onMouseUp; - - if (isFunction_1(onMouseUp)) { - var _nextState4 = _this.getMouseInfo(e); - - onMouseUp(_nextState4, e); - } - }; - - _this.handleTouchMove = function (e) { - if (e.changedTouches != null && e.changedTouches.length > 0) { - _this.handleMouseMove(e.changedTouches[0]); - } - }; - - _this.handleTouchStart = function (e) { - if (e.changedTouches != null && e.changedTouches.length > 0) { - _this.handleMouseDown(e.changedTouches[0]); - } - }; - - _this.handleTouchEnd = function (e) { - if (e.changedTouches != null && e.changedTouches.length > 0) { - _this.handleMouseUp(e.changedTouches[0]); - } - }; - - _this.verticalCoordinatesGenerator = function (_ref10) { - var xAxis = _ref10.xAxis, - width = _ref10.width, - height = _ref10.height, - offset = _ref10.offset; - return getCoordinatesOfGrid(CartesianAxis.getTicks(_objectSpread$3(_objectSpread$3(_objectSpread$3({}, CartesianAxis.defaultProps), xAxis), {}, { - ticks: getTicksOfAxis(xAxis, true), - viewBox: { - x: 0, - y: 0, - width: width, - height: height - } - })), offset.left, offset.left + offset.width); - }; - - _this.horizontalCoordinatesGenerator = function (_ref11) { - var yAxis = _ref11.yAxis, - width = _ref11.width, - height = _ref11.height, - offset = _ref11.offset; - return getCoordinatesOfGrid(CartesianAxis.getTicks(_objectSpread$3(_objectSpread$3(_objectSpread$3({}, CartesianAxis.defaultProps), yAxis), {}, { - ticks: getTicksOfAxis(yAxis, true), - viewBox: { - x: 0, - y: 0, - width: width, - height: height - } - })), offset.top, offset.top + offset.height); - }; - - _this.axesTicksGenerator = function (axis) { - return getTicksOfAxis(axis, true); - }; - - _this.renderCursor = function (element) { - var _this$state2 = _this.state, - isTooltipActive = _this$state2.isTooltipActive, - activeCoordinate = _this$state2.activeCoordinate, - activePayload = _this$state2.activePayload, - offset = _this$state2.offset, - activeTooltipIndex = _this$state2.activeTooltipIndex; - - var tooltipEventType = _this.getTooltipEventType(); - - if (!element || !element.props.cursor || !isTooltipActive || !activeCoordinate || chartName !== 'ScatterChart' && tooltipEventType !== 'axis') { - return null; - } - - var layout = _this.props.layout; - var restProps; - var cursorComp = Curve; - - if (chartName === 'ScatterChart') { - restProps = activeCoordinate; - cursorComp = Cross; - } else if (chartName === 'BarChart') { - restProps = _this.getCursorRectangle(); - cursorComp = Rectangle; - } else if (layout === 'radial') { - var _this$getCursorPoints = _this.getCursorPoints(), - cx = _this$getCursorPoints.cx, - cy = _this$getCursorPoints.cy, - radius = _this$getCursorPoints.radius, - startAngle = _this$getCursorPoints.startAngle, - endAngle = _this$getCursorPoints.endAngle; - - restProps = { - cx: cx, - cy: cy, - startAngle: startAngle, - endAngle: endAngle, - innerRadius: radius, - outerRadius: radius - }; - cursorComp = Sector; - } else { - restProps = { - points: _this.getCursorPoints() - }; - cursorComp = Curve; - } - - var key = element.key || '_recharts-cursor'; - - var cursorProps = _objectSpread$3(_objectSpread$3(_objectSpread$3(_objectSpread$3({ - stroke: '#ccc', - pointerEvents: 'none' - }, offset), restProps), filterProps(element.props.cursor)), {}, { - payload: activePayload, - payloadIndex: activeTooltipIndex, - key: key, - className: 'recharts-tooltip-cursor' - }); - - return /*#__PURE__*/isValidElement(element.props.cursor) ? /*#__PURE__*/cloneElement(element.props.cursor, cursorProps) : /*#__PURE__*/createElement(cursorComp, cursorProps); - }; - - _this.renderPolarAxis = function (element, displayName, index) { - var axisType = get_1(element, 'type.axisType'); - - var axisMap = get_1(_this.state, "".concat(axisType, "Map")); - - var axisOption = axisMap[element.props["".concat(axisType, "Id")]]; - return /*#__PURE__*/cloneElement(element, _objectSpread$3(_objectSpread$3({}, axisOption), {}, { - className: axisType, - key: element.key || "".concat(displayName, "-").concat(index), - ticks: getTicksOfAxis(axisOption, true) - })); - }; - - _this.renderXAxis = function (element, displayName, index) { - var xAxisMap = _this.state.xAxisMap; - var axisObj = xAxisMap[element.props.xAxisId]; - return _this.renderAxis(axisObj, element, displayName, index); - }; - - _this.renderYAxis = function (element, displayName, index) { - var yAxisMap = _this.state.yAxisMap; - var axisObj = yAxisMap[element.props.yAxisId]; - return _this.renderAxis(axisObj, element, displayName, index); - }; - - _this.renderGrid = function (element) { - var _this$state3 = _this.state, - xAxisMap = _this$state3.xAxisMap, - yAxisMap = _this$state3.yAxisMap, - offset = _this$state3.offset; - var _this$props = _this.props, - width = _this$props.width, - height = _this$props.height; - var xAxis = getAnyElementOfObject(xAxisMap); - - var yAxisWithFiniteDomain = find_1(yAxisMap, function (axis) { - return every_1(axis.domain, isFinit); - }); - - var yAxis = yAxisWithFiniteDomain || getAnyElementOfObject(yAxisMap); - var props = element.props || {}; - return /*#__PURE__*/cloneElement(element, { - key: element.key || 'grid', - x: isNumber(props.x) ? props.x : offset.left, - y: isNumber(props.y) ? props.y : offset.top, - width: isNumber(props.width) ? props.width : offset.width, - height: isNumber(props.height) ? props.height : offset.height, - xAxis: xAxis, - yAxis: yAxis, - offset: offset, - chartWidth: width, - chartHeight: height, - verticalCoordinatesGenerator: props.verticalCoordinatesGenerator || _this.verticalCoordinatesGenerator, - horizontalCoordinatesGenerator: props.horizontalCoordinatesGenerator || _this.horizontalCoordinatesGenerator - }); - }; - - _this.renderPolarGrid = function (element) { - var _element$props = element.props, - radialLines = _element$props.radialLines, - polarAngles = _element$props.polarAngles, - polarRadius = _element$props.polarRadius; - var _this$state4 = _this.state, - radiusAxisMap = _this$state4.radiusAxisMap, - angleAxisMap = _this$state4.angleAxisMap; - var radiusAxis = getAnyElementOfObject(radiusAxisMap); - var angleAxis = getAnyElementOfObject(angleAxisMap); - var cx = angleAxis.cx, - cy = angleAxis.cy, - innerRadius = angleAxis.innerRadius, - outerRadius = angleAxis.outerRadius; - return /*#__PURE__*/cloneElement(element, { - polarAngles: isArray_1(polarAngles) ? polarAngles : getTicksOfAxis(angleAxis, true).map(function (entry) { - return entry.coordinate; - }), - polarRadius: isArray_1(polarRadius) ? polarRadius : getTicksOfAxis(radiusAxis, true).map(function (entry) { - return entry.coordinate; - }), - cx: cx, - cy: cy, - innerRadius: innerRadius, - outerRadius: outerRadius, - key: element.key || 'polar-grid', - radialLines: radialLines - }); - }; - - _this.renderLegend = function () { - var formattedGraphicalItems = _this.state.formattedGraphicalItems; - var _this$props2 = _this.props, - children = _this$props2.children, - width = _this$props2.width, - height = _this$props2.height; - var margin = _this.props.margin || {}; - var legendWidth = width - (margin.left || 0) - (margin.right || 0); - var props = getLegendProps({ - children: children, - formattedGraphicalItems: formattedGraphicalItems, - legendWidth: legendWidth, - legendContent: legendContent - }); - - if (!props) { - return null; - } - - var item = props.item, - otherProps = _objectWithoutProperties$2(props, ["item"]); - - return /*#__PURE__*/cloneElement(item, _objectSpread$3(_objectSpread$3({}, otherProps), {}, { - chartWidth: width, - chartHeight: height, - margin: margin, - ref: function ref(legend) { - _this.legendInstance = legend; - }, - onBBoxUpdate: _this.handleLegendBBoxUpdate - })); - }; - - _this.renderTooltip = function () { - var children = _this.props.children; - var tooltipItem = findChildByType(children, Tooltip.displayName); - - if (!tooltipItem) { - return null; - } - - var _this$state5 = _this.state, - isTooltipActive = _this$state5.isTooltipActive, - activeCoordinate = _this$state5.activeCoordinate, - activePayload = _this$state5.activePayload, - activeLabel = _this$state5.activeLabel, - offset = _this$state5.offset; - return /*#__PURE__*/cloneElement(tooltipItem, { - viewBox: _objectSpread$3(_objectSpread$3({}, offset), {}, { - x: offset.left, - y: offset.top - }), - active: isTooltipActive, - label: activeLabel, - payload: isTooltipActive ? activePayload : [], - coordinate: activeCoordinate - }); - }; - - _this.renderBrush = function (element) { - var _this$props3 = _this.props, - margin = _this$props3.margin, - data = _this$props3.data; - var _this$state6 = _this.state, - offset = _this$state6.offset, - dataStartIndex = _this$state6.dataStartIndex, - dataEndIndex = _this$state6.dataEndIndex, - updateId = _this$state6.updateId; // TODO: update brush when children update - - return /*#__PURE__*/cloneElement(element, { - key: element.key || '_recharts-brush', - onChange: combineEventHandlers(_this.handleBrushChange, null, element.props.onChange), - data: data, - x: isNumber(element.props.x) ? element.props.x : offset.left, - y: isNumber(element.props.y) ? element.props.y : offset.top + offset.height + offset.brushBottom - (margin.bottom || 0), - width: isNumber(element.props.width) ? element.props.width : offset.width, - startIndex: dataStartIndex, - endIndex: dataEndIndex, - updateId: "brush-".concat(updateId) - }); - }; - - _this.renderReferenceElement = function (element, displayName, index) { - if (!element) { - return null; - } - - var _assertThisInitialize = _assertThisInitialized$4(_this), - clipPathId = _assertThisInitialize.clipPathId; - - var _this$state7 = _this.state, - xAxisMap = _this$state7.xAxisMap, - yAxisMap = _this$state7.yAxisMap, - offset = _this$state7.offset; - var _element$props2 = element.props, - xAxisId = _element$props2.xAxisId, - yAxisId = _element$props2.yAxisId; - return /*#__PURE__*/cloneElement(element, { - key: element.key || "".concat(displayName, "-").concat(index), - xAxis: xAxisMap[xAxisId], - yAxis: yAxisMap[yAxisId], - viewBox: { - x: offset.left, - y: offset.top, - width: offset.width, - height: offset.height - }, - clipPathId: clipPathId - }); - }; - - _this.renderActivePoints = function (_ref12) { - var item = _ref12.item, - activePoint = _ref12.activePoint, - basePoint = _ref12.basePoint, - childIndex = _ref12.childIndex, - isRange = _ref12.isRange; - var result = []; - var key = item.props.key; - var _item$item$props = item.item.props, - activeDot = _item$item$props.activeDot, - dataKey = _item$item$props.dataKey; - - var dotProps = _objectSpread$3(_objectSpread$3({ - index: childIndex, - dataKey: dataKey, - cx: activePoint.x, - cy: activePoint.y, - r: 4, - fill: getMainColorOfGraphicItem(item.item), - strokeWidth: 2, - stroke: '#fff', - payload: activePoint.payload, - value: activePoint.value, - key: "".concat(key, "-activePoint-").concat(childIndex) - }, filterProps(activeDot)), adaptEventHandlers(activeDot)); - - result.push(CategoricalChartWrapper.renderActiveDot(activeDot, dotProps)); - - if (basePoint) { - result.push(CategoricalChartWrapper.renderActiveDot(activeDot, _objectSpread$3(_objectSpread$3({}, dotProps), {}, { - cx: basePoint.x, - cy: basePoint.y, - key: "".concat(key, "-basePoint-").concat(childIndex) - }))); - } else if (isRange) { - result.push(null); - } - - return result; - }; - - _this.renderGraphicChild = function (element, displayName, index) { - var item = _this.filterFormatItem(element, displayName, index); - - if (!item) { - return null; - } - - var tooltipEventType = _this.getTooltipEventType(); - - var _this$state8 = _this.state, - isTooltipActive = _this$state8.isTooltipActive, - tooltipAxis = _this$state8.tooltipAxis, - activeTooltipIndex = _this$state8.activeTooltipIndex, - activeLabel = _this$state8.activeLabel; - var children = _this.props.children; - var tooltipItem = findChildByType(children, Tooltip.displayName); - var _item$props2 = item.props, - points = _item$props2.points, - isRange = _item$props2.isRange, - baseLine = _item$props2.baseLine; - var _item$item$props2 = item.item.props, - activeDot = _item$item$props2.activeDot, - hide = _item$item$props2.hide; - var hasActive = !hide && isTooltipActive && tooltipItem && activeDot && activeTooltipIndex >= 0; - var itemEvents = {}; - - if (tooltipEventType !== 'axis' && tooltipItem && tooltipItem.props.trigger === 'click') { - itemEvents = { - onClick: combineEventHandlers(_this.handleItemMouseEnter, null, element.props.onCLick) - }; - } else if (tooltipEventType !== 'axis') { - itemEvents = { - onMouseLeave: combineEventHandlers(_this.handleItemMouseLeave, null, element.props.onMouseLeave), - onMouseEnter: combineEventHandlers(_this.handleItemMouseEnter, null, element.props.onMouseEnter) - }; - } - - var graphicalItem = /*#__PURE__*/cloneElement(element, _objectSpread$3(_objectSpread$3({}, item.props), itemEvents)); - - function findWithPayload(entry) { - // TODO needs to verify dataKey is Function - return typeof tooltipAxis.dataKey === 'function' ? tooltipAxis.dataKey(entry.payload) : null; - } - - if (hasActive) { - var activePoint, basePoint; - - if (tooltipAxis.dataKey && !tooltipAxis.allowDuplicatedCategory) { - // number transform to string - var specifiedKey = typeof tooltipAxis.dataKey === 'function' ? findWithPayload : 'payload.'.concat(tooltipAxis.dataKey.toString()); - activePoint = findEntryInArray(points, specifiedKey, activeLabel); - basePoint = isRange && baseLine && findEntryInArray(baseLine, specifiedKey, activeLabel); - } else { - activePoint = points[activeTooltipIndex]; - basePoint = isRange && baseLine && baseLine[activeTooltipIndex]; - } - - if (!isNil_1(activePoint)) { - return [graphicalItem].concat(_toConsumableArray(_this.renderActivePoints({ - item: item, - activePoint: activePoint, - basePoint: basePoint, - childIndex: activeTooltipIndex, - isRange: isRange - }))); - } - } - - if (isRange) { - return [graphicalItem, null, null]; - } - - return [graphicalItem, null]; - }; - - _this.renderCustomized = function (element, displayName, index) { - return /*#__PURE__*/cloneElement(element, _objectSpread$3(_objectSpread$3({ - key: "recharts-customized-".concat(index) - }, _this.props), _this.state)); - }; - - _this.uniqueChartId = isNil_1(_props.id) ? uniqueId('recharts') : _props.id; - _this.clipPathId = "".concat(_this.uniqueChartId, "-clip"); - - if (_props.throttleDelay) { - _this.triggeredAfterMouseMove = throttle_1(_this.triggeredAfterMouseMove, _props.throttleDelay); - } - - _this.state = {}; - return _this; - } - /* eslint-disable react/no-did-mount-set-state */ - - - _createClass$4(CategoricalChartWrapper, [{ - key: "componentDidMount", - value: function componentDidMount() { - if (!isNil_1(this.props.syncId)) { - this.addListener(); - } - } - }, { - key: "componentDidUpdate", - value: function componentDidUpdate(prevProps) { - // add syncId - if (isNil_1(prevProps.syncId) && !isNil_1(this.props.syncId)) { - this.addListener(); - } // remove syncId - - - if (!isNil_1(prevProps.syncId) && isNil_1(this.props.syncId)) { - this.removeListener(); - } - } - }, { - key: "componentWillUnmount", - value: function componentWillUnmount() { - this.clearDeferId(); - - if (!isNil_1(this.props.syncId)) { - this.removeListener(); - } - - this.cancelThrottledTriggerAfterMouseMove(); - } - }, { - key: "cancelThrottledTriggerAfterMouseMove", - value: function cancelThrottledTriggerAfterMouseMove() { - if (typeof this.triggeredAfterMouseMove.cancel === 'function') { - this.triggeredAfterMouseMove.cancel(); - } - } - }, { - key: "getTooltipEventType", - value: function getTooltipEventType() { - var tooltipItem = findChildByType(this.props.children, Tooltip.displayName); - - if (tooltipItem && isBoolean_1(tooltipItem.props.shared)) { - var eventType = tooltipItem.props.shared ? 'axis' : 'item'; - return validateTooltipEventTypes.indexOf(eventType) >= 0 ? eventType : defaultTooltipEventType; - } - - return defaultTooltipEventType; - } - /** - * Get the information of mouse in chart, return null when the mouse is not in the chart - * @param {Object} event The event object - * @return {Object} Mouse data - */ - - }, { - key: "getMouseInfo", - value: function getMouseInfo(event) { - if (!this.container) { - return null; - } - - var containerOffset = getOffset(this.container); - var e = calculateChartCoordinate(event, containerOffset); - var rangeObj = this.inRange(e.chartX, e.chartY); - - if (!rangeObj) { - return null; - } - - var _this$state9 = this.state, - xAxisMap = _this$state9.xAxisMap, - yAxisMap = _this$state9.yAxisMap; - var tooltipEventType = this.getTooltipEventType(); - - if (tooltipEventType !== 'axis' && xAxisMap && yAxisMap) { - var xScale = getAnyElementOfObject(xAxisMap).scale; - var yScale = getAnyElementOfObject(yAxisMap).scale; - var xValue = xScale && xScale.invert ? xScale.invert(e.chartX) : null; - var yValue = yScale && yScale.invert ? yScale.invert(e.chartY) : null; - return _objectSpread$3(_objectSpread$3({}, e), {}, { - xValue: xValue, - yValue: yValue - }); - } - - var toolTipData = getTooltipData(this.state, this.props.data, this.props.layout, rangeObj); - - if (toolTipData) { - return _objectSpread$3(_objectSpread$3({}, e), toolTipData); - } - - return null; - } - }, { - key: "getCursorRectangle", - value: function getCursorRectangle() { - var layout = this.props.layout; - var _this$state10 = this.state, - activeCoordinate = _this$state10.activeCoordinate, - offset = _this$state10.offset, - tooltipAxisBandSize = _this$state10.tooltipAxisBandSize; - var halfSize = tooltipAxisBandSize / 2; - return { - stroke: 'none', - fill: '#ccc', - x: layout === 'horizontal' ? activeCoordinate.x - halfSize : offset.left + 0.5, - y: layout === 'horizontal' ? offset.top + 0.5 : activeCoordinate.y - halfSize, - width: layout === 'horizontal' ? tooltipAxisBandSize : offset.width - 1, - height: layout === 'horizontal' ? offset.height - 1 : tooltipAxisBandSize - }; - } - }, { - key: "getCursorPoints", - value: function getCursorPoints() { - var layout = this.props.layout; - var _this$state11 = this.state, - activeCoordinate = _this$state11.activeCoordinate, - offset = _this$state11.offset; - var x1, y1, x2, y2; - - if (layout === 'horizontal') { - x1 = activeCoordinate.x; - x2 = x1; - y1 = offset.top; - y2 = offset.top + offset.height; - } else if (layout === 'vertical') { - y1 = activeCoordinate.y; - y2 = y1; - x1 = offset.left; - x2 = offset.left + offset.width; - } else if (!isNil_1(activeCoordinate.cx) || !isNil_1(activeCoordinate.cy)) { - if (layout === 'centric') { - var cx = activeCoordinate.cx, - cy = activeCoordinate.cy, - innerRadius = activeCoordinate.innerRadius, - outerRadius = activeCoordinate.outerRadius, - angle = activeCoordinate.angle; - var innerPoint = polarToCartesian(cx, cy, innerRadius, angle); - var outerPoint = polarToCartesian(cx, cy, outerRadius, angle); - x1 = innerPoint.x; - y1 = innerPoint.y; - x2 = outerPoint.x; - y2 = outerPoint.y; - } else { - var _cx = activeCoordinate.cx, - _cy = activeCoordinate.cy, - radius = activeCoordinate.radius, - startAngle = activeCoordinate.startAngle, - endAngle = activeCoordinate.endAngle; - var startPoint = polarToCartesian(_cx, _cy, radius, startAngle); - var endPoint = polarToCartesian(_cx, _cy, radius, endAngle); - return { - points: [startPoint, endPoint], - cx: _cx, - cy: _cy, - radius: radius, - startAngle: startAngle, - endAngle: endAngle - }; - } - } - - return [{ - x: x1, - y: y1 - }, { - x: x2, - y: y2 - }]; - } - }, { - key: "inRange", - value: function inRange(x, y) { - var layout = this.props.layout; - - if (layout === 'horizontal' || layout === 'vertical') { - var offset = this.state.offset; - var isInRange = x >= offset.left && x <= offset.left + offset.width && y >= offset.top && y <= offset.top + offset.height; - return isInRange ? { - x: x, - y: y - } : null; - } - - var _this$state12 = this.state, - angleAxisMap = _this$state12.angleAxisMap, - radiusAxisMap = _this$state12.radiusAxisMap; - - if (angleAxisMap && radiusAxisMap) { - var angleAxis = getAnyElementOfObject(angleAxisMap); - return inRangeOfSector({ - x: x, - y: y - }, angleAxis); - } - - return null; - } - }, { - key: "parseEventsOfWrapper", - value: function parseEventsOfWrapper() { - var children = this.props.children; - var tooltipEventType = this.getTooltipEventType(); - var tooltipItem = findChildByType(children, Tooltip.displayName); - var tooltipEvents = {}; - - if (tooltipItem && tooltipEventType === 'axis') { - if (tooltipItem.props.trigger === 'click') { - tooltipEvents = { - onClick: this.handleClick - }; - } else { - tooltipEvents = { - onMouseEnter: this.handleMouseEnter, - onMouseMove: this.handleMouseMove, - onMouseLeave: this.handleMouseLeave, - onTouchMove: this.handleTouchMove, - onTouchStart: this.handleTouchStart, - onTouchEnd: this.handleTouchEnd - }; - } - } - - var outerEvents = adaptEventHandlers(this.props, this.handleOuterEvent); - return _objectSpread$3(_objectSpread$3({}, outerEvents), tooltipEvents); - } - /* eslint-disable no-underscore-dangle */ - - }, { - key: "addListener", - value: function addListener() { - eventCenter.on(SYNC_EVENT, this.handleReceiveSyncEvent); - - if (eventCenter.setMaxListeners && eventCenter._maxListeners) { - eventCenter.setMaxListeners(eventCenter._maxListeners + 1); - } - } - }, { - key: "removeListener", - value: function removeListener() { - eventCenter.removeListener(SYNC_EVENT, this.handleReceiveSyncEvent); - - if (eventCenter.setMaxListeners && eventCenter._maxListeners) { - eventCenter.setMaxListeners(eventCenter._maxListeners - 1); - } - } - }, { - key: "triggerSyncEvent", - value: function triggerSyncEvent(data) { - var syncId = this.props.syncId; - - if (!isNil_1(syncId)) { - eventCenter.emit(SYNC_EVENT, syncId, this.uniqueChartId, data); - } - } - }, { - key: "applySyncEvent", - value: function applySyncEvent(data) { - var _this$props4 = this.props, - layout = _this$props4.layout, - syncMethod = _this$props4.syncMethod; - var updateId = this.state.updateId; - var dataStartIndex = data.dataStartIndex, - dataEndIndex = data.dataEndIndex; - - if (!isNil_1(data.dataStartIndex) || !isNil_1(data.dataEndIndex)) { - this.setState(_objectSpread$3({ - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex - }, updateStateOfAxisMapsOffsetAndStackGroups({ - props: this.props, - dataStartIndex: dataStartIndex, - dataEndIndex: dataEndIndex, - updateId: updateId - }, this.state))); - } else if (!isNil_1(data.activeTooltipIndex)) { - var chartX = data.chartX, - chartY = data.chartY; - var activeTooltipIndex = data.activeTooltipIndex; - var _this$state13 = this.state, - offset = _this$state13.offset, - tooltipTicks = _this$state13.tooltipTicks; - - if (!offset) { - return; - } - - if (typeof syncMethod === 'function') { - // Call a callback function. If there is an application specific algorithm - activeTooltipIndex = syncMethod(tooltipTicks, data); - } else if (syncMethod === 'value') { - // Set activeTooltipIndex to the index with the same value as data.activeLabel - // For loop instead of findIndex because the latter is very slow in some browsers - activeTooltipIndex = -1; // in case we cannot find the element - - for (var i = 0; i < tooltipTicks.length; i++) { - if (tooltipTicks[i].value === data.activeLabel) { - activeTooltipIndex = i; - break; - } - } - } - - var viewBox = _objectSpread$3(_objectSpread$3({}, offset), {}, { - x: offset.left, - y: offset.top - }); // When a categotical chart is combined with another chart, the value of chartX - // and chartY may beyond the boundaries. - - - var validateChartX = Math.min(chartX, viewBox.x + viewBox.width); - var validateChartY = Math.min(chartY, viewBox.y + viewBox.height); - var activeLabel = tooltipTicks[activeTooltipIndex] && tooltipTicks[activeTooltipIndex].value; - var activePayload = getTooltipContent(this.state, this.props.data, activeTooltipIndex); - var activeCoordinate = tooltipTicks[activeTooltipIndex] ? { - x: layout === 'horizontal' ? tooltipTicks[activeTooltipIndex].coordinate : validateChartX, - y: layout === 'horizontal' ? validateChartY : tooltipTicks[activeTooltipIndex].coordinate - } : originCoordinate; - this.setState(_objectSpread$3(_objectSpread$3({}, data), {}, { - activeLabel: activeLabel, - activeCoordinate: activeCoordinate, - activePayload: activePayload, - activeTooltipIndex: activeTooltipIndex - })); - } else { - this.setState(data); - } - } - }, { - key: "filterFormatItem", - value: function filterFormatItem(item, displayName, childIndex) { - var formattedGraphicalItems = this.state.formattedGraphicalItems; - - for (var i = 0, len = formattedGraphicalItems.length; i < len; i++) { - var entry = formattedGraphicalItems[i]; - - if (entry.item === item || entry.props.key === item.key || displayName === getDisplayName(entry.item.type) && childIndex === entry.childIndex) { - return entry; - } - } - - return null; - } - }, { - key: "renderAxis", - value: - /** - * Draw axis - * @param {Object} axisOptions The options of axis - * @param {Object} element The axis element - * @param {String} displayName The display name of axis - * @param {Number} index The index of element - * @return {ReactElement} The instance of x-axes - */ - function renderAxis(axisOptions, element, displayName, index) { - var _this$props5 = this.props, - width = _this$props5.width, - height = _this$props5.height; - return /*#__PURE__*/React__default.createElement(CartesianAxis, _extends$4({}, axisOptions, { - className: "recharts-".concat(axisOptions.axisType, " ").concat(axisOptions.axisType), - key: element.key || "".concat(displayName, "-").concat(index), - viewBox: { - x: 0, - y: 0, - width: width, - height: height - }, - ticksGenerator: this.axesTicksGenerator - })); - } - /** - * Draw grid - * @param {ReactElement} element the grid item - * @return {ReactElement} The instance of grid - */ - - }, { - key: "renderClipPath", - value: function renderClipPath() { - var clipPathId = this.clipPathId; - var _this$state$offset = this.state.offset, - left = _this$state$offset.left, - top = _this$state$offset.top, - height = _this$state$offset.height, - width = _this$state$offset.width; - return /*#__PURE__*/React__default.createElement("defs", null, /*#__PURE__*/React__default.createElement("clipPath", { - id: clipPathId - }, /*#__PURE__*/React__default.createElement("rect", { - x: left, - y: top, - height: height, - width: width - }))); - } - }, { - key: "getXScales", - value: function getXScales() { - var xAxisMap = this.state.xAxisMap; - return xAxisMap ? Object.entries(xAxisMap).reduce(function (res, _ref13) { - var _ref14 = _slicedToArray$1(_ref13, 2), - axisId = _ref14[0], - axisProps = _ref14[1]; - - return _objectSpread$3(_objectSpread$3({}, res), {}, _defineProperty$3({}, axisId, axisProps.scale)); - }, {}) : null; - } - }, { - key: "getYScales", - value: function getYScales() { - var yAxisMap = this.state.yAxisMap; - return yAxisMap ? Object.entries(yAxisMap).reduce(function (res, _ref15) { - var _ref16 = _slicedToArray$1(_ref15, 2), - axisId = _ref16[0], - axisProps = _ref16[1]; - - return _objectSpread$3(_objectSpread$3({}, res), {}, _defineProperty$3({}, axisId, axisProps.scale)); - }, {}) : null; - } - }, { - key: "getXScaleByAxisId", - value: function getXScaleByAxisId(axisId) { - var _this$state$xAxisMap, _this$state$xAxisMap$; - - return (_this$state$xAxisMap = this.state.xAxisMap) === null || _this$state$xAxisMap === void 0 ? void 0 : (_this$state$xAxisMap$ = _this$state$xAxisMap[axisId]) === null || _this$state$xAxisMap$ === void 0 ? void 0 : _this$state$xAxisMap$.scale; - } - }, { - key: "getYScaleByAxisId", - value: function getYScaleByAxisId(axisId) { - var _this$state$yAxisMap, _this$state$yAxisMap$; - - return (_this$state$yAxisMap = this.state.yAxisMap) === null || _this$state$yAxisMap === void 0 ? void 0 : (_this$state$yAxisMap$ = _this$state$yAxisMap[axisId]) === null || _this$state$yAxisMap$ === void 0 ? void 0 : _this$state$yAxisMap$.scale; - } - }, { - key: "getItemByXY", - value: function getItemByXY(chartXY) { - var formattedGraphicalItems = this.state.formattedGraphicalItems; - - if (formattedGraphicalItems && formattedGraphicalItems.length) { - for (var i = 0, len = formattedGraphicalItems.length; i < len; i++) { - var graphicalItem = formattedGraphicalItems[i]; - var props = graphicalItem.props, - item = graphicalItem.item; - var itemDisplayName = getDisplayName(item.type); - - if (itemDisplayName === 'Bar') { - var activeBarItem = (props.data || []).find(function (entry) { - return isInRectangle(chartXY, entry); - }); - - if (activeBarItem) { - return { - graphicalItem: graphicalItem, - payload: activeBarItem - }; - } - } else if (itemDisplayName === 'RadialBar') { - var _activeBarItem = (props.data || []).find(function (entry) { - return inRangeOfSector(chartXY, entry); - }); - - if (_activeBarItem) { - return { - graphicalItem: graphicalItem, - payload: _activeBarItem - }; - } - } - } - } - - return null; - } - }, { - key: "render", - value: function render() { - var _this2 = this; - - if (!validateWidthHeight(this)) { - return null; - } - - var _this$props6 = this.props, - children = _this$props6.children, - className = _this$props6.className, - width = _this$props6.width, - height = _this$props6.height, - style = _this$props6.style, - compact = _this$props6.compact, - title = _this$props6.title, - desc = _this$props6.desc, - others = _objectWithoutProperties$2(_this$props6, ["children", "className", "width", "height", "style", "compact", "title", "desc"]); - - var attrs = filterProps(others); - var map = { - CartesianGrid: { - handler: this.renderGrid, - once: true - }, - ReferenceArea: { - handler: this.renderReferenceElement - }, - ReferenceLine: { - handler: this.renderReferenceElement - }, - ReferenceDot: { - handler: this.renderReferenceElement - }, - XAxis: { - handler: this.renderXAxis - }, - YAxis: { - handler: this.renderYAxis - }, - Brush: { - handler: this.renderBrush, - once: true - }, - Bar: { - handler: this.renderGraphicChild - }, - Line: { - handler: this.renderGraphicChild - }, - Area: { - handler: this.renderGraphicChild - }, - Radar: { - handler: this.renderGraphicChild - }, - RadialBar: { - handler: this.renderGraphicChild - }, - Scatter: { - handler: this.renderGraphicChild - }, - Pie: { - handler: this.renderGraphicChild - }, - Funnel: { - handler: this.renderGraphicChild - }, - Tooltip: { - handler: this.renderCursor, - once: true - }, - PolarGrid: { - handler: this.renderPolarGrid, - once: true - }, - PolarAngleAxis: { - handler: this.renderPolarAxis - }, - PolarRadiusAxis: { - handler: this.renderPolarAxis - }, - Customized: { - handler: this.renderCustomized - } - }; // The "compact" mode is mainly used as the panorama within Brush - - if (compact) { - return /*#__PURE__*/React__default.createElement(Surface, _extends$4({}, attrs, { - width: width, - height: height, - title: title, - desc: desc - }), this.renderClipPath(), renderByOrder(children, map)); - } - - var events = this.parseEventsOfWrapper(); - return /*#__PURE__*/React__default.createElement("div", _extends$4({ - className: classNames('recharts-wrapper', className), - style: _objectSpread$3({ - position: 'relative', - cursor: 'default', - width: width, - height: height - }, style) - }, events, { - ref: function ref(node) { - _this2.container = node; - } - }), /*#__PURE__*/React__default.createElement(Surface, _extends$4({}, attrs, { - width: width, - height: height, - title: title, - desc: desc - }), this.renderClipPath(), renderByOrder(children, map)), this.renderLegend(), this.renderTooltip()); - } - }]); - - return CategoricalChartWrapper; - }(Component), _class.displayName = chartName, _class.defaultProps = _objectSpread$3({ - layout: 'horizontal', - stackOffset: 'none', - barCategoryGap: '10%', - barGap: 4, - margin: { - top: 5, - right: 5, - bottom: 5, - left: 5 - }, - reverseStackOrder: false, - syncMethod: 'index' - }, defaultProps), _class.getDerivedStateFromProps = function (nextProps, prevState) { - var data = nextProps.data, - children = nextProps.children, - width = nextProps.width, - height = nextProps.height, - layout = nextProps.layout, - stackOffset = nextProps.stackOffset, - margin = nextProps.margin; - - if (isNil_1(prevState.updateId)) { - var defaultState = createDefaultState(nextProps); - return _objectSpread$3(_objectSpread$3(_objectSpread$3({}, defaultState), {}, { - updateId: 0 - }, updateStateOfAxisMapsOffsetAndStackGroups(_objectSpread$3(_objectSpread$3({ - props: nextProps - }, defaultState), {}, { - updateId: 0 - }), prevState)), {}, { - prevData: data, - prevWidth: width, - prevHeight: height, - prevLayout: layout, - prevStackOffset: stackOffset, - prevMargin: margin, - prevChildren: children - }); - } - - if (data !== prevState.prevData || width !== prevState.prevWidth || height !== prevState.prevHeight || layout !== prevState.prevLayout || stackOffset !== prevState.prevStackOffset || !shallowEqual(margin, prevState.prevMargin)) { - var _defaultState = createDefaultState(nextProps); // Fixes https://github.com/recharts/recharts/issues/2143 - - - var keepFromPrevState = { - // (chartX, chartY) are (0,0) in default state, but we want to keep the last mouse position to avoid - // any flickering - chartX: prevState.chartX, - chartY: prevState.chartY, - // The tooltip should stay active when it was active in the previous render. If this is not - // the case, the tooltip disappears and immediately re-appears, causing a flickering effect - isTooltipActive: prevState.isTooltipActive - }; - - var updatesToState = _objectSpread$3(_objectSpread$3({}, getTooltipData(prevState, data, layout)), {}, { - updateId: prevState.updateId + 1 - }); - - var newState = _objectSpread$3(_objectSpread$3(_objectSpread$3({}, _defaultState), keepFromPrevState), updatesToState); - - return _objectSpread$3(_objectSpread$3(_objectSpread$3({}, newState), updateStateOfAxisMapsOffsetAndStackGroups(_objectSpread$3({ - props: nextProps - }, newState), prevState)), {}, { - prevData: data, - prevWidth: width, - prevHeight: height, - prevLayout: layout, - prevStackOffset: stackOffset, - prevMargin: margin, - prevChildren: children - }); - } - - if (!isChildrenEqual(children, prevState.prevChildren)) { - // update configuration in chilren - var hasGlobalData = !isNil_1(data); - var newUpdateId = hasGlobalData ? prevState.updateId : prevState.updateId + 1; - return _objectSpread$3(_objectSpread$3({ - updateId: newUpdateId - }, updateStateOfAxisMapsOffsetAndStackGroups(_objectSpread$3(_objectSpread$3({ - props: nextProps - }, prevState), {}, { - updateId: newUpdateId - }), prevState)), {}, { - prevChildren: children - }); - } - - return null; - }, _class.renderActiveDot = function (option, props) { - var dot; - - if ( /*#__PURE__*/isValidElement(option)) { - dot = /*#__PURE__*/cloneElement(option, props); - } else if (isFunction_1(option)) { - dot = option(props); - } else { - dot = /*#__PURE__*/React__default.createElement(Dot, props); - } - - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-active-dot", - key: props.key - }, dot); - }, _temp; -}; - -/** - * @fileOverview Line Chart - */ -var LineChart = generateCategoricalChart({ - chartName: 'LineChart', - GraphicalChild: Line, - axisComponents: [{ - axisType: 'xAxis', - AxisComp: XAxis - }, { - axisType: 'yAxis', - AxisComp: YAxis - }], - formatAxisMap: formatAxisMap -}); - -/** - * @fileOverview Bar Chart - */ -var BarChart = generateCategoricalChart({ - chartName: 'BarChart', - GraphicalChild: Bar, - defaultTooltipEventType: 'axis', - validateTooltipEventTypes: ['axis', 'item'], - axisComponents: [{ - axisType: 'xAxis', - AxisComp: XAxis - }, { - axisType: 'yAxis', - AxisComp: YAxis - }], - formatAxisMap: formatAxisMap -}); - -/** - * @fileOverview Pie Chart - */ -var PieChart = generateCategoricalChart({ - chartName: 'PieChart', - GraphicalChild: Pie, - validateTooltipEventTypes: ['item'], - defaultTooltipEventType: 'item', - legendContent: 'children', - axisComponents: [{ - axisType: 'angleAxis', - AxisComp: PolarAngleAxis - }, { - axisType: 'radiusAxis', - AxisComp: PolarRadiusAxis - }], - formatAxisMap: formatAxisMap$1, - defaultProps: { - layout: 'centric', - startAngle: 0, - endAngle: 360, - cx: '50%', - cy: '50%', - innerRadius: 0, - outerRadius: '80%' - } -}); - -/** - * A specialized version of `_.forEach` for arrays without support for - * iteratee shorthands. - * - * @private - * @param {Array} [array] The array to iterate over. - * @param {Function} iteratee The function invoked per iteration. - * @returns {Array} Returns `array`. - */ - -function arrayEach$1(array, iteratee) { - var index = -1, - length = array == null ? 0 : array.length; - - while (++index < length) { - if (iteratee(array[index], index, array) === false) { - break; - } - } - return array; -} - -var _arrayEach = arrayEach$1; - -var baseAssignValue$1 = _baseAssignValue, - eq = eq_1; - -/** Used for built-in method references. */ -var objectProto$2 = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$2 = objectProto$2.hasOwnProperty; - -/** - * Assigns `value` to `key` of `object` if the existing value is not equivalent - * using [`SameValueZero`](http://ecma-international.org/ecma-262/7.0/#sec-samevaluezero) - * for equality comparisons. - * - * @private - * @param {Object} object The object to modify. - * @param {string} key The key of the property to assign. - * @param {*} value The value to assign. - */ -function assignValue$2(object, key, value) { - var objValue = object[key]; - if (!(hasOwnProperty$2.call(object, key) && eq(objValue, value)) || - (value === undefined && !(key in object))) { - baseAssignValue$1(object, key, value); - } -} - -var _assignValue = assignValue$2; - -var assignValue$1 = _assignValue, - baseAssignValue = _baseAssignValue; - -/** - * Copies properties of `source` to `object`. - * - * @private - * @param {Object} source The object to copy properties from. - * @param {Array} props The property identifiers to copy. - * @param {Object} [object={}] The object to copy properties to. - * @param {Function} [customizer] The function to customize copied values. - * @returns {Object} Returns `object`. - */ -function copyObject$5(source, props, object, customizer) { - var isNew = !object; - object || (object = {}); - - var index = -1, - length = props.length; - - while (++index < length) { - var key = props[index]; - - var newValue = customizer - ? customizer(object[key], source[key], key, object, source) - : undefined; - - if (newValue === undefined) { - newValue = source[key]; - } - if (isNew) { - baseAssignValue(object, key, newValue); - } else { - assignValue$1(object, key, newValue); - } - } - return object; -} - -var _copyObject = copyObject$5; - -var copyObject$4 = _copyObject, - keys$1 = keys_1; - -/** - * The base implementation of `_.assign` without support for multiple sources - * or `customizer` functions. - * - * @private - * @param {Object} object The destination object. - * @param {Object} source The source object. - * @returns {Object} Returns `object`. - */ -function baseAssign$1(object, source) { - return object && copyObject$4(source, keys$1(source), object); -} - -var _baseAssign = baseAssign$1; - -/** - * This function is like - * [`Object.keys`](http://ecma-international.org/ecma-262/7.0/#sec-object.keys) - * except that it includes inherited enumerable properties. - * - * @private - * @param {Object} object The object to query. - * @returns {Array} Returns the array of property names. - */ - -function nativeKeysIn$1(object) { - var result = []; - if (object != null) { - for (var key in Object(object)) { - result.push(key); - } - } - return result; -} - -var _nativeKeysIn = nativeKeysIn$1; - -var isObject$2 = isObject_1$1, - isPrototype$1 = _isPrototype, - nativeKeysIn = _nativeKeysIn; - -/** Used for built-in method references. */ -var objectProto$1 = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty$1 = objectProto$1.hasOwnProperty; - -/** - * The base implementation of `_.keysIn` which doesn't treat sparse arrays as dense. - * - * @private - * @param {Object} object The object to query. - * @returns {Array} Returns the array of property names. - */ -function baseKeysIn$1(object) { - if (!isObject$2(object)) { - return nativeKeysIn(object); - } - var isProto = isPrototype$1(object), - result = []; - - for (var key in object) { - if (!(key == 'constructor' && (isProto || !hasOwnProperty$1.call(object, key)))) { - result.push(key); - } - } - return result; -} - -var _baseKeysIn = baseKeysIn$1; - -var arrayLikeKeys = _arrayLikeKeys, - baseKeysIn = _baseKeysIn, - isArrayLike = isArrayLike_1; - -/** - * Creates an array of the own and inherited enumerable property names of `object`. - * - * **Note:** Non-object values are coerced to objects. - * - * @static - * @memberOf _ - * @since 3.0.0 - * @category Object - * @param {Object} object The object to query. - * @returns {Array} Returns the array of property names. - * @example - * - * function Foo() { - * this.a = 1; - * this.b = 2; - * } - * - * Foo.prototype.c = 3; - * - * _.keysIn(new Foo); - * // => ['a', 'b', 'c'] (iteration order is not guaranteed) - */ -function keysIn$3(object) { - return isArrayLike(object) ? arrayLikeKeys(object, true) : baseKeysIn(object); -} - -var keysIn_1 = keysIn$3; - -var copyObject$3 = _copyObject, - keysIn$2 = keysIn_1; - -/** - * The base implementation of `_.assignIn` without support for multiple sources - * or `customizer` functions. - * - * @private - * @param {Object} object The destination object. - * @param {Object} source The source object. - * @returns {Object} Returns `object`. - */ -function baseAssignIn$1(object, source) { - return object && copyObject$3(source, keysIn$2(source), object); -} - -var _baseAssignIn = baseAssignIn$1; - -var _cloneBuffer = {exports: {}}; - -(function (module, exports) { - var root = _root$1; - - /** Detect free variable `exports`. */ - var freeExports = exports && !exports.nodeType && exports; - - /** Detect free variable `module`. */ - var freeModule = freeExports && 'object' == 'object' && module && !module.nodeType && module; - - /** Detect the popular CommonJS extension `module.exports`. */ - var moduleExports = freeModule && freeModule.exports === freeExports; - - /** Built-in value references. */ - var Buffer = moduleExports ? root.Buffer : undefined, - allocUnsafe = Buffer ? Buffer.allocUnsafe : undefined; - - /** - * Creates a clone of `buffer`. - * - * @private - * @param {Buffer} buffer The buffer to clone. - * @param {boolean} [isDeep] Specify a deep clone. - * @returns {Buffer} Returns the cloned buffer. - */ - function cloneBuffer(buffer, isDeep) { - if (isDeep) { - return buffer.slice(); - } - var length = buffer.length, - result = allocUnsafe ? allocUnsafe(length) : new buffer.constructor(length); - - buffer.copy(result); - return result; - } - - module.exports = cloneBuffer; -} (_cloneBuffer, _cloneBuffer.exports)); - -/** - * Copies the values of `source` to `array`. - * - * @private - * @param {Array} source The array to copy values from. - * @param {Array} [array=[]] The array to copy values to. - * @returns {Array} Returns `array`. - */ - -function copyArray$1(source, array) { - var index = -1, - length = source.length; - - array || (array = Array(length)); - while (++index < length) { - array[index] = source[index]; - } - return array; -} - -var _copyArray = copyArray$1; - -var copyObject$2 = _copyObject, - getSymbols$1 = _getSymbols; - -/** - * Copies own symbols of `source` to `object`. - * - * @private - * @param {Object} source The object to copy symbols from. - * @param {Object} [object={}] The object to copy symbols to. - * @returns {Object} Returns `object`. - */ -function copySymbols$1(source, object) { - return copyObject$2(source, getSymbols$1(source), object); -} - -var _copySymbols = copySymbols$1; - -var arrayPush = _arrayPush, - getPrototype$1 = _getPrototype, - getSymbols = _getSymbols, - stubArray = stubArray_1; - -/* Built-in method references for those with the same name as other `lodash` methods. */ -var nativeGetSymbols = Object.getOwnPropertySymbols; - -/** - * Creates an array of the own and inherited enumerable symbols of `object`. - * - * @private - * @param {Object} object The object to query. - * @returns {Array} Returns the array of symbols. - */ -var getSymbolsIn$2 = !nativeGetSymbols ? stubArray : function(object) { - var result = []; - while (object) { - arrayPush(result, getSymbols(object)); - object = getPrototype$1(object); - } - return result; -}; - -var _getSymbolsIn = getSymbolsIn$2; - -var copyObject$1 = _copyObject, - getSymbolsIn$1 = _getSymbolsIn; - -/** - * Copies own and inherited symbols of `source` to `object`. - * - * @private - * @param {Object} source The object to copy symbols from. - * @param {Object} [object={}] The object to copy symbols to. - * @returns {Object} Returns `object`. - */ -function copySymbolsIn$1(source, object) { - return copyObject$1(source, getSymbolsIn$1(source), object); -} - -var _copySymbolsIn = copySymbolsIn$1; - -var baseGetAllKeys = _baseGetAllKeys, - getSymbolsIn = _getSymbolsIn, - keysIn$1 = keysIn_1; - -/** - * Creates an array of own and inherited enumerable property names and - * symbols of `object`. - * - * @private - * @param {Object} object The object to query. - * @returns {Array} Returns the array of property names and symbols. - */ -function getAllKeysIn$2(object) { - return baseGetAllKeys(object, keysIn$1, getSymbolsIn); -} - -var _getAllKeysIn = getAllKeysIn$2; - -/** Used for built-in method references. */ - -var objectProto = Object.prototype; - -/** Used to check objects for own properties. */ -var hasOwnProperty = objectProto.hasOwnProperty; - -/** - * Initializes an array clone. - * - * @private - * @param {Array} array The array to clone. - * @returns {Array} Returns the initialized clone. - */ -function initCloneArray$1(array) { - var length = array.length, - result = new array.constructor(length); - - // Add properties assigned by `RegExp#exec`. - if (length && typeof array[0] == 'string' && hasOwnProperty.call(array, 'index')) { - result.index = array.index; - result.input = array.input; - } - return result; -} - -var _initCloneArray = initCloneArray$1; - -var Uint8Array = _Uint8Array; - -/** - * Creates a clone of `arrayBuffer`. - * - * @private - * @param {ArrayBuffer} arrayBuffer The array buffer to clone. - * @returns {ArrayBuffer} Returns the cloned array buffer. - */ -function cloneArrayBuffer$3(arrayBuffer) { - var result = new arrayBuffer.constructor(arrayBuffer.byteLength); - new Uint8Array(result).set(new Uint8Array(arrayBuffer)); - return result; -} - -var _cloneArrayBuffer = cloneArrayBuffer$3; - -var cloneArrayBuffer$2 = _cloneArrayBuffer; - -/** - * Creates a clone of `dataView`. - * - * @private - * @param {Object} dataView The data view to clone. - * @param {boolean} [isDeep] Specify a deep clone. - * @returns {Object} Returns the cloned data view. - */ -function cloneDataView$1(dataView, isDeep) { - var buffer = isDeep ? cloneArrayBuffer$2(dataView.buffer) : dataView.buffer; - return new dataView.constructor(buffer, dataView.byteOffset, dataView.byteLength); -} - -var _cloneDataView = cloneDataView$1; - -/** Used to match `RegExp` flags from their coerced string values. */ - -var reFlags = /\w*$/; - -/** - * Creates a clone of `regexp`. - * - * @private - * @param {Object} regexp The regexp to clone. - * @returns {Object} Returns the cloned regexp. - */ -function cloneRegExp$1(regexp) { - var result = new regexp.constructor(regexp.source, reFlags.exec(regexp)); - result.lastIndex = regexp.lastIndex; - return result; -} - -var _cloneRegExp = cloneRegExp$1; - -var Symbol$1 = _Symbol$1; - -/** Used to convert symbols to primitives and strings. */ -var symbolProto = Symbol$1 ? Symbol$1.prototype : undefined, - symbolValueOf = symbolProto ? symbolProto.valueOf : undefined; - -/** - * Creates a clone of the `symbol` object. - * - * @private - * @param {Object} symbol The symbol object to clone. - * @returns {Object} Returns the cloned symbol object. - */ -function cloneSymbol$1(symbol) { - return symbolValueOf ? Object(symbolValueOf.call(symbol)) : {}; -} - -var _cloneSymbol = cloneSymbol$1; - -var cloneArrayBuffer$1 = _cloneArrayBuffer; - -/** - * Creates a clone of `typedArray`. - * - * @private - * @param {Object} typedArray The typed array to clone. - * @param {boolean} [isDeep] Specify a deep clone. - * @returns {Object} Returns the cloned typed array. - */ -function cloneTypedArray$1(typedArray, isDeep) { - var buffer = isDeep ? cloneArrayBuffer$1(typedArray.buffer) : typedArray.buffer; - return new typedArray.constructor(buffer, typedArray.byteOffset, typedArray.length); -} - -var _cloneTypedArray = cloneTypedArray$1; - -var cloneArrayBuffer = _cloneArrayBuffer, - cloneDataView = _cloneDataView, - cloneRegExp = _cloneRegExp, - cloneSymbol = _cloneSymbol, - cloneTypedArray = _cloneTypedArray; - -/** `Object#toString` result references. */ -var boolTag$1 = '[object Boolean]', - dateTag$1 = '[object Date]', - mapTag$2 = '[object Map]', - numberTag$1 = '[object Number]', - regexpTag$1 = '[object RegExp]', - setTag$2 = '[object Set]', - stringTag$1 = '[object String]', - symbolTag$1 = '[object Symbol]'; - -var arrayBufferTag$1 = '[object ArrayBuffer]', - dataViewTag$1 = '[object DataView]', - float32Tag$1 = '[object Float32Array]', - float64Tag$1 = '[object Float64Array]', - int8Tag$1 = '[object Int8Array]', - int16Tag$1 = '[object Int16Array]', - int32Tag$1 = '[object Int32Array]', - uint8Tag$1 = '[object Uint8Array]', - uint8ClampedTag$1 = '[object Uint8ClampedArray]', - uint16Tag$1 = '[object Uint16Array]', - uint32Tag$1 = '[object Uint32Array]'; - -/** - * Initializes an object clone based on its `toStringTag`. - * - * **Note:** This function only supports cloning values with tags of - * `Boolean`, `Date`, `Error`, `Map`, `Number`, `RegExp`, `Set`, or `String`. - * - * @private - * @param {Object} object The object to clone. - * @param {string} tag The `toStringTag` of the object to clone. - * @param {boolean} [isDeep] Specify a deep clone. - * @returns {Object} Returns the initialized clone. - */ -function initCloneByTag$1(object, tag, isDeep) { - var Ctor = object.constructor; - switch (tag) { - case arrayBufferTag$1: - return cloneArrayBuffer(object); - - case boolTag$1: - case dateTag$1: - return new Ctor(+object); - - case dataViewTag$1: - return cloneDataView(object, isDeep); - - case float32Tag$1: case float64Tag$1: - case int8Tag$1: case int16Tag$1: case int32Tag$1: - case uint8Tag$1: case uint8ClampedTag$1: case uint16Tag$1: case uint32Tag$1: - return cloneTypedArray(object, isDeep); - - case mapTag$2: - return new Ctor; - - case numberTag$1: - case stringTag$1: - return new Ctor(object); - - case regexpTag$1: - return cloneRegExp(object); - - case setTag$2: - return new Ctor; - - case symbolTag$1: - return cloneSymbol(object); - } -} - -var _initCloneByTag = initCloneByTag$1; - -var isObject$1 = isObject_1$1; - -/** Built-in value references. */ -var objectCreate = Object.create; - -/** - * The base implementation of `_.create` without support for assigning - * properties to the created object. - * - * @private - * @param {Object} proto The object to inherit from. - * @returns {Object} Returns the new object. - */ -var baseCreate$1 = (function() { - function object() {} - return function(proto) { - if (!isObject$1(proto)) { - return {}; - } - if (objectCreate) { - return objectCreate(proto); - } - object.prototype = proto; - var result = new object; - object.prototype = undefined; - return result; - }; -}()); - -var _baseCreate = baseCreate$1; - -var baseCreate = _baseCreate, - getPrototype = _getPrototype, - isPrototype = _isPrototype; - -/** - * Initializes an object clone. - * - * @private - * @param {Object} object The object to clone. - * @returns {Object} Returns the initialized clone. - */ -function initCloneObject$1(object) { - return (typeof object.constructor == 'function' && !isPrototype(object)) - ? baseCreate(getPrototype(object)) - : {}; -} - -var _initCloneObject = initCloneObject$1; - -var getTag$2 = _getTag, - isObjectLike$1 = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var mapTag$1 = '[object Map]'; - -/** - * The base implementation of `_.isMap` without Node.js optimizations. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a map, else `false`. - */ -function baseIsMap$1(value) { - return isObjectLike$1(value) && getTag$2(value) == mapTag$1; -} - -var _baseIsMap = baseIsMap$1; - -var baseIsMap = _baseIsMap, - baseUnary$1 = _baseUnary, - nodeUtil$1 = _nodeUtil.exports; - -/* Node.js helper references. */ -var nodeIsMap = nodeUtil$1 && nodeUtil$1.isMap; - -/** - * Checks if `value` is classified as a `Map` object. - * - * @static - * @memberOf _ - * @since 4.3.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a map, else `false`. - * @example - * - * _.isMap(new Map); - * // => true - * - * _.isMap(new WeakMap); - * // => false - */ -var isMap$1 = nodeIsMap ? baseUnary$1(nodeIsMap) : baseIsMap; - -var isMap_1 = isMap$1; - -var getTag$1 = _getTag, - isObjectLike = isObjectLike_1$1; - -/** `Object#toString` result references. */ -var setTag$1 = '[object Set]'; - -/** - * The base implementation of `_.isSet` without Node.js optimizations. - * - * @private - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a set, else `false`. - */ -function baseIsSet$1(value) { - return isObjectLike(value) && getTag$1(value) == setTag$1; -} - -var _baseIsSet = baseIsSet$1; - -var baseIsSet = _baseIsSet, - baseUnary = _baseUnary, - nodeUtil = _nodeUtil.exports; - -/* Node.js helper references. */ -var nodeIsSet = nodeUtil && nodeUtil.isSet; - -/** - * Checks if `value` is classified as a `Set` object. - * - * @static - * @memberOf _ - * @since 4.3.0 - * @category Lang - * @param {*} value The value to check. - * @returns {boolean} Returns `true` if `value` is a set, else `false`. - * @example - * - * _.isSet(new Set); - * // => true - * - * _.isSet(new WeakSet); - * // => false - */ -var isSet$1 = nodeIsSet ? baseUnary(nodeIsSet) : baseIsSet; - -var isSet_1 = isSet$1; - -var Stack = _Stack, - arrayEach = _arrayEach, - assignValue = _assignValue, - baseAssign = _baseAssign, - baseAssignIn = _baseAssignIn, - cloneBuffer = _cloneBuffer.exports, - copyArray = _copyArray, - copySymbols = _copySymbols, - copySymbolsIn = _copySymbolsIn, - getAllKeys = _getAllKeys, - getAllKeysIn$1 = _getAllKeysIn, - getTag = _getTag, - initCloneArray = _initCloneArray, - initCloneByTag = _initCloneByTag, - initCloneObject = _initCloneObject, - isArray = isArray_1, - isBuffer = isBuffer$3.exports, - isMap = isMap_1, - isObject = isObject_1$1, - isSet = isSet_1, - keys = keys_1, - keysIn = keysIn_1; - -/** Used to compose bitmasks for cloning. */ -var CLONE_DEEP_FLAG$1 = 1, - CLONE_FLAT_FLAG$1 = 2, - CLONE_SYMBOLS_FLAG$1 = 4; - -/** `Object#toString` result references. */ -var argsTag = '[object Arguments]', - arrayTag = '[object Array]', - boolTag = '[object Boolean]', - dateTag = '[object Date]', - errorTag = '[object Error]', - funcTag = '[object Function]', - genTag = '[object GeneratorFunction]', - mapTag = '[object Map]', - numberTag = '[object Number]', - objectTag = '[object Object]', - regexpTag = '[object RegExp]', - setTag = '[object Set]', - stringTag = '[object String]', - symbolTag = '[object Symbol]', - weakMapTag = '[object WeakMap]'; - -var arrayBufferTag = '[object ArrayBuffer]', - dataViewTag = '[object DataView]', - float32Tag = '[object Float32Array]', - float64Tag = '[object Float64Array]', - int8Tag = '[object Int8Array]', - int16Tag = '[object Int16Array]', - int32Tag = '[object Int32Array]', - uint8Tag = '[object Uint8Array]', - uint8ClampedTag = '[object Uint8ClampedArray]', - uint16Tag = '[object Uint16Array]', - uint32Tag = '[object Uint32Array]'; - -/** Used to identify `toStringTag` values supported by `_.clone`. */ -var cloneableTags = {}; -cloneableTags[argsTag] = cloneableTags[arrayTag] = -cloneableTags[arrayBufferTag] = cloneableTags[dataViewTag] = -cloneableTags[boolTag] = cloneableTags[dateTag] = -cloneableTags[float32Tag] = cloneableTags[float64Tag] = -cloneableTags[int8Tag] = cloneableTags[int16Tag] = -cloneableTags[int32Tag] = cloneableTags[mapTag] = -cloneableTags[numberTag] = cloneableTags[objectTag] = -cloneableTags[regexpTag] = cloneableTags[setTag] = -cloneableTags[stringTag] = cloneableTags[symbolTag] = -cloneableTags[uint8Tag] = cloneableTags[uint8ClampedTag] = -cloneableTags[uint16Tag] = cloneableTags[uint32Tag] = true; -cloneableTags[errorTag] = cloneableTags[funcTag] = -cloneableTags[weakMapTag] = false; - -/** - * The base implementation of `_.clone` and `_.cloneDeep` which tracks - * traversed objects. - * - * @private - * @param {*} value The value to clone. - * @param {boolean} bitmask The bitmask flags. - * 1 - Deep clone - * 2 - Flatten inherited properties - * 4 - Clone symbols - * @param {Function} [customizer] The function to customize cloning. - * @param {string} [key] The key of `value`. - * @param {Object} [object] The parent object of `value`. - * @param {Object} [stack] Tracks traversed objects and their clone counterparts. - * @returns {*} Returns the cloned value. - */ -function baseClone$1(value, bitmask, customizer, key, object, stack) { - var result, - isDeep = bitmask & CLONE_DEEP_FLAG$1, - isFlat = bitmask & CLONE_FLAT_FLAG$1, - isFull = bitmask & CLONE_SYMBOLS_FLAG$1; - - if (customizer) { - result = object ? customizer(value, key, object, stack) : customizer(value); - } - if (result !== undefined) { - return result; - } - if (!isObject(value)) { - return value; - } - var isArr = isArray(value); - if (isArr) { - result = initCloneArray(value); - if (!isDeep) { - return copyArray(value, result); - } - } else { - var tag = getTag(value), - isFunc = tag == funcTag || tag == genTag; - - if (isBuffer(value)) { - return cloneBuffer(value, isDeep); - } - if (tag == objectTag || tag == argsTag || (isFunc && !object)) { - result = (isFlat || isFunc) ? {} : initCloneObject(value); - if (!isDeep) { - return isFlat - ? copySymbolsIn(value, baseAssignIn(result, value)) - : copySymbols(value, baseAssign(result, value)); - } - } else { - if (!cloneableTags[tag]) { - return object ? value : {}; - } - result = initCloneByTag(value, tag, isDeep); - } - } - // Check for circular references and return its corresponding clone. - stack || (stack = new Stack); - var stacked = stack.get(value); - if (stacked) { - return stacked; - } - stack.set(value, result); - - if (isSet(value)) { - value.forEach(function(subValue) { - result.add(baseClone$1(subValue, bitmask, customizer, subValue, value, stack)); - }); - } else if (isMap(value)) { - value.forEach(function(subValue, key) { - result.set(key, baseClone$1(subValue, bitmask, customizer, key, value, stack)); - }); - } - - var keysFunc = isFull - ? (isFlat ? getAllKeysIn$1 : getAllKeys) - : (isFlat ? keysIn : keys); - - var props = isArr ? undefined : keysFunc(value); - arrayEach(props || value, function(subValue, key) { - if (props) { - key = subValue; - subValue = value[key]; - } - // Recursively populate clone (susceptible to call stack limits). - assignValue(result, key, baseClone$1(subValue, bitmask, customizer, key, value, stack)); - }); - return result; -} - -var _baseClone = baseClone$1; - -var baseGet = _baseGet, - baseSlice = _baseSlice; - -/** - * Gets the parent value at `path` of `object`. - * - * @private - * @param {Object} object The object to query. - * @param {Array} path The path to get the parent value of. - * @returns {*} Returns the parent value. - */ -function parent$1(object, path) { - return path.length < 2 ? object : baseGet(object, baseSlice(path, 0, -1)); -} - -var _parent = parent$1; - -var castPath$1 = _castPath, - last = last_1, - parent = _parent, - toKey = _toKey; - -/** - * The base implementation of `_.unset`. - * - * @private - * @param {Object} object The object to modify. - * @param {Array|string} path The property path to unset. - * @returns {boolean} Returns `true` if the property is deleted, else `false`. - */ -function baseUnset$1(object, path) { - path = castPath$1(path, object); - object = parent(object, path); - return object == null || delete object[toKey(last(path))]; -} - -var _baseUnset = baseUnset$1; - -var isPlainObject = isPlainObject_1; - -/** - * Used by `_.omit` to customize its `_.cloneDeep` use to only clone plain - * objects. - * - * @private - * @param {*} value The value to inspect. - * @param {string} key The key of the property to inspect. - * @returns {*} Returns the uncloned value or `undefined` to defer cloning to `_.cloneDeep`. - */ -function customOmitClone$1(value) { - return isPlainObject(value) ? undefined : value; -} - -var _customOmitClone = customOmitClone$1; - -var baseFlatten = _baseFlatten; - -/** - * Flattens `array` a single level deep. - * - * @static - * @memberOf _ - * @since 0.1.0 - * @category Array - * @param {Array} array The array to flatten. - * @returns {Array} Returns the new flattened array. - * @example - * - * _.flatten([1, [2, [3, [4]], 5]]); - * // => [1, 2, [3, [4]], 5] - */ -function flatten$1(array) { - var length = array == null ? 0 : array.length; - return length ? baseFlatten(array, 1) : []; -} - -var flatten_1 = flatten$1; - -var flatten = flatten_1, - overRest = _overRest, - setToString = _setToString; - -/** - * A specialized version of `baseRest` which flattens the rest array. - * - * @private - * @param {Function} func The function to apply a rest parameter to. - * @returns {Function} Returns the new function. - */ -function flatRest$1(func) { - return setToString(overRest(func, undefined, flatten), func + ''); -} - -var _flatRest = flatRest$1; - -var arrayMap = _arrayMap, - baseClone = _baseClone, - baseUnset = _baseUnset, - castPath = _castPath, - copyObject = _copyObject, - customOmitClone = _customOmitClone, - flatRest = _flatRest, - getAllKeysIn = _getAllKeysIn; - -/** Used to compose bitmasks for cloning. */ -var CLONE_DEEP_FLAG = 1, - CLONE_FLAT_FLAG = 2, - CLONE_SYMBOLS_FLAG = 4; - -/** - * The opposite of `_.pick`; this method creates an object composed of the - * own and inherited enumerable property paths of `object` that are not omitted. - * - * **Note:** This method is considerably slower than `_.pick`. - * - * @static - * @since 0.1.0 - * @memberOf _ - * @category Object - * @param {Object} object The source object. - * @param {...(string|string[])} [paths] The property paths to omit. - * @returns {Object} Returns the new object. - * @example - * - * var object = { 'a': 1, 'b': '2', 'c': 3 }; - * - * _.omit(object, ['a', 'c']); - * // => { 'b': '2' } - */ -var omit = flatRest(function(object, paths) { - var result = {}; - if (object == null) { - return result; - } - var isDeep = false; - paths = arrayMap(paths, function(path) { - path = castPath(path, object); - isDeep || (isDeep = path.length > 1); - return path; - }); - copyObject(object, getAllKeysIn(object), result); - if (isDeep) { - result = baseClone(result, CLONE_DEEP_FLAG | CLONE_FLAT_FLAG | CLONE_SYMBOLS_FLAG, customOmitClone); - } - var length = paths.length; - while (length--) { - baseUnset(result, paths[length]); - } - return result; -}); - -var omit_1 = omit; - -var COLOR_PANEL = ['#1890FF', '#66B5FF', '#41D9C7', '#2FC25B', '#6EDB8F', '#9AE65C', '#FACC14', '#E6965C', '#57AD71', '#223273', '#738AE6', '#7564CC', '#8543E0', '#A877ED', '#5C8EE6', '#13C2C2', '#70E0E0', '#5CA3E6', '#3436C7', '#8082FF', '#DD81E6', '#F04864', '#FA7D92', '#D598D9']; - -function _typeof$3(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$3 = function _typeof(obj) { return typeof obj; }; } else { _typeof$3 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$3(obj); } - -function _extends$3() { _extends$3 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$3.apply(this, arguments); } - -function _objectWithoutProperties$1(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose$1(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose$1(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _classCallCheck$3(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$3(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$3(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$3(Constructor.prototype, protoProps); if (staticProps) _defineProperties$3(Constructor, staticProps); return Constructor; } - -function _inherits$3(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$3(subClass, superClass); } - -function _setPrototypeOf$3(o, p) { _setPrototypeOf$3 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$3(o, p); } - -function _createSuper$3(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$3(); return function _createSuperInternal() { var Super = _getPrototypeOf$3(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$3(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$3(this, result); }; } - -function _possibleConstructorReturn$3(self, call) { if (call && (_typeof$3(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$3(self); } - -function _assertThisInitialized$3(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$3() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$3(o) { _getPrototypeOf$3 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$3(o); } - -function ownKeys$2(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$2(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$2(Object(source), true).forEach(function (key) { _defineProperty$2(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$2(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$2(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } -var NODE_VALUE_KEY = 'value'; - -var computeNode = function computeNode(_ref) { - var _objectSpread2; - - var depth = _ref.depth, - node = _ref.node, - index = _ref.index, - valueKey = _ref.valueKey; - var children = node.children; - var childDepth = depth + 1; - var computedChildren = children && children.length ? children.map(function (child, i) { - return computeNode({ - depth: childDepth, - node: child, - index: i, - valueKey: valueKey - }); - }) : null; - var nodeValue; - - if (children && children.length) { - nodeValue = computedChildren.reduce(function (result, child) { - return result + child[NODE_VALUE_KEY]; - }, 0); - } else { - // TODO need to verify valueKey - nodeValue = _isNaN(node[valueKey]) || node[valueKey] <= 0 ? 0 : node[valueKey]; - } - - return _objectSpread$2(_objectSpread$2({}, node), {}, (_objectSpread2 = { - children: computedChildren - }, _defineProperty$2(_objectSpread2, NODE_VALUE_KEY, nodeValue), _defineProperty$2(_objectSpread2, "depth", depth), _defineProperty$2(_objectSpread2, "index", index), _objectSpread2)); -}; - -var filterRect = function filterRect(node) { - return { - x: node.x, - y: node.y, - width: node.width, - height: node.height - }; -}; // Compute the area for each child based on value & scale. - - -var getAreaOfChildren = function getAreaOfChildren(children, areaValueRatio) { - var ratio = areaValueRatio < 0 ? 0 : areaValueRatio; - return children.map(function (child) { - var area = child[NODE_VALUE_KEY] * ratio; - return _objectSpread$2(_objectSpread$2({}, child), {}, { - area: _isNaN(area) || area <= 0 ? 0 : area - }); - }); -}; // Computes the score for the specified row, as the worst aspect ratio. - - -var getWorstScore = function getWorstScore(row, parentSize, aspectRatio) { - var parentArea = parentSize * parentSize; - var rowArea = row.area * row.area; - - var _row$reduce = row.reduce(function (result, child) { - return { - min: Math.min(result.min, child.area), - max: Math.max(result.max, child.area) - }; - }, { - min: Infinity, - max: 0 - }), - min = _row$reduce.min, - max = _row$reduce.max; - - return rowArea ? Math.max(parentArea * max * aspectRatio / rowArea, rowArea / (parentArea * min * aspectRatio)) : Infinity; -}; - -var horizontalPosition = function horizontalPosition(row, parentSize, parentRect, isFlush) { - var rowHeight = parentSize ? Math.round(row.area / parentSize) : 0; - - if (isFlush || rowHeight > parentRect.height) { - rowHeight = parentRect.height; - } - - var curX = parentRect.x; - var child; - - for (var _i = 0, len = row.length; _i < len; _i++) { - child = row[_i]; - child.x = curX; - child.y = parentRect.y; - child.height = rowHeight; - child.width = Math.min(rowHeight ? Math.round(child.area / rowHeight) : 0, parentRect.x + parentRect.width - curX); - curX += child.width; - } // what's z - - - child.z = true; // add the remain x to the last one of row - - child.width += parentRect.x + parentRect.width - curX; - return _objectSpread$2(_objectSpread$2({}, parentRect), {}, { - y: parentRect.y + rowHeight, - height: parentRect.height - rowHeight - }); -}; - -var verticalPosition = function verticalPosition(row, parentSize, parentRect, isFlush) { - var rowWidth = parentSize ? Math.round(row.area / parentSize) : 0; - - if (isFlush || rowWidth > parentRect.width) { - rowWidth = parentRect.width; - } - - var curY = parentRect.y; - var child; - - for (var _i2 = 0, len = row.length; _i2 < len; _i2++) { - child = row[_i2]; - child.x = parentRect.x; - child.y = curY; - child.width = rowWidth; - child.height = Math.min(rowWidth ? Math.round(child.area / rowWidth) : 0, parentRect.y + parentRect.height - curY); - curY += child.height; - } - - if (child) { - child.z = false; - child.height += parentRect.y + parentRect.height - curY; - } - - return _objectSpread$2(_objectSpread$2({}, parentRect), {}, { - x: parentRect.x + rowWidth, - width: parentRect.width - rowWidth - }); -}; - -var position = function position(row, parentSize, parentRect, isFlush) { - if (parentSize === parentRect.width) { - return horizontalPosition(row, parentSize, parentRect, isFlush); - } - - return verticalPosition(row, parentSize, parentRect, isFlush); -}; // Recursively arranges the specified node's children into squarified rows. - - -var squarify = function squarify(node, aspectRatio) { - var children = node.children; - - if (children && children.length) { - var rect = filterRect(node); // maybe a bug - - var row = []; - var best = Infinity; // the best row score so far - - var child, score; // the current row score - - var size = Math.min(rect.width, rect.height); // initial orientation - - var scaleChildren = getAreaOfChildren(children, rect.width * rect.height / node[NODE_VALUE_KEY]); - var tempChildren = scaleChildren.slice(); - row.area = 0; - - while (tempChildren.length > 0) { - // row first - // eslint-disable-next-line prefer-destructuring - row.push(child = tempChildren[0]); - row.area += child.area; - score = getWorstScore(row, size, aspectRatio); - - if (score <= best) { - // continue with this orientation - tempChildren.shift(); - best = score; - } else { - // abort, and try a different orientation - row.area -= row.pop().area; - rect = position(row, size, rect, false); - size = Math.min(rect.width, rect.height); - row.length = row.area = 0; - best = Infinity; - } - } - - if (row.length) { - rect = position(row, size, rect, true); - row.length = row.area = 0; - } - - return _objectSpread$2(_objectSpread$2({}, node), {}, { - children: scaleChildren.map(function (c) { - return squarify(c, aspectRatio); - }) - }); - } - - return node; -}; - -var defaultState = { - isTooltipActive: false, - isAnimationFinished: false, - activeNode: null, - formatRoot: null, - currentRoot: null, - nestIndex: [] -}; -var Treemap = /*#__PURE__*/function (_PureComponent) { - _inherits$3(Treemap, _PureComponent); - - var _super = _createSuper$3(Treemap); - - function Treemap() { - var _this; - - _classCallCheck$3(this, Treemap); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = _objectSpread$2({}, defaultState); - - _this.handleAnimationEnd = function () { - var onAnimationEnd = _this.props.onAnimationEnd; - - _this.setState({ - isAnimationFinished: true - }); - - if (isFunction_1(onAnimationEnd)) { - onAnimationEnd(); - } - }; - - _this.handleAnimationStart = function () { - var onAnimationStart = _this.props.onAnimationStart; - - _this.setState({ - isAnimationFinished: false - }); - - if (isFunction_1(onAnimationStart)) { - onAnimationStart(); - } - }; - - return _this; - } - - _createClass$3(Treemap, [{ - key: "handleMouseEnter", - value: function handleMouseEnter(node, e) { - var _this$props = this.props, - onMouseEnter = _this$props.onMouseEnter, - children = _this$props.children; - var tooltipItem = findChildByType(children, Tooltip.displayName); - - if (tooltipItem) { - this.setState({ - isTooltipActive: true, - activeNode: node - }, function () { - if (onMouseEnter) { - onMouseEnter(node, e); - } - }); - } else if (onMouseEnter) { - onMouseEnter(node, e); - } - } - }, { - key: "handleMouseLeave", - value: function handleMouseLeave(node, e) { - var _this$props2 = this.props, - onMouseLeave = _this$props2.onMouseLeave, - children = _this$props2.children; - var tooltipItem = findChildByType(children, Tooltip.displayName); - - if (tooltipItem) { - this.setState({ - isTooltipActive: false, - activeNode: null - }, function () { - if (onMouseLeave) { - onMouseLeave(node, e); - } - }); - } else if (onMouseLeave) { - onMouseLeave(node, e); - } - } - }, { - key: "handleClick", - value: function handleClick(node) { - var _this$props3 = this.props, - onClick = _this$props3.onClick, - type = _this$props3.type; - - if (type === 'nest' && node.children) { - var _this$props4 = this.props, - width = _this$props4.width, - height = _this$props4.height, - dataKey = _this$props4.dataKey, - aspectRatio = _this$props4.aspectRatio; - var root = computeNode({ - depth: 0, - node: _objectSpread$2(_objectSpread$2({}, node), {}, { - x: 0, - y: 0, - width: width, - height: height - }), - index: 0, - valueKey: dataKey - }); - var formatRoot = squarify(root, aspectRatio); - var nestIndex = this.state.nestIndex; - nestIndex.push(node); - this.setState({ - formatRoot: formatRoot, - currentRoot: root, - nestIndex: nestIndex - }); - } - - if (onClick) { - onClick(node); - } - } - }, { - key: "handleNestIndex", - value: function handleNestIndex(node, i) { - var nestIndex = this.state.nestIndex; - var _this$props5 = this.props, - width = _this$props5.width, - height = _this$props5.height, - dataKey = _this$props5.dataKey, - aspectRatio = _this$props5.aspectRatio; - var root = computeNode({ - depth: 0, - node: _objectSpread$2(_objectSpread$2({}, node), {}, { - x: 0, - y: 0, - width: width, - height: height - }), - index: 0, - valueKey: dataKey - }); - var formatRoot = squarify(root, aspectRatio); - nestIndex = nestIndex.slice(0, i + 1); - this.setState({ - formatRoot: formatRoot, - currentRoot: node, - nestIndex: nestIndex - }); - } - }, { - key: "renderItem", - value: function renderItem(content, nodeProps, isLeaf) { - var _this2 = this; - - var _this$props6 = this.props, - isAnimationActive = _this$props6.isAnimationActive, - animationBegin = _this$props6.animationBegin, - animationDuration = _this$props6.animationDuration, - animationEasing = _this$props6.animationEasing, - isUpdateAnimationActive = _this$props6.isUpdateAnimationActive, - type = _this$props6.type, - animationId = _this$props6.animationId, - colorPanel = _this$props6.colorPanel; - var isAnimationFinished = this.state.isAnimationFinished; - var width = nodeProps.width, - height = nodeProps.height, - x = nodeProps.x, - y = nodeProps.y, - depth = nodeProps.depth; - var translateX = parseInt("".concat((Math.random() * 2 - 1) * width), 10); - var event = {}; - - if (isLeaf || type === 'nest') { - event = { - onMouseEnter: this.handleMouseEnter.bind(this, nodeProps), - onMouseLeave: this.handleMouseLeave.bind(this, nodeProps), - onClick: this.handleClick.bind(this, nodeProps) - }; - } - - if (!isAnimationActive) { - return /*#__PURE__*/React__default.createElement(Layer, event, this.constructor.renderContentItem(content, _objectSpread$2(_objectSpread$2({}, nodeProps), {}, { - isAnimationActive: false, - isUpdateAnimationActive: false, - width: width, - height: height, - x: x, - y: y - }), type, colorPanel)); - } - - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - key: "treemap-".concat(animationId), - from: { - x: x, - y: y, - width: width, - height: height - }, - to: { - x: x, - y: y, - width: width, - height: height - }, - onAnimationStart: this.handleAnimationStart, - onAnimationEnd: this.handleAnimationEnd - }, function (_ref2) { - var currX = _ref2.x, - currY = _ref2.y, - currWidth = _ref2.width, - currHeight = _ref2.height; - return /*#__PURE__*/React__default.createElement(Animate, { - from: "translate(".concat(translateX, "px, ").concat(translateX, "px)"), - to: "translate(0, 0)", - attributeName: "transform", - begin: animationBegin, - easing: animationEasing, - isActive: isAnimationActive, - duration: animationDuration - }, /*#__PURE__*/React__default.createElement(Layer, event, function () { - // when animation Duration , only render depth=1 nodes - if (depth > 2 && !isAnimationFinished) { - return null; - } - - return _this2.constructor.renderContentItem(content, _objectSpread$2(_objectSpread$2({}, nodeProps), {}, { - isAnimationActive: isAnimationActive, - isUpdateAnimationActive: !isUpdateAnimationActive, - width: currWidth, - height: currHeight, - x: currX, - y: currY - }), type, colorPanel); - }())); - }); - } - }, { - key: "renderNode", - value: function renderNode(root, node, i) { - var _this3 = this; - - var _this$props7 = this.props, - content = _this$props7.content, - type = _this$props7.type; - - var nodeProps = _objectSpread$2(_objectSpread$2(_objectSpread$2({}, filterProps(this.props)), node), {}, { - root: root - }); - - var isLeaf = !node.children || !node.children.length; - var currentRoot = this.state.currentRoot; - var isCurrentRootChild = (currentRoot.children || []).filter(function (item) { - return item.depth === node.depth && item.name === node.name; - }); - - if (!isCurrentRootChild.length && root.depth && type === 'nest') { - return null; - } - - return ( - /*#__PURE__*/ - // eslint-disable-next-line react/no-array-index-key - React__default.createElement(Layer, { - key: "recharts-treemap-node-".concat(i), - className: "recharts-treemap-depth-".concat(node.depth) - }, this.renderItem(content, nodeProps, isLeaf), node.children && node.children.length ? node.children.map(function (child, index) { - return _this3.renderNode(node, child, index); - }) : null) - ); - } - }, { - key: "renderAllNodes", - value: function renderAllNodes() { - var formatRoot = this.state.formatRoot; - - if (!formatRoot) { - return null; - } - - return this.renderNode(formatRoot, formatRoot, 0); - } - }, { - key: "renderTooltip", - value: function renderTooltip() { - var _this$props8 = this.props, - children = _this$props8.children, - nameKey = _this$props8.nameKey; - var tooltipItem = findChildByType(children, Tooltip.displayName); - - if (!tooltipItem) { - return null; - } - - var _this$props9 = this.props, - width = _this$props9.width, - height = _this$props9.height; - var _this$state = this.state, - isTooltipActive = _this$state.isTooltipActive, - activeNode = _this$state.activeNode; - var viewBox = { - x: 0, - y: 0, - width: width, - height: height - }; - var coordinate = activeNode ? { - x: activeNode.x + activeNode.width / 2, - y: activeNode.y + activeNode.height / 2 - } : null; - var payload = isTooltipActive && activeNode ? [{ - payload: activeNode, - name: getValueByDataKey(activeNode, nameKey, ''), - value: getValueByDataKey(activeNode, NODE_VALUE_KEY) - }] : []; - return /*#__PURE__*/React__default.cloneElement(tooltipItem, { - viewBox: viewBox, - active: isTooltipActive, - coordinate: coordinate, - label: '', - payload: payload - }); - } // render nest treemap - - }, { - key: "renderNestIndex", - value: function renderNestIndex() { - var _this4 = this; - - var _this$props10 = this.props, - nameKey = _this$props10.nameKey, - nestIndexContent = _this$props10.nestIndexContent; - var nestIndex = this.state.nestIndex; - return /*#__PURE__*/React__default.createElement("div", { - className: "recharts-treemap-nest-index-wrapper", - style: { - marginTop: '8px', - textAlign: 'center' - } - }, nestIndex.map(function (item, i) { - // TODO need to verify nameKey type - var name = get_1(item, nameKey, 'root'); - - var content = null; - - if ( /*#__PURE__*/React__default.isValidElement(nestIndexContent)) { - content = /*#__PURE__*/React__default.cloneElement(nestIndexContent, item, i); - } - - if (isFunction_1(nestIndexContent)) { - content = nestIndexContent(item, i); - } else { - content = name; - } - - return ( - /*#__PURE__*/ - // eslint-disable-next-line jsx-a11y/click-events-have-key-events - React__default.createElement("div", { - onClick: _this4.handleNestIndex.bind(_this4, item, i), - key: "nest-index-".concat(uniqueId()), - className: "recharts-treemap-nest-index-box", - style: { - cursor: 'pointer', - display: 'inline-block', - padding: '0 7px', - background: '#000', - color: '#fff', - marginRight: '3px' - } - }, content) - ); - })); - } - }, { - key: "render", - value: function render() { - if (!validateWidthHeight(this)) { - return null; - } - - var _this$props11 = this.props, - width = _this$props11.width, - height = _this$props11.height, - className = _this$props11.className, - style = _this$props11.style, - children = _this$props11.children, - type = _this$props11.type, - others = _objectWithoutProperties$1(_this$props11, ["width", "height", "className", "style", "children", "type"]); - - var attrs = filterProps(others); - return /*#__PURE__*/React__default.createElement("div", { - className: classNames('recharts-wrapper', className), - style: _objectSpread$2(_objectSpread$2({}, style), {}, { - position: 'relative', - cursor: 'default', - width: width, - height: height - }) - }, /*#__PURE__*/React__default.createElement(Surface, _extends$3({}, attrs, { - width: width, - height: type === 'nest' ? height - 30 : height - }), this.renderAllNodes(), filterSvgElements(children)), this.renderTooltip(), type === 'nest' && this.renderNestIndex()); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.data !== prevState.prevData || nextProps.type !== prevState.prevType || nextProps.width !== prevState.prevWidth || nextProps.height !== prevState.prevHeight || nextProps.dataKey !== prevState.prevDataKey || nextProps.aspectRatio !== prevState.prevAspectRatio) { - var root = computeNode({ - depth: 0, - node: { - children: nextProps.data, - x: 0, - y: 0, - width: nextProps.width, - height: nextProps.height - }, - index: 0, - valueKey: nextProps.dataKey - }); - var formatRoot = squarify(root, nextProps.aspectRatio); - return _objectSpread$2(_objectSpread$2({}, prevState), {}, { - formatRoot: formatRoot, - currentRoot: root, - nestIndex: [root], - prevAspectRatio: nextProps.aspectRatio, - prevData: nextProps.data, - prevWidth: nextProps.width, - prevHeight: nextProps.height, - prevDataKey: nextProps.dataKey, - prevType: nextProps.type - }); - } - - return null; - } - }, { - key: "renderContentItem", - value: function renderContentItem(content, nodeProps, type, colorPanel) { - if ( /*#__PURE__*/React__default.isValidElement(content)) { - return /*#__PURE__*/React__default.cloneElement(content, nodeProps); - } - - if (isFunction_1(content)) { - return content(nodeProps); - } // optimize default shape - - - var x = nodeProps.x, - y = nodeProps.y, - width = nodeProps.width, - height = nodeProps.height, - index = nodeProps.index; - var arrow = null; - - if (width > 10 && height > 10 && nodeProps.children && type === 'nest') { - arrow = /*#__PURE__*/React__default.createElement(Polygon, { - points: [{ - x: x + 2, - y: y + height / 2 - }, { - x: x + 6, - y: y + height / 2 + 3 - }, { - x: x + 2, - y: y + height / 2 + 6 - }] - }); - } - - var text = null; - var nameSize = getStringSize(nodeProps.name); - - if (width > 20 && height > 20 && nameSize.width < width && nameSize.height < height) { - text = /*#__PURE__*/React__default.createElement("text", { - x: x + 8, - y: y + height / 2 + 7, - fontSize: 14 - }, nodeProps.name); - } - - var colors = colorPanel || COLOR_PANEL; - return /*#__PURE__*/React__default.createElement("g", null, /*#__PURE__*/React__default.createElement(Rectangle, _extends$3({ - fill: nodeProps.depth < 2 ? colors[index % colors.length] : 'rgba(255,255,255,0)', - stroke: "#fff" - }, omit_1(nodeProps, 'children'))), arrow, text); - } - }]); - - return Treemap; -}(PureComponent); -Treemap.displayName = 'Treemap'; -Treemap.defaultProps = { - aspectRatio: 0.5 * (1 + Math.sqrt(5)), - dataKey: 'value', - type: 'flat', - isAnimationActive: !Global.isSsr, - isUpdateAnimationActive: !Global.isSsr, - animationBegin: 0, - animationDuration: 1500, - animationEasing: 'linear' -}; - -/** - * The base implementation of `_.sum` and `_.sumBy` without support for - * iteratee shorthands. - * - * @private - * @param {Array} array The array to iterate over. - * @param {Function} iteratee The function invoked per iteration. - * @returns {number} Returns the sum. - */ - -function baseSum$1(array, iteratee) { - var result, - index = -1, - length = array.length; - - while (++index < length) { - var current = iteratee(array[index]); - if (current !== undefined) { - result = result === undefined ? current : (result + current); - } - } - return result; -} - -var _baseSum = baseSum$1; - -var baseIteratee = _baseIteratee, - baseSum = _baseSum; - -/** - * This method is like `_.sum` except that it accepts `iteratee` which is - * invoked for each element in `array` to generate the value to be summed. - * The iteratee is invoked with one argument: (value). - * - * @static - * @memberOf _ - * @since 4.0.0 - * @category Math - * @param {Array} array The array to iterate over. - * @param {Function} [iteratee=_.identity] The iteratee invoked per element. - * @returns {number} Returns the sum. - * @example - * - * var objects = [{ 'n': 4 }, { 'n': 2 }, { 'n': 8 }, { 'n': 6 }]; - * - * _.sumBy(objects, function(o) { return o.n; }); - * // => 20 - * - * // The `_.property` iteratee shorthand. - * _.sumBy(objects, 'n'); - * // => 20 - */ -function sumBy(array, iteratee) { - return (array && array.length) - ? baseSum(array, baseIteratee(iteratee)) - : 0; -} - -var sumBy_1 = sumBy; - -function _typeof$2(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$2 = function _typeof(obj) { return typeof obj; }; } else { _typeof$2 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$2(obj); } - -function _objectWithoutProperties(source, excluded) { if (source == null) return {}; var target = _objectWithoutPropertiesLoose(source, excluded); var key, i; if (Object.getOwnPropertySymbols) { var sourceSymbolKeys = Object.getOwnPropertySymbols(source); for (i = 0; i < sourceSymbolKeys.length; i++) { key = sourceSymbolKeys[i]; if (excluded.indexOf(key) >= 0) continue; if (!Object.prototype.propertyIsEnumerable.call(source, key)) continue; target[key] = source[key]; } } return target; } - -function _objectWithoutPropertiesLoose(source, excluded) { if (source == null) return {}; var target = {}; var sourceKeys = Object.keys(source); var key, i; for (i = 0; i < sourceKeys.length; i++) { key = sourceKeys[i]; if (excluded.indexOf(key) >= 0) continue; target[key] = source[key]; } return target; } - -function _extends$2() { _extends$2 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$2.apply(this, arguments); } - -function _classCallCheck$2(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$2(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$2(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$2(Constructor.prototype, protoProps); if (staticProps) _defineProperties$2(Constructor, staticProps); return Constructor; } - -function _inherits$2(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$2(subClass, superClass); } - -function _setPrototypeOf$2(o, p) { _setPrototypeOf$2 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$2(o, p); } - -function _createSuper$2(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$2(); return function _createSuperInternal() { var Super = _getPrototypeOf$2(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$2(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$2(this, result); }; } - -function _possibleConstructorReturn$2(self, call) { if (call && (_typeof$2(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$2(self); } - -function _assertThisInitialized$2(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$2() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$2(o) { _getPrototypeOf$2 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$2(o); } - -function ownKeys$1(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread$1(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys$1(Object(source), true).forEach(function (key) { _defineProperty$1(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys$1(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty$1(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } -var defaultCoordinateOfTooltip = { - x: 0, - y: 0 -}; - -var interpolationGenerator = function interpolationGenerator(a, b) { - var ka = +a; - var kb = b - ka; - return function (t) { - return ka + kb * t; - }; -}; - -var centerY = function centerY(node) { - return node.y + node.dy / 2; -}; - -var getValue = function getValue(entry) { - return entry && entry.value || 0; -}; - -var getSumOfIds = function getSumOfIds(links, ids) { - return ids.reduce(function (result, id) { - return result + getValue(links[id]); - }, 0); -}; - -var getSumWithWeightedSource = function getSumWithWeightedSource(tree, links, ids) { - return ids.reduce(function (result, id) { - var link = links[id]; - var sourceNode = tree[link.source]; - return result + centerY(sourceNode) * getValue(links[id]); - }, 0); -}; - -var getSumWithWeightedTarget = function getSumWithWeightedTarget(tree, links, ids) { - return ids.reduce(function (result, id) { - var link = links[id]; - var targetNode = tree[link.target]; - return result + centerY(targetNode) * getValue(links[id]); - }, 0); -}; - -var ascendingY = function ascendingY(a, b) { - return a.y - b.y; -}; - -var searchTargetsAndSources = function searchTargetsAndSources(links, id) { - var sourceNodes = []; - var sourceLinks = []; - var targetNodes = []; - var targetLinks = []; - - for (var i = 0, len = links.length; i < len; i++) { - var link = links[i]; - - if (link.source === id) { - targetNodes.push(link.target); - targetLinks.push(i); - } - - if (link.target === id) { - sourceNodes.push(link.source); - sourceLinks.push(i); - } - } - - return { - sourceNodes: sourceNodes, - sourceLinks: sourceLinks, - targetLinks: targetLinks, - targetNodes: targetNodes - }; -}; - -var updateDepthOfTargets = function updateDepthOfTargets(tree, curNode) { - var targetNodes = curNode.targetNodes; - - for (var i = 0, len = targetNodes.length; i < len; i++) { - var target = tree[targetNodes[i]]; - - if (target) { - target.depth = Math.max(curNode.depth + 1, target.depth); - updateDepthOfTargets(tree, target); - } - } -}; - -var getNodesTree = function getNodesTree(_ref, width, nodeWidth) { - var nodes = _ref.nodes, - links = _ref.links; - var tree = nodes.map(function (entry, index) { - var result = searchTargetsAndSources(links, index); - return _objectSpread$1(_objectSpread$1(_objectSpread$1({}, entry), result), {}, { - value: Math.max(getSumOfIds(links, result.sourceLinks), getSumOfIds(links, result.targetLinks)), - depth: 0 - }); - }); - - for (var i = 0, len = tree.length; i < len; i++) { - var node = tree[i]; - - if (!node.sourceNodes.length) { - updateDepthOfTargets(tree, node); - } - } - - var maxDepth = maxBy_1(tree, function (entry) { - return entry.depth; - }).depth; - - if (maxDepth >= 1) { - var childWidth = (width - nodeWidth) / maxDepth; - - for (var _i = 0, _len = tree.length; _i < _len; _i++) { - var _node = tree[_i]; - - if (!_node.targetNodes.length) { - _node.depth = maxDepth; - } - - _node.x = _node.depth * childWidth; - _node.dx = nodeWidth; - } - } - - return { - tree: tree, - maxDepth: maxDepth - }; -}; - -var getDepthTree = function getDepthTree(tree) { - var result = []; - - for (var i = 0, len = tree.length; i < len; i++) { - var node = tree[i]; - - if (!result[node.depth]) { - result[node.depth] = []; - } - - result[node.depth].push(node); - } - - return result; -}; - -var updateYOfTree = function updateYOfTree(depthTree, height, nodePadding, links) { - var yRatio = min_1(depthTree.map(function (nodes) { - return (height - (nodes.length - 1) * nodePadding) / sumBy_1(nodes, getValue); - })); - - for (var d = 0, maxDepth = depthTree.length; d < maxDepth; d++) { - for (var i = 0, len = depthTree[d].length; i < len; i++) { - var node = depthTree[d][i]; - node.y = i; - node.dy = node.value * yRatio; - } - } - - return links.map(function (link) { - return _objectSpread$1(_objectSpread$1({}, link), {}, { - dy: getValue(link) * yRatio - }); - }); -}; - -var resolveCollisions = function resolveCollisions(depthTree, height, nodePadding) { - for (var i = 0, len = depthTree.length; i < len; i++) { - var nodes = depthTree[i]; - var n = nodes.length; // Sort by the value of y - - nodes.sort(ascendingY); - var y0 = 0; - - for (var j = 0; j < n; j++) { - var node = nodes[j]; - var dy = y0 - node.y; - - if (dy > 0) { - node.y += dy; - } - - y0 = node.y + node.dy + nodePadding; - } - - y0 = height + nodePadding; - - for (var _j = n - 1; _j >= 0; _j--) { - var _node2 = nodes[_j]; - - var _dy = _node2.y + _node2.dy + nodePadding - y0; - - if (_dy > 0) { - _node2.y -= _dy; - y0 = _node2.y; - } else { - break; - } - } - } -}; - -var relaxLeftToRight = function relaxLeftToRight(tree, depthTree, links, alpha) { - for (var i = 0, maxDepth = depthTree.length; i < maxDepth; i++) { - var nodes = depthTree[i]; - - for (var j = 0, len = nodes.length; j < len; j++) { - var node = nodes[j]; - - if (node.sourceLinks.length) { - var sourceSum = getSumOfIds(links, node.sourceLinks); - var weightedSum = getSumWithWeightedSource(tree, links, node.sourceLinks); - var y = weightedSum / sourceSum; - node.y += (y - centerY(node)) * alpha; - } - } - } -}; - -var relaxRightToLeft = function relaxRightToLeft(tree, depthTree, links, alpha) { - for (var i = depthTree.length - 1; i >= 0; i--) { - var nodes = depthTree[i]; - - for (var j = 0, len = nodes.length; j < len; j++) { - var node = nodes[j]; - - if (node.targetLinks.length) { - var targetSum = getSumOfIds(links, node.targetLinks); - var weightedSum = getSumWithWeightedTarget(tree, links, node.targetLinks); - var y = weightedSum / targetSum; - node.y += (y - centerY(node)) * alpha; - } - } - } -}; - -var updateYOfLinks = function updateYOfLinks(tree, links) { - for (var i = 0, len = tree.length; i < len; i++) { - var node = tree[i]; - var sy = 0; - var ty = 0; - node.targetLinks.sort(function (a, b) { - return tree[links[a].target].y - tree[links[b].target].y; - }); - node.sourceLinks.sort(function (a, b) { - return tree[links[a].source].y - tree[links[b].source].y; - }); - - for (var j = 0, tLen = node.targetLinks.length; j < tLen; j++) { - var link = links[node.targetLinks[j]]; - - if (link) { - link.sy = sy; - sy += link.dy; - } - } - - for (var _j2 = 0, sLen = node.sourceLinks.length; _j2 < sLen; _j2++) { - var _link = links[node.sourceLinks[_j2]]; - - if (_link) { - _link.ty = ty; - ty += _link.dy; - } - } - } -}; - -var computeData = function computeData(_ref2) { - var data = _ref2.data, - width = _ref2.width, - height = _ref2.height, - iterations = _ref2.iterations, - nodeWidth = _ref2.nodeWidth, - nodePadding = _ref2.nodePadding; - var links = data.links; - - var _getNodesTree = getNodesTree(data, width, nodeWidth), - tree = _getNodesTree.tree; - - var depthTree = getDepthTree(tree); - var newLinks = updateYOfTree(depthTree, height, nodePadding, links); - resolveCollisions(depthTree, height, nodePadding); - var alpha = 1; - - for (var i = 1; i <= iterations; i++) { - relaxRightToLeft(tree, depthTree, newLinks, alpha *= 0.99); - resolveCollisions(depthTree, height, nodePadding); - relaxLeftToRight(tree, depthTree, newLinks, alpha); - resolveCollisions(depthTree, height, nodePadding); - } - - updateYOfLinks(tree, newLinks); - return { - nodes: tree, - links: newLinks - }; -}; - -var getCoordinateOfTooltip = function getCoordinateOfTooltip(el, type) { - if (type === 'node') { - return { - x: el.x + el.width / 2, - y: el.y + el.height / 2 - }; - } - - return { - x: (el.sourceX + el.targetX) / 2, - y: (el.sourceY + el.targetY) / 2 - }; -}; - -var getPayloadOfTooltip = function getPayloadOfTooltip(el, type, nameKey) { - var payload = el.payload; - - if (type === 'node') { - return [{ - payload: el, - name: getValueByDataKey(payload, nameKey, ''), - value: getValueByDataKey(payload, 'value') - }]; - } - - if (payload.source && payload.target) { - var sourceName = getValueByDataKey(payload.source, nameKey, ''); - var targetName = getValueByDataKey(payload.target, nameKey, ''); - return [{ - payload: el, - name: "".concat(sourceName, " - ").concat(targetName), - value: getValueByDataKey(payload, 'value') - }]; - } - - return []; -}; - -var Sankey = /*#__PURE__*/function (_PureComponent) { - _inherits$2(Sankey, _PureComponent); - - var _super = _createSuper$2(Sankey); - - function Sankey() { - var _this; - - _classCallCheck$2(this, Sankey); - - for (var _len2 = arguments.length, args = new Array(_len2), _key = 0; _key < _len2; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - activeElement: null, - activeElementType: null, - isTooltipActive: false, - nodes: [], - links: [] - }; - return _this; - } - - _createClass$2(Sankey, [{ - key: "handleMouseEnter", - value: function handleMouseEnter(el, type, e) { - var _this$props = this.props, - onMouseEnter = _this$props.onMouseEnter, - children = _this$props.children; - var tooltipItem = findChildByType(children, Tooltip.displayName); - - if (tooltipItem) { - this.setState(function (prev) { - if (tooltipItem.props.trigger === 'hover') { - return _objectSpread$1(_objectSpread$1({}, prev), {}, { - activeElement: el, - activeElementType: type, - isTooltipActive: true - }); - } - - return prev; - }, function () { - if (onMouseEnter) { - onMouseEnter(el, type, e); - } - }); - } else if (onMouseEnter) { - onMouseEnter(el, type, e); - } - } - }, { - key: "handleMouseLeave", - value: function handleMouseLeave(el, type, e) { - var _this$props2 = this.props, - onMouseLeave = _this$props2.onMouseLeave, - children = _this$props2.children; - var tooltipItem = findChildByType(children, Tooltip.displayName); - - if (tooltipItem) { - this.setState(function (prev) { - if (tooltipItem.props.trigger === 'hover') { - return _objectSpread$1(_objectSpread$1({}, prev), {}, { - activeElement: undefined, - activeElementType: undefined, - isTooltipActive: false - }); - } - - return prev; - }, function () { - if (onMouseLeave) { - onMouseLeave(el, type, e); - } - }); - } else if (onMouseLeave) { - onMouseLeave(el, type, e); - } - } - }, { - key: "handleClick", - value: function handleClick(el, type, e) { - var _this$props3 = this.props, - onClick = _this$props3.onClick, - children = _this$props3.children; - var tooltipItem = findChildByType(children, Tooltip.displayName); - - if (tooltipItem && tooltipItem.props.trigger === 'click') { - if (this.state.isTooltipActive) { - this.setState(function (prev) { - return _objectSpread$1(_objectSpread$1({}, prev), {}, { - activeElement: undefined, - activeElementType: undefined, - isTooltipActive: false - }); - }); - } else { - this.setState(function (prev) { - return _objectSpread$1(_objectSpread$1({}, prev), {}, { - activeElement: el, - activeElementType: type, - isTooltipActive: true - }); - }); - } - } - - if (onClick) onClick(el, type, e); - } - }, { - key: "renderLinks", - value: function renderLinks(links, nodes) { - var _this2 = this; - - var _this$props4 = this.props, - linkCurvature = _this$props4.linkCurvature, - linkContent = _this$props4.link, - margin = _this$props4.margin; - var top = get_1(margin, 'top') || 0; - var left = get_1(margin, 'left') || 0; - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-sankey-links", - key: "recharts-sankey-links" - }, links.map(function (link, i) { - var sourceRelativeY = link.sy, - targetRelativeY = link.ty, - linkWidth = link.dy; - var source = nodes[link.source]; - var target = nodes[link.target]; - var sourceX = source.x + source.dx + left; - var targetX = target.x + left; - var interpolationFunc = interpolationGenerator(sourceX, targetX); - var sourceControlX = interpolationFunc(linkCurvature); - var targetControlX = interpolationFunc(1 - linkCurvature); - var sourceY = source.y + sourceRelativeY + linkWidth / 2 + top; - var targetY = target.y + targetRelativeY + linkWidth / 2 + top; - - var linkProps = _objectSpread$1({ - sourceX: sourceX, - targetX: targetX, - sourceY: sourceY, - targetY: targetY, - sourceControlX: sourceControlX, - targetControlX: targetControlX, - sourceRelativeY: sourceRelativeY, - targetRelativeY: targetRelativeY, - linkWidth: linkWidth, - index: i, - payload: _objectSpread$1(_objectSpread$1({}, link), {}, { - source: source, - target: target - }) - }, filterProps(linkContent)); - - var events = { - onMouseEnter: _this2.handleMouseEnter.bind(_this2, linkProps, 'link'), - onMouseLeave: _this2.handleMouseLeave.bind(_this2, linkProps, 'link'), - onClick: _this2.handleClick.bind(_this2, linkProps, 'link') - }; - return ( - /*#__PURE__*/ - // eslint-disable-next-line react/no-array-index-key - React__default.createElement(Layer, _extends$2({ - key: "link".concat(i) - }, events), _this2.constructor.renderLinkItem(linkContent, linkProps)) - ); - })); - } - }, { - key: "renderNodes", - value: function renderNodes(nodes) { - var _this3 = this; - - var _this$props5 = this.props, - nodeContent = _this$props5.node, - margin = _this$props5.margin; - var top = get_1(margin, 'top') || 0; - var left = get_1(margin, 'left') || 0; - return /*#__PURE__*/React__default.createElement(Layer, { - className: "recharts-sankey-nodes", - key: "recharts-sankey-nodes" - }, nodes.map(function (node, i) { - var x = node.x, - y = node.y, - dx = node.dx, - dy = node.dy; - - var nodeProps = _objectSpread$1(_objectSpread$1({}, filterProps(nodeContent)), {}, { - x: x + left, - y: y + top, - width: dx, - height: dy, - index: i, - payload: node - }); - - var events = { - onMouseEnter: _this3.handleMouseEnter.bind(_this3, nodeProps, 'node'), - onMouseLeave: _this3.handleMouseLeave.bind(_this3, nodeProps, 'node'), - onClick: _this3.handleClick.bind(_this3, nodeProps, 'node') - }; - return ( - /*#__PURE__*/ - // eslint-disable-next-line react/no-array-index-key - React__default.createElement(Layer, _extends$2({ - key: "node".concat(i) - }, events), _this3.constructor.renderNodeItem(nodeContent, nodeProps)) - ); - })); - } - }, { - key: "renderTooltip", - value: function renderTooltip() { - var _this$props6 = this.props, - children = _this$props6.children, - width = _this$props6.width, - height = _this$props6.height, - nameKey = _this$props6.nameKey; - var tooltipItem = findChildByType(children, Tooltip.displayName); - - if (!tooltipItem) { - return null; - } - - var _this$state = this.state, - isTooltipActive = _this$state.isTooltipActive, - activeElement = _this$state.activeElement, - activeElementType = _this$state.activeElementType; - var viewBox = { - x: 0, - y: 0, - width: width, - height: height - }; - var coordinate = activeElement ? getCoordinateOfTooltip(activeElement, activeElementType) : defaultCoordinateOfTooltip; - var payload = activeElement ? getPayloadOfTooltip(activeElement, activeElementType, nameKey) : []; - return /*#__PURE__*/React__default.cloneElement(tooltipItem, { - viewBox: viewBox, - active: isTooltipActive, - coordinate: coordinate, - label: '', - payload: payload - }); - } - }, { - key: "render", - value: function render() { - if (!validateWidthHeight(this)) { - return null; - } - - var _this$props7 = this.props, - width = _this$props7.width, - height = _this$props7.height, - className = _this$props7.className, - style = _this$props7.style, - children = _this$props7.children, - others = _objectWithoutProperties(_this$props7, ["width", "height", "className", "style", "children"]); - - var _this$state2 = this.state, - links = _this$state2.links, - nodes = _this$state2.nodes; - var attrs = filterProps(others); - return /*#__PURE__*/React__default.createElement("div", { - className: classNames('recharts-wrapper', className), - style: _objectSpread$1(_objectSpread$1({}, style), {}, { - position: 'relative', - cursor: 'default', - width: width, - height: height - }) - }, /*#__PURE__*/React__default.createElement(Surface, _extends$2({}, attrs, { - width: width, - height: height - }), filterSvgElements(children), this.renderLinks(links, nodes), this.renderNodes(nodes)), this.renderTooltip()); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - var data = nextProps.data, - width = nextProps.width, - height = nextProps.height, - margin = nextProps.margin, - iterations = nextProps.iterations, - nodeWidth = nextProps.nodeWidth, - nodePadding = nextProps.nodePadding; - - if (data !== prevState.prevData || width !== prevState.prevWidth || height !== prevState.prevHeight || !shallowEqual(margin, prevState.prevMargin) || iterations !== prevState.prevIterations || nodeWidth !== prevState.prevNodeWidth || nodePadding !== prevState.prevNodePadding) { - var contentWidth = width - (margin && margin.left || 0) - (margin && margin.right || 0); - var contentHeight = height - (margin && margin.top || 0) - (margin && margin.bottom || 0); - - var _computeData = computeData({ - data: data, - width: contentWidth, - height: contentHeight, - iterations: iterations, - nodeWidth: nodeWidth, - nodePadding: nodePadding - }), - links = _computeData.links, - nodes = _computeData.nodes; - - return _objectSpread$1(_objectSpread$1({}, prevState), {}, { - nodes: nodes, - links: links, - prevData: data, - prevWidth: iterations, - prevHeight: height, - prevMargin: margin, - prevNodePadding: nodePadding, - prevNodeWidth: nodeWidth, - prevIterations: iterations - }); - } - - return null; - } - }, { - key: "renderLinkItem", - value: function renderLinkItem(option, props) { - if ( /*#__PURE__*/React__default.isValidElement(option)) { - return /*#__PURE__*/React__default.cloneElement(option, props); - } - - if (isFunction_1(option)) { - return option(props); - } - - var sourceX = props.sourceX, - sourceY = props.sourceY, - sourceControlX = props.sourceControlX, - targetX = props.targetX, - targetY = props.targetY, - targetControlX = props.targetControlX, - linkWidth = props.linkWidth, - others = _objectWithoutProperties(props, ["sourceX", "sourceY", "sourceControlX", "targetX", "targetY", "targetControlX", "linkWidth"]); - - return /*#__PURE__*/React__default.createElement("path", _extends$2({ - className: "recharts-sankey-link", - d: "\n M".concat(sourceX, ",").concat(sourceY, "\n C").concat(sourceControlX, ",").concat(sourceY, " ").concat(targetControlX, ",").concat(targetY, " ").concat(targetX, ",").concat(targetY, "\n "), - fill: "none", - stroke: "#333", - strokeWidth: linkWidth, - strokeOpacity: "0.2" - }, filterProps(others))); - } - }, { - key: "renderNodeItem", - value: function renderNodeItem(option, props) { - if ( /*#__PURE__*/React__default.isValidElement(option)) { - return /*#__PURE__*/React__default.cloneElement(option, props); - } - - if (isFunction_1(option)) { - return option(props); - } - - return /*#__PURE__*/React__default.createElement(Rectangle, _extends$2({ - className: "recharts-sankey-node", - fill: "#0088fe", - fillOpacity: "0.8" - }, filterProps(props))); - } - }]); - - return Sankey; -}(PureComponent); -Sankey.displayName = 'Sankey'; -Sankey.defaultProps = { - nameKey: 'name', - dataKey: 'value', - nodePadding: 10, - nodeWidth: 10, - linkCurvature: 0.5, - iterations: 32, - margin: { - top: 5, - right: 5, - bottom: 5, - left: 5 - } -}; - -/** - * @fileOverview Radar Chart - */ -var RadarChart = generateCategoricalChart({ - chartName: 'RadarChart', - GraphicalChild: Radar, - axisComponents: [{ - axisType: 'angleAxis', - AxisComp: PolarAngleAxis - }, { - axisType: 'radiusAxis', - AxisComp: PolarRadiusAxis - }], - formatAxisMap: formatAxisMap$1, - defaultProps: { - layout: 'centric', - startAngle: 90, - endAngle: -270, - cx: '50%', - cy: '50%', - innerRadius: 0, - outerRadius: '80%' - } -}); - -/** - * @fileOverview Scatter Chart - */ -var ScatterChart = generateCategoricalChart({ - chartName: 'ScatterChart', - GraphicalChild: Scatter, - defaultTooltipEventType: 'item', - validateTooltipEventTypes: ['item'], - axisComponents: [{ - axisType: 'xAxis', - AxisComp: XAxis - }, { - axisType: 'yAxis', - AxisComp: YAxis - }, { - axisType: 'zAxis', - AxisComp: ZAxis - }], - formatAxisMap: formatAxisMap -}); - -/** - * @fileOverview Area Chart - */ -var AreaChart = generateCategoricalChart({ - chartName: 'AreaChart', - GraphicalChild: Area, - axisComponents: [{ - axisType: 'xAxis', - AxisComp: XAxis - }, { - axisType: 'yAxis', - AxisComp: YAxis - }], - formatAxisMap: formatAxisMap -}); - -/** - * @fileOverview Radar Bar Chart - */ -var RadialBarChart = generateCategoricalChart({ - chartName: 'RadialBarChart', - GraphicalChild: RadialBar, - legendContent: 'children', - defaultTooltipEventType: 'axis', - validateTooltipEventTypes: ['axis', 'item'], - axisComponents: [{ - axisType: 'angleAxis', - AxisComp: PolarAngleAxis - }, { - axisType: 'radiusAxis', - AxisComp: PolarRadiusAxis - }], - formatAxisMap: formatAxisMap$1, - defaultProps: { - layout: 'radial', - startAngle: 0, - endAngle: 360, - cx: '50%', - cy: '50%', - innerRadius: 0, - outerRadius: '80%' - } -}); - -/** - * @fileOverview Composed Chart - */ -var ComposedChart = generateCategoricalChart({ - chartName: 'ComposedChart', - GraphicalChild: [Line, Area, Bar, Scatter], - axisComponents: [{ - axisType: 'xAxis', - AxisComp: XAxis - }, { - axisType: 'yAxis', - AxisComp: YAxis - }, { - axisType: 'zAxis', - AxisComp: ZAxis - }], - formatAxisMap: formatAxisMap -}); - -function _typeof$1(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof$1 = function _typeof(obj) { return typeof obj; }; } else { _typeof$1 = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof$1(obj); } - -function _extends$1() { _extends$1 = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends$1.apply(this, arguments); } - -function _classCallCheck$1(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties$1(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass$1(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties$1(Constructor.prototype, protoProps); if (staticProps) _defineProperties$1(Constructor, staticProps); return Constructor; } - -function _inherits$1(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf$1(subClass, superClass); } - -function _setPrototypeOf$1(o, p) { _setPrototypeOf$1 = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf$1(o, p); } - -function _createSuper$1(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct$1(); return function _createSuperInternal() { var Super = _getPrototypeOf$1(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf$1(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn$1(this, result); }; } - -function _possibleConstructorReturn$1(self, call) { if (call && (_typeof$1(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized$1(self); } - -function _assertThisInitialized$1(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct$1() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf$1(o) { _getPrototypeOf$1 = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf$1(o); } - -var getTrapezoidPath = function getTrapezoidPath(x, y, upperWidth, lowerWidth, height) { - var widthGap = upperWidth - lowerWidth; - var path; - path = "M ".concat(x, ",").concat(y); - path += "L ".concat(x + upperWidth, ",").concat(y); - path += "L ".concat(x + upperWidth - widthGap / 2, ",").concat(y + height); - path += "L ".concat(x + upperWidth - widthGap / 2 - lowerWidth, ",").concat(y + height); - path += "L ".concat(x, ",").concat(y, " Z"); - return path; -}; - -var Trapezoid = /*#__PURE__*/function (_PureComponent) { - _inherits$1(Trapezoid, _PureComponent); - - var _super = _createSuper$1(Trapezoid); - - function Trapezoid() { - var _this; - - _classCallCheck$1(this, Trapezoid); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - totalLength: -1 - }; - _this.node = void 0; - return _this; - } - - _createClass$1(Trapezoid, [{ - key: "componentDidMount", - value: - /* eslint-disable react/no-did-mount-set-state */ - function componentDidMount() { - if (this.node && this.node.getTotalLength) { - try { - var totalLength = this.node.getTotalLength(); - - if (totalLength) { - this.setState({ - totalLength: totalLength - }); - } - } catch (err) {// calculate total length error - } - } - } - }, { - key: "render", - value: function render() { - var _this2 = this; - - var _this$props = this.props, - x = _this$props.x, - y = _this$props.y, - upperWidth = _this$props.upperWidth, - lowerWidth = _this$props.lowerWidth, - height = _this$props.height, - className = _this$props.className; - var totalLength = this.state.totalLength; - var _this$props2 = this.props, - animationEasing = _this$props2.animationEasing, - animationDuration = _this$props2.animationDuration, - animationBegin = _this$props2.animationBegin, - isUpdateAnimationActive = _this$props2.isUpdateAnimationActive; - - if (x !== +x || y !== +y || upperWidth !== +upperWidth || lowerWidth !== +lowerWidth || height !== +height || upperWidth === 0 && lowerWidth === 0 || height === 0) { - return null; - } - - var layerClass = classNames('recharts-trapezoid', className); - - if (!isUpdateAnimationActive) { - return /*#__PURE__*/React__default.createElement("g", null, /*#__PURE__*/React__default.createElement("path", _extends$1({}, filterProps(this.props, true), { - className: layerClass, - d: getTrapezoidPath(x, y, upperWidth, lowerWidth, height) - }))); - } - - return /*#__PURE__*/React__default.createElement(Animate, { - canBegin: totalLength > 0, - from: { - upperWidth: 0, - lowerWidth: 0, - height: height, - x: x, - y: y - }, - to: { - upperWidth: upperWidth, - lowerWidth: lowerWidth, - height: height, - x: x, - y: y - }, - duration: animationDuration, - animationEasing: animationEasing, - isActive: isUpdateAnimationActive - }, function (_ref) { - var currUpperWidth = _ref.upperWidth, - currLowerWidth = _ref.lowerWidth, - currHeight = _ref.height, - currX = _ref.x, - currY = _ref.y; - return /*#__PURE__*/React__default.createElement(Animate, { - canBegin: totalLength > 0, - from: "0px ".concat(totalLength === -1 ? 1 : totalLength, "px"), - to: "".concat(totalLength, "px 0px"), - attributeName: "strokeDasharray", - begin: animationBegin, - duration: animationDuration, - easing: animationEasing - }, /*#__PURE__*/React__default.createElement("path", _extends$1({}, filterProps(_this2.props, true), { - className: layerClass, - d: getTrapezoidPath(currX, currY, currUpperWidth, currLowerWidth, currHeight), - ref: function ref(node) { - _this2.node = node; - } - }))); - }); - } - }]); - - return Trapezoid; -}(PureComponent); -Trapezoid.defaultProps = { - x: 0, - y: 0, - upperWidth: 0, - lowerWidth: 0, - height: 0, - isUpdateAnimationActive: false, - animationBegin: 0, - animationDuration: 1500, - animationEasing: 'ease' -}; - -function _typeof(obj) { "@babel/helpers - typeof"; if (typeof Symbol === "function" && typeof Symbol.iterator === "symbol") { _typeof = function _typeof(obj) { return typeof obj; }; } else { _typeof = function _typeof(obj) { return obj && typeof Symbol === "function" && obj.constructor === Symbol && obj !== Symbol.prototype ? "symbol" : typeof obj; }; } return _typeof(obj); } - -function _slicedToArray(arr, i) { return _arrayWithHoles(arr) || _iterableToArrayLimit(arr, i) || _unsupportedIterableToArray(arr, i) || _nonIterableRest(); } - -function _nonIterableRest() { throw new TypeError("Invalid attempt to destructure non-iterable instance.\nIn order to be iterable, non-array objects must have a [Symbol.iterator]() method."); } - -function _unsupportedIterableToArray(o, minLen) { if (!o) return; if (typeof o === "string") return _arrayLikeToArray(o, minLen); var n = Object.prototype.toString.call(o).slice(8, -1); if (n === "Object" && o.constructor) n = o.constructor.name; if (n === "Map" || n === "Set") return Array.from(o); if (n === "Arguments" || /^(?:Ui|I)nt(?:8|16|32)(?:Clamped)?Array$/.test(n)) return _arrayLikeToArray(o, minLen); } - -function _arrayLikeToArray(arr, len) { if (len == null || len > arr.length) len = arr.length; for (var i = 0, arr2 = new Array(len); i < len; i++) { arr2[i] = arr[i]; } return arr2; } - -function _iterableToArrayLimit(arr, i) { if (typeof Symbol === "undefined" || !(Symbol.iterator in Object(arr))) return; var _arr = []; var _n = true; var _d = false; var _e = undefined; try { for (var _i = arr[Symbol.iterator](), _s; !(_n = (_s = _i.next()).done); _n = true) { _arr.push(_s.value); if (i && _arr.length === i) break; } } catch (err) { _d = true; _e = err; } finally { try { if (!_n && _i["return"] != null) _i["return"](); } finally { if (_d) throw _e; } } return _arr; } - -function _arrayWithHoles(arr) { if (Array.isArray(arr)) return arr; } - -function _extends() { _extends = Object.assign || function (target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i]; for (var key in source) { if (Object.prototype.hasOwnProperty.call(source, key)) { target[key] = source[key]; } } } return target; }; return _extends.apply(this, arguments); } - -function ownKeys(object, enumerableOnly) { var keys = Object.keys(object); if (Object.getOwnPropertySymbols) { var symbols = Object.getOwnPropertySymbols(object); if (enumerableOnly) symbols = symbols.filter(function (sym) { return Object.getOwnPropertyDescriptor(object, sym).enumerable; }); keys.push.apply(keys, symbols); } return keys; } - -function _objectSpread(target) { for (var i = 1; i < arguments.length; i++) { var source = arguments[i] != null ? arguments[i] : {}; if (i % 2) { ownKeys(Object(source), true).forEach(function (key) { _defineProperty(target, key, source[key]); }); } else if (Object.getOwnPropertyDescriptors) { Object.defineProperties(target, Object.getOwnPropertyDescriptors(source)); } else { ownKeys(Object(source)).forEach(function (key) { Object.defineProperty(target, key, Object.getOwnPropertyDescriptor(source, key)); }); } } return target; } - -function _defineProperty(obj, key, value) { if (key in obj) { Object.defineProperty(obj, key, { value: value, enumerable: true, configurable: true, writable: true }); } else { obj[key] = value; } return obj; } - -function _classCallCheck(instance, Constructor) { if (!(instance instanceof Constructor)) { throw new TypeError("Cannot call a class as a function"); } } - -function _defineProperties(target, props) { for (var i = 0; i < props.length; i++) { var descriptor = props[i]; descriptor.enumerable = descriptor.enumerable || false; descriptor.configurable = true; if ("value" in descriptor) descriptor.writable = true; Object.defineProperty(target, descriptor.key, descriptor); } } - -function _createClass(Constructor, protoProps, staticProps) { if (protoProps) _defineProperties(Constructor.prototype, protoProps); if (staticProps) _defineProperties(Constructor, staticProps); return Constructor; } - -function _inherits(subClass, superClass) { if (typeof superClass !== "function" && superClass !== null) { throw new TypeError("Super expression must either be null or a function"); } subClass.prototype = Object.create(superClass && superClass.prototype, { constructor: { value: subClass, writable: true, configurable: true } }); if (superClass) _setPrototypeOf(subClass, superClass); } - -function _setPrototypeOf(o, p) { _setPrototypeOf = Object.setPrototypeOf || function _setPrototypeOf(o, p) { o.__proto__ = p; return o; }; return _setPrototypeOf(o, p); } - -function _createSuper(Derived) { var hasNativeReflectConstruct = _isNativeReflectConstruct(); return function _createSuperInternal() { var Super = _getPrototypeOf(Derived), result; if (hasNativeReflectConstruct) { var NewTarget = _getPrototypeOf(this).constructor; result = Reflect.construct(Super, arguments, NewTarget); } else { result = Super.apply(this, arguments); } return _possibleConstructorReturn(this, result); }; } - -function _possibleConstructorReturn(self, call) { if (call && (_typeof(call) === "object" || typeof call === "function")) { return call; } return _assertThisInitialized(self); } - -function _assertThisInitialized(self) { if (self === void 0) { throw new ReferenceError("this hasn't been initialised - super() hasn't been called"); } return self; } - -function _isNativeReflectConstruct() { if (typeof Reflect === "undefined" || !Reflect.construct) return false; if (Reflect.construct.sham) return false; if (typeof Proxy === "function") return true; try { Boolean.prototype.valueOf.call(Reflect.construct(Boolean, [], function () {})); return true; } catch (e) { return false; } } - -function _getPrototypeOf(o) { _getPrototypeOf = Object.setPrototypeOf ? Object.getPrototypeOf : function _getPrototypeOf(o) { return o.__proto__ || Object.getPrototypeOf(o); }; return _getPrototypeOf(o); } -var Funnel = /*#__PURE__*/function (_PureComponent) { - _inherits(Funnel, _PureComponent); - - var _super = _createSuper(Funnel); - - function Funnel() { - var _this; - - _classCallCheck(this, Funnel); - - for (var _len = arguments.length, args = new Array(_len), _key = 0; _key < _len; _key++) { - args[_key] = arguments[_key]; - } - - _this = _super.call.apply(_super, [this].concat(args)); - _this.state = { - isAnimationFinished: false - }; - - _this.handleAnimationEnd = function () { - var onAnimationEnd = _this.props.onAnimationEnd; - - _this.setState({ - isAnimationFinished: true - }); - - if (isFunction_1(onAnimationEnd)) { - onAnimationEnd(); - } - }; - - _this.handleAnimationStart = function () { - var onAnimationStart = _this.props.onAnimationStart; - - _this.setState({ - isAnimationFinished: false - }); - - if (isFunction_1(onAnimationStart)) { - onAnimationStart(); - } - }; - - return _this; - } - - _createClass(Funnel, [{ - key: "isActiveIndex", - value: function isActiveIndex(i) { - var activeIndex = this.props.activeIndex; - - if (Array.isArray(activeIndex)) { - return activeIndex.indexOf(i) !== -1; - } - - return i === activeIndex; - } - }, { - key: "renderTrapezoidsStatically", - value: function renderTrapezoidsStatically(trapezoids) { - var _this2 = this; - - var activeShape = this.props.activeShape; - return trapezoids.map(function (entry, i) { - var trapezoidOptions = _this2.isActiveIndex(i) ? activeShape : null; - - var trapezoidProps = _objectSpread(_objectSpread({}, entry), {}, { - stroke: entry.stroke - }); - - return /*#__PURE__*/React__default.createElement(Layer, _extends({ - className: "recharts-funnel-trapezoid" - }, adaptEventsOfChild(_this2.props, entry, i), { - key: "trapezoid-".concat(i) // eslint-disable-line react/no-array-index-key - - }), Funnel.renderTrapezoidItem(trapezoidOptions, trapezoidProps)); - }); - } - }, { - key: "renderTrapezoidsWithAnimation", - value: function renderTrapezoidsWithAnimation() { - var _this3 = this; - - var _this$props = this.props, - trapezoids = _this$props.trapezoids, - isAnimationActive = _this$props.isAnimationActive, - animationBegin = _this$props.animationBegin, - animationDuration = _this$props.animationDuration, - animationEasing = _this$props.animationEasing, - animationId = _this$props.animationId; - var prevTrapezoids = this.state.prevTrapezoids; - return /*#__PURE__*/React__default.createElement(Animate, { - begin: animationBegin, - duration: animationDuration, - isActive: isAnimationActive, - easing: animationEasing, - from: { - t: 0 - }, - to: { - t: 1 - }, - key: "funnel-".concat(animationId), - onAnimationStart: this.handleAnimationStart, - onAnimationEnd: this.handleAnimationEnd - }, function (_ref) { - var t = _ref.t; - var stepData = trapezoids.map(function (entry, index) { - var prev = prevTrapezoids && prevTrapezoids[index]; - - if (prev) { - var _interpolatorX = interpolateNumber$2(prev.x, entry.x); - - var _interpolatorY = interpolateNumber$2(prev.y, entry.y); - - var _interpolatorUpperWidth = interpolateNumber$2(prev.upperWidth, entry.upperWidth); - - var _interpolatorLowerWidth = interpolateNumber$2(prev.lowerWidth, entry.lowerWidth); - - var _interpolatorHeight = interpolateNumber$2(prev.height, entry.height); - - return _objectSpread(_objectSpread({}, entry), {}, { - x: _interpolatorX(t), - y: _interpolatorY(t), - upperWidth: _interpolatorUpperWidth(t), - lowerWidth: _interpolatorLowerWidth(t), - height: _interpolatorHeight(t) - }); - } - - var interpolatorX = interpolateNumber$2(entry.x + entry.upperWidth / 2, entry.x); - var interpolatorY = interpolateNumber$2(entry.y + entry.height / 2, entry.y); - var interpolatorUpperWidth = interpolateNumber$2(0, entry.upperWidth); - var interpolatorLowerWidth = interpolateNumber$2(0, entry.lowerWidth); - var interpolatorHeight = interpolateNumber$2(0, entry.height); - return _objectSpread(_objectSpread({}, entry), {}, { - x: interpolatorX(t), - y: interpolatorY(t), - upperWidth: interpolatorUpperWidth(t), - lowerWidth: interpolatorLowerWidth(t), - height: interpolatorHeight(t) - }); - }); - return /*#__PURE__*/React__default.createElement(Layer, null, _this3.renderTrapezoidsStatically(stepData)); - }); - } - }, { - key: "renderTrapezoids", - value: function renderTrapezoids() { - var _this$props2 = this.props, - trapezoids = _this$props2.trapezoids, - isAnimationActive = _this$props2.isAnimationActive; - var prevTrapezoids = this.state.prevTrapezoids; - - if (isAnimationActive && trapezoids && trapezoids.length && (!prevTrapezoids || !isEqual_1(prevTrapezoids, trapezoids))) { - return this.renderTrapezoidsWithAnimation(); - } - - return this.renderTrapezoidsStatically(trapezoids); - } - }, { - key: "render", - value: function render() { - var _this$props3 = this.props, - hide = _this$props3.hide, - trapezoids = _this$props3.trapezoids, - className = _this$props3.className, - isAnimationActive = _this$props3.isAnimationActive; - var isAnimationFinished = this.state.isAnimationFinished; - - if (hide || !trapezoids || !trapezoids.length) { - return null; - } - - var layerClass = classNames('recharts-trapezoids', className); - return /*#__PURE__*/React__default.createElement(Layer, { - className: layerClass - }, this.renderTrapezoids(), (!isAnimationActive || isAnimationFinished) && LabelList.renderCallByParent(this.props, trapezoids)); - } - }], [{ - key: "getDerivedStateFromProps", - value: function getDerivedStateFromProps(nextProps, prevState) { - if (nextProps.animationId !== prevState.prevAnimationId) { - return { - prevAnimationId: nextProps.animationId, - curTrapezoids: nextProps.trapezoids, - prevTrapezoids: prevState.curTrapezoids - }; - } - - if (nextProps.trapezoids !== prevState.curTrapezoids) { - return { - curTrapezoids: nextProps.trapezoids - }; - } - - return null; - } - }, { - key: "renderTrapezoidItem", - value: function renderTrapezoidItem(option, props) { - if ( /*#__PURE__*/React__default.isValidElement(option)) { - return /*#__PURE__*/React__default.cloneElement(option, props); - } - - if (isFunction_1(option)) { - return option(props); - } - - if (isPlainObject_1(option)) { - return /*#__PURE__*/React__default.createElement(Trapezoid, _extends({}, props, option)); - } - - return /*#__PURE__*/React__default.createElement(Trapezoid, props); - } - }]); - - return Funnel; -}(PureComponent); -Funnel.displayName = 'Funnel'; -Funnel.defaultProps = { - stroke: '#fff', - fill: '#808080', - legendType: 'rect', - labelLine: true, - hide: false, - isAnimationActive: !Global.isSsr, - animationBegin: 400, - animationDuration: 1500, - animationEasing: 'ease', - nameKey: 'name', - lastShapeType: 'triangle' -}; - -Funnel.getRealFunnelData = function (item) { - var _item$props = item.props, - data = _item$props.data, - children = _item$props.children; - var presentationProps = filterProps(item.props); - var cells = findAllByType(children, Cell.displayName); - - if (data && data.length) { - return data.map(function (entry, index) { - return _objectSpread(_objectSpread(_objectSpread({ - payload: entry - }, presentationProps), entry), cells && cells[index] && cells[index].props); - }); - } - - if (cells && cells.length) { - return cells.map(function (cell) { - return _objectSpread(_objectSpread({}, presentationProps), cell.props); - }); - } - - return []; -}; - -Funnel.getRealWidthHeight = function (item, offset) { - var customWidth = item.props.width; - var width = offset.width, - height = offset.height, - left = offset.left, - right = offset.right, - top = offset.top, - bottom = offset.bottom; - var realHeight = height; - var realWidth = width; - - if (isNumber_1(customWidth)) { - realWidth = customWidth; - } else if (isString_1(customWidth)) { - realWidth = realWidth * parseFloat(customWidth) / 100; - } - - return { - realWidth: realWidth - left - right - 50, - realHeight: realHeight - bottom - top, - offsetX: (width - realWidth) / 2, - offsetY: (height - realHeight) / 2 - }; -}; - -Funnel.getComposedData = function (_ref2) { - var item = _ref2.item, - offset = _ref2.offset; - var funnelData = Funnel.getRealFunnelData(item); - var _item$props2 = item.props, - dataKey = _item$props2.dataKey, - nameKey = _item$props2.nameKey, - tooltipType = _item$props2.tooltipType, - lastShapeType = _item$props2.lastShapeType, - reversed = _item$props2.reversed; - var left = offset.left, - top = offset.top; - - var _Funnel$getRealWidthH = Funnel.getRealWidthHeight(item, offset), - realHeight = _Funnel$getRealWidthH.realHeight, - realWidth = _Funnel$getRealWidthH.realWidth, - offsetX = _Funnel$getRealWidthH.offsetX, - offsetY = _Funnel$getRealWidthH.offsetY; - - var maxValue = Math.max.apply(null, funnelData.map(function (entry) { - return getValueByDataKey(entry, dataKey, 0); - })); - var len = funnelData.length; - var rowHeight = realHeight / len; - var parentViewBox = { - x: offset.left, - y: offset.top, - width: offset.width, - height: offset.height - }; - var trapezoids = funnelData.map(function (entry, i) { - var rawVal = getValueByDataKey(entry, dataKey, 0); - var name = getValueByDataKey(entry, nameKey, i); - var val = rawVal; - var nextVal; - - if (i !== len - 1) { - nextVal = getValueByDataKey(funnelData[i + 1], dataKey, 0); - - if (nextVal instanceof Array) { - var _nextVal = nextVal; - - var _nextVal2 = _slicedToArray(_nextVal, 1); - - nextVal = _nextVal2[0]; - } - } else if (rawVal instanceof Array && rawVal.length === 2) { - var _rawVal = _slicedToArray(rawVal, 2); - - val = _rawVal[0]; - nextVal = _rawVal[1]; - } else if (lastShapeType === 'rectangle') { - nextVal = val; - } else { - nextVal = 0; - } - - var x = (maxValue - val) * realWidth / (2 * maxValue) + top + 25 + offsetX; - var y = rowHeight * i + left + offsetY; - var upperWidth = val / maxValue * realWidth; - var lowerWidth = nextVal / maxValue * realWidth; - var tooltipPayload = [{ - name: name, - value: val, - payload: entry, - dataKey: dataKey, - type: tooltipType - }]; - var tooltipPosition = { - x: x + upperWidth / 2, - y: y + rowHeight / 2 - }; - return _objectSpread(_objectSpread({ - x: x, - y: y, - width: Math.max(upperWidth, lowerWidth), - upperWidth: upperWidth, - lowerWidth: lowerWidth, - height: rowHeight, - name: name, - val: val, - tooltipPayload: tooltipPayload, - tooltipPosition: tooltipPosition - }, omit_1(entry, 'width')), {}, { - payload: entry, - parentViewBox: parentViewBox, - labelViewBox: { - x: x + (upperWidth - lowerWidth) / 4, - y: y, - width: Math.abs(upperWidth - lowerWidth) / 2 + Math.min(upperWidth, lowerWidth), - height: rowHeight - } - }); - }); - - if (reversed) { - trapezoids = trapezoids.map(function (entry, index) { - var newY = entry.y - index * rowHeight + (len - 1 - index) * rowHeight; - return _objectSpread(_objectSpread({}, entry), {}, { - upperWidth: entry.lowerWidth, - lowerWidth: entry.upperWidth, - x: entry.x - (entry.lowerWidth - entry.upperWidth) / 2, - y: entry.y - index * rowHeight + (len - 1 - index) * rowHeight, - tooltipPosition: _objectSpread(_objectSpread({}, entry.tooltipPosition), {}, { - y: newY + rowHeight / 2 - }), - labelViewBox: _objectSpread(_objectSpread({}, entry.labelViewBox), {}, { - y: newY - }) - }); - }); - } - - return { - trapezoids: trapezoids, - data: funnelData - }; -}; - -/** - * @fileOverview Funnel Chart - */ -var FunnelChart = generateCategoricalChart({ - chartName: 'FunnelChart', - GraphicalChild: Funnel, - validateTooltipEventTypes: ['item'], - defaultTooltipEventType: 'item', - axisComponents: [], - defaultProps: { - layout: 'centric' - } -}); - -// "export type" declarations on separate lines are in use - -var Recharts = /*#__PURE__*/Object.freeze({ - __proto__: null, - Surface: Surface, - Layer: Layer, - Legend: Legend, - Tooltip: Tooltip, - ResponsiveContainer: ResponsiveContainer, - Cell: Cell, - Text: Text, - Label: Label, - LabelList: LabelList, - Customized: Customized, - Sector: Sector, - Curve: Curve, - Rectangle: Rectangle, - Polygon: Polygon, - Dot: Dot, - Cross: Cross, - Symbols: Symbols, - PolarGrid: PolarGrid, - PolarRadiusAxis: PolarRadiusAxis, - PolarAngleAxis: PolarAngleAxis, - Pie: Pie, - Radar: Radar, - RadialBar: RadialBar, - Brush: Brush, - ReferenceLine: ReferenceLine, - ReferenceDot: ReferenceDot, - ReferenceArea: ReferenceArea, - CartesianAxis: CartesianAxis, - CartesianGrid: CartesianGrid, - Line: Line, - Area: Area, - Bar: Bar, - Scatter: Scatter, - XAxis: XAxis, - YAxis: YAxis, - ZAxis: ZAxis, - ErrorBar: ErrorBar, - LineChart: LineChart, - BarChart: BarChart, - PieChart: PieChart, - Treemap: Treemap, - Sankey: Sankey, - RadarChart: RadarChart, - ScatterChart: ScatterChart, - AreaChart: AreaChart, - RadialBarChart: RadialBarChart, - ComposedChart: ComposedChart, - Funnel: Funnel, - FunnelChart: FunnelChart, - Trapezoid: Trapezoid, - Global: Global -}); - -const htmlTags = "a abbr address area article aside audio b base bdi bdo big blockquote body br button canvas caption cite code col colgroup data datalist dd del details dfn dialog div dl dt em embed fieldset figcaption figure footer form h1 h2 h3 h4 h5 h6 head header hr html i iframe img input ins kbd keygen label legend li link main map mark menu menuitem meta meter nav noscript object ol optgroup option output p param picture pre progress q rp rt ruby s samp script section select small source span strong style sub summary sup table tbody td textarea tfoot th thead time title tr track u ul var video wbr".split(' '); -const svgTags = "circle clipPath defs ellipse g line linearGradient mask path pattern polygon polyline radialGradient rect stop svg text tspan".split(' '); -function StaticHtml({ html, ...props }) { - if ('text' in html) { - return html.text; - } - else if ('element' in html) { - const [tag, attrsList, cs] = html.element; - const attrs = {}; - for (const [k, v] of attrsList) { - attrs[k] = v; - } - const children = cs.map(html => StaticHtml({ html })); - if (tag === "hr") { - // React is greatly concerned by
s having children. - return React.createElement("hr", null); - } - if (tag === "InteractiveCode") { - if (!("fmt" in attrs)) { - return React.createElement("span", { className: "red" }, - "Give ", - tag, - " a fmt attr."); - } - const fmt = attrs.fmt; - return React.createElement(InteractiveCode, { fmt: fmt }); - } - if (htmlTags.includes(tag) || svgTags.includes(tag)) { - if (children.length === 0) { - return React.createElement(tag, attrs); - } - else { - return React.createElement(tag, attrs, children); - } - } - else if (tag in Recharts) { - const component = Recharts[tag]; - return React.createElement(component, attrs, children); - } - return React.createElement("span", { className: "red" }, - "Unknown component ", - tag); - } - else { - throw new Error(`Unexpected ${html}`); - } -} -/** Quick way of creating time-dependent animations from widgets. - * Eventually this will get replaced with fancy RPC stuff to stream frames. - */ -function AnimatedHtml(props) { - const { frames } = props; - const framesPerSecond = props.framesPerSecond ?? 10; - if (framesPerSecond <= 0 || framesPerSecond > 60) { - throw new Error(`Invalid fps ${framesPerSecond}. Should be between 0 and 60.`); - } - const [t, setT] = React.useState(0); - React.useEffect(() => { - window.setTimeout(() => setT(t + 1), 1000.0 / framesPerSecond); - }, [t]); - const frame = frames[t % frames.length]; - return StaticHtml({ html: frame }); -} -function staticHtml (props) { - if ("frames" in props) { - return AnimatedHtml(props); - } - else { - return StaticHtml(props); - } -} - -export { AnimatedHtml, StaticHtml, staticHtml as default }; diff --git a/widget/package-lock.json b/widget/package-lock.json index 9088e58..de6a94d 100644 --- a/widget/package-lock.json +++ b/widget/package-lock.json @@ -9,23 +9,27 @@ "version": "0.1.0", "license": "Apache-2.0", "dependencies": { - "@leanprover/infoview": "^0.2.0", + "@leanprover/infoview": "~0.4.2", + "@penrose/core": "~1.3.1-alpha.363", + "@svgdotjs/svg.js": "^3.1.2", "react": "^18.2.0", "react-dom": "^18.2.0", - "recharts": "^2.1.16", + "recharts": "^2.2.0", + "use-resize-observer": "^9.1.0", "vscode-languageserver-protocol": "^3.17.2" }, "devDependencies": { - "@rollup/plugin-commonjs": "^22.0.2", - "@rollup/plugin-node-resolve": "^14.1.0", - "@rollup/plugin-replace": "^4.0.0", - "@rollup/plugin-typescript": "^8.5.0", - "@types/react": "^17.0.39", - "@types/react-dom": "^17.0.11", - "rollup": "^2.79.0", - "rollup-plugin-string": "^3.0.0", - "tslib": "^2.4.0", - "typescript": "^4.8.3" + "@rollup/plugin-commonjs": "^24.0.0", + "@rollup/plugin-node-resolve": "^15.0.1", + "@rollup/plugin-replace": "^5.0.2", + "@rollup/plugin-terser": "^0.3.0", + "@rollup/plugin-typescript": "^11.0.0", + "@types/react": "^18.0.26", + "@types/react-dom": "^18.0.10", + "glob": "^8.0.3", + "rollup": "^3.9.1", + "tslib": "^2.4.1", + "typescript": "^4.9.4" } }, "node_modules/@babel/runtime": { @@ -39,123 +43,266 @@ "node": ">=6.9.0" } }, + "node_modules/@datastructures-js/heap": { + "version": "3.2.0", + "resolved": "https://registry.npmjs.org/@datastructures-js/heap/-/heap-3.2.0.tgz", + "integrity": "sha512-FcU5ZAyb+VIOZz1HABsJUsbJi2ZyUDO7aoe97hq4d3tK3z8nMgwdxf5bO0gafR0ExFi18YTspntqHLzt4XOgnA==" + }, + "node_modules/@datastructures-js/queue": { + "version": "4.2.3", + "resolved": "https://registry.npmjs.org/@datastructures-js/queue/-/queue-4.2.3.tgz", + "integrity": "sha512-GWVMorC/xi2V2ta+Z/CPgPGHL2ZJozcj48g7y2nIX5GIGZGRrbShSHgvMViJwHJurUzJYOdIdRZnWDRrROFwJA==" + }, + "node_modules/@jridgewell/gen-mapping": { + "version": "0.3.2", + "resolved": "https://registry.npmjs.org/@jridgewell/gen-mapping/-/gen-mapping-0.3.2.tgz", + "integrity": "sha512-mh65xKQAzI6iBcFzwv28KVWSmCkdRBWoOh+bYQGW3+6OZvbbN3TqMGo5hqYxQniRcH9F2VZIoJCm4pa3BPDK/A==", + "dev": true, + "dependencies": { + "@jridgewell/set-array": "^1.0.1", + "@jridgewell/sourcemap-codec": "^1.4.10", + "@jridgewell/trace-mapping": "^0.3.9" + }, + "engines": { + "node": ">=6.0.0" + } + }, + "node_modules/@jridgewell/resolve-uri": { + "version": "3.1.0", + "resolved": "https://registry.npmjs.org/@jridgewell/resolve-uri/-/resolve-uri-3.1.0.tgz", + "integrity": "sha512-F2msla3tad+Mfht5cJq7LSXcdudKTWCVYUgw6pLFOOHSTtZlj6SWNYAp+AhuqLmWdBO2X5hPrLcu8cVP8fy28w==", + "dev": true, + "engines": { + "node": ">=6.0.0" + } + }, + "node_modules/@jridgewell/set-array": { + "version": "1.1.2", + "resolved": "https://registry.npmjs.org/@jridgewell/set-array/-/set-array-1.1.2.tgz", + "integrity": "sha512-xnkseuNADM0gt2bs+BvhO0p78Mk762YnZdsuzFV018NoG1Sj1SCQvpSqa7XUaTam5vAGasABV9qXASMKnFMwMw==", + "dev": true, + "engines": { + "node": ">=6.0.0" + } + }, + "node_modules/@jridgewell/source-map": { + "version": "0.3.2", + "resolved": "https://registry.npmjs.org/@jridgewell/source-map/-/source-map-0.3.2.tgz", + "integrity": "sha512-m7O9o2uR8k2ObDysZYzdfhb08VuEml5oWGiosa1VdaPZ/A6QyPkAJuwN0Q1lhULOf6B7MtQmHENS743hWtCrgw==", + "dev": true, + "dependencies": { + "@jridgewell/gen-mapping": "^0.3.0", + "@jridgewell/trace-mapping": "^0.3.9" + } + }, + "node_modules/@jridgewell/sourcemap-codec": { + "version": "1.4.14", + "resolved": "https://registry.npmjs.org/@jridgewell/sourcemap-codec/-/sourcemap-codec-1.4.14.tgz", + "integrity": "sha512-XPSJHWmi394fuUuzDnGz1wiKqWfo1yXecHQMRf2l6hztTO+nPru658AyDngaBe7isIxEkRsPR3FZh+s7iVa4Uw==", + "dev": true + }, + "node_modules/@jridgewell/trace-mapping": { + "version": "0.3.17", + "resolved": "https://registry.npmjs.org/@jridgewell/trace-mapping/-/trace-mapping-0.3.17.tgz", + "integrity": "sha512-MCNzAp77qzKca9+W/+I0+sEpaUnZoeasnghNeVc41VZCEKaCH73Vq3BZZ/SzWIgrqE4H4ceI+p+b6C0mHf9T4g==", + "dev": true, + "dependencies": { + "@jridgewell/resolve-uri": "3.1.0", + "@jridgewell/sourcemap-codec": "1.4.14" + } + }, + "node_modules/@juggle/resize-observer": { + "version": "3.4.0", + "resolved": "https://registry.npmjs.org/@juggle/resize-observer/-/resize-observer-3.4.0.tgz", + "integrity": "sha512-dfLbk+PwWvFzSxwk3n5ySL0hfBog779o8h68wK/7/APo/7cgyWp5jcXockbxdk5kFRkbeXWm4Fbi9FrdN381sA==" + }, "node_modules/@leanprover/infoview": { - "version": "0.2.0", - "resolved": "https://registry.npmjs.org/@leanprover/infoview/-/infoview-0.2.0.tgz", - "integrity": "sha512-QrN0YV/WieVHf/WeMQpni8Lo6WQRQ8lQ3ImSoAyZd4aToJYzTV4r2dh3PtjwZ65/c/GMWF8O6o8sssGWCEZYwA==", + "version": "0.4.2", + "resolved": "https://registry.npmjs.org/@leanprover/infoview/-/infoview-0.4.2.tgz", + "integrity": "sha512-Bj+q/7n1xlnpRd4xsJ75uikZPwW5id/G4G1JPqhbbfMacTwAgzLNFaqJhzBrIUuF5nxJYGKRqRIeFtkalGxJ0g==", "dependencies": { - "@leanprover/infoview-api": "^0.1.0", - "@vscode/codicons": "0.0.28", - "marked": "^4.0.17", + "@leanprover/infoview-api": "~0.2.1", + "@vscode/codicons": "^0.0.32", + "es-module-shims": "^1.6.2", + "marked": "^4.2.2", "react-fast-compare": "^3.2.0", "tachyons": "^4.12.0", "vscode-languageserver-protocol": "^3.17.2" } }, "node_modules/@leanprover/infoview-api": { - "version": "0.1.0", - "resolved": "https://registry.npmjs.org/@leanprover/infoview-api/-/infoview-api-0.1.0.tgz", - "integrity": "sha512-vNMFrk5U1UlWhHp7G1e33mG7hJHyIErb3INgsVITbEW8CzOlAUrc83Pp3ARZykIc8++WJhT7WH+lALo8nO7G0Q==" + "version": "0.2.1", + "resolved": "https://registry.npmjs.org/@leanprover/infoview-api/-/infoview-api-0.2.1.tgz", + "integrity": "sha512-4sYdwOhUsa5wfvo/ZsCbcm8fBWcrATciZq0sWfmi5NRbIyZ+c2QjTm6D9CeYPCNvz9yvD1KBp/2+hKEZ8SOHkA==" + }, + "node_modules/@penrose/core": { + "version": "1.3.1-alpha.363", + "resolved": "https://registry.npmjs.org/@penrose/core/-/core-1.3.1-alpha.363.tgz", + "integrity": "sha512-rO5OkjT7+y0k8Rdg+DW6fxOC23V1pcH0+9FDSmJeM5lqfAW3phO2mQtiHI0VbcwMURDcH7wEGLVnpS7RVf1d0g==", + "dependencies": { + "@datastructures-js/heap": "^3.2.0", + "@datastructures-js/queue": "^4.1.3", + "consola": "^2.15.2", + "immutable": "^4.0.0-rc.12", + "lodash": "^4.17.15", + "mathjax-full": "^3.0.1", + "moo": "^0.5.1", + "nearley": "^2.20.1", + "poly-partition": "^1.0.2", + "recursive-diff": "^1.0.8", + "seedrandom": "^3.0.5", + "true-myth": "^4.1.0" + } }, "node_modules/@rollup/plugin-commonjs": { - "version": "22.0.2", - "resolved": "https://registry.npmjs.org/@rollup/plugin-commonjs/-/plugin-commonjs-22.0.2.tgz", - "integrity": "sha512-//NdP6iIwPbMTcazYsiBMbJW7gfmpHom33u1beiIoHDEM0Q9clvtQB1T0efvMqHeKsGohiHo97BCPCkBXdscwg==", + "version": "24.0.0", + "resolved": "https://registry.npmjs.org/@rollup/plugin-commonjs/-/plugin-commonjs-24.0.0.tgz", + "integrity": "sha512-0w0wyykzdyRRPHOb0cQt14mIBLujfAv6GgP6g8nvg/iBxEm112t3YPPq+Buqe2+imvElTka+bjNlJ/gB56TD8g==", "dev": true, "dependencies": { - "@rollup/pluginutils": "^3.1.0", + "@rollup/pluginutils": "^5.0.1", "commondir": "^1.0.1", - "estree-walker": "^2.0.1", - "glob": "^7.1.6", - "is-reference": "^1.2.1", - "magic-string": "^0.25.7", - "resolve": "^1.17.0" + "estree-walker": "^2.0.2", + "glob": "^8.0.3", + "is-reference": "1.2.1", + "magic-string": "^0.27.0" }, "engines": { - "node": ">= 12.0.0" + "node": ">=14.0.0" }, "peerDependencies": { - "rollup": "^2.68.0" + "rollup": "^2.68.0||^3.0.0" + }, + "peerDependenciesMeta": { + "rollup": { + "optional": true + } } }, "node_modules/@rollup/plugin-node-resolve": { - "version": "14.1.0", - "resolved": "https://registry.npmjs.org/@rollup/plugin-node-resolve/-/plugin-node-resolve-14.1.0.tgz", - "integrity": "sha512-5G2niJroNCz/1zqwXtk0t9+twOSDlG00k1Wfd7bkbbXmwg8H8dvgHdIWAun53Ps/rckfvOC7scDBjuGFg5OaWw==", + "version": "15.0.1", + "resolved": "https://registry.npmjs.org/@rollup/plugin-node-resolve/-/plugin-node-resolve-15.0.1.tgz", + "integrity": "sha512-ReY88T7JhJjeRVbfCyNj+NXAG3IIsVMsX9b5/9jC98dRP8/yxlZdz7mHZbHk5zHr24wZZICS5AcXsFZAXYUQEg==", "dev": true, "dependencies": { - "@rollup/pluginutils": "^3.1.0", - "@types/resolve": "1.17.1", + "@rollup/pluginutils": "^5.0.1", + "@types/resolve": "1.20.2", "deepmerge": "^4.2.2", - "is-builtin-module": "^3.1.0", + "is-builtin-module": "^3.2.0", "is-module": "^1.0.0", - "resolve": "^1.19.0" + "resolve": "^1.22.1" }, "engines": { - "node": ">= 10.0.0" + "node": ">=14.0.0" }, "peerDependencies": { - "rollup": "^2.78.0" + "rollup": "^2.78.0||^3.0.0" + }, + "peerDependenciesMeta": { + "rollup": { + "optional": true + } } }, "node_modules/@rollup/plugin-replace": { - "version": "4.0.0", - "resolved": "https://registry.npmjs.org/@rollup/plugin-replace/-/plugin-replace-4.0.0.tgz", - "integrity": "sha512-+rumQFiaNac9y64OHtkHGmdjm7us9bo1PlbgQfdihQtuNxzjpaB064HbRnewUOggLQxVCCyINfStkgmBeQpv1g==", + "version": "5.0.2", + "resolved": "https://registry.npmjs.org/@rollup/plugin-replace/-/plugin-replace-5.0.2.tgz", + "integrity": "sha512-M9YXNekv/C/iHHK+cvORzfRYfPbq0RDD8r0G+bMiTXjNGKulPnCT9O3Ss46WfhI6ZOCgApOP7xAdmCQJ+U2LAA==", "dev": true, "dependencies": { - "@rollup/pluginutils": "^3.1.0", - "magic-string": "^0.25.7" + "@rollup/pluginutils": "^5.0.1", + "magic-string": "^0.27.0" + }, + "engines": { + "node": ">=14.0.0" }, "peerDependencies": { - "rollup": "^1.20.0 || ^2.0.0" + "rollup": "^1.20.0||^2.0.0||^3.0.0" + }, + "peerDependenciesMeta": { + "rollup": { + "optional": true + } + } + }, + "node_modules/@rollup/plugin-terser": { + "version": "0.3.0", + "resolved": "https://registry.npmjs.org/@rollup/plugin-terser/-/plugin-terser-0.3.0.tgz", + "integrity": "sha512-mYTkNW9KjOscS/3QWU5LfOKsR3/fAAVDaqcAe2TZ7ng6pN46f+C7FOZbITuIW/neA+PhcjoKl7yMyB3XcmA4gw==", + "dev": true, + "dependencies": { + "serialize-javascript": "^6.0.0", + "smob": "^0.0.6", + "terser": "^5.15.1" + }, + "engines": { + "node": ">=14.0.0" + }, + "peerDependencies": { + "rollup": "^2.x || ^3.x" + }, + "peerDependenciesMeta": { + "rollup": { + "optional": true + } } }, "node_modules/@rollup/plugin-typescript": { - "version": "8.5.0", - "resolved": "https://registry.npmjs.org/@rollup/plugin-typescript/-/plugin-typescript-8.5.0.tgz", - "integrity": "sha512-wMv1/scv0m/rXx21wD2IsBbJFba8wGF3ErJIr6IKRfRj49S85Lszbxb4DCo8iILpluTjk2GAAu9CoZt4G3ppgQ==", + "version": "11.0.0", + "resolved": "https://registry.npmjs.org/@rollup/plugin-typescript/-/plugin-typescript-11.0.0.tgz", + "integrity": "sha512-goPyCWBiimk1iJgSTgsehFD5OOFHiAknrRJjqFCudcW8JtWiBlK284Xnn4flqMqg6YAjVG/EE+3aVzrL5qNSzQ==", "dev": true, "dependencies": { - "@rollup/pluginutils": "^3.1.0", - "resolve": "^1.17.0" + "@rollup/pluginutils": "^5.0.1", + "resolve": "^1.22.1" }, "engines": { - "node": ">=8.0.0" + "node": ">=14.0.0" }, "peerDependencies": { - "rollup": "^2.14.0", + "rollup": "^2.14.0||^3.0.0", "tslib": "*", "typescript": ">=3.7.0" }, "peerDependenciesMeta": { + "rollup": { + "optional": true + }, "tslib": { "optional": true } } }, "node_modules/@rollup/pluginutils": { - "version": "3.1.0", - "resolved": "https://registry.npmjs.org/@rollup/pluginutils/-/pluginutils-3.1.0.tgz", - "integrity": "sha512-GksZ6pr6TpIjHm8h9lSQ8pi8BE9VeubNT0OMJ3B5uZJ8pz73NPiqOtCog/x2/QzM1ENChPKxMDhiQuRHsqc+lg==", + "version": "5.0.2", + "resolved": "https://registry.npmjs.org/@rollup/pluginutils/-/pluginutils-5.0.2.tgz", + "integrity": "sha512-pTd9rIsP92h+B6wWwFbW8RkZv4hiR/xKsqre4SIuAOaOEQRxi0lqLke9k2/7WegC85GgUs9pjmOjCUi3In4vwA==", "dev": true, "dependencies": { - "@types/estree": "0.0.39", - "estree-walker": "^1.0.1", - "picomatch": "^2.2.2" + "@types/estree": "^1.0.0", + "estree-walker": "^2.0.2", + "picomatch": "^2.3.1" }, "engines": { - "node": ">= 8.0.0" + "node": ">=14.0.0" }, "peerDependencies": { - "rollup": "^1.20.0||^2.0.0" + "rollup": "^1.20.0||^2.0.0||^3.0.0" + }, + "peerDependenciesMeta": { + "rollup": { + "optional": true + } } }, - "node_modules/@rollup/pluginutils/node_modules/estree-walker": { - "version": "1.0.1", - "resolved": "https://registry.npmjs.org/estree-walker/-/estree-walker-1.0.1.tgz", - "integrity": "sha512-1fMXF3YP4pZZVozF8j/ZLfvnR8NSIljt56UhbZ5PeeDmmGHpgpdwQt7ITlGvYaQukCvuBRMLEiKiYC+oeIg4cg==", - "dev": true + "node_modules/@svgdotjs/svg.js": { + "version": "3.1.2", + "resolved": "https://registry.npmjs.org/@svgdotjs/svg.js/-/svg.js-3.1.2.tgz", + "integrity": "sha512-0ZCWTiuRjCXT2SUoVDiu+8DLuRQlkxZbO680Y2QkV7jNsULzJajrx6A3MxA/IBQg6UGV5csgPZ8w7U57hZSK+A==", + "funding": { + "type": "github", + "url": "https://github.com/sponsors/Fuzzyma" + } }, "node_modules/@types/d3-color": { "version": "2.0.3", @@ -197,15 +344,9 @@ "integrity": "sha512-9MVYlmIgmRR31C5b4FVSWtuMmBHh2mOWQYfl7XAYOa8dsnb7iEmUmRSWSFgXFtkjxO65d7hTUHQC+RhR/9IWFg==" }, "node_modules/@types/estree": { - "version": "0.0.39", - "resolved": "https://registry.npmjs.org/@types/estree/-/estree-0.0.39.tgz", - "integrity": "sha512-EYNwp3bU+98cpU4lAWYYL7Zz+2gryWH1qbdDTidVd6hkiR6weksdbMadyXKXNPEkQFhXM+hVO9ZygomHXp+AIw==", - "dev": true - }, - "node_modules/@types/node": { - "version": "18.11.9", - "resolved": "https://registry.npmjs.org/@types/node/-/node-18.11.9.tgz", - "integrity": "sha512-CRpX21/kGdzjOpFsZSkcrXMGIBWMGNIHXXBVFSH+ggkftxg+XYP20TESbh+zFvFj3EQOl5byk0HTRn1IL6hbqg==", + "version": "1.0.0", + "resolved": "https://registry.npmjs.org/@types/estree/-/estree-1.0.0.tgz", + "integrity": "sha512-WulqXMDUTYAXCjZnk6JtIHPigp55cVtDgDrO2gHRwhyJto21+1zbVCtOYB2L1F9w4qCQ0rOGWBnBe0FNTiEJIQ==", "dev": true }, "node_modules/@types/prop-types": { @@ -215,9 +356,9 @@ "dev": true }, "node_modules/@types/react": { - "version": "17.0.52", - "resolved": "https://registry.npmjs.org/@types/react/-/react-17.0.52.tgz", - "integrity": "sha512-vwk8QqVODi0VaZZpDXQCmEmiOuyjEFPY7Ttaw5vjM112LOq37yz1CDJGrRJwA1fYEq4Iitd5rnjd1yWAc/bT+A==", + "version": "18.0.26", + "resolved": "https://registry.npmjs.org/@types/react/-/react-18.0.26.tgz", + "integrity": "sha512-hCR3PJQsAIXyxhTNSiDFY//LhnMZWpNNr5etoCqx/iUfGc5gXWtQR2Phl908jVR6uPXacojQWTg4qRpkxTuGug==", "dev": true, "dependencies": { "@types/prop-types": "*", @@ -226,22 +367,19 @@ } }, "node_modules/@types/react-dom": { - "version": "17.0.18", - "resolved": "https://registry.npmjs.org/@types/react-dom/-/react-dom-17.0.18.tgz", - "integrity": "sha512-rLVtIfbwyur2iFKykP2w0pl/1unw26b5td16d5xMgp7/yjTHomkyxPYChFoCr/FtEX1lN9wY6lFj1qvKdS5kDw==", + "version": "18.0.10", + "resolved": "https://registry.npmjs.org/@types/react-dom/-/react-dom-18.0.10.tgz", + "integrity": "sha512-E42GW/JA4Qv15wQdqJq8DL4JhNpB3prJgjgapN3qJT9K2zO5IIAQh4VXvCEDupoqAwnz0cY4RlXeC/ajX5SFHg==", "dev": true, "dependencies": { - "@types/react": "^17" + "@types/react": "*" } }, "node_modules/@types/resolve": { - "version": "1.17.1", - "resolved": "https://registry.npmjs.org/@types/resolve/-/resolve-1.17.1.tgz", - "integrity": "sha512-yy7HuzQhj0dhGpD8RLXSZWEkLsV9ibvxvi6EiJ3bkqLAO1RGo0WbkWQiwpRlSFymTJRz0d3k5LM3kkx8ArDbLw==", - "dev": true, - "dependencies": { - "@types/node": "*" - } + "version": "1.20.2", + "resolved": "https://registry.npmjs.org/@types/resolve/-/resolve-1.20.2.tgz", + "integrity": "sha512-60BCwRFOZCQhDncwQdxxeOEEkbc5dIMccYLwbxsS4TUNeVECQ/pBJ0j09mrHOl/JJvpRPGwO9SvE4nR2Nb/a4Q==", + "dev": true }, "node_modules/@types/scheduler": { "version": "0.16.2", @@ -250,9 +388,21 @@ "dev": true }, "node_modules/@vscode/codicons": { - "version": "0.0.28", - "resolved": "https://registry.npmjs.org/@vscode/codicons/-/codicons-0.0.28.tgz", - "integrity": "sha512-p8zph8Tflh6KUirDCVOti8tr/BhngQx9Z6QXSKBJgPTZMxiKmP6eF75AKDopUeffp7X80Vdo48nv4kdW9d73Dg==" + "version": "0.0.32", + "resolved": "https://registry.npmjs.org/@vscode/codicons/-/codicons-0.0.32.tgz", + "integrity": "sha512-3lgSTWhAzzWN/EPURoY4ZDBEA80OPmnaknNujA3qnI4Iu7AONWd9xF3iE4L+4prIe8E3TUnLQ4pxoaFTEEZNwg==" + }, + "node_modules/acorn": { + "version": "8.8.1", + "resolved": "https://registry.npmjs.org/acorn/-/acorn-8.8.1.tgz", + "integrity": "sha512-7zFpHzhnqYKrkYdUjF1HI1bzd0VygEGX8lFk4k5zVMqHEoES+P+7TKI+EvLO9WVMJ8eekdO0aDEK044xTXwPPA==", + "dev": true, + "bin": { + "acorn": "bin/acorn" + }, + "engines": { + "node": ">=0.4.0" + } }, "node_modules/balanced-match": { "version": "1.0.2", @@ -261,15 +411,20 @@ "dev": true }, "node_modules/brace-expansion": { - "version": "1.1.11", - "resolved": "https://registry.npmjs.org/brace-expansion/-/brace-expansion-1.1.11.tgz", - "integrity": "sha512-iCuPHDFgrHX7H2vEI/5xpz07zSHB00TpugqhmYtVmMO6518mCuRMoOYFldEBl0g187ufozdaHgWKcYFb61qGiA==", + "version": "2.0.1", + "resolved": "https://registry.npmjs.org/brace-expansion/-/brace-expansion-2.0.1.tgz", + "integrity": "sha512-XnAIvQ8eM+kC6aULx6wuQiwVsnzsi9d3WxzV3FpWTGA19F621kwdbsAcFKXgKUHZWsy+mY6iL1sHTxWEFCytDA==", "dev": true, "dependencies": { - "balanced-match": "^1.0.0", - "concat-map": "0.0.1" + "balanced-match": "^1.0.0" } }, + "node_modules/buffer-from": { + "version": "1.1.2", + "resolved": "https://registry.npmjs.org/buffer-from/-/buffer-from-1.1.2.tgz", + "integrity": "sha512-E+XQCRwSbaaiChtv6k6Dwgc+bx+Bs6vuKJHHl5kox/BaKbhiXzqQOwK4cO22yElGp2OCmjwVhT3HmxgyPGnJfQ==", + "dev": true + }, "node_modules/builtin-modules": { "version": "3.3.0", "resolved": "https://registry.npmjs.org/builtin-modules/-/builtin-modules-3.3.0.tgz", @@ -287,17 +442,21 @@ "resolved": "https://registry.npmjs.org/classnames/-/classnames-2.3.2.tgz", "integrity": "sha512-CSbhY4cFEJRe6/GQzIk5qXZ4Jeg5pcsP7b5peFSDpffpe1cqjASH/n9UTjBwOp6XpMSTwQ8Za2K5V02ueA7Tmw==" }, + "node_modules/commander": { + "version": "2.20.3", + "resolved": "https://registry.npmjs.org/commander/-/commander-2.20.3.tgz", + "integrity": "sha512-GpVkmM8vF2vQUkj2LvZmD35JxeJOLCwJ9cUkugyk2nuhbv3+mJvpLYYt+0+USMxE+oj+ey/lJEnhZw75x/OMcQ==" + }, "node_modules/commondir": { "version": "1.0.1", "resolved": "https://registry.npmjs.org/commondir/-/commondir-1.0.1.tgz", "integrity": "sha512-W9pAhw0ja1Edb5GVdIF1mjZw/ASI0AlShXM83UUGe2DVr5TdAPEA1OA8m/g8zWp9x6On7gqufY+FatDbC3MDQg==", "dev": true }, - "node_modules/concat-map": { - "version": "0.0.1", - "resolved": "https://registry.npmjs.org/concat-map/-/concat-map-0.0.1.tgz", - "integrity": "sha512-/Srv4dswyQNBfohGpz9o6Yb3Gz3SrUDqBH5rTuhGR7ahtlbYKnVxw2bCFMRljaA7EXHaXZ8wsHdodFvbkhKmqg==", - "dev": true + "node_modules/consola": { + "version": "2.15.3", + "resolved": "https://registry.npmjs.org/consola/-/consola-2.15.3.tgz", + "integrity": "sha512-9vAdYbHj6x2fLKC4+oPH0kFzY/orMZyG2Aj+kNylHxKGJ/Ed4dpNyAQYwJOdqO4zdM7XpVHmyejQDcQHrnuXbw==" }, "node_modules/css-unit-converter": { "version": "1.1.2", @@ -391,6 +550,11 @@ "node": ">=0.10.0" } }, + "node_modules/discontinuous-range": { + "version": "1.0.0", + "resolved": "https://registry.npmjs.org/discontinuous-range/-/discontinuous-range-1.0.0.tgz", + "integrity": "sha512-c68LpLbO+7kP/b1Hr1qs8/BJ09F5khZGTxqxZuhzxpmwJKOgRFHJWIb9/KmqnqHhLdO55aOxFH/EGBvUQbL/RQ==" + }, "node_modules/dom-helpers": { "version": "3.4.0", "resolved": "https://registry.npmjs.org/dom-helpers/-/dom-helpers-3.4.0.tgz", @@ -399,6 +563,19 @@ "@babel/runtime": "^7.1.2" } }, + "node_modules/es-module-shims": { + "version": "1.6.3", + "resolved": "https://registry.npmjs.org/es-module-shims/-/es-module-shims-1.6.3.tgz", + "integrity": "sha512-+BQyPRZczeV9JH/17X1nu1GbD+SZvdPKD4Nrt2S61J94A2yc8DpWBlzv9KgF9cOXUZKifEShy8/qLelSVNo/rA==" + }, + "node_modules/esm": { + "version": "3.2.25", + "resolved": "https://registry.npmjs.org/esm/-/esm-3.2.25.tgz", + "integrity": "sha512-U1suiZ2oDVWv4zPO56S0NcR5QriEahGtdN2OR6FiOG4WJvcjBVFB0qI4+eKoWFH483PKGuLuu6V8Z4T5g63UVA==", + "engines": { + "node": ">=6" + } + }, "node_modules/estree-walker": { "version": "2.0.2", "resolved": "https://registry.npmjs.org/estree-walker/-/estree-walker-2.0.2.tgz", @@ -421,20 +598,6 @@ "integrity": "sha512-OO0pH2lK6a0hZnAdau5ItzHPI6pUlvI7jMVnxUQRtw4owF2wk8lOSabtGDCTP4Ggrg2MbGnWO9X8K1t4+fGMDw==", "dev": true }, - "node_modules/fsevents": { - "version": "2.3.2", - "resolved": "https://registry.npmjs.org/fsevents/-/fsevents-2.3.2.tgz", - "integrity": "sha512-xiqMQR4xAeHTuB9uWm+fFRcIOgKBMiOBP+eXiyT7jsgVCq1bkVygt00oASowB7EdtpOHaaPgKt812P9ab+DDKA==", - "dev": true, - "hasInstallScript": true, - "optional": true, - "os": [ - "darwin" - ], - "engines": { - "node": "^8.16.0 || ^10.6.0 || >=11.0.0" - } - }, "node_modules/function-bind": { "version": "1.1.1", "resolved": "https://registry.npmjs.org/function-bind/-/function-bind-1.1.1.tgz", @@ -442,20 +605,19 @@ "dev": true }, "node_modules/glob": { - "version": "7.2.3", - "resolved": "https://registry.npmjs.org/glob/-/glob-7.2.3.tgz", - "integrity": "sha512-nFR0zLpU2YCaRxwoCJvL6UvCH2JFyFVIvwTLsIf21AuHlMskA1hhTdk+LlYJtOlYt9v6dvszD2BGRqBL+iQK9Q==", + "version": "8.0.3", + "resolved": "https://registry.npmjs.org/glob/-/glob-8.0.3.tgz", + "integrity": "sha512-ull455NHSHI/Y1FqGaaYFaLGkNMMJbavMrEGFXG/PGrg6y7sutWHUHrz6gy6WEBH6akM1M414dWKCNs+IhKdiQ==", "dev": true, "dependencies": { "fs.realpath": "^1.0.0", "inflight": "^1.0.4", "inherits": "2", - "minimatch": "^3.1.1", - "once": "^1.3.0", - "path-is-absolute": "^1.0.0" + "minimatch": "^5.0.1", + "once": "^1.3.0" }, "engines": { - "node": "*" + "node": ">=12" }, "funding": { "url": "https://github.com/sponsors/isaacs" @@ -473,6 +635,11 @@ "node": ">= 0.4.0" } }, + "node_modules/immutable": { + "version": "4.2.2", + "resolved": "https://registry.npmjs.org/immutable/-/immutable-4.2.2.tgz", + "integrity": "sha512-fTMKDwtbvO5tldky9QZ2fMX7slR0mYpY5nbnFWYp0fOzDhHqhgIw9KoYgxLWsoNTS9ZHGauHj18DTyEw6BK3Og==" + }, "node_modules/inflight": { "version": "1.0.6", "resolved": "https://registry.npmjs.org/inflight/-/inflight-1.0.6.tgz", @@ -558,18 +725,21 @@ } }, "node_modules/magic-string": { - "version": "0.25.9", - "resolved": "https://registry.npmjs.org/magic-string/-/magic-string-0.25.9.tgz", - "integrity": "sha512-RmF0AsMzgt25qzqqLc1+MbHmhdx0ojF2Fvs4XnOqz2ZOBXzzkEwc/dJQZCYHAn7v1jbVOjAZfK8msRn4BxO4VQ==", + "version": "0.27.0", + "resolved": "https://registry.npmjs.org/magic-string/-/magic-string-0.27.0.tgz", + "integrity": "sha512-8UnnX2PeRAPZuN12svgR9j7M1uWMovg/CEnIwIG0LFkXSJJe4PdfUGiTGl8V9bsBHFUtfVINcSyYxd7q+kx9fA==", "dev": true, "dependencies": { - "sourcemap-codec": "^1.4.8" + "@jridgewell/sourcemap-codec": "^1.4.13" + }, + "engines": { + "node": ">=12" } }, "node_modules/marked": { - "version": "4.2.2", - "resolved": "https://registry.npmjs.org/marked/-/marked-4.2.2.tgz", - "integrity": "sha512-JjBTFTAvuTgANXx82a5vzK9JLSMoV6V3LBVn4Uhdso6t7vXrGx7g1Cd2r6NYSsxrYbQGFCMqBDhFHyK5q2UvcQ==", + "version": "4.2.12", + "resolved": "https://registry.npmjs.org/marked/-/marked-4.2.12.tgz", + "integrity": "sha512-yr8hSKa3Fv4D3jdZmtMMPghgVt6TWbk86WQaWhDloQjRSQhMMYCAro7jP7VDJrjjdV8pxVxMssXS8B8Y5DZ5aw==", "bin": { "marked": "bin/marked.js" }, @@ -577,16 +747,63 @@ "node": ">= 12" } }, + "node_modules/mathjax-full": { + "version": "3.2.2", + "resolved": "https://registry.npmjs.org/mathjax-full/-/mathjax-full-3.2.2.tgz", + "integrity": "sha512-+LfG9Fik+OuI8SLwsiR02IVdjcnRCy5MufYLi0C3TdMT56L/pjB0alMVGgoWJF8pN9Rc7FESycZB9BMNWIid5w==", + "dependencies": { + "esm": "^3.2.25", + "mhchemparser": "^4.1.0", + "mj-context-menu": "^0.6.1", + "speech-rule-engine": "^4.0.6" + } + }, + "node_modules/mhchemparser": { + "version": "4.1.1", + "resolved": "https://registry.npmjs.org/mhchemparser/-/mhchemparser-4.1.1.tgz", + "integrity": "sha512-R75CUN6O6e1t8bgailrF1qPq+HhVeFTM3XQ0uzI+mXTybmphy3b6h4NbLOYhemViQ3lUs+6CKRkC3Ws1TlYREA==" + }, "node_modules/minimatch": { - "version": "3.1.2", - "resolved": "https://registry.npmjs.org/minimatch/-/minimatch-3.1.2.tgz", - "integrity": "sha512-J7p63hRiAjw1NDEww1W7i37+ByIrOWO5XQQAzZ3VOcL0PNybwpfmV/N05zFAzwQ9USyEcX6t3UO+K5aqBQOIHw==", + "version": "5.1.2", + "resolved": "https://registry.npmjs.org/minimatch/-/minimatch-5.1.2.tgz", + "integrity": "sha512-bNH9mmM9qsJ2X4r2Nat1B//1dJVcn3+iBLa3IgqJ7EbGaDNepL9QSHOxN4ng33s52VMMhhIfgCYDk3C4ZmlDAg==", "dev": true, "dependencies": { - "brace-expansion": "^1.1.7" + "brace-expansion": "^2.0.1" }, "engines": { - "node": "*" + "node": ">=10" + } + }, + "node_modules/mj-context-menu": { + "version": "0.6.1", + "resolved": "https://registry.npmjs.org/mj-context-menu/-/mj-context-menu-0.6.1.tgz", + "integrity": "sha512-7NO5s6n10TIV96d4g2uDpG7ZDpIhMh0QNfGdJw/W47JswFcosz457wqz/b5sAKvl12sxINGFCn80NZHKwxQEXA==" + }, + "node_modules/moo": { + "version": "0.5.2", + "resolved": "https://registry.npmjs.org/moo/-/moo-0.5.2.tgz", + "integrity": "sha512-iSAJLHYKnX41mKcJKjqvnAN9sf0LMDTXDEvFv+ffuRR9a1MIuXLjMNL6EsnDHSkKLTWNqQQ5uo61P4EbU4NU+Q==" + }, + "node_modules/nearley": { + "version": "2.20.1", + "resolved": "https://registry.npmjs.org/nearley/-/nearley-2.20.1.tgz", + "integrity": "sha512-+Mc8UaAebFzgV+KpI5n7DasuuQCHA89dmwm7JXw3TV43ukfNQ9DnBH3Mdb2g/I4Fdxc26pwimBWvjIw0UAILSQ==", + "dependencies": { + "commander": "^2.19.0", + "moo": "^0.5.0", + "railroad-diagrams": "^1.0.0", + "randexp": "0.4.6" + }, + "bin": { + "nearley-railroad": "bin/nearley-railroad.js", + "nearley-test": "bin/nearley-test.js", + "nearley-unparse": "bin/nearley-unparse.js", + "nearleyc": "bin/nearleyc.js" + }, + "funding": { + "type": "individual", + "url": "https://nearley.js.org/#give-to-nearley" } }, "node_modules/object-assign": { @@ -606,15 +823,6 @@ "wrappy": "1" } }, - "node_modules/path-is-absolute": { - "version": "1.0.1", - "resolved": "https://registry.npmjs.org/path-is-absolute/-/path-is-absolute-1.0.1.tgz", - "integrity": "sha512-AVbw3UJ2e9bq64vSaS9Am0fje1Pa8pbGqTTsmXfaIiMpnr5DlDhfJOuLj9Sf95ZPVDAUerDfEk88MPmPe7UCQg==", - "dev": true, - "engines": { - "node": ">=0.10.0" - } - }, "node_modules/path-parse": { "version": "1.0.7", "resolved": "https://registry.npmjs.org/path-parse/-/path-parse-1.0.7.tgz", @@ -633,6 +841,11 @@ "url": "https://github.com/sponsors/jonschlinkert" } }, + "node_modules/poly-partition": { + "version": "1.0.2", + "resolved": "https://registry.npmjs.org/poly-partition/-/poly-partition-1.0.2.tgz", + "integrity": "sha512-uUqyuUq1n9wk3ZruUm7LdZiqcJ+RwBf3iPxdPe+kn6U3kn1H8jtH9EKg6I1NDG8pomnnxvuXoTpcusCWyemc2A==" + }, "node_modules/postcss-value-parser": { "version": "3.3.1", "resolved": "https://registry.npmjs.org/postcss-value-parser/-/postcss-value-parser-3.3.1.tgz", @@ -648,6 +861,32 @@ "react-is": "^16.13.1" } }, + "node_modules/railroad-diagrams": { + "version": "1.0.0", + "resolved": "https://registry.npmjs.org/railroad-diagrams/-/railroad-diagrams-1.0.0.tgz", + "integrity": "sha512-cz93DjNeLY0idrCNOH6PviZGRN9GJhsdm9hpn1YCS879fj4W+x5IFJhhkRZcwVgMmFF7R82UA/7Oh+R8lLZg6A==" + }, + "node_modules/randexp": { + "version": "0.4.6", + "resolved": "https://registry.npmjs.org/randexp/-/randexp-0.4.6.tgz", + "integrity": "sha512-80WNmd9DA0tmZrw9qQa62GPPWfuXJknrmVmLcxvq4uZBdYqb1wYoKTmnlGUchvVWe0XiLupYkBoXVOxz3C8DYQ==", + "dependencies": { + "discontinuous-range": "1.0.0", + "ret": "~0.1.10" + }, + "engines": { + "node": ">=0.12" + } + }, + "node_modules/randombytes": { + "version": "2.1.0", + "resolved": "https://registry.npmjs.org/randombytes/-/randombytes-2.1.0.tgz", + "integrity": "sha512-vYl3iOX+4CKUWuxGi9Ukhie6fsqXqS9FE2Zaic4tNFD2N2QQaXOMFbuKK4QmDHC0JO6B1Zp41J0LpT0oR68amQ==", + "dev": true, + "dependencies": { + "safe-buffer": "^5.1.0" + } + }, "node_modules/react": { "version": "18.2.0", "resolved": "https://registry.npmjs.org/react/-/react-18.2.0.tgz", @@ -728,9 +967,9 @@ } }, "node_modules/recharts": { - "version": "2.1.16", - "resolved": "https://registry.npmjs.org/recharts/-/recharts-2.1.16.tgz", - "integrity": "sha512-aYn1plTjYzRCo3UGxtWsduslwYd+Cuww3h/YAAEoRdGe0LRnBgYgaXSlVrNFkWOOSXrBavpmnli9h7pvRuk5wg==", + "version": "2.2.0", + "resolved": "https://registry.npmjs.org/recharts/-/recharts-2.2.0.tgz", + "integrity": "sha512-/uRJ0oaESGyz//PgAzvrwXEhrKaNha1ELLysEMRklbnsddiVQsSNicP7DWiz8qFcyYXy3BrDqrUjiLiVRTSMtA==", "dependencies": { "@types/d3-interpolate": "^2.0.0", "@types/d3-scale": "^3.0.0", @@ -764,6 +1003,11 @@ "decimal.js-light": "^2.4.1" } }, + "node_modules/recursive-diff": { + "version": "1.0.8", + "resolved": "https://registry.npmjs.org/recursive-diff/-/recursive-diff-1.0.8.tgz", + "integrity": "sha512-WfUcVao/HRsZZjI96PjcefTZ87rsxrgryIlWJB2UrHs4v3kFFzvMOm801a1WoqG+IAsLKJPsGlvubR74EDqGWQ==" + }, "node_modules/reduce-css-calc": { "version": "2.1.8", "resolved": "https://registry.npmjs.org/reduce-css-calc/-/reduce-css-calc-2.1.8.tgz", @@ -795,58 +1039,117 @@ "url": "https://github.com/sponsors/ljharb" } }, + "node_modules/ret": { + "version": "0.1.15", + "resolved": "https://registry.npmjs.org/ret/-/ret-0.1.15.tgz", + "integrity": "sha512-TTlYpa+OL+vMMNG24xSlQGEJ3B/RzEfUlLct7b5G/ytav+wPrplCpVMFuwzXbkecJrb6IYo1iFb0S9v37754mg==", + "engines": { + "node": ">=0.12" + } + }, "node_modules/rollup": { - "version": "2.79.1", - "resolved": "https://registry.npmjs.org/rollup/-/rollup-2.79.1.tgz", - "integrity": "sha512-uKxbd0IhMZOhjAiD5oAFp7BqvkA4Dv47qpOCtaNvng4HBwdbWtdOh8f5nZNuk2rp51PMGk3bzfWu5oayNEuYnw==", + "version": "3.9.1", + "resolved": "https://registry.npmjs.org/rollup/-/rollup-3.9.1.tgz", + "integrity": "sha512-GswCYHXftN8ZKGVgQhTFUJB/NBXxrRGgO2NCy6E8s1rwEJ4Q9/VttNqcYfEvx4dTo4j58YqdC3OVztPzlKSX8w==", "dev": true, "bin": { "rollup": "dist/bin/rollup" }, "engines": { - "node": ">=10.0.0" + "node": ">=14.18.0", + "npm": ">=8.0.0" }, "optionalDependencies": { "fsevents": "~2.3.2" } }, - "node_modules/rollup-plugin-string": { - "version": "3.0.0", - "resolved": "https://registry.npmjs.org/rollup-plugin-string/-/rollup-plugin-string-3.0.0.tgz", - "integrity": "sha512-vqyzgn9QefAgeKi+Y4A7jETeIAU1zQmS6VotH6bzm/zmUQEnYkpIGRaOBPY41oiWYV4JyBoGAaBjYMYuv+6wVw==", + "node_modules/safe-buffer": { + "version": "5.2.1", + "resolved": "https://registry.npmjs.org/safe-buffer/-/safe-buffer-5.2.1.tgz", + "integrity": "sha512-rp3So07KcdmmKbGvgaNxQSJr7bGVSVk5S9Eq1F+ppbRo70+YeaDxkw5Dd8NPN+GD6bjnYm2VuPuCXmpuYvmCXQ==", "dev": true, + "funding": [ + { + "type": "github", + "url": "https://github.com/sponsors/feross" + }, + { + "type": "patreon", + "url": "https://www.patreon.com/feross" + }, + { + "type": "consulting", + "url": "https://feross.org/support" + } + ] + }, + "node_modules/scheduler": { + "version": "0.23.0", + "resolved": "https://registry.npmjs.org/scheduler/-/scheduler-0.23.0.tgz", + "integrity": "sha512-CtuThmgHNg7zIZWAXi3AsyIzA3n4xx7aNyjwC2VJldO2LMVDhFK+63xGqq6CsJH4rTAt6/M+N4GhZiDYPx9eUw==", "dependencies": { - "rollup-pluginutils": "^2.4.1" + "loose-envify": "^1.1.0" } }, - "node_modules/rollup-pluginutils": { - "version": "2.8.2", - "resolved": "https://registry.npmjs.org/rollup-pluginutils/-/rollup-pluginutils-2.8.2.tgz", - "integrity": "sha512-EEp9NhnUkwY8aif6bxgovPHMoMoNr2FulJziTndpt5H9RdwC47GSGuII9XxpSdzVGM0GWrNPHV6ie1LTNJPaLQ==", + "node_modules/seedrandom": { + "version": "3.0.5", + "resolved": "https://registry.npmjs.org/seedrandom/-/seedrandom-3.0.5.tgz", + "integrity": "sha512-8OwmbklUNzwezjGInmZ+2clQmExQPvomqjL7LFqOYqtmuxRgQYqOD3mHaU+MvZn5FLUeVxVfQjwLZW/n/JFuqg==" + }, + "node_modules/serialize-javascript": { + "version": "6.0.0", + "resolved": "https://registry.npmjs.org/serialize-javascript/-/serialize-javascript-6.0.0.tgz", + "integrity": "sha512-Qr3TosvguFt8ePWqsvRfrKyQXIiW+nGbYpy8XK24NQHE83caxWt+mIymTT19DGFbNWNLfEwsrkSmN64lVWB9ag==", "dev": true, "dependencies": { - "estree-walker": "^0.6.1" + "randombytes": "^2.1.0" } }, - "node_modules/rollup-pluginutils/node_modules/estree-walker": { - "version": "0.6.1", - "resolved": "https://registry.npmjs.org/estree-walker/-/estree-walker-0.6.1.tgz", - "integrity": "sha512-SqmZANLWS0mnatqbSfRP5g8OXZC12Fgg1IwNtLsyHDzJizORW4khDfjPqJZsemPWBB2uqykUah5YpQ6epsqC/w==", + "node_modules/smob": { + "version": "0.0.6", + "resolved": "https://registry.npmjs.org/smob/-/smob-0.0.6.tgz", + "integrity": "sha512-V21+XeNni+tTyiST1MHsa84AQhT1aFZipzPpOFAVB8DkHzwJyjjAmt9bgwnuZiZWnIbMo2duE29wybxv/7HWUw==", "dev": true }, - "node_modules/scheduler": { - "version": "0.23.0", - "resolved": "https://registry.npmjs.org/scheduler/-/scheduler-0.23.0.tgz", - "integrity": "sha512-CtuThmgHNg7zIZWAXi3AsyIzA3n4xx7aNyjwC2VJldO2LMVDhFK+63xGqq6CsJH4rTAt6/M+N4GhZiDYPx9eUw==", + "node_modules/source-map": { + "version": "0.6.1", + "resolved": "https://registry.npmjs.org/source-map/-/source-map-0.6.1.tgz", + "integrity": "sha512-UjgapumWlbMhkBgzT7Ykc5YXUT46F0iKu8SGXq0bcwP5dz/h0Plj6enJqjz1Zbq2l5WaqYnrVbwWOWMyF3F47g==", + "dev": true, + "engines": { + "node": ">=0.10.0" + } + }, + "node_modules/source-map-support": { + "version": "0.5.21", + "resolved": "https://registry.npmjs.org/source-map-support/-/source-map-support-0.5.21.tgz", + "integrity": "sha512-uBHU3L3czsIyYXKX88fdrGovxdSCoTGDRZ6SYXtSRxLZUzHg5P/66Ht6uoUlHu9EZod+inXhKo3qQgwXUT/y1w==", + "dev": true, "dependencies": { - "loose-envify": "^1.1.0" + "buffer-from": "^1.0.0", + "source-map": "^0.6.0" } }, - "node_modules/sourcemap-codec": { - "version": "1.4.8", - "resolved": "https://registry.npmjs.org/sourcemap-codec/-/sourcemap-codec-1.4.8.tgz", - "integrity": "sha512-9NykojV5Uih4lgo5So5dtw+f0JgJX30KCNI8gwhz2J9A15wD0Ml6tjHKwf6fTSa6fAdVBdZeNOs9eJ71qCk8vA==", - "dev": true + "node_modules/speech-rule-engine": { + "version": "4.0.7", + "resolved": "https://registry.npmjs.org/speech-rule-engine/-/speech-rule-engine-4.0.7.tgz", + "integrity": "sha512-sJrL3/wHzNwJRLBdf6CjJWIlxC04iYKkyXvYSVsWVOiC2DSkHmxsqOhEeMsBA9XK+CHuNcsdkbFDnoUfAsmp9g==", + "dependencies": { + "commander": "9.2.0", + "wicked-good-xpath": "1.3.0", + "xmldom-sre": "0.1.31" + }, + "bin": { + "sre": "bin/sre" + } + }, + "node_modules/speech-rule-engine/node_modules/commander": { + "version": "9.2.0", + "resolved": "https://registry.npmjs.org/commander/-/commander-9.2.0.tgz", + "integrity": "sha512-e2i4wANQiSXgnrBlIatyHtP1odfUp0BbV5Y5nEGbxtIrStkEOAAzCUirvLBNXHLr7kwLvJl6V+4V3XV9x7Wd9w==", + "engines": { + "node": "^12.20.0 || >=14" + } }, "node_modules/supports-preserve-symlinks-flag": { "version": "1.0.0", @@ -865,6 +1168,32 @@ "resolved": "https://registry.npmjs.org/tachyons/-/tachyons-4.12.0.tgz", "integrity": "sha512-2nA2IrYFy3raCM9fxJ2KODRGHVSZNTW3BR0YnlGsLUf1DA3pk3YfWZ/DdfbnZK6zLZS+jUenlUGJsKcA5fUiZg==" }, + "node_modules/terser": { + "version": "5.16.1", + "resolved": "https://registry.npmjs.org/terser/-/terser-5.16.1.tgz", + "integrity": "sha512-xvQfyfA1ayT0qdK47zskQgRZeWLoOQ8JQ6mIgRGVNwZKdQMU+5FkCBjmv4QjcrTzyZquRw2FVtlJSRUmMKQslw==", + "dev": true, + "dependencies": { + "@jridgewell/source-map": "^0.3.2", + "acorn": "^8.5.0", + "commander": "^2.20.0", + "source-map-support": "~0.5.20" + }, + "bin": { + "terser": "bin/terser" + }, + "engines": { + "node": ">=10" + } + }, + "node_modules/true-myth": { + "version": "4.1.1", + "resolved": "https://registry.npmjs.org/true-myth/-/true-myth-4.1.1.tgz", + "integrity": "sha512-rqy30BSpxPznbbTcAcci90oZ1YR4DqvKcNXNerG5gQBU2v4jk0cygheiul5J6ExIMrgDVuanv/MkGfqZbKrNNg==", + "engines": { + "node": "10.* || >= 12.*" + } + }, "node_modules/tslib": { "version": "2.4.1", "resolved": "https://registry.npmjs.org/tslib/-/tslib-2.4.1.tgz", @@ -872,9 +1201,9 @@ "dev": true }, "node_modules/typescript": { - "version": "4.8.4", - "resolved": "https://registry.npmjs.org/typescript/-/typescript-4.8.4.tgz", - "integrity": "sha512-QCh+85mCy+h0IGff8r5XWzOVSbBO+KfeYrMQh7NJ58QujwcE22u+NUSmUxqF+un70P9GXKxa2HCNiTTMJknyjQ==", + "version": "4.9.4", + "resolved": "https://registry.npmjs.org/typescript/-/typescript-4.9.4.tgz", + "integrity": "sha512-Uz+dTXYzxXXbsFpM86Wh3dKCxrQqUcVMxwU54orwlJjOpO3ao8L7j5lH+dWfTwgCwIuM9GQ2kvVotzYJMXTBZg==", "dev": true, "bin": { "tsc": "bin/tsc", @@ -884,6 +1213,18 @@ "node": ">=4.2.0" } }, + "node_modules/use-resize-observer": { + "version": "9.1.0", + "resolved": "https://registry.npmjs.org/use-resize-observer/-/use-resize-observer-9.1.0.tgz", + "integrity": "sha512-R25VqO9Wb3asSD4eqtcxk8sJalvIOYBqS8MNZlpDSQ4l4xMQxC/J7Id9HoTqPq8FwULIn0PVW+OAqF2dyYbjow==", + "dependencies": { + "@juggle/resize-observer": "^3.3.1" + }, + "peerDependencies": { + "react": "16.8.0 - 18", + "react-dom": "16.8.0 - 18" + } + }, "node_modules/vscode-jsonrpc": { "version": "8.0.2", "resolved": "https://registry.npmjs.org/vscode-jsonrpc/-/vscode-jsonrpc-8.0.2.tgz", @@ -906,11 +1247,24 @@ "resolved": "https://registry.npmjs.org/vscode-languageserver-types/-/vscode-languageserver-types-3.17.2.tgz", "integrity": "sha512-zHhCWatviizPIq9B7Vh9uvrH6x3sK8itC84HkamnBWoDFJtzBf7SWlpLCZUit72b3os45h6RWQNC9xHRDF8dRA==" }, + "node_modules/wicked-good-xpath": { + "version": "1.3.0", + "resolved": "https://registry.npmjs.org/wicked-good-xpath/-/wicked-good-xpath-1.3.0.tgz", + "integrity": "sha512-Gd9+TUn5nXdwj/hFsPVx5cuHHiF5Bwuc30jZ4+ronF1qHK5O7HD0sgmXWSEgwKquT3ClLoKPVbO6qGwVwLzvAw==" + }, "node_modules/wrappy": { "version": "1.0.2", "resolved": "https://registry.npmjs.org/wrappy/-/wrappy-1.0.2.tgz", "integrity": "sha512-l4Sp/DRseor9wL6EvV2+TuQn63dMkPjZ/sp9XkghTEbV9KlPS1xUsZ3u7/IQO4wxtcFB4bgpQPRcR3QCvezPcQ==", "dev": true + }, + "node_modules/xmldom-sre": { + "version": "0.1.31", + "resolved": "https://registry.npmjs.org/xmldom-sre/-/xmldom-sre-0.1.31.tgz", + "integrity": "sha512-f9s+fUkX04BxQf+7mMWAp5zk61pciie+fFLC9hX9UVvCeJQfNHRHXpeo5MPcR0EUf57PYLdt+ZO4f3Ipk2oZUw==", + "engines": { + "node": ">=0.1" + } } }, "dependencies": { @@ -922,92 +1276,183 @@ "regenerator-runtime": "^0.13.10" } }, + "@datastructures-js/heap": { + "version": "3.2.0", + "resolved": "https://registry.npmjs.org/@datastructures-js/heap/-/heap-3.2.0.tgz", + "integrity": "sha512-FcU5ZAyb+VIOZz1HABsJUsbJi2ZyUDO7aoe97hq4d3tK3z8nMgwdxf5bO0gafR0ExFi18YTspntqHLzt4XOgnA==" + }, + "@datastructures-js/queue": { + "version": "4.2.3", + "resolved": "https://registry.npmjs.org/@datastructures-js/queue/-/queue-4.2.3.tgz", + "integrity": "sha512-GWVMorC/xi2V2ta+Z/CPgPGHL2ZJozcj48g7y2nIX5GIGZGRrbShSHgvMViJwHJurUzJYOdIdRZnWDRrROFwJA==" + }, + "@jridgewell/gen-mapping": { + "version": "0.3.2", + "resolved": "https://registry.npmjs.org/@jridgewell/gen-mapping/-/gen-mapping-0.3.2.tgz", + "integrity": "sha512-mh65xKQAzI6iBcFzwv28KVWSmCkdRBWoOh+bYQGW3+6OZvbbN3TqMGo5hqYxQniRcH9F2VZIoJCm4pa3BPDK/A==", + "dev": true, + "requires": { + "@jridgewell/set-array": "^1.0.1", + "@jridgewell/sourcemap-codec": "^1.4.10", + "@jridgewell/trace-mapping": "^0.3.9" + } + }, + "@jridgewell/resolve-uri": { + "version": "3.1.0", + "resolved": "https://registry.npmjs.org/@jridgewell/resolve-uri/-/resolve-uri-3.1.0.tgz", + "integrity": "sha512-F2msla3tad+Mfht5cJq7LSXcdudKTWCVYUgw6pLFOOHSTtZlj6SWNYAp+AhuqLmWdBO2X5hPrLcu8cVP8fy28w==", + "dev": true + }, + "@jridgewell/set-array": { + "version": "1.1.2", + "resolved": "https://registry.npmjs.org/@jridgewell/set-array/-/set-array-1.1.2.tgz", + "integrity": "sha512-xnkseuNADM0gt2bs+BvhO0p78Mk762YnZdsuzFV018NoG1Sj1SCQvpSqa7XUaTam5vAGasABV9qXASMKnFMwMw==", + "dev": true + }, + "@jridgewell/source-map": { + "version": "0.3.2", + "resolved": "https://registry.npmjs.org/@jridgewell/source-map/-/source-map-0.3.2.tgz", + "integrity": "sha512-m7O9o2uR8k2ObDysZYzdfhb08VuEml5oWGiosa1VdaPZ/A6QyPkAJuwN0Q1lhULOf6B7MtQmHENS743hWtCrgw==", + "dev": true, + "requires": { + "@jridgewell/gen-mapping": "^0.3.0", + "@jridgewell/trace-mapping": "^0.3.9" + } + }, + "@jridgewell/sourcemap-codec": { + "version": "1.4.14", + "resolved": "https://registry.npmjs.org/@jridgewell/sourcemap-codec/-/sourcemap-codec-1.4.14.tgz", + "integrity": "sha512-XPSJHWmi394fuUuzDnGz1wiKqWfo1yXecHQMRf2l6hztTO+nPru658AyDngaBe7isIxEkRsPR3FZh+s7iVa4Uw==", + "dev": true + }, + "@jridgewell/trace-mapping": { + "version": "0.3.17", + "resolved": "https://registry.npmjs.org/@jridgewell/trace-mapping/-/trace-mapping-0.3.17.tgz", + "integrity": "sha512-MCNzAp77qzKca9+W/+I0+sEpaUnZoeasnghNeVc41VZCEKaCH73Vq3BZZ/SzWIgrqE4H4ceI+p+b6C0mHf9T4g==", + "dev": true, + "requires": { + "@jridgewell/resolve-uri": "3.1.0", + "@jridgewell/sourcemap-codec": "1.4.14" + } + }, + "@juggle/resize-observer": { + "version": "3.4.0", + "resolved": "https://registry.npmjs.org/@juggle/resize-observer/-/resize-observer-3.4.0.tgz", + "integrity": "sha512-dfLbk+PwWvFzSxwk3n5ySL0hfBog779o8h68wK/7/APo/7cgyWp5jcXockbxdk5kFRkbeXWm4Fbi9FrdN381sA==" + }, "@leanprover/infoview": { - "version": "0.2.0", - "resolved": "https://registry.npmjs.org/@leanprover/infoview/-/infoview-0.2.0.tgz", - "integrity": "sha512-QrN0YV/WieVHf/WeMQpni8Lo6WQRQ8lQ3ImSoAyZd4aToJYzTV4r2dh3PtjwZ65/c/GMWF8O6o8sssGWCEZYwA==", + "version": "0.4.2", + "resolved": "https://registry.npmjs.org/@leanprover/infoview/-/infoview-0.4.2.tgz", + "integrity": "sha512-Bj+q/7n1xlnpRd4xsJ75uikZPwW5id/G4G1JPqhbbfMacTwAgzLNFaqJhzBrIUuF5nxJYGKRqRIeFtkalGxJ0g==", "requires": { - "@leanprover/infoview-api": "^0.1.0", - "@vscode/codicons": "0.0.28", - "marked": "^4.0.17", + "@leanprover/infoview-api": "~0.2.1", + "@vscode/codicons": "^0.0.32", + "es-module-shims": "^1.6.2", + "marked": "^4.2.2", "react-fast-compare": "^3.2.0", "tachyons": "^4.12.0", "vscode-languageserver-protocol": "^3.17.2" } }, "@leanprover/infoview-api": { - "version": "0.1.0", - "resolved": "https://registry.npmjs.org/@leanprover/infoview-api/-/infoview-api-0.1.0.tgz", - "integrity": "sha512-vNMFrk5U1UlWhHp7G1e33mG7hJHyIErb3INgsVITbEW8CzOlAUrc83Pp3ARZykIc8++WJhT7WH+lALo8nO7G0Q==" + "version": "0.2.1", + "resolved": "https://registry.npmjs.org/@leanprover/infoview-api/-/infoview-api-0.2.1.tgz", + "integrity": "sha512-4sYdwOhUsa5wfvo/ZsCbcm8fBWcrATciZq0sWfmi5NRbIyZ+c2QjTm6D9CeYPCNvz9yvD1KBp/2+hKEZ8SOHkA==" + }, + "@penrose/core": { + "version": "1.3.1-alpha.363", + "resolved": "https://registry.npmjs.org/@penrose/core/-/core-1.3.1-alpha.363.tgz", + "integrity": "sha512-rO5OkjT7+y0k8Rdg+DW6fxOC23V1pcH0+9FDSmJeM5lqfAW3phO2mQtiHI0VbcwMURDcH7wEGLVnpS7RVf1d0g==", + "requires": { + "@datastructures-js/heap": "^3.2.0", + "@datastructures-js/queue": "^4.1.3", + "consola": "^2.15.2", + "immutable": "^4.0.0-rc.12", + "lodash": "^4.17.15", + "mathjax-full": "^3.0.1", + "moo": "^0.5.1", + "nearley": "^2.20.1", + "poly-partition": "^1.0.2", + "recursive-diff": "^1.0.8", + "seedrandom": "^3.0.5", + "true-myth": "^4.1.0" + } }, "@rollup/plugin-commonjs": { - "version": "22.0.2", - "resolved": "https://registry.npmjs.org/@rollup/plugin-commonjs/-/plugin-commonjs-22.0.2.tgz", - "integrity": "sha512-//NdP6iIwPbMTcazYsiBMbJW7gfmpHom33u1beiIoHDEM0Q9clvtQB1T0efvMqHeKsGohiHo97BCPCkBXdscwg==", + "version": "24.0.0", + "resolved": "https://registry.npmjs.org/@rollup/plugin-commonjs/-/plugin-commonjs-24.0.0.tgz", + "integrity": "sha512-0w0wyykzdyRRPHOb0cQt14mIBLujfAv6GgP6g8nvg/iBxEm112t3YPPq+Buqe2+imvElTka+bjNlJ/gB56TD8g==", "dev": true, "requires": { - "@rollup/pluginutils": "^3.1.0", + "@rollup/pluginutils": "^5.0.1", "commondir": "^1.0.1", - "estree-walker": "^2.0.1", - "glob": "^7.1.6", - "is-reference": "^1.2.1", - "magic-string": "^0.25.7", - "resolve": "^1.17.0" + "estree-walker": "^2.0.2", + "glob": "^8.0.3", + "is-reference": "1.2.1", + "magic-string": "^0.27.0" } }, "@rollup/plugin-node-resolve": { - "version": "14.1.0", - "resolved": "https://registry.npmjs.org/@rollup/plugin-node-resolve/-/plugin-node-resolve-14.1.0.tgz", - "integrity": "sha512-5G2niJroNCz/1zqwXtk0t9+twOSDlG00k1Wfd7bkbbXmwg8H8dvgHdIWAun53Ps/rckfvOC7scDBjuGFg5OaWw==", + "version": "15.0.1", + "resolved": "https://registry.npmjs.org/@rollup/plugin-node-resolve/-/plugin-node-resolve-15.0.1.tgz", + "integrity": "sha512-ReY88T7JhJjeRVbfCyNj+NXAG3IIsVMsX9b5/9jC98dRP8/yxlZdz7mHZbHk5zHr24wZZICS5AcXsFZAXYUQEg==", "dev": true, "requires": { - "@rollup/pluginutils": "^3.1.0", - "@types/resolve": "1.17.1", + "@rollup/pluginutils": "^5.0.1", + "@types/resolve": "1.20.2", "deepmerge": "^4.2.2", - "is-builtin-module": "^3.1.0", + "is-builtin-module": "^3.2.0", "is-module": "^1.0.0", - "resolve": "^1.19.0" + "resolve": "^1.22.1" } }, "@rollup/plugin-replace": { - "version": "4.0.0", - "resolved": "https://registry.npmjs.org/@rollup/plugin-replace/-/plugin-replace-4.0.0.tgz", - "integrity": "sha512-+rumQFiaNac9y64OHtkHGmdjm7us9bo1PlbgQfdihQtuNxzjpaB064HbRnewUOggLQxVCCyINfStkgmBeQpv1g==", + "version": "5.0.2", + "resolved": "https://registry.npmjs.org/@rollup/plugin-replace/-/plugin-replace-5.0.2.tgz", + "integrity": "sha512-M9YXNekv/C/iHHK+cvORzfRYfPbq0RDD8r0G+bMiTXjNGKulPnCT9O3Ss46WfhI6ZOCgApOP7xAdmCQJ+U2LAA==", "dev": true, "requires": { - "@rollup/pluginutils": "^3.1.0", - "magic-string": "^0.25.7" + "@rollup/pluginutils": "^5.0.1", + "magic-string": "^0.27.0" + } + }, + "@rollup/plugin-terser": { + "version": "0.3.0", + "resolved": "https://registry.npmjs.org/@rollup/plugin-terser/-/plugin-terser-0.3.0.tgz", + "integrity": "sha512-mYTkNW9KjOscS/3QWU5LfOKsR3/fAAVDaqcAe2TZ7ng6pN46f+C7FOZbITuIW/neA+PhcjoKl7yMyB3XcmA4gw==", + "dev": true, + "requires": { + "serialize-javascript": "^6.0.0", + "smob": "^0.0.6", + "terser": "^5.15.1" } }, "@rollup/plugin-typescript": { - "version": "8.5.0", - "resolved": "https://registry.npmjs.org/@rollup/plugin-typescript/-/plugin-typescript-8.5.0.tgz", - "integrity": "sha512-wMv1/scv0m/rXx21wD2IsBbJFba8wGF3ErJIr6IKRfRj49S85Lszbxb4DCo8iILpluTjk2GAAu9CoZt4G3ppgQ==", + "version": "11.0.0", + "resolved": "https://registry.npmjs.org/@rollup/plugin-typescript/-/plugin-typescript-11.0.0.tgz", + "integrity": "sha512-goPyCWBiimk1iJgSTgsehFD5OOFHiAknrRJjqFCudcW8JtWiBlK284Xnn4flqMqg6YAjVG/EE+3aVzrL5qNSzQ==", "dev": true, "requires": { - "@rollup/pluginutils": "^3.1.0", - "resolve": "^1.17.0" + "@rollup/pluginutils": "^5.0.1", + "resolve": "^1.22.1" } }, "@rollup/pluginutils": { - "version": "3.1.0", - "resolved": "https://registry.npmjs.org/@rollup/pluginutils/-/pluginutils-3.1.0.tgz", - "integrity": "sha512-GksZ6pr6TpIjHm8h9lSQ8pi8BE9VeubNT0OMJ3B5uZJ8pz73NPiqOtCog/x2/QzM1ENChPKxMDhiQuRHsqc+lg==", + "version": "5.0.2", + "resolved": "https://registry.npmjs.org/@rollup/pluginutils/-/pluginutils-5.0.2.tgz", + "integrity": "sha512-pTd9rIsP92h+B6wWwFbW8RkZv4hiR/xKsqre4SIuAOaOEQRxi0lqLke9k2/7WegC85GgUs9pjmOjCUi3In4vwA==", "dev": true, "requires": { - "@types/estree": "0.0.39", - "estree-walker": "^1.0.1", - "picomatch": "^2.2.2" - }, - "dependencies": { - "estree-walker": { - "version": "1.0.1", - "resolved": "https://registry.npmjs.org/estree-walker/-/estree-walker-1.0.1.tgz", - "integrity": "sha512-1fMXF3YP4pZZVozF8j/ZLfvnR8NSIljt56UhbZ5PeeDmmGHpgpdwQt7ITlGvYaQukCvuBRMLEiKiYC+oeIg4cg==", - "dev": true - } + "@types/estree": "^1.0.0", + "estree-walker": "^2.0.2", + "picomatch": "^2.3.1" } }, + "@svgdotjs/svg.js": { + "version": "3.1.2", + "resolved": "https://registry.npmjs.org/@svgdotjs/svg.js/-/svg.js-3.1.2.tgz", + "integrity": "sha512-0ZCWTiuRjCXT2SUoVDiu+8DLuRQlkxZbO680Y2QkV7jNsULzJajrx6A3MxA/IBQg6UGV5csgPZ8w7U57hZSK+A==" + }, "@types/d3-color": { "version": "2.0.3", "resolved": "https://registry.npmjs.org/@types/d3-color/-/d3-color-2.0.3.tgz", @@ -1048,15 +1493,9 @@ "integrity": "sha512-9MVYlmIgmRR31C5b4FVSWtuMmBHh2mOWQYfl7XAYOa8dsnb7iEmUmRSWSFgXFtkjxO65d7hTUHQC+RhR/9IWFg==" }, "@types/estree": { - "version": "0.0.39", - "resolved": "https://registry.npmjs.org/@types/estree/-/estree-0.0.39.tgz", - "integrity": "sha512-EYNwp3bU+98cpU4lAWYYL7Zz+2gryWH1qbdDTidVd6hkiR6weksdbMadyXKXNPEkQFhXM+hVO9ZygomHXp+AIw==", - "dev": true - }, - "@types/node": { - "version": "18.11.9", - "resolved": "https://registry.npmjs.org/@types/node/-/node-18.11.9.tgz", - "integrity": "sha512-CRpX21/kGdzjOpFsZSkcrXMGIBWMGNIHXXBVFSH+ggkftxg+XYP20TESbh+zFvFj3EQOl5byk0HTRn1IL6hbqg==", + "version": "1.0.0", + "resolved": "https://registry.npmjs.org/@types/estree/-/estree-1.0.0.tgz", + "integrity": "sha512-WulqXMDUTYAXCjZnk6JtIHPigp55cVtDgDrO2gHRwhyJto21+1zbVCtOYB2L1F9w4qCQ0rOGWBnBe0FNTiEJIQ==", "dev": true }, "@types/prop-types": { @@ -1066,9 +1505,9 @@ "dev": true }, "@types/react": { - "version": "17.0.52", - "resolved": "https://registry.npmjs.org/@types/react/-/react-17.0.52.tgz", - "integrity": "sha512-vwk8QqVODi0VaZZpDXQCmEmiOuyjEFPY7Ttaw5vjM112LOq37yz1CDJGrRJwA1fYEq4Iitd5rnjd1yWAc/bT+A==", + "version": "18.0.26", + "resolved": "https://registry.npmjs.org/@types/react/-/react-18.0.26.tgz", + "integrity": "sha512-hCR3PJQsAIXyxhTNSiDFY//LhnMZWpNNr5etoCqx/iUfGc5gXWtQR2Phl908jVR6uPXacojQWTg4qRpkxTuGug==", "dev": true, "requires": { "@types/prop-types": "*", @@ -1077,22 +1516,19 @@ } }, "@types/react-dom": { - "version": "17.0.18", - "resolved": "https://registry.npmjs.org/@types/react-dom/-/react-dom-17.0.18.tgz", - "integrity": "sha512-rLVtIfbwyur2iFKykP2w0pl/1unw26b5td16d5xMgp7/yjTHomkyxPYChFoCr/FtEX1lN9wY6lFj1qvKdS5kDw==", + "version": "18.0.10", + "resolved": "https://registry.npmjs.org/@types/react-dom/-/react-dom-18.0.10.tgz", + "integrity": "sha512-E42GW/JA4Qv15wQdqJq8DL4JhNpB3prJgjgapN3qJT9K2zO5IIAQh4VXvCEDupoqAwnz0cY4RlXeC/ajX5SFHg==", "dev": true, "requires": { - "@types/react": "^17" + "@types/react": "*" } }, "@types/resolve": { - "version": "1.17.1", - "resolved": "https://registry.npmjs.org/@types/resolve/-/resolve-1.17.1.tgz", - "integrity": "sha512-yy7HuzQhj0dhGpD8RLXSZWEkLsV9ibvxvi6EiJ3bkqLAO1RGo0WbkWQiwpRlSFymTJRz0d3k5LM3kkx8ArDbLw==", - "dev": true, - "requires": { - "@types/node": "*" - } + "version": "1.20.2", + "resolved": "https://registry.npmjs.org/@types/resolve/-/resolve-1.20.2.tgz", + "integrity": "sha512-60BCwRFOZCQhDncwQdxxeOEEkbc5dIMccYLwbxsS4TUNeVECQ/pBJ0j09mrHOl/JJvpRPGwO9SvE4nR2Nb/a4Q==", + "dev": true }, "@types/scheduler": { "version": "0.16.2", @@ -1101,9 +1537,15 @@ "dev": true }, "@vscode/codicons": { - "version": "0.0.28", - "resolved": "https://registry.npmjs.org/@vscode/codicons/-/codicons-0.0.28.tgz", - "integrity": "sha512-p8zph8Tflh6KUirDCVOti8tr/BhngQx9Z6QXSKBJgPTZMxiKmP6eF75AKDopUeffp7X80Vdo48nv4kdW9d73Dg==" + "version": "0.0.32", + "resolved": "https://registry.npmjs.org/@vscode/codicons/-/codicons-0.0.32.tgz", + "integrity": "sha512-3lgSTWhAzzWN/EPURoY4ZDBEA80OPmnaknNujA3qnI4Iu7AONWd9xF3iE4L+4prIe8E3TUnLQ4pxoaFTEEZNwg==" + }, + "acorn": { + "version": "8.8.1", + "resolved": "https://registry.npmjs.org/acorn/-/acorn-8.8.1.tgz", + "integrity": "sha512-7zFpHzhnqYKrkYdUjF1HI1bzd0VygEGX8lFk4k5zVMqHEoES+P+7TKI+EvLO9WVMJ8eekdO0aDEK044xTXwPPA==", + "dev": true }, "balanced-match": { "version": "1.0.2", @@ -1112,15 +1554,20 @@ "dev": true }, "brace-expansion": { - "version": "1.1.11", - "resolved": "https://registry.npmjs.org/brace-expansion/-/brace-expansion-1.1.11.tgz", - "integrity": "sha512-iCuPHDFgrHX7H2vEI/5xpz07zSHB00TpugqhmYtVmMO6518mCuRMoOYFldEBl0g187ufozdaHgWKcYFb61qGiA==", + "version": "2.0.1", + "resolved": "https://registry.npmjs.org/brace-expansion/-/brace-expansion-2.0.1.tgz", + "integrity": "sha512-XnAIvQ8eM+kC6aULx6wuQiwVsnzsi9d3WxzV3FpWTGA19F621kwdbsAcFKXgKUHZWsy+mY6iL1sHTxWEFCytDA==", "dev": true, "requires": { - "balanced-match": "^1.0.0", - "concat-map": "0.0.1" + "balanced-match": "^1.0.0" } }, + "buffer-from": { + "version": "1.1.2", + "resolved": "https://registry.npmjs.org/buffer-from/-/buffer-from-1.1.2.tgz", + "integrity": "sha512-E+XQCRwSbaaiChtv6k6Dwgc+bx+Bs6vuKJHHl5kox/BaKbhiXzqQOwK4cO22yElGp2OCmjwVhT3HmxgyPGnJfQ==", + "dev": true + }, "builtin-modules": { "version": "3.3.0", "resolved": "https://registry.npmjs.org/builtin-modules/-/builtin-modules-3.3.0.tgz", @@ -1132,17 +1579,21 @@ "resolved": "https://registry.npmjs.org/classnames/-/classnames-2.3.2.tgz", "integrity": "sha512-CSbhY4cFEJRe6/GQzIk5qXZ4Jeg5pcsP7b5peFSDpffpe1cqjASH/n9UTjBwOp6XpMSTwQ8Za2K5V02ueA7Tmw==" }, + "commander": { + "version": "2.20.3", + "resolved": "https://registry.npmjs.org/commander/-/commander-2.20.3.tgz", + "integrity": "sha512-GpVkmM8vF2vQUkj2LvZmD35JxeJOLCwJ9cUkugyk2nuhbv3+mJvpLYYt+0+USMxE+oj+ey/lJEnhZw75x/OMcQ==" + }, "commondir": { "version": "1.0.1", "resolved": "https://registry.npmjs.org/commondir/-/commondir-1.0.1.tgz", "integrity": "sha512-W9pAhw0ja1Edb5GVdIF1mjZw/ASI0AlShXM83UUGe2DVr5TdAPEA1OA8m/g8zWp9x6On7gqufY+FatDbC3MDQg==", "dev": true }, - "concat-map": { - "version": "0.0.1", - "resolved": "https://registry.npmjs.org/concat-map/-/concat-map-0.0.1.tgz", - "integrity": "sha512-/Srv4dswyQNBfohGpz9o6Yb3Gz3SrUDqBH5rTuhGR7ahtlbYKnVxw2bCFMRljaA7EXHaXZ8wsHdodFvbkhKmqg==", - "dev": true + "consola": { + "version": "2.15.3", + "resolved": "https://registry.npmjs.org/consola/-/consola-2.15.3.tgz", + "integrity": "sha512-9vAdYbHj6x2fLKC4+oPH0kFzY/orMZyG2Aj+kNylHxKGJ/Ed4dpNyAQYwJOdqO4zdM7XpVHmyejQDcQHrnuXbw==" }, "css-unit-converter": { "version": "1.1.2", @@ -1233,6 +1684,11 @@ "integrity": "sha512-FJ3UgI4gIl+PHZm53knsuSFpE+nESMr7M4v9QcgB7S63Kj/6WqMiFQJpBBYz1Pt+66bZpP3Q7Lye0Oo9MPKEdg==", "dev": true }, + "discontinuous-range": { + "version": "1.0.0", + "resolved": "https://registry.npmjs.org/discontinuous-range/-/discontinuous-range-1.0.0.tgz", + "integrity": "sha512-c68LpLbO+7kP/b1Hr1qs8/BJ09F5khZGTxqxZuhzxpmwJKOgRFHJWIb9/KmqnqHhLdO55aOxFH/EGBvUQbL/RQ==" + }, "dom-helpers": { "version": "3.4.0", "resolved": "https://registry.npmjs.org/dom-helpers/-/dom-helpers-3.4.0.tgz", @@ -1241,6 +1697,16 @@ "@babel/runtime": "^7.1.2" } }, + "es-module-shims": { + "version": "1.6.3", + "resolved": "https://registry.npmjs.org/es-module-shims/-/es-module-shims-1.6.3.tgz", + "integrity": "sha512-+BQyPRZczeV9JH/17X1nu1GbD+SZvdPKD4Nrt2S61J94A2yc8DpWBlzv9KgF9cOXUZKifEShy8/qLelSVNo/rA==" + }, + "esm": { + "version": "3.2.25", + "resolved": "https://registry.npmjs.org/esm/-/esm-3.2.25.tgz", + "integrity": "sha512-U1suiZ2oDVWv4zPO56S0NcR5QriEahGtdN2OR6FiOG4WJvcjBVFB0qI4+eKoWFH483PKGuLuu6V8Z4T5g63UVA==" + }, "estree-walker": { "version": "2.0.2", "resolved": "https://registry.npmjs.org/estree-walker/-/estree-walker-2.0.2.tgz", @@ -1263,13 +1729,6 @@ "integrity": "sha512-OO0pH2lK6a0hZnAdau5ItzHPI6pUlvI7jMVnxUQRtw4owF2wk8lOSabtGDCTP4Ggrg2MbGnWO9X8K1t4+fGMDw==", "dev": true }, - "fsevents": { - "version": "2.3.2", - "resolved": "https://registry.npmjs.org/fsevents/-/fsevents-2.3.2.tgz", - "integrity": "sha512-xiqMQR4xAeHTuB9uWm+fFRcIOgKBMiOBP+eXiyT7jsgVCq1bkVygt00oASowB7EdtpOHaaPgKt812P9ab+DDKA==", - "dev": true, - "optional": true - }, "function-bind": { "version": "1.1.1", "resolved": "https://registry.npmjs.org/function-bind/-/function-bind-1.1.1.tgz", @@ -1277,17 +1736,16 @@ "dev": true }, "glob": { - "version": "7.2.3", - "resolved": "https://registry.npmjs.org/glob/-/glob-7.2.3.tgz", - "integrity": "sha512-nFR0zLpU2YCaRxwoCJvL6UvCH2JFyFVIvwTLsIf21AuHlMskA1hhTdk+LlYJtOlYt9v6dvszD2BGRqBL+iQK9Q==", + "version": "8.0.3", + "resolved": "https://registry.npmjs.org/glob/-/glob-8.0.3.tgz", + "integrity": "sha512-ull455NHSHI/Y1FqGaaYFaLGkNMMJbavMrEGFXG/PGrg6y7sutWHUHrz6gy6WEBH6akM1M414dWKCNs+IhKdiQ==", "dev": true, "requires": { "fs.realpath": "^1.0.0", "inflight": "^1.0.4", "inherits": "2", - "minimatch": "^3.1.1", - "once": "^1.3.0", - "path-is-absolute": "^1.0.0" + "minimatch": "^5.0.1", + "once": "^1.3.0" } }, "has": { @@ -1299,6 +1757,11 @@ "function-bind": "^1.1.1" } }, + "immutable": { + "version": "4.2.2", + "resolved": "https://registry.npmjs.org/immutable/-/immutable-4.2.2.tgz", + "integrity": "sha512-fTMKDwtbvO5tldky9QZ2fMX7slR0mYpY5nbnFWYp0fOzDhHqhgIw9KoYgxLWsoNTS9ZHGauHj18DTyEw6BK3Og==" + }, "inflight": { "version": "1.0.6", "resolved": "https://registry.npmjs.org/inflight/-/inflight-1.0.6.tgz", @@ -1372,26 +1835,63 @@ } }, "magic-string": { - "version": "0.25.9", - "resolved": "https://registry.npmjs.org/magic-string/-/magic-string-0.25.9.tgz", - "integrity": "sha512-RmF0AsMzgt25qzqqLc1+MbHmhdx0ojF2Fvs4XnOqz2ZOBXzzkEwc/dJQZCYHAn7v1jbVOjAZfK8msRn4BxO4VQ==", + "version": "0.27.0", + "resolved": "https://registry.npmjs.org/magic-string/-/magic-string-0.27.0.tgz", + "integrity": "sha512-8UnnX2PeRAPZuN12svgR9j7M1uWMovg/CEnIwIG0LFkXSJJe4PdfUGiTGl8V9bsBHFUtfVINcSyYxd7q+kx9fA==", "dev": true, "requires": { - "sourcemap-codec": "^1.4.8" + "@jridgewell/sourcemap-codec": "^1.4.13" } }, "marked": { - "version": "4.2.2", - "resolved": "https://registry.npmjs.org/marked/-/marked-4.2.2.tgz", - "integrity": "sha512-JjBTFTAvuTgANXx82a5vzK9JLSMoV6V3LBVn4Uhdso6t7vXrGx7g1Cd2r6NYSsxrYbQGFCMqBDhFHyK5q2UvcQ==" + "version": "4.2.12", + "resolved": "https://registry.npmjs.org/marked/-/marked-4.2.12.tgz", + "integrity": "sha512-yr8hSKa3Fv4D3jdZmtMMPghgVt6TWbk86WQaWhDloQjRSQhMMYCAro7jP7VDJrjjdV8pxVxMssXS8B8Y5DZ5aw==" + }, + "mathjax-full": { + "version": "3.2.2", + "resolved": "https://registry.npmjs.org/mathjax-full/-/mathjax-full-3.2.2.tgz", + "integrity": "sha512-+LfG9Fik+OuI8SLwsiR02IVdjcnRCy5MufYLi0C3TdMT56L/pjB0alMVGgoWJF8pN9Rc7FESycZB9BMNWIid5w==", + "requires": { + "esm": "^3.2.25", + "mhchemparser": "^4.1.0", + "mj-context-menu": "^0.6.1", + "speech-rule-engine": "^4.0.6" + } + }, + "mhchemparser": { + "version": "4.1.1", + "resolved": "https://registry.npmjs.org/mhchemparser/-/mhchemparser-4.1.1.tgz", + "integrity": "sha512-R75CUN6O6e1t8bgailrF1qPq+HhVeFTM3XQ0uzI+mXTybmphy3b6h4NbLOYhemViQ3lUs+6CKRkC3Ws1TlYREA==" }, "minimatch": { - "version": "3.1.2", - "resolved": "https://registry.npmjs.org/minimatch/-/minimatch-3.1.2.tgz", - "integrity": "sha512-J7p63hRiAjw1NDEww1W7i37+ByIrOWO5XQQAzZ3VOcL0PNybwpfmV/N05zFAzwQ9USyEcX6t3UO+K5aqBQOIHw==", + "version": "5.1.2", + "resolved": "https://registry.npmjs.org/minimatch/-/minimatch-5.1.2.tgz", + "integrity": "sha512-bNH9mmM9qsJ2X4r2Nat1B//1dJVcn3+iBLa3IgqJ7EbGaDNepL9QSHOxN4ng33s52VMMhhIfgCYDk3C4ZmlDAg==", "dev": true, "requires": { - "brace-expansion": "^1.1.7" + "brace-expansion": "^2.0.1" + } + }, + "mj-context-menu": { + "version": "0.6.1", + "resolved": "https://registry.npmjs.org/mj-context-menu/-/mj-context-menu-0.6.1.tgz", + "integrity": "sha512-7NO5s6n10TIV96d4g2uDpG7ZDpIhMh0QNfGdJw/W47JswFcosz457wqz/b5sAKvl12sxINGFCn80NZHKwxQEXA==" + }, + "moo": { + "version": "0.5.2", + "resolved": "https://registry.npmjs.org/moo/-/moo-0.5.2.tgz", + "integrity": "sha512-iSAJLHYKnX41mKcJKjqvnAN9sf0LMDTXDEvFv+ffuRR9a1MIuXLjMNL6EsnDHSkKLTWNqQQ5uo61P4EbU4NU+Q==" + }, + "nearley": { + "version": "2.20.1", + "resolved": "https://registry.npmjs.org/nearley/-/nearley-2.20.1.tgz", + "integrity": "sha512-+Mc8UaAebFzgV+KpI5n7DasuuQCHA89dmwm7JXw3TV43ukfNQ9DnBH3Mdb2g/I4Fdxc26pwimBWvjIw0UAILSQ==", + "requires": { + "commander": "^2.19.0", + "moo": "^0.5.0", + "railroad-diagrams": "^1.0.0", + "randexp": "0.4.6" } }, "object-assign": { @@ -1408,12 +1908,6 @@ "wrappy": "1" } }, - "path-is-absolute": { - "version": "1.0.1", - "resolved": "https://registry.npmjs.org/path-is-absolute/-/path-is-absolute-1.0.1.tgz", - "integrity": "sha512-AVbw3UJ2e9bq64vSaS9Am0fje1Pa8pbGqTTsmXfaIiMpnr5DlDhfJOuLj9Sf95ZPVDAUerDfEk88MPmPe7UCQg==", - "dev": true - }, "path-parse": { "version": "1.0.7", "resolved": "https://registry.npmjs.org/path-parse/-/path-parse-1.0.7.tgz", @@ -1426,6 +1920,11 @@ "integrity": "sha512-JU3teHTNjmE2VCGFzuY8EXzCDVwEqB2a8fsIvwaStHhAWJEeVd1o1QD80CU6+ZdEXXSLbSsuLwJjkCBWqRQUVA==", "dev": true }, + "poly-partition": { + "version": "1.0.2", + "resolved": "https://registry.npmjs.org/poly-partition/-/poly-partition-1.0.2.tgz", + "integrity": "sha512-uUqyuUq1n9wk3ZruUm7LdZiqcJ+RwBf3iPxdPe+kn6U3kn1H8jtH9EKg6I1NDG8pomnnxvuXoTpcusCWyemc2A==" + }, "postcss-value-parser": { "version": "3.3.1", "resolved": "https://registry.npmjs.org/postcss-value-parser/-/postcss-value-parser-3.3.1.tgz", @@ -1441,6 +1940,29 @@ "react-is": "^16.13.1" } }, + "railroad-diagrams": { + "version": "1.0.0", + "resolved": "https://registry.npmjs.org/railroad-diagrams/-/railroad-diagrams-1.0.0.tgz", + "integrity": "sha512-cz93DjNeLY0idrCNOH6PviZGRN9GJhsdm9hpn1YCS879fj4W+x5IFJhhkRZcwVgMmFF7R82UA/7Oh+R8lLZg6A==" + }, + "randexp": { + "version": "0.4.6", + "resolved": "https://registry.npmjs.org/randexp/-/randexp-0.4.6.tgz", + "integrity": "sha512-80WNmd9DA0tmZrw9qQa62GPPWfuXJknrmVmLcxvq4uZBdYqb1wYoKTmnlGUchvVWe0XiLupYkBoXVOxz3C8DYQ==", + "requires": { + "discontinuous-range": "1.0.0", + "ret": "~0.1.10" + } + }, + "randombytes": { + "version": "2.1.0", + "resolved": "https://registry.npmjs.org/randombytes/-/randombytes-2.1.0.tgz", + "integrity": "sha512-vYl3iOX+4CKUWuxGi9Ukhie6fsqXqS9FE2Zaic4tNFD2N2QQaXOMFbuKK4QmDHC0JO6B1Zp41J0LpT0oR68amQ==", + "dev": true, + "requires": { + "safe-buffer": "^5.1.0" + } + }, "react": { "version": "18.2.0", "resolved": "https://registry.npmjs.org/react/-/react-18.2.0.tgz", @@ -1502,9 +2024,9 @@ } }, "recharts": { - "version": "2.1.16", - "resolved": "https://registry.npmjs.org/recharts/-/recharts-2.1.16.tgz", - "integrity": "sha512-aYn1plTjYzRCo3UGxtWsduslwYd+Cuww3h/YAAEoRdGe0LRnBgYgaXSlVrNFkWOOSXrBavpmnli9h7pvRuk5wg==", + "version": "2.2.0", + "resolved": "https://registry.npmjs.org/recharts/-/recharts-2.2.0.tgz", + "integrity": "sha512-/uRJ0oaESGyz//PgAzvrwXEhrKaNha1ELLysEMRklbnsddiVQsSNicP7DWiz8qFcyYXy3BrDqrUjiLiVRTSMtA==", "requires": { "@types/d3-interpolate": "^2.0.0", "@types/d3-scale": "^3.0.0", @@ -1530,6 +2052,11 @@ "decimal.js-light": "^2.4.1" } }, + "recursive-diff": { + "version": "1.0.8", + "resolved": "https://registry.npmjs.org/recursive-diff/-/recursive-diff-1.0.8.tgz", + "integrity": "sha512-WfUcVao/HRsZZjI96PjcefTZ87rsxrgryIlWJB2UrHs4v3kFFzvMOm801a1WoqG+IAsLKJPsGlvubR74EDqGWQ==" + }, "reduce-css-calc": { "version": "2.1.8", "resolved": "https://registry.npmjs.org/reduce-css-calc/-/reduce-css-calc-2.1.8.tgz", @@ -1555,40 +2082,25 @@ "supports-preserve-symlinks-flag": "^1.0.0" } }, + "ret": { + "version": "0.1.15", + "resolved": "https://registry.npmjs.org/ret/-/ret-0.1.15.tgz", + "integrity": "sha512-TTlYpa+OL+vMMNG24xSlQGEJ3B/RzEfUlLct7b5G/ytav+wPrplCpVMFuwzXbkecJrb6IYo1iFb0S9v37754mg==" + }, "rollup": { - "version": "2.79.1", - "resolved": "https://registry.npmjs.org/rollup/-/rollup-2.79.1.tgz", - "integrity": "sha512-uKxbd0IhMZOhjAiD5oAFp7BqvkA4Dv47qpOCtaNvng4HBwdbWtdOh8f5nZNuk2rp51PMGk3bzfWu5oayNEuYnw==", + "version": "3.9.1", + "resolved": "https://registry.npmjs.org/rollup/-/rollup-3.9.1.tgz", + "integrity": "sha512-GswCYHXftN8ZKGVgQhTFUJB/NBXxrRGgO2NCy6E8s1rwEJ4Q9/VttNqcYfEvx4dTo4j58YqdC3OVztPzlKSX8w==", "dev": true, "requires": { "fsevents": "~2.3.2" } }, - "rollup-plugin-string": { - "version": "3.0.0", - "resolved": "https://registry.npmjs.org/rollup-plugin-string/-/rollup-plugin-string-3.0.0.tgz", - "integrity": "sha512-vqyzgn9QefAgeKi+Y4A7jETeIAU1zQmS6VotH6bzm/zmUQEnYkpIGRaOBPY41oiWYV4JyBoGAaBjYMYuv+6wVw==", - "dev": true, - "requires": { - "rollup-pluginutils": "^2.4.1" - } - }, - "rollup-pluginutils": { - "version": "2.8.2", - "resolved": "https://registry.npmjs.org/rollup-pluginutils/-/rollup-pluginutils-2.8.2.tgz", - "integrity": "sha512-EEp9NhnUkwY8aif6bxgovPHMoMoNr2FulJziTndpt5H9RdwC47GSGuII9XxpSdzVGM0GWrNPHV6ie1LTNJPaLQ==", - "dev": true, - "requires": { - "estree-walker": "^0.6.1" - }, - "dependencies": { - "estree-walker": { - "version": "0.6.1", - "resolved": "https://registry.npmjs.org/estree-walker/-/estree-walker-0.6.1.tgz", - "integrity": "sha512-SqmZANLWS0mnatqbSfRP5g8OXZC12Fgg1IwNtLsyHDzJizORW4khDfjPqJZsemPWBB2uqykUah5YpQ6epsqC/w==", - "dev": true - } - } + "safe-buffer": { + "version": "5.2.1", + "resolved": "https://registry.npmjs.org/safe-buffer/-/safe-buffer-5.2.1.tgz", + "integrity": "sha512-rp3So07KcdmmKbGvgaNxQSJr7bGVSVk5S9Eq1F+ppbRo70+YeaDxkw5Dd8NPN+GD6bjnYm2VuPuCXmpuYvmCXQ==", + "dev": true }, "scheduler": { "version": "0.23.0", @@ -1598,12 +2110,59 @@ "loose-envify": "^1.1.0" } }, - "sourcemap-codec": { - "version": "1.4.8", - "resolved": "https://registry.npmjs.org/sourcemap-codec/-/sourcemap-codec-1.4.8.tgz", - "integrity": "sha512-9NykojV5Uih4lgo5So5dtw+f0JgJX30KCNI8gwhz2J9A15wD0Ml6tjHKwf6fTSa6fAdVBdZeNOs9eJ71qCk8vA==", + "seedrandom": { + "version": "3.0.5", + "resolved": "https://registry.npmjs.org/seedrandom/-/seedrandom-3.0.5.tgz", + "integrity": "sha512-8OwmbklUNzwezjGInmZ+2clQmExQPvomqjL7LFqOYqtmuxRgQYqOD3mHaU+MvZn5FLUeVxVfQjwLZW/n/JFuqg==" + }, + "serialize-javascript": { + "version": "6.0.0", + "resolved": "https://registry.npmjs.org/serialize-javascript/-/serialize-javascript-6.0.0.tgz", + "integrity": "sha512-Qr3TosvguFt8ePWqsvRfrKyQXIiW+nGbYpy8XK24NQHE83caxWt+mIymTT19DGFbNWNLfEwsrkSmN64lVWB9ag==", + "dev": true, + "requires": { + "randombytes": "^2.1.0" + } + }, + "smob": { + "version": "0.0.6", + "resolved": "https://registry.npmjs.org/smob/-/smob-0.0.6.tgz", + "integrity": "sha512-V21+XeNni+tTyiST1MHsa84AQhT1aFZipzPpOFAVB8DkHzwJyjjAmt9bgwnuZiZWnIbMo2duE29wybxv/7HWUw==", "dev": true }, + "source-map": { + "version": "0.6.1", + "resolved": "https://registry.npmjs.org/source-map/-/source-map-0.6.1.tgz", + "integrity": "sha512-UjgapumWlbMhkBgzT7Ykc5YXUT46F0iKu8SGXq0bcwP5dz/h0Plj6enJqjz1Zbq2l5WaqYnrVbwWOWMyF3F47g==", + "dev": true + }, + "source-map-support": { + "version": "0.5.21", + "resolved": "https://registry.npmjs.org/source-map-support/-/source-map-support-0.5.21.tgz", + "integrity": "sha512-uBHU3L3czsIyYXKX88fdrGovxdSCoTGDRZ6SYXtSRxLZUzHg5P/66Ht6uoUlHu9EZod+inXhKo3qQgwXUT/y1w==", + "dev": true, + "requires": { + "buffer-from": "^1.0.0", + "source-map": "^0.6.0" + } + }, + "speech-rule-engine": { + "version": "4.0.7", + "resolved": "https://registry.npmjs.org/speech-rule-engine/-/speech-rule-engine-4.0.7.tgz", + "integrity": "sha512-sJrL3/wHzNwJRLBdf6CjJWIlxC04iYKkyXvYSVsWVOiC2DSkHmxsqOhEeMsBA9XK+CHuNcsdkbFDnoUfAsmp9g==", + "requires": { + "commander": "9.2.0", + "wicked-good-xpath": "1.3.0", + "xmldom-sre": "0.1.31" + }, + "dependencies": { + "commander": { + "version": "9.2.0", + "resolved": "https://registry.npmjs.org/commander/-/commander-9.2.0.tgz", + "integrity": "sha512-e2i4wANQiSXgnrBlIatyHtP1odfUp0BbV5Y5nEGbxtIrStkEOAAzCUirvLBNXHLr7kwLvJl6V+4V3XV9x7Wd9w==" + } + } + }, "supports-preserve-symlinks-flag": { "version": "1.0.0", "resolved": "https://registry.npmjs.org/supports-preserve-symlinks-flag/-/supports-preserve-symlinks-flag-1.0.0.tgz", @@ -1615,6 +2174,23 @@ "resolved": "https://registry.npmjs.org/tachyons/-/tachyons-4.12.0.tgz", "integrity": "sha512-2nA2IrYFy3raCM9fxJ2KODRGHVSZNTW3BR0YnlGsLUf1DA3pk3YfWZ/DdfbnZK6zLZS+jUenlUGJsKcA5fUiZg==" }, + "terser": { + "version": "5.16.1", + "resolved": "https://registry.npmjs.org/terser/-/terser-5.16.1.tgz", + "integrity": "sha512-xvQfyfA1ayT0qdK47zskQgRZeWLoOQ8JQ6mIgRGVNwZKdQMU+5FkCBjmv4QjcrTzyZquRw2FVtlJSRUmMKQslw==", + "dev": true, + "requires": { + "@jridgewell/source-map": "^0.3.2", + "acorn": "^8.5.0", + "commander": "^2.20.0", + "source-map-support": "~0.5.20" + } + }, + "true-myth": { + "version": "4.1.1", + "resolved": "https://registry.npmjs.org/true-myth/-/true-myth-4.1.1.tgz", + "integrity": "sha512-rqy30BSpxPznbbTcAcci90oZ1YR4DqvKcNXNerG5gQBU2v4jk0cygheiul5J6ExIMrgDVuanv/MkGfqZbKrNNg==" + }, "tslib": { "version": "2.4.1", "resolved": "https://registry.npmjs.org/tslib/-/tslib-2.4.1.tgz", @@ -1622,11 +2198,19 @@ "dev": true }, "typescript": { - "version": "4.8.4", - "resolved": "https://registry.npmjs.org/typescript/-/typescript-4.8.4.tgz", - "integrity": "sha512-QCh+85mCy+h0IGff8r5XWzOVSbBO+KfeYrMQh7NJ58QujwcE22u+NUSmUxqF+un70P9GXKxa2HCNiTTMJknyjQ==", + "version": "4.9.4", + "resolved": "https://registry.npmjs.org/typescript/-/typescript-4.9.4.tgz", + "integrity": "sha512-Uz+dTXYzxXXbsFpM86Wh3dKCxrQqUcVMxwU54orwlJjOpO3ao8L7j5lH+dWfTwgCwIuM9GQ2kvVotzYJMXTBZg==", "dev": true }, + "use-resize-observer": { + "version": "9.1.0", + "resolved": "https://registry.npmjs.org/use-resize-observer/-/use-resize-observer-9.1.0.tgz", + "integrity": "sha512-R25VqO9Wb3asSD4eqtcxk8sJalvIOYBqS8MNZlpDSQ4l4xMQxC/J7Id9HoTqPq8FwULIn0PVW+OAqF2dyYbjow==", + "requires": { + "@juggle/resize-observer": "^3.3.1" + } + }, "vscode-jsonrpc": { "version": "8.0.2", "resolved": "https://registry.npmjs.org/vscode-jsonrpc/-/vscode-jsonrpc-8.0.2.tgz", @@ -1646,11 +2230,21 @@ "resolved": "https://registry.npmjs.org/vscode-languageserver-types/-/vscode-languageserver-types-3.17.2.tgz", "integrity": "sha512-zHhCWatviizPIq9B7Vh9uvrH6x3sK8itC84HkamnBWoDFJtzBf7SWlpLCZUit72b3os45h6RWQNC9xHRDF8dRA==" }, + "wicked-good-xpath": { + "version": "1.3.0", + "resolved": "https://registry.npmjs.org/wicked-good-xpath/-/wicked-good-xpath-1.3.0.tgz", + "integrity": "sha512-Gd9+TUn5nXdwj/hFsPVx5cuHHiF5Bwuc30jZ4+ronF1qHK5O7HD0sgmXWSEgwKquT3ClLoKPVbO6qGwVwLzvAw==" + }, "wrappy": { "version": "1.0.2", "resolved": "https://registry.npmjs.org/wrappy/-/wrappy-1.0.2.tgz", "integrity": "sha512-l4Sp/DRseor9wL6EvV2+TuQn63dMkPjZ/sp9XkghTEbV9KlPS1xUsZ3u7/IQO4wxtcFB4bgpQPRcR3QCvezPcQ==", "dev": true + }, + "xmldom-sre": { + "version": "0.1.31", + "resolved": "https://registry.npmjs.org/xmldom-sre/-/xmldom-sre-0.1.31.tgz", + "integrity": "sha512-f9s+fUkX04BxQf+7mMWAp5zk61pciie+fFLC9hX9UVvCeJQfNHRHXpeo5MPcR0EUf57PYLdt+ZO4f3Ipk2oZUw==" } } } diff --git a/widget/package-lock.json.trace b/widget/package-lock.json.trace index fc4557a..0ce562a 100644 --- a/widget/package-lock.json.trace +++ b/widget/package-lock.json.trace @@ -1 +1 @@ -2327248813061066255 \ No newline at end of file +15925424785280637440 \ No newline at end of file diff --git a/widget/package.json b/widget/package.json index ad2319c..e76ba8a 100644 --- a/widget/package.json +++ b/widget/package.json @@ -8,23 +8,28 @@ }, "author": "", "license": "Apache-2.0", + "type": "module", "dependencies": { - "@leanprover/infoview": "^0.2.0", + "@leanprover/infoview": "~0.4.2", + "@penrose/core": "~1.3.1-alpha.363", + "@svgdotjs/svg.js": "^3.1.2", "react": "^18.2.0", "react-dom": "^18.2.0", - "recharts": "^2.1.16", + "recharts": "^2.2.0", + "use-resize-observer": "^9.1.0", "vscode-languageserver-protocol": "^3.17.2" }, "devDependencies": { - "@rollup/plugin-commonjs": "^22.0.2", - "@rollup/plugin-node-resolve": "^14.1.0", - "@rollup/plugin-replace": "^4.0.0", - "@rollup/plugin-typescript": "^8.5.0", - "@types/react": "^17.0.39", - "@types/react-dom": "^17.0.11", - "rollup": "^2.79.0", - "rollup-plugin-string": "^3.0.0", - "tslib": "^2.4.0", - "typescript": "^4.8.3" + "@rollup/plugin-commonjs": "^24.0.0", + "@rollup/plugin-node-resolve": "^15.0.1", + "@rollup/plugin-replace": "^5.0.2", + "@rollup/plugin-terser": "^0.3.0", + "@rollup/plugin-typescript": "^11.0.0", + "@types/react": "^18.0.26", + "@types/react-dom": "^18.0.10", + "glob": "^8.0.3", + "rollup": "^3.9.1", + "tslib": "^2.4.1", + "typescript": "^4.9.4" } } diff --git a/widget/rollup.config.js b/widget/rollup.config.js index 0295636..6f20426 100644 --- a/widget/rollup.config.js +++ b/widget/rollup.config.js @@ -1,42 +1,39 @@ -import { nodeResolve} from '@rollup/plugin-node-resolve' +import glob from 'glob'; +import { nodeResolve } from '@rollup/plugin-node-resolve' import typescript from '@rollup/plugin-typescript'; import commonjs from '@rollup/plugin-commonjs'; import replace from '@rollup/plugin-replace'; -import { string } from 'rollup-plugin-string'; +import terser from '@rollup/plugin-terser'; + +/** @type {(_: any) => import('rollup').RollupOptions} */ export default cliArgs => { const tsxName = cliArgs.tsxName; - if (tsxName === undefined) - throw new Error("Please specify the .tsx to build with --tsxName ."); - // We delete the custom argument so that Rollup does not try to process it and complain. - delete cliArgs.tsxName; + if (tsxName !== undefined) + // We delete the custom argument so that Rollup does not try to process it and complain. + delete cliArgs.tsxName; - return { - input: [ - `src/${tsxName}.tsx` - ], + const inputs = tsxName ? + [ `src/${tsxName}.tsx` ] : + glob.sync('src/**/*.tsx') + + const configForInput = fname => ({ + input: fname, output: { - dir: "dist", - format: "es", + dir: '../build/js', + format: 'es', // Hax: apparently setting `global` makes some CommonJS modules work ¯\_(ツ)_/¯ - intro: "const global = window;" + intro: 'const global = window;' }, external: [ 'react', 'react-dom', + 'react/jsx-runtime', '@leanprover/infoview', ], plugins: [ - string({ - include: [ - '**/*.dsl', - '**/*.sty', - '**/*.sub' - ] - }), typescript({ - tsconfig: "./tsconfig.json", + tsconfig: './tsconfig.json', outputToFilesystem: false, - sourceMap: false }), nodeResolve({ browser: true @@ -59,5 +56,12 @@ export default cliArgs => { 'cluster', 'module', 'net', 'readline', 'repl', 'tls', 'fs', 'crypto', 'perf_hooks', ], }), + terser(), ], -}} + }) + + // For why we must return an array of configs rather than use an array in `input`, see + // https://github.com/rollup/rollup/issues/2756. This is pretty suboptimal as every build + // rechecks all TS files, so the process is quadratic. + return inputs.map(configForInput) +} diff --git a/widget/src/compat.tsx b/widget/src/compat.tsx new file mode 100644 index 0000000..eefe5f2 --- /dev/null +++ b/widget/src/compat.tsx @@ -0,0 +1,69 @@ +import * as React from 'react' +import { mapRpcError, useAsync, DynamicComponent, PanelWidgetProps, RpcContext } + from '@leanprover/infoview' +import { Range } from 'vscode-languageserver-protocol' + +interface MetaWidgetProps extends PanelWidgetProps { + infoId: string +} + +interface PanelWidgetInstance { + id: string + srcHash: string + props: any + infoId: string + range?: Range +} + +export default function(props_: MetaWidgetProps): JSX.Element { + const {pos, infoId, ...props} = props_ + + const rs = React.useContext(RpcContext) + const st = useAsync(async () => { + const ws: { widgets: PanelWidgetInstance[] } = await rs.call('WidgetKit.getPanelWidgets', {pos}) + const ret = [] + for (const w of ws.widgets) { + if (w.infoId === infoId) { + ret.push(w) + } + } + return ret + }, [rs, infoId, pos]) + + // Hide the '▶ Widget Name' element + const ref = React.useRef(null) + React.useLayoutEffect(() => { + if (ref.current === null) return + const parent = ref.current.parentElement + if (parent === null) return + for (const c of parent.children) { + if (c.tagName !== 'SUMMARY') continue + if (!(c instanceof HTMLElement)) continue + c.style.display = 'none' + break + } + }, []) + + const [st0, setSt0] = React.useState(undefined) + React.useEffect(() => { + if (st.state === 'resolved') + setSt0(st.value) + }, [st.state]) + + let inner = <> + if (st0 !== undefined) + inner = <> + {st0.map(w => + )} + + else if (st.state === 'rejected') + inner = <>Error: {mapRpcError(st.error).message} + else + inner = <>Loading.. + + return
{inner}
+} diff --git a/widget/src/exprPresentation.tsx b/widget/src/exprPresentation.tsx new file mode 100644 index 0000000..974a167 --- /dev/null +++ b/widget/src/exprPresentation.tsx @@ -0,0 +1,65 @@ +import * as React from 'react' +import { useAsync, RpcContext, RpcSessionAtPos, RpcPtr, Name, DocumentPosition, mapRpcError } + from '@leanprover/infoview' +import HtmlDisplay, { Html } from './htmlDisplay' +import InteractiveExpr from './interactiveExpr' + +type ExprWithCtx = RpcPtr<'WidgetKit.ExprWithCtx'> + +interface PresenterId { + name: Name + userName: string +} + +async function applicableExprPresenters(rs: RpcSessionAtPos, expr: ExprWithCtx): + Promise { + const ret: any = await rs.call('WidgetKit.applicableExprPresenters', { expr }) + return ret.presenters +} + +async function getExprPresentation(rs: RpcSessionAtPos, expr: ExprWithCtx, name: Name): + Promise { + return await rs.call('WidgetKit.getExprPresentation', { expr, name }) +} + +interface ExprPresentationUsingProps { + pos: DocumentPosition + expr: ExprWithCtx + name: Name +} + +// TODO: a UseAsync component which displays resolved/loading/error like we do in every component + +/** Display the given expression using the `ExprPresenter` registered at `name`. */ +function ExprPresentationUsing({pos, expr, name}: ExprPresentationUsingProps): JSX.Element { + const rs = React.useContext(RpcContext) + const st = useAsync(() => getExprPresentation(rs, expr, name), [rs, expr, name]) + return st.state === 'resolved' ? + : st.state === 'loading' ? <>Loading.. + : <>Error: {mapRpcError(st.error).message} +} + +/** Display the given expression using an `ExprPresenter`. The server is queried for registered + * `ExprPresenter`s. A dropdown is shown allowing the user to select which of these should be used + * to display the expression. */ +export default function({pos, expr}: {pos: DocumentPosition, expr: ExprWithCtx}): JSX.Element { + const rs = React.useContext(RpcContext) + const st = useAsync(() => applicableExprPresenters(rs, expr), [rs, expr]) + const [selection, setSelection] = React.useState(undefined) + + if (st.state === 'rejected') + return <>Error: {mapRpcError(st.error).message} + else if (st.state === 'resolved' && 0 < st.value.length) + // For explanation of flow-root see https://stackoverflow.com/a/32301823 + return
+ {selection && selection !== 'none' ? + : + } + +
+ else + return +} diff --git a/widget/src/htmlDisplay.tsx b/widget/src/htmlDisplay.tsx new file mode 100644 index 0000000..23fcfb8 --- /dev/null +++ b/widget/src/htmlDisplay.tsx @@ -0,0 +1,49 @@ +/* +Copyright (c) 2022 E.W.Ayers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: E.W.Ayers +*/ + +import * as React from 'react'; +import { DynamicComponent, DocumentPosition } from '@leanprover/infoview'; + +type HtmlAttribute = [string, any] + +export type Html = + | { element: [string, HtmlAttribute[], Html[]] } + | { text: string } + | { component: [string, any, Html[]] } + +export default function HtmlDisplay({pos, html} : {pos: DocumentPosition, html: Html}): JSX.Element { + if ('text' in html) { + return <>{html.text} + } else if ('element' in html) { + const [tag, attrsList, cs] = html.element + const attrs: any = {} + for (const [k,v] of attrsList) { + attrs[k] = v + } + const children = cs.map(html => HtmlDisplay({ pos, html })) + if (tag === "hr") { + // React is greatly concerned by
s having children. + return
+ } else if (children.length === 0) { + return React.createElement(tag, attrs) + } else { + return React.createElement(tag, attrs, children) + } + } else if ('component' in html) { + const [hash, props, cs] = html.component + const children = cs.map(html => HtmlDisplay({ pos, html })) + const dynProps = {...props, pos} + if (children.length === 0) { + return + } else { + return + <>{children} + + } + } else { + return Unknown HTML: {JSON.stringify(html)} + } +} diff --git a/widget/src/htmlDisplayPanel.tsx b/widget/src/htmlDisplayPanel.tsx new file mode 100644 index 0000000..5f2a6ba --- /dev/null +++ b/widget/src/htmlDisplayPanel.tsx @@ -0,0 +1,10 @@ +import { DocumentPosition } from "@leanprover/infoview" +import HtmlDisplay, { Html } from "./htmlDisplay" + +export default function HtmlDisplayPanel({pos, html} : {pos: DocumentPosition, html: Html}): + JSX.Element { + return
+ HTML Display + +
+} diff --git a/widget/src/interactiveExpr.tsx b/widget/src/interactiveExpr.tsx new file mode 100644 index 0000000..7ac49f9 --- /dev/null +++ b/widget/src/interactiveExpr.tsx @@ -0,0 +1,15 @@ +import { mapRpcError, useAsync, InteractiveCode, RpcContext, RpcPtr } from '@leanprover/infoview' +import * as React from 'react' + +type ExprWithCtx = RpcPtr<'WidgetKit.ExprWithCtx'> + +export default function({expr}: {expr: ExprWithCtx}): JSX.Element { + const rs = React.useContext(RpcContext) + const st = useAsync(() => rs.call('WidgetKit.ppExprTagged', {expr}), [expr]) + if (st.state === 'resolved') + return + else if (st.state === 'rejected') + return <>Error: ${mapRpcError(st.error).message} + else + return <>Loading.. +} diff --git a/widget/src/interactiveSvg.tsx b/widget/src/interactiveSvg.tsx index c42412a..66f2d19 100644 --- a/widget/src/interactiveSvg.tsx +++ b/widget/src/interactiveSvg.tsx @@ -105,7 +105,7 @@ export function Svg(props : UpdateResult) { function handleMouseEvent(e : MouseEvent) { console.log(e) const id = e.target.id || undefined - const data = e.target.data || undefined + const data = e.target.data || undefined console.log(id) console.log(data) if (e.type === "mouseup" || e.type === "mousedown") { @@ -123,4 +123,4 @@ export function Svg(props : UpdateResult) { } -export default Svg \ No newline at end of file +export default Svg diff --git a/widget/src/penroseCanvas.tsx b/widget/src/penroseCanvas.tsx new file mode 100644 index 0000000..6664712 --- /dev/null +++ b/widget/src/penroseCanvas.tsx @@ -0,0 +1,303 @@ +/* +Copyright (c) 2022 Wojciech Nawrocki. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Wojciech Nawrocki +*/ +import * as React from "react" +import * as ReactDOM from "react-dom" +import * as penrose from "@penrose/core" +import * as SVG from "@svgdotjs/svg.js" +import useResizeObserver from "use-resize-observer"; + +/** See [here](https://penrose.gitbook.io/penrose/#what-makes-up-a-penrose-program) for explanation. */ +export interface PenroseTrio { + dsl: string + sty: string + sub: string +} + +/** Compute the hash of a Penrose trio. */ +async function hashTrio({dsl, sty, sub}: PenroseTrio): Promise { + // https://developer.mozilla.org/en-US/docs/Web/API/TextEncoder/encodeInto#buffer_sizing + const data = new Uint8Array(3 * (dsl.length + sty.length + sub.length)) + const encoder = new TextEncoder() + let dataView = data + const {written} = encoder.encodeInto(dsl, dataView) + dataView = data.subarray(written) + const {written: written2} = encoder.encodeInto(sty, dataView) + dataView = data.subarray(written2) + encoder.encodeInto(sub, dataView) + const digest = await crypto.subtle.digest("SHA-1", data) + const digestArr = Array.from(new Uint8Array(digest)) + // https://developer.mozilla.org/en-US/docs/Web/API/SubtleCrypto/digest#converting_a_digest_to_a_hex_string + return digestArr.map(b => b.toString(16).padStart(2, '0')).join('') +} + +/** The compile -> optimize -> prepare SVG sequence is not cheap (on the order of 1s for a simple + * diagram), so we cache its SVG outputs. */ +const diagramSvgCache = new Map() + +function svgNumberToNumber (x: SVG.NumberAlias): number { + let y: string | number + if (x instanceof Number) y = x.valueOf() + else y = x as any + + if (typeof y === 'string') return parseFloat(y) + else return y +} + +async function deleteCachedTrio(trio: PenroseTrio) { + const hash = await hashTrio(trio) + diagramSvgCache.delete(hash) +} + +async function renderPenroseTrio(trio: PenroseTrio, maxOptSteps: number): Promise { + const hash = await hashTrio(trio) + if (diagramSvgCache.has(hash)) return diagramSvgCache.get(hash)! + const {dsl, sty, sub} = trio + const compileRes = await penrose.compileTrio({ + domain: dsl, + style: sty, + substance: sub, + variation: '' + }) + if (compileRes.isErr()) throw new Error(penrose.showError(compileRes.error)) + const state = await penrose.prepareState(compileRes.value) + const stateRes = penrose.stepUntilConvergence(state, maxOptSteps) + if (stateRes.isErr()) throw new Error(penrose.showError(stateRes.error)) + if (!penrose.stateConverged(stateRes.value)) + console.warn(`Diagram failed to converge in ${maxOptSteps} steps`) + const svg = await penrose.RenderStatic(stateRes.value, async path => path) + + // The canvas is usually too large, so trim the SVG as a postprocessing step + const obj = SVG.SVG(svg) + const view = obj.viewbox() + let minX = view.width, maxX = 0, minY = view.height, maxY = 0 + + obj.each((i, children) => { + const child = children[i] + minX = Math.min(minX, svgNumberToNumber(child.x())) + maxX = Math.max(maxX, svgNumberToNumber(child.x()) + svgNumberToNumber(child.width())) + minY = Math.min(minY, svgNumberToNumber(child.y())) + maxY = Math.max(maxY, svgNumberToNumber(child.y()) + svgNumberToNumber(child.height())) + }) + + const padding = 10 + const newX = minX - padding, newY = minY - padding, + newW = (maxX - minX) + padding, newH = (maxY - minY) + padding + const newSvg = obj.viewbox(newX, newY, newW, newH).width(newW).height(newH) + diagramSvgCache.set(hash, newSvg.node) + + return newSvg.node +} + +/** Return all elements in a Penrose-generated SVG which have names corresponding to objects in the + * substance program. These are detected by looking for strings in the elements' `textContent`s. */ +function getPenroseNamedElements(svg: SVG.Svg): Map { + const res = new Map() + for (const child of svg.find('g, rect')) { + if (!child.node.textContent) continue + const groups = child.node.textContent.match(/`(\w+)`.textBox/) + if (!groups) continue + const name = groups[1] + res.set(name, child) + } + return res +} + +export interface PenroseCanvasProps { + trio: PenroseTrio + maxOptSteps: number + embedNodes: Map +} + +interface InnerWithContainerProps extends PenroseCanvasProps { + diagramWidth: number + hiddenDiv: HTMLDivElement +} + +interface InnerWithEmbedsProps extends Omit { + embeds: Map +} + +function InnerWithEmbeds({trio: {dsl, sty, sub}, maxOptSteps, embeds, diagramWidth, hiddenDiv}: + InnerWithEmbedsProps): JSX.Element { + sty = sty + +` +canvas { + width = ${diagramWidth} + height = ${diagramWidth} +} +` + + const cssColourToRgba = (col: string, alpha: number = 255) => { + if (col.startsWith('#')) { + const gps = col.match(/\w\w/g) + if (!gps) throw new Error(`cannot parse colour '${col}'`) + const [r, g, b] = gps.map(x => parseInt(x, 16)) + return `rgba(${r}/255,${g}/255,${b}/255,${alpha}/255)` + } else throw new Error(`cannot parse colour '${col}'`) + } + + const boxCol = cssColourToRgba( + getComputedStyle(document.documentElement) + .getPropertyValue('--vscode-editorHoverWidget-background')) + + for (const [name, elt] of embeds) { + // NOTE(WN): this getBoundingClientRect call is not easy to remove, but it might be okay; + // the dimensions of this rect should never change unless the whole diagram changes. + const rect = elt.getBoundingClientRect() + + // KC's hack: https://github.com/penrose/penrose/issues/1057#issuecomment-1164313880 + // NOTE: To manipulate objects directly, could look at impl of `compileTrio`. `compileSubstance` + // exposes the objects from the subtance program in the `Env` type. + sty = sty + +` +forall Targettable \`${name}\` { +override \`${name}\`.textBox.width = ${Math.ceil(rect.width)} +override \`${name}\`.textBox.height = ${Math.ceil(rect.height)} +override \`${name}\`.textBox.fillColor = ${boxCol} +} +` + } + + const [element, setElement] = React.useState(
Drawing..
) + const [svg, setSvg] = React.useState() + + const render = async () => { + try { + const svg = await renderPenroseTrio({dsl, sty, sub}, maxOptSteps) + setElement(<> + { + setElement(
Drawing..
) + void deleteCachedTrio({dsl, sty, sub}).then(() => render()) + }} /> +
{ + if (!ref) return + if (ref.firstChild) ref.replaceChild(svg, ref.firstChild) + else ref.appendChild(svg) + setSvg(svg) + }} /> + ) + } catch(ex: any) { + setElement(
Error while drawing: {ex.toString()}
) + } + } + + React.useEffect(() => void render(), [dsl, sty, sub, maxOptSteps, embeds]) + + // Position embeds over nodes in the SVG + React.useEffect(() => { + if (!svg) return + + /** The SVG boxes that we can draw interactive elements over. */ + const obj = SVG.SVG(svg) + const diagramBoxes = getPenroseNamedElements(obj) + for (const [name, embedElt] of embeds) { + const gElt = diagramBoxes.get(name) + if (!gElt) throw new Error(`Could not find object named '${name}' in the diagram.`) + // Note: this calculation assumes that one SVG user unit is one pixel. We achieve + // this by setting the `viewBox` width/height to the `` width/height. + const userY = svgNumberToNumber(gElt.y()) - obj.viewbox().y + const userX = svgNumberToNumber(gElt.x()) - obj.viewbox().x + embedElt.style.top = `${userY}px` + embedElt.style.left = `${userX}px` + } + + hiddenDiv.style.visibility = 'visible' + }, [svg]) + + return element +} + +function InnerWithContainer(props: InnerWithContainerProps): JSX.Element { + const {embedNodes, diagramWidth, hiddenDiv} = props + + interface EmbedData { + elt: HTMLDivElement | undefined + portal: React.ReactPortal + } + + const [embeds, setEmbeds] = React.useState>(new Map()) + // This is set once when all the embeds have been drawn + const [embedsFinal, setEmbedsFinal] = React.useState>() + + // Create portals for the embedded nodes as children of `hiddenDiv`; they will update `embeds` + // when rendered + React.useEffect(() => { + const newEmbeds: Map = new Map() + for (const [name, nd] of embedNodes) { + const div =
{ + if (!newDiv) return + setEmbeds(embeds => { + const newEmbeds: Map = new Map() + let changed = false + for (const [eName, data] of embeds) { + if (eName === name && data.elt !== newDiv) { + changed = true + newEmbeds.set(eName, {...data, elt: newDiv}) + } else newEmbeds.set(eName, data) + } + return changed ? newEmbeds : embeds + }) + }}>{nd}
+ const portal = ReactDOM.createPortal(div, hiddenDiv, name) + const data: EmbedData = { + elt: undefined, + portal + } + newEmbeds.set(name, data) + } + setEmbeds(newEmbeds) + // `deps` must have constant size so we can't do a deeper comparison + }, [embedNodes, hiddenDiv, diagramWidth]) + + React.useEffect(() => { + const embedDivs = new Map() + for (const [name, {elt}] of embeds) { + if (!elt) return + embedDivs.set(name, elt) + } + if (embedDivs.size !== embedNodes.size) return + setEmbedsFinal(embedDivs) + }, [embeds, embedNodes]) + + return <> + {embedsFinal && } + {Array.from(embeds.values()).map(({portal}) => portal)} + +} + +/** Renders an interactive [Penrose](https://github.com/penrose/penrose) diagram with the specified + * trio. The Penrose optimizer is ran for at most `maxOptSteps`, a heuristic for when to stop trying. + * + * For every `[name, nd]` in `embedNodes` we locate an object with the same `name` in the substance + * program, adjust the style program so that the object's dimensions match those of `nd`, and draw + * the React node `nd` over `name` in the SVG diagram. */ +export function PenroseCanvas(props: PenroseCanvasProps): JSX.Element { + /* The implementation: do some work, store results, pass results as props to a nested component. + * As opposed to doing everything in one component, nested components don't have to handle + * partial results and are drawn atomically. The flow is: + * - `PenroseCanvas` waits for the container div and an invisible `hiddenDiv`. + * - `InnerWithContainer` uses `hiddenDiv` to render embedded nodes without displaying them, + * so that their sizes can be measured. + * - `InnerWithEmbeds` adjusts the style program to match the sizes of embeds, draws the diagram, + * positions the embeds over it, and finally makes them visible. */ + + const { ref: containerRef, width = 1 } = useResizeObserver({ + round: Math.ceil, + }) + // TODO(WN): debounce changes to this; it's a lot of computation to keep redrawing the diagram + // when the infoview is being resized. + const diagramWidth = Math.max(400, width) + const [hiddenDiv, setHiddenDiv] = React.useState(null) + return
+ {1 < width && hiddenDiv && + } +
+
+} diff --git a/widget/src/penroseDisplay.tsx b/widget/src/penroseDisplay.tsx new file mode 100644 index 0000000..3f6d2b0 --- /dev/null +++ b/widget/src/penroseDisplay.tsx @@ -0,0 +1,35 @@ +/* +Copyright (c) 2022 Wojciech Nawrocki. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Wojciech Nawrocki +*/ +import * as React from 'react'; + +import { PenroseCanvas } from './penroseCanvas'; +import { DocumentPosition } from '@leanprover/infoview/infoview/util'; +import { Html, default as HtmlDisplay } from './htmlDisplay'; + +interface DiagramData { + embeds: [string, Html][], + dsl: string, + sty: string, + sub: string, +} + +export default function(props: DiagramData & {pos: DocumentPosition}): JSX.Element { + const {embeds, dsl, sty, sub} = props + const mkElt = (html: Html): JSX.Element => +
+ +
+ + const embedNodes = + embeds.reduce( + (acc, [nm, cwi]) => acc.set(nm, mkElt(cwi)), + new Map()) + + return +} diff --git a/widget/src/presentSelection.tsx b/widget/src/presentSelection.tsx new file mode 100644 index 0000000..4067cba --- /dev/null +++ b/widget/src/presentSelection.tsx @@ -0,0 +1,50 @@ +import { DocumentPosition, GoalsLocation, InteractiveGoal, mapRpcError, PanelWidgetProps, + RpcContext, RpcPtr, useAsync } from "@leanprover/infoview"; +import * as React from "react"; +import ExprPresentation from "./exprPresentation"; + +function findGoalForLocation(goals: InteractiveGoal[], loc: GoalsLocation): InteractiveGoal { + for (const g of goals) { + if (g.mvarId === loc.mvarId) return g + } + throw new Error(`Could not find goal for location ${JSON.stringify(loc)}`) +} + +type ExprWithCtx = RpcPtr<'WidgetKit.ExprWithCtx'> + +interface GoalsLocationsToExprsResponse { + exprs: ExprWithCtx[] +} + +/** + * Display the expression corresponding to a given `GoalsLocation` using {@link ExprPresentation}. + */ +function GoalsLocationPresentation({pos, goals, loc}: + {pos: DocumentPosition, goals: InteractiveGoal[], loc: GoalsLocation}) { + const rs = React.useContext(RpcContext) + const st = useAsync(async () => { + const g = findGoalForLocation(goals, loc) + if (g.ctx === undefined) throw new Error(`Lean server 1.1.2 or newer is required.`) + const ret: GoalsLocationsToExprsResponse = + await rs.call('WidgetKit.goalsLocationsToExprs', {locations: [[g.ctx, loc]]}) + return ret.exprs[0] + }, [rs, goals, loc]) + + if (st.state === 'loading') + return <>Loading.. + else if (st.state === 'rejected') + return <>Error: {mapRpcError(st.error).message} + else + return +} + +export default function(props: PanelWidgetProps) { + return
+ Selected expressions + {props.selectedLocations.length > 0 ? + props.selectedLocations.map(loc => + ) : + <>Nothing selected. You can use shift-click to select expressions in the goal.} +
+} diff --git a/widget/src/staticHtml.tsx b/widget/src/staticHtml.tsx index e35df5c..b929d9c 100644 --- a/widget/src/staticHtml.tsx +++ b/widget/src/staticHtml.tsx @@ -2,6 +2,9 @@ import * as React from 'react'; import * as Recharts from 'recharts'; import { InteractiveCode, CodeWithInfos } from '@leanprover/infoview'; +// TODO: The HTML part of this is now htmlDisplay.tsx. The rest should be split into a Recharts +// component and an AnimatedHtml component. + const htmlTags = "a abbr address area article aside audio b base bdi bdo big blockquote body br button canvas caption cite code col colgroup data datalist dd del details dfn dialog div dl dt em embed fieldset figcaption figure footer form h1 h2 h3 h4 h5 h6 head header hr html i iframe img input ins kbd keygen label legend li link main map mark menu menuitem meta meter nav noscript object ol optgroup option output p param picture pre progress q rp rt ruby s samp script section select small source span strong style sub summary sup table tbody td textarea tfoot th thead time title tr track u ul var video wbr".split(' ') const svgTags = "circle clipPath defs ellipse g line linearGradient mask path pattern polygon polyline radialGradient rect stop svg text tspan".split(' ') @@ -73,4 +76,4 @@ export default function(props : any) { } else { return StaticHtml(props) } -} \ No newline at end of file +} diff --git a/widget/tsconfig.json b/widget/tsconfig.json index bcc25dd..b48d520 100644 --- a/widget/tsconfig.json +++ b/widget/tsconfig.json @@ -1,11 +1,11 @@ { "compilerOptions": { "target": "esnext", - "jsx": "react", + "jsx": "react-jsx", "module": "esnext", "moduleResolution": "node", - "sourceMap": true, - "outDir": "dist/", + "sourceMap": false, + "outDir": "../build/js", "allowSyntheticDefaultImports": true, "esModuleInterop": true, "forceConsistentCasingInFileNames": true,