Skip to content

Commit

Permalink
feat: edge labels for GraphDisplay (#86)
Browse files Browse the repository at this point in the history
* feat(DigraphDisplay): edge labels

* fix(Html): specify type

* feat(DigraphDisplay): better defaults

* chore: reorganize

* chore: release notes
  • Loading branch information
Vtec234 authored Nov 4, 2024
1 parent 4d58e47 commit 81dfa90
Show file tree
Hide file tree
Showing 9 changed files with 203 additions and 127 deletions.
1 change: 1 addition & 0 deletions ProofWidgets.lean
Original file line number Diff line number Diff line change
@@ -1,6 +1,7 @@
import ProofWidgets.Compat
import ProofWidgets.Component.Basic
import ProofWidgets.Component.FilterDetails
import ProofWidgets.Component.GraphDisplay
import ProofWidgets.Component.HtmlDisplay
import ProofWidgets.Component.InteractiveSvg
import ProofWidgets.Component.MakeEditLink
Expand Down
5 changes: 4 additions & 1 deletion ProofWidgets/Component/FilterDetails.lean
Original file line number Diff line number Diff line change
@@ -1,6 +1,7 @@
import ProofWidgets.Data.Html

open Lean ProofWidgets
namespace ProofWidgets
open Lean

/-- Props for the `FilterDetails` component. -/
structure FilterDetailsProps where
Expand All @@ -20,3 +21,5 @@ that allows you to switch between filtered and unfiltered states. -/
@[widget_module]
def FilterDetails : Component FilterDetailsProps where
javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "filterDetails.js"

end ProofWidgets
Original file line number Diff line number Diff line change
@@ -1,28 +1,32 @@
import ProofWidgets.Component.Basic
import ProofWidgets.Data.Html

namespace ProofWidgets.DigraphDisplay
namespace ProofWidgets.GraphDisplay
open Lean Server Jsx

/-- A themed `<circle>` SVG element, with optional extra attributes. -/
def mkCircle (attrs : Array (String × Json) := #[]) : Html :=
<circle
r={5}
fill="var(--vscode-editor-background)"
stroke="var(--vscode-editor-foreground)"
strokeWidth={.num 1.5}
{...attrs}
/>

structure Vertex where
/-- Identifier for this vertex. Must be unique. -/
id : String
/-- The label is drawn at the vertex position.
This must be an SVG element.
Use `<foreignObject>` to draw non-SVG elements. -/
label : Html :=
<circle
r={5}
className="dim"
fill="var(--vscode-editor-background)"
stroke="var(--vscode-editor-foreground)"
strokeWidth={.num 1.5}
/>
label : Html := mkCircle
/-- Radius of a circle bounding this vertex.
Used to place incident edge endpoints. -/
radius : Float := 5
/-- Details are shown below the graph display
after the vertex label has been clicked. -/
after the vertex label has been clicked.
See also `Props.showDetails`. -/
details? : Option Html := none
deriving Inhabited, RpcEncodable

Expand All @@ -31,15 +35,16 @@ structure Edge where
source : String
/-- Target vertex. Must match the `id` of one of the vertices. -/
target : String
/-- Attributes to set on the SVG `<line>` element representing this edge. -/
attrs : Json := json% {
className: "dim",
fill: "var(--vscode-editor-foreground)",
stroke: "var(--vscode-editor-foreground)",
strokeWidth: 2
}
/-- Extra attributes to set on the SVG `<line>` element representing this edge.
See also `Props.defaultEdgeAttrs`. -/
attrs : Array (String × Json) := #[]
/-- If present, the label is shown over the edge midpoint.
This must be an SVG element.
Use `<foreignObject>` to draw non-SVG elements. -/
label? : Option Html := none
/-- Details are shown below the graph display
after the edge has been clicked. -/
after the edge has been clicked.
See also `Props.showDetails`. -/
details? : Option Html := none
deriving Inhabited, RpcEncodable

Expand Down Expand Up @@ -99,19 +104,29 @@ inductive ForceParams where

structure Props where
vertices : Array Vertex
/-- At most one edge may exist between any two vertices.
Self-loops are allowed,
but (TODO) are currently not rendered well. -/
edges : Array Edge
/-- Attributes to set by default on `<line>` elements representing edges. -/
defaultEdgeAttrs : Array (String × Json) := #[
("fill", "var(--vscode-editor-foreground)"),
("stroke", "var(--vscode-editor-foreground)"),
("strokeWidth", 2),
("markerEnd", "url(#arrow)")
]
/-- Which forces to apply to the vertices.
Most force parameters are optional, using default values if not specified. -/
forces : Array ForceParams := #[ .link {}, .manyBody {}, .x {}, .y {} ]
/-- Whether to show the details box below the graph. -/
/-- Whether to show a details box below the graph. -/
showDetails : Bool := false
deriving Inhabited, RpcEncodable

end DigraphDisplay
end GraphDisplay

/-- Display a directed graph with an interactive force simulation. -/
/-- Display a graph with an interactive force simulation. -/
@[widget_module]
def DigraphDisplay : Component DigraphDisplay.Props where
def GraphDisplay : Component GraphDisplay.Props where
javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "d3Graph.js"

end ProofWidgets
13 changes: 8 additions & 5 deletions ProofWidgets/Data/Html.lean
Original file line number Diff line number Diff line change
Expand Up @@ -46,10 +46,13 @@ declare_syntax_cat jsxAttr
declare_syntax_cat jsxAttrVal

scoped syntax str : jsxAttrVal
/-- Interpolates an expression into a JSX attribute literal -/
/-- Interpolates an expression into a JSX attribute literal. -/
scoped syntax group("{" term "}") : jsxAttrVal
scoped syntax ident "=" jsxAttrVal : jsxAttr
/-- Interpolates an array of expressions into a JSX attribute literal -/
/-- Interpolates a collection into a JSX attribute literal.
For HTML tags, this should have type `Array (String × Json)`.
For `ProofWidgets.Component`s, it can be any structure `$t` whatsoever:
it is interpolated into the component's props using `{ $t with ... }` notation. -/
scoped syntax group(" {..." term "}") : jsxAttr

/-- Characters not allowed inside JSX plain text. -/
Expand Down Expand Up @@ -148,16 +151,16 @@ def transformTag (tk : Syntax) (n m : Ident) (vs : Array (TSyntax `jsxAttr))
else
let vs ← joinArrays <| ← foldInlsM vs (fun vs' => do
let vs' ← vs'.mapM (fun (k, v) =>
`(term| ($(quote <| toString k.getId), $v)))
`(term| ($(quote <| toString k.getId), ($v : Json))))
`(term| #[$vs',*]))
`(Html.element $(quote tag) $vs $children)

/-- Support for writing HTML trees directly, using XML-like angle bracket syntax. It works very
similarly to [JSX](https://react.dev/learn/writing-markup-with-jsx) in JavaScript. The syntax is
enabled using `open scoped ProofWidgets.Jsx`.
Lowercase tags are interpreted as standard HTML whereas uppercase ones are expected to be
`ProofWidgets.Component`s. -/
Lowercase tags are interpreted as standard HTML
whereas uppercase ones are expected to be `ProofWidgets.Component`s. -/
macro_rules
| `(<$n:ident $[$attrs:jsxAttr]* />%$tk) => transformTag tk n n attrs #[]
| `(<$n:ident $[$attrs:jsxAttr]* >%$tk $cs*</$m>) => transformTag tk n m attrs cs
Expand Down
71 changes: 0 additions & 71 deletions ProofWidgets/Demos/Digraph.lean

This file was deleted.

8 changes: 4 additions & 4 deletions ProofWidgets/Demos/ExprGraph.lean
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
import ProofWidgets.Component.Digraph
import ProofWidgets.Component.GraphDisplay
import ProofWidgets.Component.HtmlDisplay
import Lean.Util.FoldConsts

Expand All @@ -20,7 +20,7 @@ def elabExprGraphCmd : CommandElab := fun
let some bb := env.find? b | continue
if bb.getUsedConstantsAsSet.contains a then
edges := edges.insert (toString a, toString b)
let mut nodesWithInfos : Array DigraphDisplay.Vertex := #[]
let mut nodesWithInfos : Array GraphDisplay.Vertex := #[]
let mut maxRadius := 10
for node in nodes.toArray do
let some c := env.find? node | continue
Expand All @@ -31,7 +31,7 @@ def elabExprGraphCmd : CommandElab := fun
let node := toString node
let rx := node.length * 3
maxRadius := Nat.max maxRadius rx
let newNode : DigraphDisplay.Vertex := {
let newNode : GraphDisplay.Vertex := {
id := node
label :=
<g>
Expand All @@ -57,7 +57,7 @@ def elabExprGraphCmd : CommandElab := fun
</span>
}
nodesWithInfos := nodesWithInfos.push newNode
let html : Html := <DigraphDisplay
let html : Html := <GraphDisplay
vertices={nodesWithInfos}
edges={edges.fold (init := #[]) fun acc (a,b) => acc.push {source := a, target := b}}
forces={#[
Expand Down
115 changes: 115 additions & 0 deletions ProofWidgets/Demos/Graph.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,115 @@
import ProofWidgets.Component.GraphDisplay
import ProofWidgets.Component.HtmlDisplay

/-! ## Directed graphs with `GraphDisplay` -/

open ProofWidgets Jsx

/-! ### Basic usage -/

def mkEdge (st : String × String) : GraphDisplay.Edge := {source := st.1, target := st.2}

-- Place your cursor here.
#html <GraphDisplay
vertices={#["a", "b", "c", "d", "e", "f"].map ({id := ·})}
edges={#[("b","c"), ("d","e"), ("e","f"), ("f","d")].map mkEdge}
/>

/-! ### Custom layout
Forces acting on the vertices can be customized
in order to control graph layout. -/

def complete (n : Nat) : Array GraphDisplay.Vertex × Array GraphDisplay.Edge := Id.run do
let mut verts := #[]
let mut edges := #[]
for i in [:n] do
verts := verts.push {id := toString i}
for j in [:i] do
edges := edges.push {source := toString i, target := toString j}
edges := edges.push {source := toString j, target := toString i}
return (verts, edges)

def K₁₀ := complete 10

#html <GraphDisplay
vertices={K₁₀.1}
edges={K₁₀.2}
-- Specify forces here.
forces={#[
.link { distance? := some 150 }
]}
/>

/-! ### Vertex labels
Arbitrary SVG elements can be used as vertex labels. -/

#html <GraphDisplay
vertices={#[
{ id := "a"
-- Specify a label here.
label := <circle r={5} fill="#ff0000" />
},
{ id := "b"
-- Use `<foreignObject>` to draw non-SVG elements.
label := <foreignObject height={50} width={50}>
-- TODO: the extra `<p>` node messes up positioning
<MarkdownDisplay contents="$xyz$" />
</foreignObject>
}
]}
edges={#[{ source := "a", target := "b" }]}
/>

/-! ### Edge labels
Arbitrary SVG elements can be used as edge labels. -/

#html <GraphDisplay
vertices={#["a", "b", "c"].map ({ id := ·})}
edges={#[
{ source := "a", target := "b"
-- Specify a label here.
label? := <g>
{GraphDisplay.mkCircle #[("r", "10")]}
<text textAnchor="middle" dominantBaseline="middle">1</text>
</g>
},
{ source := "b", target := "c"
-- Use `<foreignObject>` to draw non-SVG elements.
label? := <foreignObject height={50} width={50}>
-- TODO: the extra `<p>` node messes up positioning
<MarkdownDisplay contents="$e_2$" />
</foreignObject>
}
]}
forces={#[
.link { distance? := some 100 }
]}
/>

/-! ### Extra details
A details box with extra information can be displayed below the graph.
Click on vertices and edges to view their details. -/

#html <GraphDisplay
vertices={#[
{ id := "a"
-- Specify details here.
details? := Html.text "Vertex a."
-- Add class to indicate clickability.
label := GraphDisplay.mkCircle #[("className", "dim")]
},
{ id := "b" }
]}
edges={#[
{ source := "a", target := "b"
details? := Html.text "Edge a → b."
attrs := #[("className", "dim"), ("strokeWidth", 3)]
}
]}
-- Set this to display details.
showDetails={true}
/>
7 changes: 7 additions & 0 deletions RELEASES.md
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,13 @@
This file contains work-in-progress notes for the upcoming release, as well as previous releases.
Please check the [releases](https://github.com/leanprover-community/ProofWidgets4/releases) page for the build artifacts.

v0.0.45
-------

* Renamed `DigraphDisplay` to `GraphDisplay`.
Undirected graphs can be rendered by turning off arrowheads.
* Added support for edge labels and more customization to `GraphDisplay`.

v0.0.44
-------

Expand Down
Loading

0 comments on commit 81dfa90

Please sign in to comment.