├── lean-toolchain ├── widget ├── src │ ├── recharts.tsx │ ├── index.ts │ ├── htmlDisplayPanel.tsx │ ├── goalTypePanel.tsx │ ├── interactiveExpr.tsx │ ├── makeEditLink.tsx │ ├── filterDetails.tsx │ ├── animatedHtml.tsx │ ├── cancellable.ts │ ├── ofRpcMethod.tsx │ ├── presentSelection.tsx │ ├── rbTree.tsx │ ├── penroseDisplay.tsx │ ├── htmlDisplay.tsx │ ├── exprPresentation.tsx │ ├── interactiveSvg.tsx │ └── rubiks.tsx ├── penrose │ ├── setTheory.dsl │ ├── euclidean.dsl │ ├── venn.sty │ └── euclidean.sty ├── package-lock.json.trace ├── tsconfig.json ├── package.json └── rollup.config.js ├── doc └── infoview-rbtree.png ├── .gitignore ├── lake-manifest.json ├── test ├── checkh.lean └── delab.lean ├── .vscode └── settings.json ├── ProofWidgets ├── Component │ ├── Panel │ │ ├── GoalTypePanel.lean │ │ ├── SelectionPanel.lean │ │ └── Basic.lean │ ├── FilterDetails.lean │ ├── HtmlDisplay.lean │ ├── Recharts.lean │ ├── MakeEditLink.lean │ ├── OfRpcMethod.lean │ ├── PenroseDiagram.lean │ ├── InteractiveSvg.lean │ ├── Basic.lean │ └── GraphDisplay.lean ├── Demos │ ├── Rubiks.lean │ ├── ExprPresentation.lean │ ├── Macro.lean │ ├── Jsx.lean │ ├── Plot.lean │ ├── InteractiveSvg.lean │ ├── LazyComputation.lean │ ├── Dynkin.lean │ ├── Graph │ │ ├── ExprGraph.lean │ │ ├── Basic.lean │ │ └── MVarGraph.lean │ ├── Svg.lean │ ├── Venn.lean │ ├── RbTree.lean │ ├── SelectInsertConv.lean │ └── Euclidean.lean ├── Compat.lean ├── Util.lean ├── Extra │ └── CheckHighlight.lean ├── Presentation │ └── Expr.lean ├── Cancellable.lean └── Data │ ├── Svg.lean │ └── Html.lean ├── ProofWidgets.lean ├── .github └── workflows │ └── build.yml ├── RELEASES.md ├── lakefile.lean ├── README.md └── LICENSE /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.27.0-rc1 2 | -------------------------------------------------------------------------------- /widget/src/recharts.tsx: -------------------------------------------------------------------------------- 1 | export * from 'recharts'; 2 | -------------------------------------------------------------------------------- /doc/infoview-rbtree.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/leanprover-community/ProofWidgets4/HEAD/doc/infoview-rbtree.png -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | # JavaScript build artefacts. 3 | /widget/dist 4 | /widget/node_modules 5 | /widget/*.log.json 6 | /widget/*.hash 7 | .direnv/ 8 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": [], 4 | "name": "proofwidgets", 5 | "lakeDir": ".lake"} 6 | -------------------------------------------------------------------------------- /widget/penrose/setTheory.dsl: -------------------------------------------------------------------------------- 1 | type Set 2 | 3 | predicate Not(Prop p1) 4 | predicate Intersecting(Set s1, Set s2) 5 | predicate IsSubset(Set s1, Set s2) 6 | -------------------------------------------------------------------------------- /widget/src/index.ts: -------------------------------------------------------------------------------- 1 | import { Html, renderHtml, default as HtmlDisplay } from "./htmlDisplay" 2 | 3 | export type { Html } 4 | export { renderHtml, HtmlDisplay } 5 | -------------------------------------------------------------------------------- /test/checkh.lean: -------------------------------------------------------------------------------- 1 | import ProofWidgets.Extra.CheckHighlight 2 | 3 | /-- info: Decidable.em (p : Prop) : p ∨ ¬p -/ 4 | #guard_msgs in 5 | #checkh Decidable.em 6 | 7 | /-- info: Decidable.em (p : Prop) [Decidable p] : p ∨ ¬p -/ 8 | #guard_msgs in 9 | #checkh' Decidable.em 10 | -------------------------------------------------------------------------------- /.vscode/settings.json: -------------------------------------------------------------------------------- 1 | { 2 | "editor.insertSpaces": true, 3 | "editor.tabSize": 2, 4 | "editor.rulers" : [100], 5 | "files.encoding": "utf8", 6 | "files.eol": "\n", 7 | "files.insertFinalNewline": true, 8 | "files.trimFinalNewlines": true, 9 | "files.trimTrailingWhitespace": true, 10 | "search.usePCRE2": true 11 | } -------------------------------------------------------------------------------- /widget/src/htmlDisplayPanel.tsx: -------------------------------------------------------------------------------- 1 | import { DocumentPosition } from "@leanprover/infoview" 2 | import HtmlDisplay, { Html } from "./htmlDisplay" 3 | 4 | export default function HtmlDisplayPanel({html} : {pos: DocumentPosition, html: Html}): 5 | JSX.Element { 6 | return
7 | HTML Display 8 | 9 |
10 | } 11 | -------------------------------------------------------------------------------- /widget/penrose/euclidean.dsl: -------------------------------------------------------------------------------- 1 | type Object 2 | type Point <: Object 3 | type Line <: Object 4 | type Circle <: Object 5 | 6 | predicate Between(Point a, Point b, Point c) 7 | predicate OnLine(Point a, Line L) 8 | predicate OnCircle(Point a, Circle C) 9 | predicate InCircle(Point a, Circle C) 10 | predicate CenterCircle(Point a, Circle C) 11 | predicate CirclesInter(Circle C, Circle D) 12 | 13 | predicate Emphasize(Object o) 14 | -------------------------------------------------------------------------------- /widget/package-lock.json.trace: -------------------------------------------------------------------------------- 1 | {"log": 2 | [{"message": "././widget> npm install", "level": "trace"}, 3 | {"message": 4 | "stdout:\nadded 1 package, changed 2 packages, and audited 367 packages in 2s\n\n67 packages are looking for funding\n run `npm fund` for details\n\n1 moderate severity vulnerability\n\nTo address all issues, run:\n npm audit fix\n\nRun `npm audit` for details.", 5 | "level": "trace"}], 6 | "depHash": "254600175341942808"} -------------------------------------------------------------------------------- /ProofWidgets/Component/Panel/GoalTypePanel.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | public meta import ProofWidgets.Component.Panel.Basic 4 | 5 | public meta section 6 | 7 | namespace ProofWidgets 8 | 9 | /-- Display the goal type using known `Expr` presenters. -/ 10 | @[widget_module] 11 | def GoalTypePanel : Component PanelWidgetProps where 12 | javascript := include_str ".." / ".." / ".." / ".lake" / "build" / "js" / "goalTypePanel.js" 13 | 14 | end ProofWidgets 15 | -------------------------------------------------------------------------------- /ProofWidgets/Demos/Rubiks.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | public meta import ProofWidgets.Component.HtmlDisplay 4 | 5 | public meta section 6 | 7 | open Lean ProofWidgets 8 | open scoped ProofWidgets.Jsx 9 | 10 | structure RubiksProps where 11 | seq : Array String := #[] 12 | deriving ToJson, FromJson, Inhabited 13 | 14 | @[widget_module] 15 | def Rubiks : Component RubiksProps where 16 | javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "rubiks.js" 17 | 18 | def eg := #["L", "L", "D⁻¹", "U⁻¹", "L", "D", "D", "L", "U⁻¹", "R", "D", "F", "F", "D"] 19 | 20 | #html 21 | -------------------------------------------------------------------------------- /widget/src/goalTypePanel.tsx: -------------------------------------------------------------------------------- 1 | import { PanelWidgetProps } from "@leanprover/infoview"; 2 | import { GoalsLocationPresentation } from "./presentSelection"; 3 | 4 | export default function (props: PanelWidgetProps) { 5 | if (props.goals.length === 0) return <> 6 | const g = props.goals[0] 7 | if (!g.mvarId) 8 | throw new Error(`Lean server 1.1.2 or newer is required.`) 9 | return
10 | Main goal type 11 | 15 |
16 | } 17 | -------------------------------------------------------------------------------- /widget/src/interactiveExpr.tsx: -------------------------------------------------------------------------------- 1 | import { mapRpcError, InteractiveCode, RpcPtr, useAsyncPersistent, useRpcSession } 2 | from '@leanprover/infoview' 3 | 4 | type ExprWithCtx = RpcPtr<'ProofWidgets.ExprWithCtx'> 5 | 6 | export default function ({ expr }: { expr: ExprWithCtx }): JSX.Element { 7 | const rs = useRpcSession() 8 | const st = useAsyncPersistent(() => rs.call('ProofWidgets.ppExprTagged', { expr }), [expr]) 9 | if (st.state === 'resolved') 10 | return 11 | else if (st.state === 'rejected') 12 | return <>Error: ${mapRpcError(st.error).message} 13 | else 14 | return <>Loading.. 15 | } 16 | -------------------------------------------------------------------------------- /widget/tsconfig.json: -------------------------------------------------------------------------------- 1 | { 2 | "compilerOptions": { 3 | /* Language and environment */ 4 | "target": "esnext", 5 | "jsx": "react-jsx", 6 | 7 | /* Modules */ 8 | "module": "esnext", 9 | "moduleResolution": "bundler", 10 | 11 | /* Emit code */ 12 | "sourceMap": false, 13 | "outDir": "./dist/", 14 | 15 | /* Emit declarations */ 16 | "declaration": true, 17 | 18 | "allowSyntheticDefaultImports": true, 19 | 20 | /* Interop constraints */ 21 | "esModuleInterop": true, 22 | "forceConsistentCasingInFileNames": true, 23 | 24 | /* Typechecking */ 25 | "strict": true, 26 | "isolatedModules": true, 27 | 28 | /* Completeness */ 29 | "skipLibCheck": true 30 | }, 31 | "include": ["./src/*.ts", "./src/*.tsx"] 32 | } 33 | -------------------------------------------------------------------------------- /ProofWidgets/Compat.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | public meta import Lean.Elab.InfoTree.Main 4 | 5 | public meta section 6 | 7 | namespace ProofWidgets 8 | open Lean Server Elab 9 | 10 | abbrev LazyEncodable α := StateM RpcObjectStore α 11 | 12 | -- back from exile 13 | structure ExprWithCtx where 14 | ci : Elab.ContextInfo 15 | lctx : LocalContext 16 | linsts : LocalInstances 17 | expr : Expr 18 | deriving TypeName 19 | 20 | def ExprWithCtx.runMetaM (e : ExprWithCtx) (x : Expr → MetaM α) : IO α := 21 | e.ci.runMetaM {} $ 22 | Meta.withLCtx e.lctx e.linsts (x e.expr) 23 | 24 | def ExprWithCtx.save (e : Expr) : MetaM ExprWithCtx := 25 | return { 26 | ci := { ← CommandContextInfo.save with } 27 | lctx := ← getLCtx 28 | linsts := ← Meta.getLocalInstances 29 | expr := e 30 | } 31 | 32 | end ProofWidgets 33 | -------------------------------------------------------------------------------- /ProofWidgets/Demos/ExprPresentation.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | public meta import ProofWidgets.Component.Panel.SelectionPanel 4 | public meta import ProofWidgets.Component.Panel.GoalTypePanel 5 | 6 | public meta section 7 | 8 | open ProofWidgets Jsx 9 | 10 | @[expr_presenter] 11 | def presenter : ExprPresenter where 12 | userName := "With octopodes" 13 | layoutKind := .inline 14 | present e := 15 | return 16 | {.text "🐙 "}{.text " 🐙"} 17 | 18 | 19 | example : 2 + 2 = 4 ∧ 3 + 3 = 6 := by 20 | with_panel_widgets [GoalTypePanel] 21 | -- Place cursor here. 22 | constructor 23 | rfl 24 | rfl 25 | 26 | example (_h : 2 + 2 = 5) : 2 + 2 = 4 := by 27 | with_panel_widgets [SelectionPanel] 28 | -- Place cursor here and select subexpressions in the goal with shift-click. 29 | rfl 30 | -------------------------------------------------------------------------------- /widget/penrose/venn.sty: -------------------------------------------------------------------------------- 1 | forall Set x { 2 | x.icon = Circle { 3 | strokeWidth : 0 4 | } 5 | 6 | x.textBox = Rectangle { 7 | -- string : x.label 8 | -- fontSize : "25px" 9 | cornerRadius : 5 10 | width : 50 11 | height : 30 12 | } 13 | 14 | ensure contains(x.icon, x.textBox) 15 | --encourage sameCenter(x.textBox, x.icon) 16 | -- x.textLayering = x.textBox above x.icon 17 | } 18 | 19 | forall Set x; Set y 20 | where IsSubset(x, y) { 21 | ensure disjoint(y.textBox, x.icon, 10) 22 | ensure contains(y.icon, x.icon, 5) 23 | x.icon above y.icon 24 | } 25 | 26 | forall Set x; Set y 27 | where Not(Intersecting(x, y)) { 28 | ensure disjoint(x.icon, y.icon) 29 | } 30 | 31 | forall Set x; Set y 32 | where Intersecting(x, y) { 33 | ensure overlapping(x.icon, y.icon) 34 | ensure disjoint(y.textBox, x.icon) 35 | ensure disjoint(x.textBox, y.icon) 36 | } 37 | -------------------------------------------------------------------------------- /ProofWidgets.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | public import ProofWidgets.Compat 4 | public import ProofWidgets.Component.Basic 5 | public import ProofWidgets.Component.FilterDetails 6 | public import ProofWidgets.Component.GraphDisplay 7 | public import ProofWidgets.Component.HtmlDisplay 8 | public import ProofWidgets.Component.InteractiveSvg 9 | public import ProofWidgets.Component.MakeEditLink 10 | public import ProofWidgets.Component.OfRpcMethod 11 | public import ProofWidgets.Component.Panel.Basic 12 | public import ProofWidgets.Component.Panel.GoalTypePanel 13 | public import ProofWidgets.Component.Panel.SelectionPanel 14 | public import ProofWidgets.Component.PenroseDiagram 15 | public import ProofWidgets.Component.Recharts 16 | public import ProofWidgets.Data.Html 17 | public import ProofWidgets.Data.Svg 18 | public import ProofWidgets.Presentation.Expr 19 | public import ProofWidgets.Extra.CheckHighlight 20 | -------------------------------------------------------------------------------- /widget/src/makeEditLink.tsx: -------------------------------------------------------------------------------- 1 | import * as React from 'react' 2 | import { EditorContext } from '@leanprover/infoview' 3 | import { Range, TextDocumentEdit } from 'vscode-languageserver-protocol' 4 | 5 | interface MakeEditLinkProps { 6 | edit : TextDocumentEdit 7 | newSelection? : Range 8 | title? : string 9 | } 10 | 11 | export default function(props: React.PropsWithChildren) { 12 | const ec = React.useContext(EditorContext) 13 | 14 | return { 16 | await ec.api.applyEdit({ documentChanges: [props.edit] }) 17 | // TODO: https://github.com/leanprover/vscode-lean4/issues/225 18 | if (props.newSelection) 19 | await ec.revealLocation({ uri: props.edit.textDocument.uri, range: props.newSelection }) 20 | }} 21 | > 22 | {props.children} 23 | 24 | } 25 | -------------------------------------------------------------------------------- /ProofWidgets/Component/FilterDetails.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | public meta import ProofWidgets.Data.Html 4 | 5 | public meta section 6 | 7 | namespace ProofWidgets 8 | open Lean 9 | 10 | /-- Props for the `FilterDetails` component. -/ 11 | structure FilterDetailsProps where 12 | /-- Contents of the ``. -/ 13 | summary : Html 14 | /-- What is shown in the filtered state. -/ 15 | filtered : Html 16 | /-- What is shown in the non-filtered state. -/ 17 | all : Html 18 | /-- Whether to start in the filtered state. -/ 19 | initiallyFiltered : Bool := true 20 | deriving Server.RpcEncodable 21 | 22 | /-- The `FilterDetails` component is like a `
` HTML element, 23 | but also has a filter button 24 | that allows you to switch between filtered and unfiltered states. -/ 25 | @[widget_module] 26 | def FilterDetails : Component FilterDetailsProps where 27 | javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "filterDetails.js" 28 | 29 | end ProofWidgets 30 | -------------------------------------------------------------------------------- /widget/src/filterDetails.tsx: -------------------------------------------------------------------------------- 1 | import * as React from 'react'; 2 | import HtmlDisplay, { Html } from './htmlDisplay'; 3 | 4 | interface FilterDetailsProps { 5 | summary : Html 6 | filtered : Html 7 | all : Html 8 | initiallyFiltered : boolean 9 | } 10 | 11 | export default function FilterDetails(props: FilterDetailsProps) { 12 | const [isFiltered, setFiltered] = React.useState(props.initiallyFiltered); 13 | 14 | return
15 | 16 | 17 | { e.preventDefault() }}> 18 | { setFiltered(s => !s) }} /> 22 | 23 | 24 | 25 |
26 | } 27 | -------------------------------------------------------------------------------- /widget/src/animatedHtml.tsx: -------------------------------------------------------------------------------- 1 | import * as React from 'react'; 2 | import { Html, default as HtmlDisplay } from './htmlDisplay'; 3 | 4 | interface AnimatedHtmlProps { 5 | frames: Html[] 6 | framesPerSecond?: number 7 | } 8 | 9 | /** Display the series of HTML frames as a time-dependent animation. This will only look reasonable 10 | * if the frames are related somehow, e.g. are all instances of a single component with slightly 11 | * different props. */ 12 | export default function AnimatedHtml(props: AnimatedHtmlProps) { 13 | const { frames } = props 14 | const framesPerSecond = props.framesPerSecond ?? 10 15 | if (framesPerSecond <= 0 || framesPerSecond > 60) { 16 | throw new Error(`Invalid fps ${framesPerSecond}. Should be between 0 and 60.`) 17 | } 18 | const [t, setT] = React.useState(0) 19 | React.useEffect(() => { 20 | const interval = window.setInterval(() => setT(t => t + 1), 1000.0 / framesPerSecond) 21 | return () => window.clearInterval(interval) 22 | }, [framesPerSecond]) 23 | 24 | const frame = frames[t % frames.length] 25 | 26 | return 27 | } 28 | -------------------------------------------------------------------------------- /ProofWidgets/Demos/Macro.lean: -------------------------------------------------------------------------------- 1 | module 2 | 3 | public meta import ProofWidgets.Component.HtmlDisplay 4 | 5 | public meta section 6 | 7 | -- See the `Jsx.lean` demo for more about JSX. 8 | open scoped ProofWidgets.Jsx 9 | 10 | /-! # Widgets in macros 11 | 12 | Macros may expand to commands or terms that end up saving widget info, 13 | but the syntax to which the widget info is associated must be marked as `canonical` 14 | for the widget to be displayed. 15 | -/ 16 | 17 | def Lean.SourceInfo.mkCanonical : SourceInfo → SourceInfo 18 | | .synthetic s e _ => .synthetic s e true 19 | | si => si 20 | 21 | def Lean.Syntax.mkInfoCanonical : Syntax → Syntax 22 | | .missing => .missing 23 | | .node i k a => .node i.mkCanonical k a 24 | | .atom i v => .atom i.mkCanonical v 25 | | .ident i r v p => .ident i.mkCanonical r v p 26 | 27 | def Lean.TSyntax.mkInfoCanonical : TSyntax k → TSyntax k := 28 | (.mk ·.raw.mkInfoCanonical) 29 | 30 | macro "#browse " src:term : command => 31 | Lean.TSyntax.mkInfoCanonical <$> 32 | `(#html