@@ -22,6 +22,103 @@ namespace Lean.Meta.Hint
2222
2323open Elab Tactic PrettyPrinter TryThis
2424
25+ /--
26+ A widget for a clickable link (or icon) that inserts text into the document at a given position.
27+
28+ The props to this widget are of the following form:
29+ ```json
30+ {
31+ "range" : {
32+ "start" : {"line" : 100 , "character" : 0 },
33+ "end" : {"line" : 100 , "character" : 5 }
34+ },
35+ "suggestion" : " hi" ,
36+ "acceptSuggestionProps" : {
37+ "kind" : " text" ,
38+ "hoverText" : " Displayed on hover" ,
39+ "linkText" : " Displayed as the text of the link"
40+ }
41+ }
42+ ```
43+ ... or the following form, where `codiconName` is one of the icons at
44+ https://microsoft.github.io/vscode-codicons/dist/codicon.html and `gaps` determines
45+ whether there are clickable spaces surrounding the icon:
46+ ```json
47+ {
48+ "range" : {
49+ "start" : {"line" : 100 , "character" : 0 },
50+ "end" : {"line" : 100 , "character" : 5 }
51+ },
52+ "suggestion" : " hi" ,
53+ "acceptSuggestionProps" : {
54+ "kind" : " icon" ,
55+ "hoverText" : " Displayed on hover" ,
56+ "codiconName" : " search" ,
57+ "gaps" : true
58+ }
59+ }
60+ ```
61+
62+ Note: we cannot add the `builtin_widget_module` attribute here because that would require importing
63+ `Lean.Widget.UserWidget`, which in turn imports much of `Lean.Elab` -- the module where we want to
64+ be able to use this widget. Instead, we register the attribute post-hoc when we declare the regular
65+ "Try This" widget in `Lean.Meta.Tactic.TryThis`.
66+ -/
67+ def textInsertionWidget : Widget.Module where
68+ javascript := "
69+ import * as React from 'react';
70+ import { EditorContext, EnvPosContext } from '@leanprover/infoview';
71+
72+ const e = React.createElement;
73+ export default function ({ range, suggestion, acceptSuggestionProps }) {
74+ const pos = React.useContext(EnvPosContext)
75+ const editorConnection = React.useContext(EditorContext)
76+ function onClick() {
77+ editorConnection.api.applyEdit({
78+ changes: { [pos.uri]: [{ range, newText: suggestion }] }
79+ })
80+ }
81+
82+ if (acceptSuggestionProps.kind === 'text') {
83+ return e('span', {
84+ onClick,
85+ title: acceptSuggestionProps.hoverText,
86+ className: 'link pointer dim font-code',
87+ style: { color: 'var(--vscode-textLink-foreground)' }
88+ },
89+ acceptSuggestionProps.linkText)
90+ } else if (acceptSuggestionProps.kind === 'icon') {
91+ if (acceptSuggestionProps.gaps) {
92+ const icon = e('span', {
93+ className: `codicon codicon-${acceptSuggestionProps.codiconName}`,
94+ style: {
95+ verticalAlign: 'sub',
96+ fontSize: 'var(--vscode-editor-font-size)'
97+ }
98+ })
99+ return e('span', {
100+ onClick,
101+ title: acceptSuggestionProps.hoverText,
102+ className: `link pointer dim font-code`,
103+ style: { color: 'var(--vscode-textLink-foreground)' }
104+ }, ' ', icon, ' ')
105+ } else {
106+ return e('span', {
107+ onClick,
108+ title: acceptSuggestionProps.hoverText,
109+ className: `link pointer dim font-code codicon codicon-${acceptSuggestionProps.codiconName}`,
110+ style: {
111+ color: 'var(--vscode-textLink-foreground)',
112+ verticalAlign: 'sub',
113+ fontSize: 'var(--vscode-editor-font-size)'
114+ }
115+ })
116+ }
117+
118+ }
119+ throw new Error('Unexpected `acceptSuggestionProps` kind: ' + acceptSuggestionProps.kind)
120+ }"
121+
25122/--
26123A widget for rendering code action suggestions in error messages. Generally, this widget should not
27124be used directly; instead, use `MessageData.hint`. Note that this widget is intended only for use
@@ -333,52 +430,82 @@ def mkSuggestionsMessage (suggestions : Array Suggestion) (ref : Syntax)
333430 (codeActionPrefix? : Option String) (forceList : Bool) : CoreM MessageData := do
334431 let mut msg := m!""
335432 for suggestion in suggestions do
336- if let some range := (suggestion.span?.getD ref).getRange? then
337- let { info, suggestions := suggestionArr, range := lspRange } ←
338- processSuggestions ref range #[suggestion.toTryThisSuggestion] codeActionPrefix?
339- pushInfoLeaf info
340- -- The following access is safe because
341- -- `suggestionsArr = #[suggestion.toTryThisSuggestion].map ...` (see `processSuggestions`)
342- let suggestionText := suggestionArr[0 ]!.2 .1
343- let map ← getFileMap
344- let rangeContents := map.source.extract range.start range.stop
345- let edits ← do
346- if let some msgData := suggestion.messageData? then
347- pure #[(.insert, toString <| ← msgData.format)]
348- else
349- pure <| readableDiff rangeContents suggestionText suggestion.diffGranularity
350- let mut edits := edits
351- if let some previewRange := suggestion.previewSpan? >>= Syntax.getRange? then
352- if previewRange.includes range then
353- let map ← getFileMap
354- if previewRange.start < range.start then
355- edits := #[(.skip, (map.source.extract previewRange.start range.start))] ++ edits
356- if range.stop < previewRange.stop then
357- edits := edits.push (.skip, (map.source.extract range.stop previewRange.stop ))
358- let diffJson := mkDiffJson edits
359- let json := json% {
360- diff: $diffJson,
361- suggestion: $suggestionText,
362- range: $lspRange
433+ let some range := suggestion.span?.getD ref |>.getRange?
434+ | continue
435+ let edit ← suggestion.processEdit range
436+ let suggestionText := edit.newText
437+ let ref := Syntax.ofRange <| ref.getRange?.getD range
438+ let codeActionTitleOverride? := suggestion.toCodeActionTitle?.map (· suggestionText)
439+ let codeActionTitle := codeActionTitleOverride?.getD <| (codeActionPrefix?.getD "Try this: " ) ++ suggestionText
440+ let info := Info.ofCustomInfo {
441+ stx := ref
442+ value := Dynamic.mk {
443+ edit
444+ suggestion := suggestion.toTryThisSuggestion
445+ codeActionTitle
446+ : TryThisInfo
363447 }
364- let preInfo := suggestion.preInfo?.getD ""
365- let postInfo := suggestion.postInfo?.getD ""
366- let diffString :=
367- if suggestion.diffGranularity matches .none then
368- edits.foldl (· ++ ·.2 ) ""
369- else
370- mkDiffString edits
371- let widget := MessageData.ofWidget {
372- id := ``tryThisDiffWidget
373- javascriptHash := tryThisDiffWidget.javascriptHash
374- props := return json
375- } diffString
376- let widgetMsg := m!"{preInfo}{widget}{postInfo}"
377- let suggestionMsg := if suggestions.size == 1 && !forceList then
378- m!"\n {widgetMsg}"
448+ }
449+ pushInfoLeaf info
450+ let map ← getFileMap
451+ let rangeContents := map.source.extract range.start range.stop
452+ let edits ← do
453+ if let some msgData := suggestion.messageData? then
454+ pure #[(.insert, toString <| ← msgData.format)]
379455 else
380- m!"\n " ++ MessageData.nest 2 m!"• {widgetMsg}"
381- msg := msg ++ MessageData.nestD suggestionMsg
456+ pure <| readableDiff rangeContents suggestionText suggestion.diffGranularity
457+ let mut edits := edits
458+ if let some previewRange := suggestion.previewSpan? >>= Syntax.getRange? then
459+ if previewRange.includes range then
460+ let map ← getFileMap
461+ if previewRange.start < range.start then
462+ edits := #[(.skip, (map.source.extract previewRange.start range.start))] ++ edits
463+ if range.stop < previewRange.stop then
464+ edits := edits.push (.skip, (map.source.extract range.stop previewRange.stop ))
465+ let preInfo := suggestion.preInfo?.getD ""
466+ let postInfo := suggestion.postInfo?.getD ""
467+ let isDiffSuggestion :=
468+ ! (suggestion.diffGranularity matches .none) && suggestion.messageData?.isNone
469+ || suggestion.previewSpan?.isSome
470+ let suggestionMsg :=
471+ if ! isDiffSuggestion then
472+ let applyButton := MessageData.ofWidget {
473+ id := ``textInsertionWidget
474+ javascriptHash := textInsertionWidget.javascriptHash
475+ props := return json% {
476+ range: $edit.range,
477+ suggestion: $suggestionText,
478+ acceptSuggestionProps: {
479+ kind: "text" ,
480+ hoverText: "Apply suggestion" ,
481+ linkText: "[apply]"
482+ }
483+ }
484+ } "[apply]"
485+ m!"\n {applyButton} {preInfo}{toMessageData suggestion}{postInfo}"
486+ else
487+ let diffJson := mkDiffJson edits
488+ let json := json% {
489+ diff: $diffJson,
490+ suggestion: $suggestionText,
491+ range: $edit.range
492+ }
493+ let diffString :=
494+ if suggestion.diffGranularity matches .none then
495+ edits.foldl (· ++ ·.2 ) ""
496+ else
497+ mkDiffString edits
498+ let diffWidget := MessageData.ofWidget {
499+ id := ``tryThisDiffWidget
500+ javascriptHash := tryThisDiffWidget.javascriptHash
501+ props := return json
502+ } diffString
503+ let msg := m!"{preInfo}{diffWidget}{postInfo}"
504+ if suggestions.size == 1 && !forceList then
505+ m!"\n {msg}"
506+ else
507+ m!"\n " ++ MessageData.nest 2 m!"• {msg}"
508+ msg := msg ++ MessageData.nestD suggestionMsg
382509 return msg
383510
384511/--
0 commit comments