├── .github └── workflows │ └── build.yml ├── .gitignore ├── .vscode └── settings.json ├── LICENSE ├── ProofWidgets.lean ├── ProofWidgets ├── Cancellable.lean ├── Compat.lean ├── Component │ ├── Basic.lean │ ├── FilterDetails.lean │ ├── GraphDisplay.lean │ ├── HtmlDisplay.lean │ ├── InteractiveSvg.lean │ ├── MakeEditLink.lean │ ├── OfRpcMethod.lean │ ├── Panel │ │ ├── Basic.lean │ │ ├── GoalTypePanel.lean │ │ └── SelectionPanel.lean │ ├── PenroseDiagram.lean │ └── Recharts.lean ├── Data │ ├── Html.lean │ └── Svg.lean ├── Demos │ ├── Conv.lean │ ├── Dynkin.lean │ ├── Euclidean.lean │ ├── ExprPresentation.lean │ ├── Graph │ │ ├── Basic.lean │ │ ├── ExprGraph.lean │ │ └── MVarGraph.lean │ ├── InteractiveSvg.lean │ ├── Jsx.lean │ ├── LazyComputation.lean │ ├── Macro.lean │ ├── Plot.lean │ ├── RbTree.lean │ ├── Rubiks.lean │ ├── SelectInsertConv.lean │ ├── Svg.lean │ └── Venn.lean ├── Presentation │ └── Expr.lean └── Util.lean ├── README.md ├── RELEASES.md ├── doc └── infoview-rbtree.png ├── lake-manifest.json ├── lakefile.lean ├── lean-toolchain ├── test └── delab.lean └── widget ├── package-lock.json ├── package-lock.json.hash ├── package-lock.json.trace ├── package.json ├── penrose ├── euclidean.dsl ├── euclidean.sty ├── setTheory.dsl └── venn.sty ├── rollup.config.js ├── src ├── animatedHtml.tsx ├── cancellable.ts ├── d3Graph.tsx ├── exprPresentation.tsx ├── filterDetails.tsx ├── goalTypePanel.tsx ├── htmlDisplay.tsx ├── htmlDisplayPanel.tsx ├── index.ts ├── interactiveExpr.tsx ├── interactiveSvg.tsx ├── makeEditLink.tsx ├── ofRpcMethod.tsx ├── penroseCanvas.tsx ├── penroseDisplay.tsx ├── presentSelection.tsx ├── rbTree.tsx ├── recharts.tsx └── rubiks.tsx └── tsconfig.json /.github/workflows/build.yml: -------------------------------------------------------------------------------- 1 | on: 2 | push: 3 | pull_request: 4 | 5 | name: ci 6 | 7 | jobs: 8 | build: 9 | name: ${{ matrix.name }} 10 | runs-on: ${{ matrix.os }} 11 | defaults: 12 | run: 13 | shell: ${{ matrix.shell || 'sh' }} 14 | strategy: 15 | matrix: 16 | include: 17 | - name: Build on Ubuntu 18 | os: ubuntu-latest 19 | # TODO: investigate why `lake build` deadlocks on macOS (issue #33) 20 | #- name: Build on macOS 21 | # os: macos-latest 22 | - name: Build on Windows 23 | os: windows-latest 24 | shell: msys2 {0} 25 | steps: 26 | - name: Install MSYS2 (Windows) 27 | if: matrix.os == 'windows-latest' 28 | uses: msys2/setup-msys2@v2 29 | with: 30 | path-type: inherit 31 | install: curl unzip 32 | 33 | - name: Install elan 34 | shell: bash 35 | run: | 36 | curl -sSfLO https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh 37 | chmod +x elan-init.sh 38 | ./elan-init.sh -y --default-toolchain none 39 | echo "$HOME/.elan/bin" >> $GITHUB_PATH 40 | 41 | - uses: actions/checkout@v4 42 | 43 | - name: Update NPM (Windows) 44 | shell: bash 45 | if: matrix.os == 'windows-latest' 46 | # See https://github.com/nodejs/node/issues/52682 47 | run: cd widget/ && npm install npm 48 | 49 | - name: Build package 50 | run: lake build ProofWidgets 51 | 52 | - name: Try publishing @leanprover-community/proofwidgets4 53 | if: ${{ startsWith(github.ref, 'refs/tags/v') && !contains(github.ref, '-pre') && matrix.os == 'ubuntu-latest' }} 54 | continue-on-error: true 55 | run: cd widget/ && npm publish --access=public 56 | 57 | - name: Build demos 58 | run: lake build ProofWidgets.Demos 59 | 60 | - name: Create GitHub release for tag (Ubuntu) 61 | if: github.ref_type == 'tag' && matrix.os == 'ubuntu-latest' 62 | uses: softprops/action-gh-release@v2 63 | with: 64 | prerelease: ${{ contains(github.ref, '-pre') }} 65 | 66 | - name: Upload release archive (Ubuntu) 67 | if: github.ref_type == 'tag' && matrix.os == 'ubuntu-latest' 68 | run: lake upload $RELEASE_TAG 69 | env: 70 | GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} 71 | RELEASE_TAG: ${{ github.ref_name }} 72 | 73 | - name: Test release archive (Ubuntu) 74 | if: github.ref_type == 'tag' && matrix.os == 'ubuntu-latest' 75 | run: lake clean && lake build :release && lake build --no-build 76 | env: 77 | NPM_TOKEN: ${{ secrets.NPM_TOKEN }} 78 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | # JavaScript build artefacts. 3 | /widget/dist 4 | /widget/node_modules 5 | /widget/*.log.json 6 | .direnv/ 7 | -------------------------------------------------------------------------------- /.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 | } -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | APPENDIX: How to apply the Apache License to your work. 179 | 180 | To apply the Apache License to your work, attach the following 181 | boilerplate notice, with the fields enclosed by brackets "[]" 182 | replaced with your own identifying information. (Don't include 183 | the brackets!) The text should be enclosed in the appropriate 184 | comment syntax for the file format. We also recommend that a 185 | file or class name and description of purpose be included on the 186 | same "printed page" as the copyright notice for easier 187 | identification within third-party archives. 188 | 189 | Copyright [yyyy] [name of copyright owner] 190 | 191 | Licensed under the Apache License, Version 2.0 (the "License"); 192 | you may not use this file except in compliance with the License. 193 | You may obtain a copy of the License at 194 | 195 | http://www.apache.org/licenses/LICENSE-2.0 196 | 197 | Unless required by applicable law or agreed to in writing, software 198 | distributed under the License is distributed on an "AS IS" BASIS, 199 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 200 | See the License for the specific language governing permissions and 201 | limitations under the License. 202 | -------------------------------------------------------------------------------- /ProofWidgets.lean: -------------------------------------------------------------------------------- 1 | import ProofWidgets.Compat 2 | import ProofWidgets.Component.Basic 3 | import ProofWidgets.Component.FilterDetails 4 | import ProofWidgets.Component.GraphDisplay 5 | import ProofWidgets.Component.HtmlDisplay 6 | import ProofWidgets.Component.InteractiveSvg 7 | import ProofWidgets.Component.MakeEditLink 8 | import ProofWidgets.Component.OfRpcMethod 9 | import ProofWidgets.Component.Panel.Basic 10 | import ProofWidgets.Component.Panel.GoalTypePanel 11 | import ProofWidgets.Component.Panel.SelectionPanel 12 | import ProofWidgets.Component.PenroseDiagram 13 | import ProofWidgets.Component.Recharts 14 | import ProofWidgets.Data.Html 15 | import ProofWidgets.Data.Svg 16 | import ProofWidgets.Presentation.Expr 17 | -------------------------------------------------------------------------------- /ProofWidgets/Cancellable.lean: -------------------------------------------------------------------------------- 1 | import Lean.Data.Json.FromToJson 2 | import Lean.Server.Rpc.RequestHandling 3 | import Std.Data.HashMap 4 | import ProofWidgets.Compat 5 | 6 | /-! Experimental support for cancellable RPC requests. 7 | 8 | Note: Cancellation should eventually become a feature of the core RPC protocol, 9 | and the requests map should be stored in `RequestM`, 10 | or somewhere in the server anyway. -/ 11 | 12 | namespace ProofWidgets 13 | open Lean Server Meta Std 14 | 15 | abbrev RequestId := Nat 16 | structure CancellableTask where 17 | task : Task (Except RequestError (LazyEncodable Json)) 18 | /- Note: we cannot just `IO.cancel task` because it is a result of `map`. 19 | See https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/Should.20cancelling.20a.20purely.20mapped.20task.20cancel.20the.20original.3F -/ 20 | cancel : IO Unit 21 | 22 | /-- Maps the ID of each currently executing request to its task. -/ 23 | initialize runningRequests : 24 | IO.Ref (RequestId × Std.HashMap RequestId CancellableTask) ← 25 | IO.mkRef (0, ∅) 26 | 27 | /-- Transforms a request handler returning `β` 28 | into one that returns immediately with a `RequestId`. 29 | 30 | The ID uniquely identifies the running request: 31 | its results can be retrieved using `checkRequest`, 32 | and it can be cancelled using `cancelRequest`. -/ 33 | def mkCancellable [RpcEncodable β] (handler : α → RequestM (RequestTask β)) : 34 | α → RequestM (RequestTask RequestId) := fun a => do 35 | RequestM.asTask do 36 | let t ← handler a 37 | let t' := t.mapCheap (·.map rpcEncode) 38 | runningRequests.modifyGet fun (id, m) => 39 | (id, (id+1, m.insert id ⟨t'.task, t.cancel⟩)) 40 | 41 | /-- Cancel the request with ID `rid`. 42 | Does nothing if `rid` is invalid. -/ 43 | @[server_rpc_method] 44 | def cancelRequest (rid : RequestId) : RequestM (RequestTask String) := do 45 | RequestM.asTask do 46 | let t? ← runningRequests.modifyGet fun (id, m) => (m[rid]?, (id, m.erase rid)) 47 | if let some t := t? then 48 | t.cancel 49 | return "ok" 50 | 51 | /-- The status of a running cancellable request. -/ 52 | inductive CheckRequestResponse 53 | | running 54 | | done (result : LazyEncodable Json) 55 | deriving RpcEncodable 56 | 57 | /-- Check whether a request has finished computing, 58 | and return the response if so. 59 | The request is removed from `runningRequests` the first time it is checked 60 | and found to have finished. 61 | Throws an error if the `rid` is invalid, 62 | or if the request itself threw an error. -/ 63 | /- NOTE: a notification-based version would be better than this polling-based one. 64 | But we cannot include RPC references in notifications atm; 65 | another possible addition to the RPC protocol? -/ 66 | @[server_rpc_method] 67 | def checkRequest (rid : RequestId) : RequestM (RequestTask CheckRequestResponse) := do 68 | RequestM.asTask do 69 | let (_, m) ← runningRequests.get 70 | match m[rid]? with 71 | | none => 72 | throw $ RequestError.invalidParams 73 | s!"Request '{rid}' has already finished, or the ID is invalid." 74 | | some t => 75 | if !(← IO.hasFinished t.task) then 76 | return .running 77 | runningRequests.modify fun (id, m) => (id, m.erase rid) 78 | match t.task.get with 79 | | .error e => throw e 80 | | .ok v => return .done v 81 | 82 | def cancellableSuffix : Name := `_cancellable 83 | 84 | /-- Like `server_rpc_method`, but requests for this method can be cancelled. 85 | The method should check for that using `IO.checkCanceled`. 86 | Cancellable methods are invoked differently from JavaScript: 87 | see `callCancellable` in `cancellable.ts`. -/ 88 | initialize 89 | registerBuiltinAttribute { 90 | name := `server_rpc_method_cancellable 91 | descr := "Like `server_rpc_method`, \ 92 | but requests for this method can be cancelled. \ 93 | The method should check for that using `IO.checkCanceled`. \ 94 | Cancellable methods are invoked differently from JavaScript: \ 95 | see `callCancellable` in `cancellable.ts`." 96 | applicationTime := AttributeApplicationTime.afterCompilation 97 | add := fun decl _ _ => Prod.fst <$> MetaM.run do 98 | let name := decl ++ cancellableSuffix 99 | let value ← mkAppM ``mkCancellable #[mkConst decl] 100 | addAndCompile $ .defnDecl { 101 | name 102 | levelParams := [] 103 | type := ← inferType value 104 | value 105 | hints := .opaque 106 | safety := .safe 107 | } 108 | registerRpcProcedure name 109 | } 110 | 111 | end ProofWidgets 112 | -------------------------------------------------------------------------------- /ProofWidgets/Compat.lean: -------------------------------------------------------------------------------- 1 | import Lean.Elab.InfoTree.Main 2 | 3 | namespace ProofWidgets 4 | open Lean Server Elab 5 | 6 | abbrev LazyEncodable α := StateM RpcObjectStore α 7 | 8 | -- back from exile 9 | structure ExprWithCtx where 10 | ci : Elab.ContextInfo 11 | lctx : LocalContext 12 | linsts : LocalInstances 13 | expr : Expr 14 | deriving TypeName 15 | 16 | def ExprWithCtx.runMetaM (e : ExprWithCtx) (x : Expr → MetaM α) : IO α := 17 | e.ci.runMetaM {} $ 18 | Meta.withLCtx e.lctx e.linsts (x e.expr) 19 | 20 | def ExprWithCtx.save (e : Expr) : MetaM ExprWithCtx := 21 | return { 22 | ci := { ← CommandContextInfo.save with } 23 | lctx := ← getLCtx 24 | linsts := ← Meta.getLocalInstances 25 | expr := e 26 | } 27 | 28 | end ProofWidgets 29 | -------------------------------------------------------------------------------- /ProofWidgets/Component/Basic.lean: -------------------------------------------------------------------------------- 1 | import Lean.Widget.InteractiveCode 2 | import Lean.Widget.UserWidget 3 | import ProofWidgets.Compat 4 | 5 | namespace ProofWidgets 6 | open Lean 7 | 8 | /-- A component is a widget module with a `default` or named export which is a 9 | [React component](https://react.dev/learn/your-first-component). Every component definition must 10 | be annotated with `@[widget_module]`. This makes it possible for the infoview to load the component. 11 | 12 | ## Execution environment 13 | 14 | The JS environment in which components execute provides a fixed set of libraries accessible via 15 | direct `import`, notably 16 | [`@leanprover/infoview`](https://www.npmjs.com/package/@leanprover/infoview). 17 | All [React contexts](https://react.dev/learn/passing-data-deeply-with-context) exported from 18 | `@leanprover/infoview` are usable from components. 19 | 20 | ## Lean encoding of props 21 | 22 | `Props` is the Lean representation of the type `JsProps` of 23 | [React props](https://react.dev/learn/passing-props-to-a-component) 24 | that the component expects. 25 | The export of the module specified in `«export»` should then have type 26 | `(props: JsProps & { pos: DocumentPosition }): React.ReactNode` 27 | where `DocumentPosition` is defined in `@leanprover/infoview`. 28 | `Props` is expected to have a `Lean.Server.RpcEncodable` instance 29 | specifying how to encode props as JSON. 30 | 31 | Note that by defining a `Component Props` with a specific JS implementation, 32 | you are *asserting* that `Props` is a correct representation of `JsProps`. -/ 33 | structure Component (Props : Type) extends Widget.Module where 34 | /-- Which export of the module to use as the component function. -/ 35 | «export» : String := "default" 36 | 37 | instance : Widget.ToModule (Component Props) := ⟨Component.toModule⟩ 38 | 39 | structure InteractiveCodeProps where 40 | fmt : Widget.CodeWithInfos 41 | deriving Server.RpcEncodable 42 | 43 | /-- Present pretty-printed code as interactive text. 44 | 45 | The most common use case is to instantiate this component from a `Lean.Expr`. To do so, you must 46 | eagerly pretty-print the `Expr` using `Widget.ppExprTagged`. See also `InteractiveExpr`. -/ 47 | @[widget_module] 48 | def InteractiveCode : Component InteractiveCodeProps where 49 | javascript := " 50 | import { InteractiveCode } from '@leanprover/infoview' 51 | import * as React from 'react' 52 | export default function(props) { 53 | return React.createElement(InteractiveCode, props) 54 | }" 55 | 56 | structure InteractiveExprProps where 57 | expr : Server.WithRpcRef ExprWithCtx 58 | deriving Server.RpcEncodable 59 | 60 | @[server_rpc_method] 61 | def ppExprTagged : InteractiveExprProps → Server.RequestM (Server.RequestTask Widget.CodeWithInfos) 62 | | ⟨⟨expr⟩⟩ => Server.RequestM.asTask <| expr.runMetaM Widget.ppExprTagged 63 | 64 | /-- Lazily pretty-print and present a `Lean.Expr` as interactive text. 65 | 66 | This component is preferrable over `InteractiveCode` when the `Expr` will not necessarily be 67 | displayed in the UI (e.g. it may be hidden by default), in which case laziness saves some work. 68 | On the other hand if the `Expr` will likely be shown and you are in a `MetaM` context, it is 69 | preferrable to use the eager `InteractiveCode` in order to avoid the extra client-server roundtrip 70 | needed for the pretty-printing RPC call. -/ 71 | @[widget_module] 72 | def InteractiveExpr : Component InteractiveExprProps where 73 | javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "interactiveExpr.js" 74 | 75 | structure InteractiveMessageProps where 76 | msg : Server.WithRpcRef MessageData 77 | deriving Server.RpcEncodable 78 | 79 | /-- Present a structured Lean message. -/ 80 | @[widget_module] 81 | def InteractiveMessage : Component InteractiveMessageProps where 82 | javascript := " 83 | import { InteractiveMessageData } from '@leanprover/infoview' 84 | import * as React from 'react' 85 | export default function(props) { 86 | return React.createElement(InteractiveMessageData, props) 87 | } 88 | " 89 | 90 | structure MarkdownDisplay.Props where 91 | contents : String 92 | deriving ToJson, FromJson 93 | 94 | /-- Render a given string as Markdown. 95 | LaTeX is supported with MathJax: 96 | use `$...$` for inline math, 97 | and `$$...$$` for displayed math. 98 | 99 | Example usage: 100 | ```lean 101 | 102 | ``` -/ 103 | @[widget_module] 104 | def MarkdownDisplay : Component MarkdownDisplay.Props where 105 | javascript := " 106 | import { Markdown } from '@leanprover/infoview' 107 | import * as React from 'react' 108 | export default (props) => React.createElement(Markdown, props) 109 | " 110 | 111 | end ProofWidgets 112 | 113 | /-- Construct a structured message from a ProofWidgets component. 114 | 115 | For the meaning of `alt`, see `MessageData.ofWidget`. -/ 116 | def Lean.MessageData.ofComponent [Server.RpcEncodable Props] 117 | (c : ProofWidgets.Component Props) (p : Props) (alt : String) : CoreM MessageData := do 118 | let wi ← Widget.WidgetInstance.ofHash c.javascriptHash 119 | (Server.RpcEncodable.rpcEncode p) 120 | return .ofWidget wi alt 121 | -------------------------------------------------------------------------------- /ProofWidgets/Component/FilterDetails.lean: -------------------------------------------------------------------------------- 1 | import ProofWidgets.Data.Html 2 | 3 | namespace ProofWidgets 4 | open Lean 5 | 6 | /-- Props for the `FilterDetails` component. -/ 7 | structure FilterDetailsProps where 8 | /-- Contents of the ``. -/ 9 | summary : Html 10 | /-- What is shown in the filtered state. -/ 11 | filtered : Html 12 | /-- What is shown in the non-filtered state. -/ 13 | all : Html 14 | /-- Whether to start in the filtered state. -/ 15 | initiallyFiltered : Bool := true 16 | deriving Server.RpcEncodable 17 | 18 | /-- The `FilterDetails` component is like a `
` HTML element, 19 | but also has a filter button 20 | that allows you to switch between filtered and unfiltered states. -/ 21 | @[widget_module] 22 | def FilterDetails : Component FilterDetailsProps where 23 | javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "filterDetails.js" 24 | 25 | end ProofWidgets 26 | -------------------------------------------------------------------------------- /ProofWidgets/Component/GraphDisplay.lean: -------------------------------------------------------------------------------- 1 | import ProofWidgets.Component.Basic 2 | import ProofWidgets.Data.Html 3 | 4 | namespace ProofWidgets.GraphDisplay 5 | open Lean Server Jsx 6 | 7 | /-- A themed `` SVG element, with optional extra attributes. -/ 8 | def mkCircle (attrs : Array (String × Json) := #[]) : Html := 9 | 16 | 17 | /-- A shape containing the vertex label. 18 | Used to position incident edge endpoints. 19 | The shape is assumed to be centred on the vertex position. -/ 20 | -- TODO: use `getBoundingClientRect` to dynamically compute size 21 | inductive BoundingShape where 22 | /-- A circle of fixed radius. -/ 23 | | circle (radius : Float) : BoundingShape 24 | /-- A rectangle of fixed dimensions. -/ 25 | | rect (width height : Float) : BoundingShape 26 | deriving Inhabited, FromJson, ToJson 27 | 28 | structure Vertex where 29 | /-- Identifier for this vertex. Must be unique. -/ 30 | id : String 31 | /-- The label is drawn at the vertex position. 32 | This must be an SVG element. 33 | Use `` to draw non-SVG elements. -/ 34 | label : Html := mkCircle 35 | boundingShape : BoundingShape := .circle 5 36 | /-- Details are shown below the graph display 37 | after the vertex label has been clicked. 38 | See also `Props.showDetails`. -/ 39 | details? : Option Html := none 40 | deriving Inhabited, RpcEncodable 41 | 42 | structure Edge where 43 | /-- Source vertex. Must match the `id` of one of the vertices. -/ 44 | source : String 45 | /-- Target vertex. Must match the `id` of one of the vertices. -/ 46 | target : String 47 | /-- Extra attributes to set on the SVG `` element representing this edge. 48 | See also `Props.defaultEdgeAttrs`. -/ 49 | attrs : Array (String × Json) := #[] 50 | /-- If present, the label is shown over the edge midpoint. 51 | This must be an SVG element. 52 | Use `` to draw non-SVG elements. -/ 53 | label? : Option Html := none 54 | /-- Details are shown below the graph display 55 | after the edge has been clicked. 56 | See also `Props.showDetails`. -/ 57 | details? : Option Html := none 58 | deriving Inhabited, RpcEncodable 59 | 60 | structure ForceCenterParams where 61 | x? : Option Float := none 62 | y? : Option Float := none 63 | strength? : Option Float := none 64 | deriving Inhabited, FromJson, ToJson 65 | 66 | structure ForceCollideParams where 67 | radius? : Option Float := none 68 | strength? : Option Float := none 69 | iterations? : Option Nat := none 70 | deriving Inhabited, FromJson, ToJson 71 | 72 | structure ForceLinkParams where 73 | distance? : Option Float := none 74 | strength? : Option Float := none 75 | iterations? : Option Nat := none 76 | deriving Inhabited, FromJson, ToJson 77 | 78 | structure ForceManyBodyParams where 79 | strength? : Option Float := none 80 | theta? : Option Float := none 81 | distanceMin? : Option Float := none 82 | distanceMax? : Option Float := none 83 | deriving Inhabited, FromJson, ToJson 84 | 85 | structure ForceXParams where 86 | x? : Option Float := none 87 | strength? : Option Float := none 88 | deriving Inhabited, FromJson, ToJson 89 | 90 | structure ForceYParams where 91 | y? : Option Float := none 92 | strength? : Option Float := none 93 | deriving Inhabited, FromJson, ToJson 94 | 95 | structure ForceRadialParams where 96 | radius : Float 97 | x? : Option Float := none 98 | y? : Option Float := none 99 | strength? : Option Float := none 100 | deriving Inhabited, FromJson, ToJson 101 | 102 | /-- Settings for the simulation of forces on vertices. 103 | See https://d3js.org/d3-force. -/ 104 | inductive ForceParams where 105 | | center : ForceCenterParams → ForceParams 106 | | collide : ForceCollideParams → ForceParams 107 | | link : ForceLinkParams → ForceParams 108 | | manyBody : ForceManyBodyParams → ForceParams 109 | | x : ForceXParams → ForceParams 110 | | y : ForceYParams → ForceParams 111 | | radial : ForceRadialParams → ForceParams 112 | deriving Inhabited, FromJson, ToJson 113 | 114 | structure Props where 115 | vertices : Array Vertex 116 | /-- At most one edge may exist between any two vertices. 117 | Self-loops are allowed, 118 | but (TODO) are currently not rendered well. -/ 119 | edges : Array Edge 120 | /-- Attributes to set by default on `` elements representing edges. -/ 121 | defaultEdgeAttrs : Array (String × Json) := #[ 122 | ("fill", "var(--vscode-editor-foreground)"), 123 | ("stroke", "var(--vscode-editor-foreground)"), 124 | ("strokeWidth", 2), 125 | ("markerEnd", "url(#arrow)") 126 | ] 127 | /-- Which forces to apply to the vertices. 128 | Most force parameters are optional, using default values if not specified. -/ 129 | forces : Array ForceParams := #[ .link {}, .manyBody {}, .x {}, .y {} ] 130 | /-- Whether to show a details box below the graph. -/ 131 | showDetails : Bool := false 132 | deriving Inhabited, RpcEncodable 133 | 134 | end GraphDisplay 135 | 136 | /-- Display a graph with an interactive force simulation. -/ 137 | @[widget_module] 138 | def GraphDisplay : Component GraphDisplay.Props where 139 | javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "d3Graph.js" 140 | 141 | end ProofWidgets 142 | -------------------------------------------------------------------------------- /ProofWidgets/Component/HtmlDisplay.lean: -------------------------------------------------------------------------------- 1 | import Lean.Server.Rpc.Basic 2 | import Lean.Elab.Command 3 | 4 | import ProofWidgets.Data.Html 5 | 6 | namespace ProofWidgets 7 | open Lean Server 8 | 9 | structure HtmlDisplayProps where 10 | html : Html 11 | deriving RpcEncodable 12 | 13 | @[widget_module] 14 | def HtmlDisplay : Component HtmlDisplayProps where 15 | javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "htmlDisplay.js" 16 | 17 | @[widget_module] 18 | def HtmlDisplayPanel : Component HtmlDisplayProps where 19 | javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "htmlDisplayPanel.js" 20 | 21 | open Lean Server Elab Command 22 | 23 | /-- Any term `t : α` with a `HtmlEval α` instance 24 | can be evaluated in a `#html t` command. 25 | 26 | This is analogous to how `Lean.MetaEval` supports `#eval`. -/ 27 | class HtmlEval (α : Type u) where 28 | eval : α → CommandElabM Html 29 | 30 | instance : HtmlEval Html where 31 | eval ht := pure ht 32 | 33 | instance [MonadLiftT m CommandElabM] : HtmlEval (m Html) where 34 | eval := monadLift 35 | 36 | instance : HtmlEval (CoreM Html) where 37 | eval := liftCoreM 38 | 39 | instance : HtmlEval (MetaM Html) where 40 | eval x := liftTermElabM x 41 | 42 | instance : HtmlEval (TermElabM Html) where 43 | eval := liftTermElabM 44 | 45 | namespace HtmlCommand 46 | 47 | open Elab Command 48 | 49 | unsafe def evalCommandMHtmlUnsafe (stx : Term) : TermElabM (CommandElabM Html) := do 50 | let tp := mkApp (mkConst ``CommandElabM) (mkConst ``Html) 51 | Term.evalTerm _ tp stx 52 | 53 | @[implemented_by evalCommandMHtmlUnsafe] 54 | opaque evalCommandMHtml : Term → TermElabM (CommandElabM Html) 55 | 56 | /-- Display a value of type `Html` in the infoview. 57 | 58 | The input can be a pure value 59 | or a computation in any Lean metaprogramming monad 60 | (e.g. `CommandElabM Html`). -/ 61 | syntax (name := htmlCmd) "#html " term : command 62 | 63 | @[command_elab htmlCmd] 64 | def elabHtmlCmd : CommandElab := fun 65 | | stx@`(#html $t:term) => do 66 | let htX ← liftTermElabM <| evalCommandMHtml <| ← ``(HtmlEval.eval $t) 67 | let ht ← htX 68 | liftCoreM <| Widget.savePanelWidgetInfo 69 | (hash HtmlDisplayPanel.javascript) 70 | (return json% { html: $(← rpcEncode ht) }) 71 | stx 72 | | stx => throwError "Unexpected syntax {stx}." 73 | 74 | end HtmlCommand 75 | end ProofWidgets 76 | 77 | /-- Construct a structured message from ProofWidgets HTML. 78 | 79 | For the meaning of `alt`, see `MessageData.ofWidget`. -/ 80 | def Lean.MessageData.ofHtml (h : ProofWidgets.Html) (alt : String) : CoreM MessageData := 81 | MessageData.ofComponent ProofWidgets.HtmlDisplay ⟨h⟩ alt 82 | -------------------------------------------------------------------------------- /ProofWidgets/Component/InteractiveSvg.lean: -------------------------------------------------------------------------------- 1 | import ProofWidgets.Data.Svg 2 | 3 | namespace ProofWidgets 4 | open Lean Std 5 | 6 | private def _root_.Float.toInt (x : Float) : Int := 7 | if x >= 0 then 8 | x.toUInt64.toNat 9 | else 10 | -((-x).toUInt64.toNat) 11 | 12 | namespace Svg 13 | 14 | inductive ActionKind where 15 | | timeout 16 | | mousedown 17 | | mouseup 18 | | mousemove -- [note] mouse moves only happen when mouse button is down. 19 | deriving ToJson, FromJson, DecidableEq 20 | 21 | structure Action where 22 | kind : ActionKind 23 | id : Option String 24 | data : Option Json 25 | deriving ToJson, FromJson 26 | 27 | /-- The input type `State` is any state the user wants to use and update 28 | 29 | SvgState in addition automatically handles tracking of time, selection and custom data -/ 30 | structure SvgState (State : Type) where 31 | state : State 32 | time : Float /-- time in milliseconds -/ 33 | selected : Option String 34 | mousePos : Option (Int × Int) 35 | idToData : List (String × Json) 36 | deriving ToJson, FromJson, Server.RpcEncodable 37 | 38 | structure UpdateParams (State : Type) where 39 | elapsed : Float 40 | actions : Array Action 41 | state : SvgState State 42 | mousePos : Option (Float × Float) -- TODO: change to Option (Int × Int) or do we want to support subpixel precision? 43 | deriving ToJson, FromJson 44 | 45 | structure UpdateResult (State : Type) where 46 | html : Html 47 | state : SvgState State 48 | /-- Approximate number of milliseconds to wait before calling again. -/ 49 | callbackTime : Option Float := some 33 50 | deriving Server.RpcEncodable 51 | 52 | -- maybe add title, refresh rate, initial time?, custom selection rendering 53 | structure InteractiveSvg (State : Type) where 54 | init : State 55 | frame : Svg.Frame 56 | update (time_ms Δt_ms : Float) (action : Action) 57 | (mouseStart mouseEnd : Option (Svg.Point frame)) 58 | (selectedId : Option String) (getSelectedData : (α : Type) → [FromJson α] → Option α) 59 | : State → State 60 | render (time_ms : Float) (mouseStart mouseEnd : Option (Svg.Point frame)) : State → Svg frame 61 | 62 | open Server RequestM Jsx in 63 | def InteractiveSvg.serverRpcMethod {State : Type} (isvg : InteractiveSvg State) (params : UpdateParams State) 64 | : RequestM (RequestTask (UpdateResult State)) := do 65 | 66 | -- Ideally, each action should have time and mouse position attached 67 | -- right now we just assume that all actions are uqually spaced within the frame 68 | let Δt := (params.elapsed - params.state.time) / params.actions.size.toFloat 69 | 70 | let idToData : Std.HashMap String Json := HashMap.ofList params.state.idToData 71 | 72 | let mut time := params.state.time 73 | let mut state := params.state.state 74 | let mut selected := params.state.selected 75 | 76 | let getData := λ (α : Type) [FromJson α] => do 77 | let id ← selected; 78 | let data ← idToData[id]? 79 | match fromJson? (α:=α) data with 80 | | .error _ => none 81 | | .ok val => some val 82 | 83 | 84 | let mouseStart := params.state.mousePos.map λ (i,j) => (i, j) 85 | let mouseEnd := params.mousePos.map λ (x,y) => (x.toInt, y.toInt) 86 | 87 | for action in params.actions do 88 | -- todo: interpolate mouse movenment! 89 | 90 | -- update state 91 | state := isvg.update time Δt action mouseStart mouseEnd selected getData state 92 | 93 | -- update selection 94 | if action.kind == ActionKind.mousedown then 95 | selected := action.id 96 | if action.kind == ActionKind.mouseup then 97 | selected := none 98 | 99 | -- update time 100 | time := time + Δt 101 | 102 | let mut svg := isvg.render time mouseStart mouseEnd state 103 | 104 | let svgState : SvgState State := 105 | { state := state 106 | time := params.elapsed 107 | selected := selected 108 | mousePos := mouseEnd.map λ p => p.toPixels 109 | idToData := svg.idToDataList } 110 | 111 | 112 | -- highlight selection 113 | if let some id := selected then 114 | if let some idx := svg.idToIdx[id]? then 115 | svg := { elements := svg.elements.modify idx λ e => e.setStroke (1.,1.,0.) (.px 5) } 116 | 117 | 118 | return RequestTask.pure { 119 | html :=
120 | {svg.toHtml} 121 |
, 122 | state := svgState, 123 | callbackTime := some 33, 124 | } 125 | 126 | end Svg 127 | end ProofWidgets 128 | -------------------------------------------------------------------------------- /ProofWidgets/Component/MakeEditLink.lean: -------------------------------------------------------------------------------- 1 | import Lean.Server.Utils 2 | import ProofWidgets.Component.Basic 3 | 4 | /-- Assuming that `s` is the content of a file starting at position `p`, 5 | advance `p` to the end of `s`. -/ 6 | def Lean.Lsp.Position.advance (p : Position) (s : Substring) : Position := 7 | let (nLinesAfter, lastLineUtf16Sz) := s.foldl 8 | (init := (0, 0)) 9 | fun (n, l) c => if c == '\n' then (n + 1, 0) else (n, l + c.utf16Size.toNat) 10 | { 11 | line := p.line + nLinesAfter 12 | character := (if nLinesAfter == 0 then p.character else 0) + lastLineUtf16Sz 13 | } 14 | 15 | namespace ProofWidgets 16 | open Lean 17 | 18 | structure MakeEditLinkProps where 19 | /-- The edit to perform on the file. -/ 20 | edit : Lsp.TextDocumentEdit 21 | /-- Which textual range to select after the edit. 22 | The range is interpreted in the file that `edit` applies to. 23 | If present and `start == end`, the cursor is moved to `start` and nothing is selected. 24 | If not present, the selection is not changed. -/ 25 | newSelection? : Option Lsp.Range := none 26 | /-- The `title` property, if any, to set on the displayed `` link. -/ 27 | title? : Option String := none 28 | deriving FromJson, ToJson 29 | 30 | /-- Replace `range` with `newText`. 31 | If `newSelection?` is absent, place the cursor at the end of the new text. 32 | If `newSelection?` is present, make the specified selection instead. 33 | See also `MakeEditLinkProps.ofReplaceRange`. 34 | -/ 35 | def MakeEditLinkProps.ofReplaceRange' (doc : Server.DocumentMeta) (range : Lsp.Range) 36 | (newText : String) (newSelection? : Option Lsp.Range := none) : MakeEditLinkProps := 37 | let edit := { textDocument := { uri := doc.uri, version? := doc.version } 38 | edits := #[{ range, newText }] } 39 | if newSelection?.isSome then 40 | { edit, newSelection? } 41 | else 42 | let endPos := range.start.advance newText.toSubstring 43 | { edit, newSelection? := some { start := endPos, «end» := endPos } } 44 | 45 | /-- Replace `range` with `newText`. 46 | If `newSelection?` is absent, place the cursor at the end of the new text. 47 | If `newSelection?` is present, select the range it specifies within `newText`. 48 | See also `MakeEditLinkProps.ofReplaceRange'`. -/ 49 | def MakeEditLinkProps.ofReplaceRange (doc : Server.DocumentMeta) (range : Lsp.Range) 50 | (newText : String) (newSelection? : Option (String.Pos × String.Pos) := none) : 51 | MakeEditLinkProps := 52 | ofReplaceRange' doc range newText (newSelection?.map fun (s, e) => 53 | let ps := range.start.advance (newText.toSubstring.extract 0 s) 54 | let pe := ps.advance (newText.toSubstring.extract s e) 55 | { start := ps, «end» := pe }) 56 | 57 | /-- A link that, when clicked, makes the specified edit 58 | and potentially moves the cursor 59 | or makes a selection. -/ 60 | @[widget_module] 61 | def MakeEditLink : Component MakeEditLinkProps where 62 | javascript := include_str ".." / ".." / ".lake" / "build" / "js" / "makeEditLink.js" 63 | 64 | end ProofWidgets 65 | -------------------------------------------------------------------------------- /ProofWidgets/Component/OfRpcMethod.lean: -------------------------------------------------------------------------------- 1 | import Lean.Elab.ElabRules 2 | import ProofWidgets.Component.Basic 3 | import ProofWidgets.Data.Html 4 | import ProofWidgets.Cancellable 5 | 6 | namespace ProofWidgets 7 | open Lean Server Meta Elab Term 8 | 9 | def ofRpcMethodTemplate := include_str ".." / ".." / ".lake" / "build" / "js" / "ofRpcMethod.js" 10 | 11 | /-- The elaborator `mk_rpc_widget%` allows writing certain widgets in Lean instead of JavaScript. 12 | Specifically, it translates an RPC method of type `MyProps → RequestM (RequestTask Html)` 13 | into a widget component of type `Component MyProps`. 14 | 15 | Even more specifically, we can write: 16 | ```lean 17 | open Lean Server 18 | 19 | structure MyProps where 20 | ... 21 | deriving RpcEncodable 22 | 23 | @[server_rpc_method] 24 | def MyComponent.rpc (ps : MyProps) : RequestM (RequestTask Html) := 25 | ... 26 | 27 | @[widget_module] 28 | def MyComponent : Component MyProps := 29 | mk_rpc_widget% MyComponent.rpc 30 | ``` 31 | 32 | This is convenient because we can program the logic that computes an output HTML tree 33 | given input props in Lean directly. 34 | 35 | ⚠️ However, note that there are several limitations on what such component can do 36 | compared to ones written natively in TypeScript or JavaScript: 37 | - It must be pure, i.e. cannot directly store any React state. 38 | Child components may store state as usual. 39 | - It cannot pass closures as props to the child components that it returns. 40 | For example, it is not currently possible to write click event handlers in Lean 41 | and pass them to a `