Skip to content

Commit 9d427fd

Browse files
authored
feat: "try this" messages with support for interactivity (#10524)
This PR adds support for interactivity to the combined "try this" messages that were introduced in #9966. In doing so, it moves the link to apply a suggestion to a separate `[apply]` button in front of the suggestion. Hints with diffs remain unchanged, as they did not previously support interacting with terms in the diff, either. <img width="379" height="256" alt="Suggestion with interactive message" src="https://github.com/user-attachments/assets/7838ebf6-0613-46e7-bc88-468a05acbf51" />
1 parent fe1e7d5 commit 9d427fd

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

48 files changed

+493
-455
lines changed

src/Lean/Meta/Hint.lean

Lines changed: 171 additions & 44 deletions
Original file line numberDiff line numberDiff line change
@@ -22,6 +22,103 @@ namespace Lean.Meta.Hint
2222

2323
open 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
/--
26123
A widget for rendering code action suggestions in error messages. Generally, this widget should not
27124
be 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
/--

src/Lean/Meta/Tactic/TryThis.lean

Lines changed: 26 additions & 30 deletions
Original file line numberDiff line numberDiff line change
@@ -34,6 +34,7 @@ open Lean Elab Tactic PrettyPrinter Meta Server RequestM
3434

3535
-- Because we can't use `builtin_widget_module` in `Lean.Meta.Hint`, we add the attribute here
3636
attribute [builtin_widget_module] Hint.tryThisDiffWidget
37+
attribute [builtin_widget_module] Hint.textInsertionWidget
3738

3839
/-! # Code action -/
3940

@@ -45,24 +46,17 @@ apply the replacement.
4546
let doc ← readDoc
4647
pure <| snap.infoTree.foldInfo (init := #[]) fun _ctx info result => Id.run do
4748
let .ofCustomInfo { stx, value } := info | result
48-
let some { range, suggestionTexts, codeActionPrefix? } :=
49+
let some { edit, codeActionTitle, .. } :=
4950
value.get? TryThisInfo | result
5051
let some stxRange := stx.getRange? | result
5152
let stxRange := doc.meta.text.utf8RangeToLspRange stxRange
5253
unless stxRange.start.line ≤ params.range.end.line do return result
5354
unless params.range.start.line ≤ stxRange.end.line do return result
54-
let mut result := result
55-
for h : i in *...suggestionTexts.size do
56-
let (newText, title?) := suggestionTexts[i]
57-
let title := title?.getD <| (codeActionPrefix?.getD "Try this: ") ++ newText
58-
result := result.push {
59-
eager.title := title
60-
eager.kind? := "quickfix"
61-
-- Only make the first option preferred
62-
eager.isPreferred? := if i = 0 then true else none
63-
eager.edit? := some <| .ofTextEdit doc.versionedIdentifier { range, newText }
64-
}
65-
result
55+
result.push {
56+
eager.title := codeActionTitle
57+
eager.kind? := "quickfix"
58+
eager.edit? := some <| .ofTextEdit doc.versionedIdentifier edit
59+
}
6660

6761
/-! # Formatting -/
6862

@@ -72,7 +66,7 @@ def delabToRefinableSyntax (e : Expr) : MetaM Term :=
7266

7367
/-- Delaborate `e` into a suggestion suitable for use by `refine`. -/
7468
def delabToRefinableSuggestion (e : Expr) : MetaM Suggestion :=
75-
return { suggestion := ← delabToRefinableSyntax e }
69+
return { suggestion := ← delabToRefinableSyntax e, messageData? := ← addMessageContext <| toMessageData e }
7670

7771
/-- Add a "try this" suggestion. This has two effects:
7872
@@ -187,19 +181,19 @@ Remark: We cannot determine if a tactic requires `expose_names` merely by inspec
187181
(shadowed variables are delaborated without daggers) nor the underlying `Expr` used to produce it
188182
(some inaccessible names may be implicit arguments that do not appear in the delaborated syntax).
189183
-/
190-
private def mkValidatedTactic (tac : TSyntax `tactic)
184+
private def mkValidatedTactic (tac : TSyntax `tactic) (msg : MessageData)
191185
(initialState : Tactic.SavedState) (expectedType? : Option Expr := none) :
192-
TacticM (Option (TSyntax `tactic)) := do
186+
TacticM (Option (TSyntax `tactic × MessageData)) := do
193187
try
194188
evalTacticWithState initialState tac expectedType?
195-
return some tac
189+
return some (tac, msg)
196190
catch _ =>
197191
-- Note: we must use `(expose_names; _)` and not `· expose_names; _` to avoid generating
198192
-- spurious tactic-abort exceptions, since these tactics may not close the goal
199193
let tac ← `(tactic| (expose_names; $tac))
200194
try
201195
evalTacticWithState initialState tac expectedType?
202-
return some tac
196+
return some (tac, m!"(expose_names; {msg})")
203197
catch _ =>
204198
return none
205199

@@ -231,7 +225,7 @@ private def addExactSuggestionCore (addSubgoalsMsg : Bool) (checkState? : Option
231225
let hasMVars := !mvars.isEmpty
232226
let (suggestion, messageData) ← mkExactSuggestionSyntax e (useRefine := hasMVars)
233227
let some checkState := checkState? | return .inl suggestion
234-
let some suggestion ← mkValidatedTactic suggestion checkState
228+
let some (suggestion, messageData) ← mkValidatedTactic suggestion messageData checkState
235229
| let messageData := m!"(expose_names; {messageData})"
236230
return .inr <| mkFailedToMakeTacticMsg m!"a {if hasMVars then "partial " else ""}proof" messageData
237231
let postInfo? ← if !addSubgoalsMsg || mvars.isEmpty then pure none else
@@ -241,7 +235,7 @@ private def addExactSuggestionCore (addSubgoalsMsg : Bool) (checkState? : Option
241235
let e ← withExposedNames <| PrettyPrinter.ppExpr (← instantiateMVars (← g.getType))
242236
str := str ++ Format.pretty ("\n-- ⊢ " ++ e)
243237
pure str
244-
return .inl { suggestion := suggestion, postInfo? }
238+
return .inl { suggestion := suggestion, postInfo?, messageData? := messageData }
245239

246240
/-- Add an `exact e` or `refine e` suggestion.
247241
@@ -374,11 +368,12 @@ def addHaveSuggestion (ref : Syntax) (h? : Option Name) (t? : Option Expr) (e :
374368
pure (← `(tactic| let $(mkIdent h):ident := $estx), m!"let {h} := {e}")
375369
pure (tac, ← addMessageContext msg)
376370
if let some checkState := checkState? then
377-
let some tac' ← mkValidatedTactic tac checkState
371+
let some (tac', msg') ← mkValidatedTactic tac msg checkState
378372
| logInfo <| mkFailedToMakeTacticMsg "a proof" msg
379373
return
380374
tac := tac'
381-
addSuggestion ref (s := { suggestion := tac }) origSpan?
375+
msg := msg'
376+
addSuggestion ref (s := { suggestion := tac, messageData? := msg }) origSpan?
382377

383378
open Lean.Parser.Tactic
384379
open Lean.Syntax
@@ -400,7 +395,7 @@ def addRewriteSuggestion (ref : Syntax) (rules : List (Expr × Bool))
400395
(type? : LOption Expr := .undef) (loc? : Option Expr := none)
401396
(origSpan? : Option Syntax := none) (checkState? : Option Tactic.SavedState := none) :
402397
TacticM Unit := do
403-
let mut (tac, tacMsg, extraMsg, extraStr) ← withExposedNames do
398+
let mut (tac, tacMsg, extraMsg) ← withExposedNames do
404399
let rulesStx := TSepArray.ofElems <| ← rules.toArray.mapM fun ⟨e, symm⟩ => do
405400
let t ← delabToRefinableSyntax e
406401
if symm then `(rwRule| ← $t:term) else `(rwRule| $t:term)
@@ -423,22 +418,23 @@ def addRewriteSuggestion (ref : Syntax) (rules : List (Expr × Bool))
423418
else
424419
m!"rw {rulesMsg}"
425420

426-
let (extraMsg, extraStr)
421+
let extraMsg ←
427422
match type? with
428423
| .some type =>
429-
pure (← addMessageContext m!"\n-- {type}", s!"\n-- {← PrettyPrinter.ppExpr type}")
430-
| .none => pure (m!"\n-- no goals", "\n-- no goals")
431-
| .undef => pure (m!"", "")
432-
return (tac, tacMsg, extraMsg, extraStr)
424+
pure <| ← addMessageContext m!"\n-- {type}"
425+
| .none => pure m!"\n-- no goals"
426+
| .undef => pure m!""
427+
return (tac, tacMsg, extraMsg)
433428

434429
if let some checkState := checkState? then
435430
let type? := match type? with
436431
| .some type => some type
437432
| _ => none
438-
let some tac' ← mkValidatedTactic tac checkState type?
433+
let some (tac', tacMsg') ← mkValidatedTactic tac tacMsg checkState type?
439434
| tacMsg := m!"(expose_names; {tacMsg})"
440435
logInfo <| mkFailedToMakeTacticMsg "an applicable rewrite lemma" (tacMsg ++ extraMsg)
441436
return
442437
tac := tac'
443-
addSuggestion ref (s := { suggestion := tac, postInfo? := extraStr })
438+
tacMsg := tacMsg'
439+
addSuggestion ref (s := { suggestion := tac, messageData? := tacMsg ++ extraMsg })
444440
origSpan?

0 commit comments

Comments
 (0)