diff --git a/src/Lean/Meta/Hint.lean b/src/Lean/Meta/Hint.lean index 7ca0051ebe79..10a001738508 100644 --- a/src/Lean/Meta/Hint.lean +++ b/src/Lean/Meta/Hint.lean @@ -22,6 +22,103 @@ namespace Lean.Meta.Hint open Elab Tactic PrettyPrinter TryThis +/-- +A widget for a clickable link (or icon) that inserts text into the document at a given position. + +The props to this widget are of the following form: +```json +{ + "range": { + "start": {"line": 100, "character": 0}, + "end": {"line": 100, "character": 5} + }, + "suggestion": "hi", + "acceptSuggestionProps": { + "kind": "text", + "hoverText": "Displayed on hover", + "linkText": "Displayed as the text of the link" + } +} +``` +... or the following form, where `codiconName` is one of the icons at +https://microsoft.github.io/vscode-codicons/dist/codicon.html and `gaps` determines +whether there are clickable spaces surrounding the icon: +```json +{ + "range": { + "start": {"line": 100, "character": 0}, + "end": {"line": 100, "character": 5} + }, + "suggestion": "hi", + "acceptSuggestionProps": { + "kind": "icon", + "hoverText": "Displayed on hover", + "codiconName": "search", + "gaps": true + } +} +``` + +Note: we cannot add the `builtin_widget_module` attribute here because that would require importing +`Lean.Widget.UserWidget`, which in turn imports much of `Lean.Elab` -- the module where we want to +be able to use this widget. Instead, we register the attribute post-hoc when we declare the regular +"Try This" widget in `Lean.Meta.Tactic.TryThis`. +-/ +def textInsertionWidget : Widget.Module where + javascript := " +import * as React from 'react'; +import { EditorContext, EnvPosContext } from '@leanprover/infoview'; + +const e = React.createElement; +export default function ({ range, suggestion, acceptSuggestionProps }) { + const pos = React.useContext(EnvPosContext) + const editorConnection = React.useContext(EditorContext) + function onClick() { + editorConnection.api.applyEdit({ + changes: { [pos.uri]: [{ range, newText: suggestion }] } + }) + } + + if (acceptSuggestionProps.kind === 'text') { + return e('span', { + onClick, + title: acceptSuggestionProps.hoverText, + className: 'link pointer dim font-code', + style: { color: 'var(--vscode-textLink-foreground)' } + }, + acceptSuggestionProps.linkText) + } else if (acceptSuggestionProps.kind === 'icon') { + if (acceptSuggestionProps.gaps) { + const icon = e('span', { + className: `codicon codicon-${acceptSuggestionProps.codiconName}`, + style: { + verticalAlign: 'sub', + fontSize: 'var(--vscode-editor-font-size)' + } + }) + return e('span', { + onClick, + title: acceptSuggestionProps.hoverText, + className: `link pointer dim font-code`, + style: { color: 'var(--vscode-textLink-foreground)' } + }, ' ', icon, ' ') + } else { + return e('span', { + onClick, + title: acceptSuggestionProps.hoverText, + className: `link pointer dim font-code codicon codicon-${acceptSuggestionProps.codiconName}`, + style: { + color: 'var(--vscode-textLink-foreground)', + verticalAlign: 'sub', + fontSize: 'var(--vscode-editor-font-size)' + } + }) + } + + } + throw new Error('Unexpected `acceptSuggestionProps` kind: ' + acceptSuggestionProps.kind) +}" + /-- A widget for rendering code action suggestions in error messages. Generally, this widget should not 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) (codeActionPrefix? : Option String) (forceList : Bool) : CoreM MessageData := do let mut msg := m!"" for suggestion in suggestions do - if let some range := (suggestion.span?.getD ref).getRange? then - let { info, suggestions := suggestionArr, range := lspRange } ← - processSuggestions ref range #[suggestion.toTryThisSuggestion] codeActionPrefix? - pushInfoLeaf info - -- The following access is safe because - -- `suggestionsArr = #[suggestion.toTryThisSuggestion].map ...` (see `processSuggestions`) - let suggestionText := suggestionArr[0]!.2.1 - let map ← getFileMap - let rangeContents := map.source.extract range.start range.stop - let edits ← do - if let some msgData := suggestion.messageData? then - pure #[(.insert, toString <| ← msgData.format)] - else - pure <| readableDiff rangeContents suggestionText suggestion.diffGranularity - let mut edits := edits - if let some previewRange := suggestion.previewSpan? >>= Syntax.getRange? then - if previewRange.includes range then - let map ← getFileMap - if previewRange.start < range.start then - edits := #[(.skip, (map.source.extract previewRange.start range.start))] ++ edits - if range.stop < previewRange.stop then - edits := edits.push (.skip, (map.source.extract range.stop previewRange.stop)) - let diffJson := mkDiffJson edits - let json := json% { - diff: $diffJson, - suggestion: $suggestionText, - range: $lspRange + let some range := suggestion.span?.getD ref |>.getRange? + | continue + let edit ← suggestion.processEdit range + let suggestionText := edit.newText + let ref := Syntax.ofRange <| ref.getRange?.getD range + let codeActionTitleOverride? := suggestion.toCodeActionTitle?.map (· suggestionText) + let codeActionTitle := codeActionTitleOverride?.getD <| (codeActionPrefix?.getD "Try this: ") ++ suggestionText + let info := Info.ofCustomInfo { + stx := ref + value := Dynamic.mk { + edit + suggestion := suggestion.toTryThisSuggestion + codeActionTitle + : TryThisInfo } - let preInfo := suggestion.preInfo?.getD "" - let postInfo := suggestion.postInfo?.getD "" - let diffString := - if suggestion.diffGranularity matches .none then - edits.foldl (· ++ ·.2) "" - else - mkDiffString edits - let widget := MessageData.ofWidget { - id := ``tryThisDiffWidget - javascriptHash := tryThisDiffWidget.javascriptHash - props := return json - } diffString - let widgetMsg := m!"{preInfo}{widget}{postInfo}" - let suggestionMsg := if suggestions.size == 1 && !forceList then - m!"\n{widgetMsg}" + } + pushInfoLeaf info + let map ← getFileMap + let rangeContents := map.source.extract range.start range.stop + let edits ← do + if let some msgData := suggestion.messageData? then + pure #[(.insert, toString <| ← msgData.format)] else - m!"\n" ++ MessageData.nest 2 m!"• {widgetMsg}" - msg := msg ++ MessageData.nestD suggestionMsg + pure <| readableDiff rangeContents suggestionText suggestion.diffGranularity + let mut edits := edits + if let some previewRange := suggestion.previewSpan? >>= Syntax.getRange? then + if previewRange.includes range then + let map ← getFileMap + if previewRange.start < range.start then + edits := #[(.skip, (map.source.extract previewRange.start range.start))] ++ edits + if range.stop < previewRange.stop then + edits := edits.push (.skip, (map.source.extract range.stop previewRange.stop)) + let preInfo := suggestion.preInfo?.getD "" + let postInfo := suggestion.postInfo?.getD "" + let isDiffSuggestion := + ! (suggestion.diffGranularity matches .none) && suggestion.messageData?.isNone + || suggestion.previewSpan?.isSome + let suggestionMsg := + if ! isDiffSuggestion then + let applyButton := MessageData.ofWidget { + id := ``textInsertionWidget + javascriptHash := textInsertionWidget.javascriptHash + props := return json% { + range: $edit.range, + suggestion: $suggestionText, + acceptSuggestionProps: { + kind: "text", + hoverText: "Apply suggestion", + linkText: "[apply]" + } + } + } "[apply]" + m!"\n{applyButton} {preInfo}{toMessageData suggestion}{postInfo}" + else + let diffJson := mkDiffJson edits + let json := json% { + diff: $diffJson, + suggestion: $suggestionText, + range: $edit.range + } + let diffString := + if suggestion.diffGranularity matches .none then + edits.foldl (· ++ ·.2) "" + else + mkDiffString edits + let diffWidget := MessageData.ofWidget { + id := ``tryThisDiffWidget + javascriptHash := tryThisDiffWidget.javascriptHash + props := return json + } diffString + let msg := m!"{preInfo}{diffWidget}{postInfo}" + if suggestions.size == 1 && !forceList then + m!"\n{msg}" + else + m!"\n" ++ MessageData.nest 2 m!"• {msg}" + msg := msg ++ MessageData.nestD suggestionMsg return msg /-- diff --git a/src/Lean/Meta/Tactic/TryThis.lean b/src/Lean/Meta/Tactic/TryThis.lean index cfbc22fca9cc..df636c49035d 100644 --- a/src/Lean/Meta/Tactic/TryThis.lean +++ b/src/Lean/Meta/Tactic/TryThis.lean @@ -34,6 +34,7 @@ open Lean Elab Tactic PrettyPrinter Meta Server RequestM -- Because we can't use `builtin_widget_module` in `Lean.Meta.Hint`, we add the attribute here attribute [builtin_widget_module] Hint.tryThisDiffWidget +attribute [builtin_widget_module] Hint.textInsertionWidget /-! # Code action -/ @@ -45,24 +46,17 @@ apply the replacement. let doc ← readDoc pure <| snap.infoTree.foldInfo (init := #[]) fun _ctx info result => Id.run do let .ofCustomInfo { stx, value } := info | result - let some { range, suggestionTexts, codeActionPrefix? } := + let some { edit, codeActionTitle, .. } := value.get? TryThisInfo | result let some stxRange := stx.getRange? | result let stxRange := doc.meta.text.utf8RangeToLspRange stxRange unless stxRange.start.line ≤ params.range.end.line do return result unless params.range.start.line ≤ stxRange.end.line do return result - let mut result := result - for h : i in *...suggestionTexts.size do - let (newText, title?) := suggestionTexts[i] - let title := title?.getD <| (codeActionPrefix?.getD "Try this: ") ++ newText - result := result.push { - eager.title := title - eager.kind? := "quickfix" - -- Only make the first option preferred - eager.isPreferred? := if i = 0 then true else none - eager.edit? := some <| .ofTextEdit doc.versionedIdentifier { range, newText } - } - result + result.push { + eager.title := codeActionTitle + eager.kind? := "quickfix" + eager.edit? := some <| .ofTextEdit doc.versionedIdentifier edit + } /-! # Formatting -/ @@ -72,7 +66,7 @@ def delabToRefinableSyntax (e : Expr) : MetaM Term := /-- Delaborate `e` into a suggestion suitable for use by `refine`. -/ def delabToRefinableSuggestion (e : Expr) : MetaM Suggestion := - return { suggestion := ← delabToRefinableSyntax e } + return { suggestion := ← delabToRefinableSyntax e, messageData? := ← addMessageContext <| toMessageData e } /-- Add a "try this" suggestion. This has two effects: @@ -187,19 +181,19 @@ Remark: We cannot determine if a tactic requires `expose_names` merely by inspec (shadowed variables are delaborated without daggers) nor the underlying `Expr` used to produce it (some inaccessible names may be implicit arguments that do not appear in the delaborated syntax). -/ -private def mkValidatedTactic (tac : TSyntax `tactic) +private def mkValidatedTactic (tac : TSyntax `tactic) (msg : MessageData) (initialState : Tactic.SavedState) (expectedType? : Option Expr := none) : - TacticM (Option (TSyntax `tactic)) := do + TacticM (Option (TSyntax `tactic × MessageData)) := do try evalTacticWithState initialState tac expectedType? - return some tac + return some (tac, msg) catch _ => -- Note: we must use `(expose_names; _)` and not `· expose_names; _` to avoid generating -- spurious tactic-abort exceptions, since these tactics may not close the goal let tac ← `(tactic| (expose_names; $tac)) try evalTacticWithState initialState tac expectedType? - return some tac + return some (tac, m!"(expose_names; {msg})") catch _ => return none @@ -231,7 +225,7 @@ private def addExactSuggestionCore (addSubgoalsMsg : Bool) (checkState? : Option let hasMVars := !mvars.isEmpty let (suggestion, messageData) ← mkExactSuggestionSyntax e (useRefine := hasMVars) let some checkState := checkState? | return .inl suggestion - let some suggestion ← mkValidatedTactic suggestion checkState + let some (suggestion, messageData) ← mkValidatedTactic suggestion messageData checkState | let messageData := m!"(expose_names; {messageData})" return .inr <| mkFailedToMakeTacticMsg m!"a {if hasMVars then "partial " else ""}proof" messageData let postInfo? ← if !addSubgoalsMsg || mvars.isEmpty then pure none else @@ -241,7 +235,7 @@ private def addExactSuggestionCore (addSubgoalsMsg : Bool) (checkState? : Option let e ← withExposedNames <| PrettyPrinter.ppExpr (← instantiateMVars (← g.getType)) str := str ++ Format.pretty ("\n-- ⊢ " ++ e) pure str - return .inl { suggestion := suggestion, postInfo? } + return .inl { suggestion := suggestion, postInfo?, messageData? := messageData } /-- Add an `exact e` or `refine e` suggestion. @@ -374,11 +368,12 @@ def addHaveSuggestion (ref : Syntax) (h? : Option Name) (t? : Option Expr) (e : pure (← `(tactic| let $(mkIdent h):ident := $estx), m!"let {h} := {e}") pure (tac, ← addMessageContext msg) if let some checkState := checkState? then - let some tac' ← mkValidatedTactic tac checkState + let some (tac', msg') ← mkValidatedTactic tac msg checkState | logInfo <| mkFailedToMakeTacticMsg "a proof" msg return tac := tac' - addSuggestion ref (s := { suggestion := tac }) origSpan? + msg := msg' + addSuggestion ref (s := { suggestion := tac, messageData? := msg }) origSpan? open Lean.Parser.Tactic open Lean.Syntax @@ -400,7 +395,7 @@ def addRewriteSuggestion (ref : Syntax) (rules : List (Expr × Bool)) (type? : LOption Expr := .undef) (loc? : Option Expr := none) (origSpan? : Option Syntax := none) (checkState? : Option Tactic.SavedState := none) : TacticM Unit := do - let mut (tac, tacMsg, extraMsg, extraStr) ← withExposedNames do + let mut (tac, tacMsg, extraMsg) ← withExposedNames do let rulesStx := TSepArray.ofElems <| ← rules.toArray.mapM fun ⟨e, symm⟩ => do let t ← delabToRefinableSyntax e if symm then `(rwRule| ← $t:term) else `(rwRule| $t:term) @@ -423,22 +418,23 @@ def addRewriteSuggestion (ref : Syntax) (rules : List (Expr × Bool)) else m!"rw {rulesMsg}" - let (extraMsg, extraStr) ← + let extraMsg ← match type? with | .some type => - pure (← addMessageContext m!"\n-- {type}", s!"\n-- {← PrettyPrinter.ppExpr type}") - | .none => pure (m!"\n-- no goals", "\n-- no goals") - | .undef => pure (m!"", "") - return (tac, tacMsg, extraMsg, extraStr) + pure <| ← addMessageContext m!"\n-- {type}" + | .none => pure m!"\n-- no goals" + | .undef => pure m!"" + return (tac, tacMsg, extraMsg) if let some checkState := checkState? then let type? := match type? with | .some type => some type | _ => none - let some tac' ← mkValidatedTactic tac checkState type? + let some (tac', tacMsg') ← mkValidatedTactic tac tacMsg checkState type? | tacMsg := m!"(expose_names; {tacMsg})" logInfo <| mkFailedToMakeTacticMsg "an applicable rewrite lemma" (tacMsg ++ extraMsg) return tac := tac' - addSuggestion ref (s := { suggestion := tac, postInfo? := extraStr }) + tacMsg := tacMsg' + addSuggestion ref (s := { suggestion := tac, messageData? := tacMsg ++ extraMsg }) origSpan? diff --git a/src/Lean/Meta/TryThis.lean b/src/Lean/Meta/TryThis.lean index f35ea08182de..44ad5ac48e2e 100644 --- a/src/Lean/Meta/TryThis.lean +++ b/src/Lean/Meta/TryThis.lean @@ -28,21 +28,6 @@ open PrettyPrinter /-! # Code action information -/ -/-- A packet of information about a "Try this" suggestion -that we store in the infotree for the associated code action to retrieve. -/ -structure TryThisInfo : Type where - /-- The textual range to be replaced by one of the suggestions. -/ - range : Lsp.Range - /-- - A list of suggestions for the user to choose from. - Each suggestion may optionally come with an override for the code action title. - -/ - suggestionTexts : Array (String × Option String) - /-- The prefix to display before the code action for a "Try this" suggestion if no custom code - action title is provided. If not provided, `"Try this: "` is used. -/ - codeActionPrefix? : Option String - deriving TypeName - /-! # `Suggestion` data -/ -- TODO: we could also support `Syntax` and `Format` @@ -175,11 +160,28 @@ attribute [deprecated "The `style?` property is not used anymore." (since := "20 /- Use `toMessageData` of the suggestion text. -/ instance : ToMessageData Suggestion where - toMessageData s := toMessageData s.suggestion + toMessageData s := s.messageData?.getD (toMessageData s.suggestion) instance : Coe SuggestionText Suggestion where coe t := { suggestion := t } +/-- +A packet of information about a "Try this" suggestion +that we store in the infotree for the associated code action to retrieve. +-/ +structure TryThisInfo : Type where + /-- Suggestion edit as it will be applied to the editor document. -/ + edit : Lsp.TextEdit + /-- Title of the code action that is displayed in the code action selection dialog. -/ + codeActionTitle : String + /-- + Original suggestion that this `TryThisInfo` was derived from. + Stored so that meta-code downstream can extract the suggestions + produced by a tactic from the `InfoTree`. + -/ + suggestion : Suggestion + deriving TypeName + /-! # Formatting -/ /-- Yields `(indent, column)` given a `FileMap` and a `String.Range`, where `indent` is the number @@ -226,54 +228,16 @@ def prettyExtra (s : SuggestionText) (w : Option Nat := none) end SuggestionText -/-- Converts a `Suggestion` to `Json` in `CoreM`. We need `CoreM` in order to pretty-print syntax. - -This also returns a `String × Option String` consisting of the pretty-printed text and any custom -code action title if `toCodeActionTitle?` is provided. - -If `w := none`, then `w := getInputWidth (← getOptions)` is used. --/ -def Suggestion.toJsonAndInfoM (s : Suggestion) (w : Option Nat := none) (indent column : Nat := 0) : - CoreM (Json × String × Option String) := do - let text ← s.suggestion.prettyExtra w indent column - let mut json := [("suggestion", (text : Json))] - if let some preInfo := s.preInfo? then json := ("preInfo", preInfo) :: json - if let some postInfo := s.postInfo? then json := ("postInfo", postInfo) :: json - return (Json.mkObj json, text, s.toCodeActionTitle?.map (· text)) - -/-- -Represents processed data for a collection of suggestions that can be passed to a widget and pushed -in an info leaf. - -It contains the following data: -* `suggestions`: tuples of the form `(j, t, p)` where `j` is JSON containing a suggestion and its - pre- and post-info, `t` is the text to be inserted by the suggestion, and `p` is the code action - prefix thereof. -* `info`: the `TryThisInfo` data corresponding to a collection of suggestions -* `range`: the range at which the suggestion is to be applied. --/ -structure ProcessedSuggestions where - suggestions : Array (Json × String × Option String) - info : Elab.Info - range : Lsp.Range +def Suggestion.pretty (s : Suggestion) (w : Option Nat := none) (indent column : Nat := 0) : + CoreM String := do + s.suggestion.prettyExtra w indent column -/-- -Processes an array of `Suggestion`s into data that can be used to construct a code-action info leaf -and "try this" widget. --/ -def processSuggestions (ref : Syntax) (range : String.Range) (suggestions : Array Suggestion) - (codeActionPrefix? : Option String) : CoreM ProcessedSuggestions := do +def Suggestion.processEdit (s : Suggestion) (range : String.Range) : CoreM Lsp.TextEdit := do let map ← getFileMap -- FIXME: this produces incorrect results when `by` is at the beginning of the line, i.e. -- replacing `tac` in `by tac`, because the next line will only be 2 space indented -- (less than `tac` which starts at column 3) let (indent, column) := getIndentAndColumn map range - let suggestions ← suggestions.mapM (·.toJsonAndInfoM (indent := indent) (column := column)) - let suggestionTexts := suggestions.map (·.2) - let ref := Syntax.ofRange <| ref.getRange?.getD range + let newText ← s.pretty (indent := indent) (column := column) let range := map.utf8RangeToLspRange range - let info := .ofCustomInfo { - stx := ref - value := Dynamic.mk { range, suggestionTexts, codeActionPrefix? : TryThisInfo } - } - return { info, suggestions, range } + return { range, newText } diff --git a/tests/lean/4591.lean.expected.out b/tests/lean/4591.lean.expected.out index 920689b7d3ae..8d179f52b1f5 100644 --- a/tests/lean/4591.lean.expected.out +++ b/tests/lean/4591.lean.expected.out @@ -1,2 +1,2 @@ Try this: - simp only [Nat.foo] + [apply] simp only [Nat.foo] diff --git a/tests/lean/interactive/builtinCodeactions.lean.expected.out b/tests/lean/interactive/builtinCodeactions.lean.expected.out index 812f99988484..3a2d20f4864e 100644 --- a/tests/lean/interactive/builtinCodeactions.lean.expected.out +++ b/tests/lean/interactive/builtinCodeactions.lean.expected.out @@ -5,7 +5,6 @@ [{"title": "Try this: simp only [AAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAAA,\n aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa] at h", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///builtinCodeactions.lean"}, diff --git a/tests/lean/interactive/structInstFieldHints.lean.expected.out b/tests/lean/interactive/structInstFieldHints.lean.expected.out index a3582b46d26f..a1ed80cc29cf 100644 --- a/tests/lean/interactive/structInstFieldHints.lean.expected.out +++ b/tests/lean/interactive/structInstFieldHints.lean.expected.out @@ -4,7 +4,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -20,7 +19,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -36,7 +34,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -52,7 +49,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -68,7 +64,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -84,7 +79,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -100,7 +94,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -116,7 +109,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -132,7 +124,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -148,7 +139,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -163,7 +153,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -178,7 +167,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -193,7 +181,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -208,7 +195,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -224,7 +210,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -239,7 +224,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -254,7 +238,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -270,7 +253,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -287,7 +269,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -303,7 +284,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -319,7 +299,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -335,7 +314,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -351,7 +329,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -366,7 +343,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -383,7 +359,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -399,7 +374,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -415,7 +389,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -432,7 +405,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -449,7 +421,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -466,7 +437,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -482,7 +452,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -498,7 +467,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -515,7 +483,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -532,7 +499,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -548,7 +514,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, @@ -564,7 +529,6 @@ "context": {"diagnostics": []}} [{"title": "Add missing fields", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///structInstFieldHints.lean"}, diff --git a/tests/lean/interactive/terminationBySuggestion.lean.expected.out b/tests/lean/interactive/terminationBySuggestion.lean.expected.out index a9edab25f06f..8a0b40ac01c4 100644 --- a/tests/lean/interactive/terminationBySuggestion.lean.expected.out +++ b/tests/lean/interactive/terminationBySuggestion.lean.expected.out @@ -4,7 +4,6 @@ "context": {"diagnostics": []}} [{"title": "Try this: termination_by (n, m)", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": @@ -20,7 +19,6 @@ "context": {"diagnostics": []}} [{"title": "Try this: termination_by structural n", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": @@ -36,7 +34,6 @@ "context": {"diagnostics": []}} [{"title": "Try this: termination_by structural x => x", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": @@ -52,7 +49,6 @@ "context": {"diagnostics": []}} [{"title": "Try this: termination_by n", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": @@ -68,7 +64,6 @@ "context": {"diagnostics": []}} [{"title": "Try this: termination_by x1 => x1", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": diff --git a/tests/lean/interactive/tryThisCodeAction.lean.expected.out b/tests/lean/interactive/tryThisCodeAction.lean.expected.out index b27369369db8..b8759acd23a6 100644 --- a/tests/lean/interactive/tryThisCodeAction.lean.expected.out +++ b/tests/lean/interactive/tryThisCodeAction.lean.expected.out @@ -4,7 +4,6 @@ "context": {"diagnostics": []}} [{"title": "Try this: simp only", "kind": "quickfix", - "isPreferred": true, "edit": {"documentChanges": [{"textDocument": {"version": 1, "uri": "file:///tryThisCodeAction.lean"}, diff --git a/tests/lean/librarySearch.lean b/tests/lean/librarySearch.lean index 1bfd3b9e4d69..5c81d60b07e1 100644 --- a/tests/lean/librarySearch.lean +++ b/tests/lean/librarySearch.lean @@ -21,14 +21,14 @@ noncomputable section /-- info: Try this: - exact Nat.lt_add_one x + [apply] exact Nat.lt_add_one x -/ #guard_msgs in example (x : Nat) : x ≠ x.succ := Nat.ne_of_lt (by apply?) /-- info: Try this: - exact Nat.zero_lt_succ 1 + [apply] exact Nat.zero_lt_succ 1 -/ #guard_msgs in example : 0 ≠ 1 + 1 := Nat.ne_of_lt (by apply?) @@ -37,28 +37,28 @@ example : 0 ≠ 1 + 1 := Nat.ne_of_lt (by exact Fin.pos') /-- info: Try this: - exact Nat.add_comm x y + [apply] exact Nat.add_comm x y -/ #guard_msgs in example (x y : Nat) : x + y = y + x := by apply? /-- info: Try this: - exact fun a => Nat.add_le_add_right a k + [apply] exact fun a => Nat.add_le_add_right a k -/ #guard_msgs in example (n m k : Nat) : n ≤ m → n + k ≤ m + k := by apply? /-- info: Try this: - exact Nat.mul_dvd_mul_left a w + [apply] exact Nat.mul_dvd_mul_left a w -/ #guard_msgs in example (_ha : a > 0) (w : b ∣ c) : a * b ∣ a * c := by apply? /-- info: Try this: - Nat.lt_add_one x + [apply] Nat.lt_add_one x -/ #guard_msgs in example : x < x + 1 := exact?% @@ -73,19 +73,19 @@ example (x y : Nat) : x ≤ y := exact?% /-- info: Try this: - exact p + [apply] exact p -/ #guard_msgs in example (P : Prop) (p : P) : P := by apply? /-- info: Try this: - exact False.elim (np p) + [apply] exact False.elim (np p) -/ #guard_msgs in example (P : Prop) (p : P) (np : ¬P) : false := by apply? /-- info: Try this: - exact h x rfl + [apply] exact h x rfl -/ #guard_msgs in example (X : Type) (P : Prop) (x : X) (h : ∀ x : X, x = x → P) : P := by apply? @@ -101,7 +101,7 @@ example (α : Prop) : α → α := by apply? /-- info: Try this: - exact Nat.add_comm a b + [apply] exact Nat.add_comm a b -/ #guard_msgs in example (a b : Nat) : a + b = b + a := @@ -109,7 +109,7 @@ by apply? /-- info: Try this: - exact Nat.mul_sub_left_distrib n m k + [apply] exact Nat.mul_sub_left_distrib n m k -/ #guard_msgs in example (n m k : Nat) : n * (m - k) = n * m - n * k := @@ -119,7 +119,7 @@ attribute [symm] Eq.symm /-- info: Try this: - exact Eq.symm (Nat.mul_sub_left_distrib n m k) + [apply] exact Eq.symm (Nat.mul_sub_left_distrib n m k) -/ #guard_msgs in example (n m k : Nat) : n * m - n * k = n * (m - k) := by @@ -127,21 +127,21 @@ example (n m k : Nat) : n * m - n * k = n * (m - k) := by /-- info: Try this: - exact eq_comm + [apply] exact eq_comm -/ #guard_msgs in example {α : Type} (x y : α) : x = y ↔ y = x := by apply? /-- info: Try this: - exact Nat.add_pos_right a hb + [apply] exact Nat.add_pos_right a hb -/ #guard_msgs in example (a b : Nat) (_ha : 0 < a) (hb : 0 < b) : 0 < a + b := by apply? /-- info: Try this: - exact Nat.add_pos_right a hb + [apply] exact Nat.add_pos_right a hb -/ #guard_msgs in -- Verify that if maxHeartbeats is 0 we don't stop immediately. @@ -152,14 +152,14 @@ section synonym /-- info: Try this: - exact Nat.add_pos_right a hb + [apply] exact Nat.add_pos_right a hb -/ #guard_msgs in example (a b : Nat) (_ha : a > 0) (hb : 0 < b) : 0 < a + b := by apply? /-- info: Try this: - exact Nat.le_of_dvd w h + [apply] exact Nat.le_of_dvd w h -/ #guard_msgs in example (a b : Nat) (h : a ∣ b) (w : b > 0) : a ≤ b := @@ -167,7 +167,7 @@ by apply? /-- info: Try this: - exact Nat.le_of_dvd w h + [apply] exact Nat.le_of_dvd w h -/ #guard_msgs in example (a b : Nat) (h : a ∣ b) (w : b > 0) : b ≥ a := by apply? @@ -175,13 +175,13 @@ example (a b : Nat) (h : a ∣ b) (w : b > 0) : b ≥ a := by apply? -- TODO: A lemma with head symbol `¬` can be used to prove `¬ p` or `⊥` /-- info: Try this: - exact Nat.not_lt_zero a + [apply] exact Nat.not_lt_zero a -/ #guard_msgs in example (a : Nat) : ¬ (a < 0) := by apply? /-- info: Try this: - exact Nat.not_succ_le_zero a h + [apply] exact Nat.not_succ_le_zero a h -/ #guard_msgs in example (a : Nat) (h : a < 0) : False := by apply? @@ -200,20 +200,20 @@ theorem lemma_with_false_in_head (a b : Nat) (_h1 : a < b) (h2 : P a) : False := /-- info: Try this: - exact lemma_with_gt_in_head a h + [apply] exact lemma_with_gt_in_head a h -/ #guard_msgs in example (a : Nat) (h : P a) : 0 > a := by apply? /-- info: Try this: - exact lemma_with_gt_in_head a h + [apply] exact lemma_with_gt_in_head a h -/ #guard_msgs in example (a : Nat) (h : P a) : a < 0 := by apply? /-- info: Try this: - exact lemma_with_false_in_head a b h1 h2 + [apply] exact lemma_with_false_in_head a b h1 h2 -/ #guard_msgs in example (a b : Nat) (h1 : a < b) (h2 : P a) : False := by apply? @@ -225,7 +225,7 @@ end synonym /-- info: Try this: - exact fun P => iff_not_self + [apply] exact fun P => iff_not_self -/ #guard_msgs in example : ∀ P : Prop, ¬(P ↔ ¬P) := by apply? @@ -241,22 +241,21 @@ theorem p_iff_q (i : Nat) : P i ↔ Q i := /-- info: Try this: - exact (p_iff_q a).mp h + [apply] exact (p_iff_q a).mp h -/ #guard_msgs in example {a : Nat} (h : P a) : Q a := by apply? /-- info: Try this: - exact (p_iff_q a).mpr h + [apply] exact (p_iff_q a).mpr h -/ #guard_msgs in example {a : Nat} (h : Q a) : P a := by apply? /-- info: Try this: - exact - (Nat.dvd_add_iff_left h₁).mpr h₂ + [apply] exact (Nat.dvd_add_iff_left h₁).mpr h₂ -/ #guard_msgs in example {a b c : Nat} (h₁ : a ∣ c) (h₂ : a ∣ b + c) : a ∣ b := by apply? @@ -271,14 +270,14 @@ axiom F (a b : Nat) : f a ≤ f b ↔ a ≤ b /-- info: Try this: - exact (F a b).mpr h + [apply] exact (F a b).mpr h -/ #guard_msgs in example (a b : Nat) (h : a ≤ b) : f a ≤ f b := by apply? /-- info: Try this: - exact L.flatten + [apply] exact L.flatten -/ #guard_msgs in example (L : List (List Nat)) : List Nat := by apply? using L @@ -306,7 +305,7 @@ theorem Bool_eq_iff {A B : Bool} : (A = B) = (A ↔ B) := /-- info: Try this: - exact Bool_eq_iff + [apply] exact Bool_eq_iff -/ #guard_msgs in theorem Bool_eq_iff2 {A B : Bool} : (A = B) = (A ↔ B) := by @@ -347,7 +346,7 @@ theorem Bool_eq_iff2 {A B : Bool} : (A = B) = (A ↔ B) := by /-- info: Try this: - refine Int.mul_ne_zero ?_ h + [apply] refine Int.mul_ne_zero ?_ h -- Remaining subgoals: -- ⊢ 2 ≠ 0 --- @@ -361,7 +360,7 @@ example {x : Int} (h : x ≠ 0) : 2 * x ≠ 0 := by -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/.60exact.3F.60.20failure.3A.20.22maximum.20recursion.20depth.20has.20been.20reached.22/near/417649319 /-- info: Try this: - exact Nat.add_comm n m + [apply] exact Nat.add_comm n m -/ #guard_msgs in example (_h : List.range 10000 = List.range 10000) (n m : Nat) : n + m = m + n := by diff --git a/tests/lean/librarySearch.lean.expected.out b/tests/lean/librarySearch.lean.expected.out index abf668aafe89..67b429dfd3dd 100644 --- a/tests/lean/librarySearch.lean.expected.out +++ b/tests/lean/librarySearch.lean.expected.out @@ -1 +1 @@ -librarySearch.lean:384:0-384:7: warning: declaration uses 'sorry' +librarySearch.lean:383:0-383:7: warning: declaration uses 'sorry' diff --git a/tests/lean/run/3501.lean b/tests/lean/run/3501.lean index 4925d0622827..54d05f8fba14 100644 --- a/tests/lean/run/3501.lean +++ b/tests/lean/run/3501.lean @@ -1,6 +1,6 @@ /-- info: Try this: - simp only [Nat.reduceMul, a] + [apply] simp only [Nat.reduceMul, a] -/ #guard_msgs in example : True := by diff --git a/tests/lean/run/3519.lean b/tests/lean/run/3519.lean index b49436e73614..260cf594f4c8 100644 --- a/tests/lean/run/3519.lean +++ b/tests/lean/run/3519.lean @@ -1,6 +1,6 @@ /-- info: Try this: - simp only [x] + [apply] simp only [x] --- warning: declaration uses 'sorry' -/ @@ -12,7 +12,7 @@ example {P : Nat → Prop} : let x := 0; P x := by /-- info: Try this: - simp_all only [x] + [apply] simp_all only [x] --- warning: declaration uses 'sorry' -/ diff --git a/tests/lean/run/3547.lean b/tests/lean/run/3547.lean index 5d2d040bc501..83c90eb1683e 100644 --- a/tests/lean/run/3547.lean +++ b/tests/lean/run/3547.lean @@ -5,7 +5,7 @@ decreasing_by decreasing_tactic /-- info: Try this: - simp only [foo] + [apply] simp only [foo] -/ #guard_msgs in def foo2 : foo 2 = 0 := by diff --git a/tests/lean/run/3710.lean b/tests/lean/run/3710.lean index 4aeef16a7603..1b23953dd543 100644 --- a/tests/lean/run/3710.lean +++ b/tests/lean/run/3710.lean @@ -24,7 +24,7 @@ theorem indicator_of_not {s : Set} {a : Nat} (h : ¬ s a) : indicator s a = 0 := /-- info: Try this: - simp only [compl_iff, singleton_iff, not_true_eq_false, not_false_eq_true, indicator_of_not] + [apply] simp only [compl_iff, singleton_iff, not_true_eq_false, not_false_eq_true, indicator_of_not] -/ #guard_msgs in theorem test : indicator (compl <| singleton 0) 0 = 0 := by diff --git a/tests/lean/run/3922.lean b/tests/lean/run/3922.lean index ea926926053e..ea62e85c606a 100644 --- a/tests/lean/run/3922.lean +++ b/tests/lean/run/3922.lean @@ -19,7 +19,7 @@ axiom r.trans {a b c : Nat} : r a b → r b c → r a c /-- info: Try this: - refine r.symm ?_ + [apply] refine r.symm ?_ -- Remaining subgoals: -- ⊢ r a c --- @@ -39,7 +39,7 @@ attribute [symm] r.symm /-- info: Try this: - exact r.trans (id (r.symm h₂)) h₁ + [apply] exact r.trans (id (r.symm h₂)) h₁ -/ #guard_msgs in example (a b c : Nat) (h₁ : r b a) (h₂ : r b c) : r c a := by diff --git a/tests/lean/run/4230.lean b/tests/lean/run/4230.lean index 070fb5593ced..f01b153d604b 100644 --- a/tests/lean/run/4230.lean +++ b/tests/lean/run/4230.lean @@ -7,7 +7,7 @@ termination_by input.size - curr /-- info: Try this: - termination_by input.size - curr + [apply] termination_by input.size - curr -/ #guard_msgs(drop warning, info) in theorem foo (curr : Nat) (input : Array Nat) (output : Array Nat) diff --git a/tests/lean/run/4251.lean b/tests/lean/run/4251.lean index 2b4b8e050085..a3074b885634 100644 --- a/tests/lean/run/4251.lean +++ b/tests/lean/run/4251.lean @@ -1,6 +1,6 @@ /-- info: Try this: - simp only [ha, Nat.reduceEqDiff, imp_self] + [apply] simp only [ha, Nat.reduceEqDiff, imp_self] -/ #guard_msgs in theorem foo₁ (a : Nat) (ha : a = 37) : diff --git a/tests/lean/run/4534.lean b/tests/lean/run/4534.lean index 41284ba4fe65..eff1b02a48f5 100644 --- a/tests/lean/run/4534.lean +++ b/tests/lean/run/4534.lean @@ -13,7 +13,7 @@ instance : MyClass Prop where /-- info: Try this: - simp only [MyClass.le_refl, MyClass.sup_of_le_left] + [apply] simp only [MyClass.le_refl, MyClass.sup_of_le_left] -/ #guard_msgs in example : MyClass.sup False False = False := by diff --git a/tests/lean/run/5407.lean b/tests/lean/run/5407.lean index e8784b7622e5..9d3042c63f9d 100644 --- a/tests/lean/run/5407.lean +++ b/tests/lean/run/5407.lean @@ -40,7 +40,7 @@ axiom a : A axiom b : B /-- info: Try this: - (expose_names; exact imp h_1 h) + [apply] (expose_names; exact imp h_1 h) -/ #guard_msgs in example : C := by @@ -66,7 +66,7 @@ example : EqExplicit (fun (f : α → β) => (fun g x => g x) f) id := by /-! Suggests `expose_names` if a name in the syntax contains macro scopes -/ /-- info: Try this: - (expose_names; exact h) + [apply] (expose_names; exact h) -/ #guard_msgs in example {P : Prop} : P → P := by @@ -79,7 +79,7 @@ opaque D : Nat → Type axiom c_of_d {n : Nat} : D n → C /-- info: Try this: - exact c_of_d x + [apply] exact c_of_d x -/ #guard_msgs in example : (n : Nat) → D n → C := by @@ -94,7 +94,7 @@ axiom option1 : A → E axiom option2 {_ : B} : E /-- info: Try this: - refine option1 ?_ + [apply] refine option1 ?_ -- Remaining subgoals: -- ⊢ A --- @@ -117,7 +117,7 @@ opaque R : A → B → Prop axiom rImp (b : B) : R a b → R a b /-- info: Try this: - (expose_names; refine rImp b ?_) + [apply] (expose_names; refine rImp b ?_) -- Remaining subgoals: -- ⊢ R a b --- @@ -132,7 +132,7 @@ example : (b : B) → R a b := by /-! `show_term` exhibits the same behavior as `exact?` -/ /-- info: Try this: - (expose_names; exact h) + [apply] (expose_names; exact h) -/ #guard_msgs in example : E → E := by @@ -154,7 +154,7 @@ example : B → E := by show_term apply option2 /-- info: Try this: - exact c_of_d d + [apply] exact c_of_d d -/ #guard_msgs in example : (n : Nat) → D n → C := by @@ -174,7 +174,7 @@ axiom eq (a' : A) : a = a' /-- info: Try this: - (expose_names; rw [eq a']) + [apply] (expose_names; rw [eq a']) -- Foo a' a' --- error: unsolved goals diff --git a/tests/lean/run/5634.lean b/tests/lean/run/5634.lean index 801f8aa2f4a0..ede2be1a7f42 100644 --- a/tests/lean/run/5634.lean +++ b/tests/lean/run/5634.lean @@ -86,14 +86,14 @@ Regression test: make sure `simpa?` reports lemmas for both the goal and the `us /-- info: Try this: - simpa only [id] using h + [apply] simpa only [id] using h -/ #guard_msgs in example (p : Prop) (h : p) : id p := by simpa? only [id] using h /-- info: Try this: - simpa only [id] using h + [apply] simpa only [id] using h -/ #guard_msgs in example (p : Prop) (h : id p) : p := by simpa? only [id] using h diff --git a/tests/lean/run/5993.lean b/tests/lean/run/5993.lean index f440757946ff..9e3531896fd1 100644 --- a/tests/lean/run/5993.lean +++ b/tests/lean/run/5993.lean @@ -12,7 +12,7 @@ No field notation notation here. Used to print `refine ?_.succ` and `refine ?_.s /-- info: Try this: - refine Nat.succ ?_ + [apply] refine Nat.succ ?_ --- info: found a partial proof, but the corresponding tactic failed: (expose_names; refine Prod.snd ?_) @@ -31,7 +31,7 @@ No field notation even under binders. (That is, be aware of delayed assignment m /-- info: Try this: - refine fun x => Nat.succ ?_ + [apply] refine fun x => Nat.succ ?_ -/ #guard_msgs in example : Nat → Nat := by diff --git a/tests/lean/run/6164.lean b/tests/lean/run/6164.lean index df92644ac1c6..d3d33db3d044 100644 --- a/tests/lean/run/6164.lean +++ b/tests/lean/run/6164.lean @@ -9,7 +9,7 @@ attribute [simp] Nat.two_mul /-- info: Try this: - simp only [Nat.two_mul] + [apply] simp only [Nat.two_mul] -/ #guard_msgs in example (n : Nat) : 123 + 2 * n = 123 + (n + n) := by @@ -22,7 +22,7 @@ example (n : Nat) : 123 + 2 * n = 123 + (n + n) := by /-- info: Try this: - dsimp only [foo] + [apply] dsimp only [foo] -/ #guard_msgs in example (n : Nat) : foo n = n + 1 := by diff --git a/tests/lean/run/6655.lean b/tests/lean/run/6655.lean index c3641d4ccacc..2b586accc401 100644 --- a/tests/lean/run/6655.lean +++ b/tests/lean/run/6655.lean @@ -18,7 +18,7 @@ Example from #6655. This used to suggest `simp only [e, d]`. -/ /-- info: Try this: - simp only [e] + [apply] simp only [e] --- trace: α : Type c : α → α @@ -43,7 +43,7 @@ Example from #6655. This used to suggest `simp only [d]`. -/ /-- info: Try this: - simp only + [apply] simp only --- warning: declaration uses 'sorry' -/ @@ -62,7 +62,7 @@ Example from comments of #6655. This used to suggest `simp only [Int.add_sub_can -/ /-- info: Try this: - simp only [Int.add_sub_cancel] + [apply] simp only [Int.add_sub_cancel] -/ #guard_msgs in example (a b : Int) : a + b - b = a := by @@ -77,10 +77,10 @@ This used to suggest `simp only [a, b] ` and `simp only [a, b]` -/ /-- info: Try this: - simp only [a] + [apply] simp only [a] --- info: Try this: - simp only + [apply] simp only -/ #guard_msgs in example : True := by @@ -96,7 +96,7 @@ Test that there is still a deficiency. This should say `simp only [e]`. -/ /-- info: Try this: - simp only [e, c] + [apply] simp only [e, c] --- trace: α : Type b : α → α diff --git a/tests/lean/run/6706.lean b/tests/lean/run/6706.lean index 703975817b5a..c8398ef5199c 100644 --- a/tests/lean/run/6706.lean +++ b/tests/lean/run/6706.lean @@ -11,7 +11,7 @@ def P := True theorem N.A.B : P := trivial /-- info: Try this: - simp only [N.A.B] + [apply] simp only [N.A.B] -/ #guard_msgs in theorem N.X.A.B : P := by @@ -19,7 +19,7 @@ theorem N.X.A.B : P := by /-- info: Try this: - simp only [_root_.N.A.B] + [apply] simp only [_root_.N.A.B] -/ #guard_msgs in theorem A : P := diff --git a/tests/lean/run/7927.lean b/tests/lean/run/7927.lean index 45e376268924..d33a88bd68ea 100644 --- a/tests/lean/run/7927.lean +++ b/tests/lean/run/7927.lean @@ -6,21 +6,21 @@ example : f true = true := by simp! /-- info: Try this: - simp! only [f] + [apply] simp! only [f] -/ #guard_msgs in example : f true = true := by simp?! /-- info: Try this: - simp_all! only [f] + [apply] simp_all! only [f] -/ #guard_msgs in example : f true = true := by simp_all?! /-- info: Try this: - dsimp! only [f] + [apply] dsimp! only [f] -/ #guard_msgs in example : f true = true := by dsimp?! diff --git a/tests/lean/run/allGoals.lean b/tests/lean/run/allGoals.lean index b4b89237d76a..20d403350ce5 100644 --- a/tests/lean/run/allGoals.lean +++ b/tests/lean/run/allGoals.lean @@ -125,8 +125,8 @@ but is expected to have type false = false --- info: Try this: - Bool.casesOn (motive := fun t => b = t → b = b) b - (fun h => Eq.symm h ▸ sorry) (fun h => Eq.symm h ▸ Eq.refl true) (Eq.refl b) + [apply] Bool.casesOn (motive := fun t => b = t → b = b) b (fun h => Eq.symm h ▸ sorry) + (fun h => Eq.symm h ▸ Eq.refl true) (Eq.refl b) -/ #guard_msgs in example (b : Bool) : b = b := by? @@ -221,7 +221,7 @@ error: Failed: `fail` tactic was invoked ⊢ True --- info: Try this: - sorry + [apply] sorry -/ #guard_msgs in example : True := by? @@ -241,7 +241,7 @@ use `set_option maxRecDepth ` to increase limit use `set_option diagnostics true` to get diagnostic information --- info: Try this: - sorry + [apply] sorry -/ #guard_msgs in example : True := by? diff --git a/tests/lean/run/grind_constProp.lean b/tests/lean/run/grind_constProp.lean index ecb4dc1a4eeb..aa8994815ba0 100644 --- a/tests/lean/run/grind_constProp.lean +++ b/tests/lean/run/grind_constProp.lean @@ -209,7 +209,7 @@ def evalExpr (e : Expr) : EvalM Val := do /-- info: Try this: - fun_induction Expr.simplify <;> grind + [apply] (fun_induction Expr.simplify) <;> grind -/ #guard_msgs (info) in example (e : Expr) : e.simplify.eval σ = e.eval σ := by @@ -311,7 +311,7 @@ theorem State.cons_le_of_eq (h₁ : σ' ≼ σ) (h₂ : σ.find? x = some v) : ( /-- info: Try this: - fun_induction join <;> grind + [apply] (fun_induction join) <;> grind -/ #guard_msgs (info) in open State in diff --git a/tests/lean/run/grind_ematch_gen_pattern.lean b/tests/lean/run/grind_ematch_gen_pattern.lean index 7d6f8dd28bb4..451725fe5f68 100644 --- a/tests/lean/run/grind_ematch_gen_pattern.lean +++ b/tests/lean/run/grind_ematch_gen_pattern.lean @@ -10,7 +10,7 @@ example (h : b = some a) : (b.pbind fun a h => some <| a + f b (by grind)) = som /-- info: Try this: - grind only [= gen Option.pbind_some', f, cases Or] + [apply] grind only [= gen Option.pbind_some', f, cases Or] -/ #guard_msgs (info) in example (h : b = some a) : (b.pbind fun a h => some <| a + f b (by grind)) = some (a + a) := by diff --git a/tests/lean/run/grind_eval_suggest.lean b/tests/lean/run/grind_eval_suggest.lean index 86301d04d3d0..36e67a2aaf97 100644 --- a/tests/lean/run/grind_eval_suggest.lean +++ b/tests/lean/run/grind_eval_suggest.lean @@ -15,10 +15,10 @@ opaque f : Nat → Nat /-- info: Try these: - • simp +arith - • simp +arith only [Nat.reduceAdd, fthm] - • grind - • grind only [= fthm] + [apply] simp +arith + [apply] simp +arith only [Nat.reduceAdd, fthm] + [apply] grind + [apply] grind only [= fthm] -/ #guard_msgs (info) in example (x : Nat) : 1 + 1 + f x = x + 2 := by @@ -26,11 +26,11 @@ example (x : Nat) : 1 + 1 + f x = x + 2 := by /-- info: Try these: - • rfl - • simp - • simp only [Nat.succ_eq_add_one, Nat.add_left_cancel_iff] - • grind - • grind only + [apply] rfl + [apply] simp + [apply] simp only [Nat.succ_eq_add_one, Nat.add_left_cancel_iff] + [apply] grind + [apply] grind only -/ #guard_msgs (info) in example (x : Nat) : x + 1 = Nat.succ x := by @@ -38,11 +38,11 @@ example (x : Nat) : x + 1 = Nat.succ x := by /-- info: Try these: - • · intros; rfl - • · intros; simp - • · intros; simp only [Nat.succ_eq_add_one, Nat.add_left_cancel_iff] - • · intros; grind - • · intros; grind only + [apply] · intros; rfl + [apply] · intros; simp + [apply] · intros; simp only [Nat.succ_eq_add_one, Nat.add_left_cancel_iff] + [apply] · intros; grind + [apply] · intros; grind only -/ #guard_msgs (info) in example (x : Nat) : True → x + 1 = Nat.succ x := by @@ -50,9 +50,9 @@ example (x : Nat) : True → x + 1 = Nat.succ x := by /-- info: Try these: - • simp_all - • grind - • grind only + [apply] simp_all + [apply] grind + [apply] grind only -/ #guard_msgs (info) in example (h : 0 + x = y) : f x = f y := by @@ -74,7 +74,7 @@ macro "simple_tac" : tactic => `(tactic| eval_suggest (intros; skip; first | ski /-- info: Try this: - simp + [apply] simp -/ #guard_msgs (info) in example : True ∧ True := by @@ -89,7 +89,7 @@ macro "simple_tac2" : tactic => `(tactic| eval_suggest (intros; (simp only [Nat. /-- info: Try this: - · intros; simp only [Nat.zero_add]; simp only [Nat.one_mul]; simp [*] + [apply] · intros; simp only [Nat.zero_add]; simp only [Nat.one_mul]; simp [*] -/ #guard_msgs (info) in example : x = 0 → 0 + 1*x = 0 := by diff --git a/tests/lean/run/grind_option.lean b/tests/lean/run/grind_option.lean index 9f63c15093b5..a5fff8b7ae22 100644 --- a/tests/lean/run/grind_option.lean +++ b/tests/lean/run/grind_option.lean @@ -9,32 +9,29 @@ variable [BEq α] {o₁ o₂ o₃ o₄ o₅ : Option α} /-- info: Try this: - grind only [= - Option.or_some, - = Option.some_beq_none, = Option.getD_or, = Option.some_or, =_ Option.or_assoc, = Option.or_assoc] + [apply] grind only [= Option.or_some, = Option.some_beq_none, = Option.getD_or, = Option.some_or, =_ Option.or_assoc, + = Option.or_assoc] -/ #guard_msgs in example : ((o₁.or (o₂.or (some x))).or (o₄.or o₅) == none) = false := by grind? /-- info: Try this: - grind only [= Option.max_none_right, - = Option.min_some_some, = Nat.min_def] + [apply] grind only [= Option.max_none_right, = Option.min_some_some, = Nat.min_def] -/ #guard_msgs in example : max (some 7) none = min (some 13) (some 7) := by grind? /-- info: Try this: - grind only [= Option.guard_apply] + [apply] grind only [= Option.guard_apply] -/ #guard_msgs in example : Option.guard (· ≤ 7) 3 = some 3 := by grind? /-- info: Try this: - grind only [= - Option.mem_bind_iff] + [apply] grind only [= Option.mem_bind_iff] -/ #guard_msgs in example {x : β} {o : Option α} {f : α → Option β} (h : a ∈ o) (h' : x ∈ f a) : x ∈ o.bind f := by grind? diff --git a/tests/lean/run/grind_pat_sel.lean b/tests/lean/run/grind_pat_sel.lean index 8583274e6166..cfbc3ee5fe93 100644 --- a/tests/lean/run/grind_pat_sel.lean +++ b/tests/lean/run/grind_pat_sel.lean @@ -64,35 +64,35 @@ opaque r : Nat → Nat → Nat /-- info: Try these: - • [grind! .] for pattern: [r #2 (f #1)] - • [grind! =>] for pattern: [f #2, f #1] + [apply] [grind! .] for pattern: [r #2 (f #1)] + [apply] [grind! =>] for pattern: [f #2, f #1] -/ #guard_msgs in @[grind!] theorem pr₅ : p (f x) → r x (f y) = y := sorry /-- info: Try these: - • [grind! .] for pattern: [f #1, f #2] - • [grind! =>] for pattern: [f #2, f #1] + [apply] [grind! .] for pattern: [f #1, f #2] + [apply] [grind! =>] for pattern: [f #2, f #1] -/ #guard_msgs in @[grind!] theorem pr₆ : p (f x) → r y (f y) = y := sorry /-- info: Try these: - • [grind .] for pattern: [r #1 (f #1), p (f #2)] - • [grind =>] for pattern: [p (f #2), r #1 (f #1)] - • [grind! .] for pattern: [f #1, f #2] - • [grind! =>] for pattern: [f #2, f #1] + [apply] [grind .] for pattern: [r #1 (f #1), p (f #2)] + [apply] [grind =>] for pattern: [p (f #2), r #1 (f #1)] + [apply] [grind! .] for pattern: [f #1, f #2] + [apply] [grind! =>] for pattern: [f #2, f #1] -/ #guard_msgs in @[grind] theorem pr₇ : p (f x) → r y (f y) = y := sorry /-- info: Try these: - • [grind =] for pattern: [r #2 (f #1)] - • [grind =>] for pattern: [p (f #2), r #2 (f #1)] - • [grind! =>] for pattern: [f #2, f #1] + [apply] [grind =] for pattern: [r #2 (f #1)] + [apply] [grind =>] for pattern: [p (f #2), r #2 (f #1)] + [apply] [grind! =>] for pattern: [f #2, f #1] -/ #guard_msgs in @[grind] theorem pr₈ : p (f x) → r x (f y) = y := sorry @@ -100,24 +100,24 @@ info: Try these: /-- info: Try these: - • [grind =] for pattern: [f (g #0)] - • [grind! .] for pattern: [g #0] + [apply] [grind =] for pattern: [f (g #0)] + [apply] [grind! .] for pattern: [g #0] -/ #guard_msgs in @[grind] axiom fg₆ : f (g x) = x /-- info: Try these: - • [grind =] for pattern: [f (g #0)] - • [grind =_] for pattern: [r #0 #0] - • [grind! .] for pattern: [g #0] + [apply] [grind =] for pattern: [f (g #0)] + [apply] [grind =_] for pattern: [r #0 #0] + [apply] [grind! .] for pattern: [g #0] -/ #guard_msgs in @[grind] axiom fg₇ : f (g x) = r x x /-- info: Try this: - [grind =] for pattern: [f #0] + [apply] [grind =] for pattern: [f #0] -/ #guard_msgs in @[grind] axiom fg₈ : f x = x diff --git a/tests/lean/run/grind_trace.lean b/tests/lean/run/grind_trace.lean index 74df4982ce1b..12ae974d454c 100644 --- a/tests/lean/run/grind_trace.lean +++ b/tests/lean/run/grind_trace.lean @@ -21,7 +21,7 @@ attribute [grind =] List.getElem?_cons_zero in /-- info: Try this: - grind only [= List.getElem?_eq_none, = List.getElem?_replicate, = List.getElem?_eq_getElem] + [apply] grind only [= List.getElem?_eq_none, = List.getElem?_replicate, = List.getElem?_eq_getElem] -/ #guard_msgs (info) in theorem getElem?_replicate' : (List.replicate n a)[m]? = if m < n then some a else none := by @@ -29,7 +29,7 @@ theorem getElem?_replicate' : (List.replicate n a)[m]? = if m < n then some a el /-- info: Try this: - grind only [= List.length_cons] + [apply] grind only [= List.length_cons] -/ #guard_msgs (info) in example : 0 < (x :: t).length := by @@ -38,9 +38,9 @@ example : 0 < (x :: t).length := by attribute [grind ext] List.ext_getElem? /-- info: Try this: - grind only [= List.getElem?_eq_some_iff, = List.length_replicate, = List.getElem?_eq_none, - = List.getElem_replicate, = Option.map_some, = Option.map_none, = List.getElem?_replicate, - = List.getElem?_eq_getElem, = List.getElem?_map] + [apply] grind only [= List.getElem?_eq_some_iff, = List.length_replicate, = List.getElem?_eq_none, + = List.getElem_replicate, = Option.map_some, = Option.map_none, = List.getElem?_replicate, + = List.getElem?_eq_getElem, = List.getElem?_map] -/ #guard_msgs (info) in theorem map_replicate' : (List.replicate n a).map f = List.replicate n (f a) := by @@ -48,7 +48,7 @@ theorem map_replicate' : (List.replicate n a).map f = List.replicate n (f a) := /-- info: Try this: - grind only [= List.getLast?_eq_some_iff, ← List.mem_concat_self] + [apply] grind only [= List.getLast?_eq_some_iff, ← List.mem_concat_self] -/ #guard_msgs (info) in theorem mem_of_getLast?_eq_some' {xs : List α} {a : α} (h : xs.getLast? = some a) : a ∈ xs := by @@ -60,7 +60,7 @@ theorem mem_of_getLast?_eq_some' {xs : List α} {a : α} (h : xs.getLast? = some /-- info: Try this: - grind only + [apply] grind only -/ #guard_msgs (info) in example : x = 0 → f x = 1 := by @@ -71,7 +71,7 @@ attribute [grind] f /-- info: Try this: - grind only [f] + [apply] grind only [f] -/ #guard_msgs (info) in example : x = 0 → f x = 1 := by @@ -85,7 +85,7 @@ grind_pattern gthm => g (g x) /-- info: Try this: - grind only [usr gthm] + [apply] grind only [usr gthm] -/ #guard_msgs (info) in example : g (g (g x)) = g x := by diff --git a/tests/lean/run/grind_try_exact.lean b/tests/lean/run/grind_try_exact.lean index 9bae4848790b..719f8049d26a 100644 --- a/tests/lean/run/grind_try_exact.lean +++ b/tests/lean/run/grind_try_exact.lean @@ -7,7 +7,7 @@ theorem Pall : Q x → P x := sorry /-- info: Try this: - exact Pall h + [apply] exact Pall h -/ #guard_msgs (info) in example (h : Q x) (_ : x > 0) : P x := by @@ -15,7 +15,7 @@ example (h : Q x) (_ : x > 0) : P x := by /-- info: Try this: - · intros; expose_names; exact Pall h + [apply] · intros; expose_names; exact Pall h -/ #guard_msgs (info) in example: Q x → True → P x := by @@ -23,7 +23,7 @@ example: Q x → True → P x := by /-- info: Try this: - · intros; expose_names; exact Pall h_1 + [apply] · intros; expose_names; exact Pall h_1 -/ #guard_msgs (info) in example: True → Q x → True → P x := by @@ -43,7 +43,7 @@ example : Q x := by /-- info: Try this: - · expose_names; exact Pall h + [apply] · expose_names; exact Pall h -/ #guard_msgs (info) in example (_ : Q x) (_ : x > 0) : P x := by diff --git a/tests/lean/run/grind_try_extend.lean b/tests/lean/run/grind_try_extend.lean index b396f77bed24..8d3a43e5eb2d 100644 --- a/tests/lean/run/grind_try_extend.lean +++ b/tests/lean/run/grind_try_extend.lean @@ -13,7 +13,7 @@ meta def evalTryApply : TryTactic := fun tac => do /-- info: Try this: - · trace "worked"; assumption + [apply] · trace "worked"; assumption -/ #guard_msgs (info) in example (h : False) : False := by @@ -28,8 +28,8 @@ elab stx:"my_try?" : tactic => do /-- info: Try these: - • · trace "worked"; assumption - • rfl + [apply] · trace "worked"; assumption + [apply] rfl -/ #guard_msgs (info) in example (a : Nat) (h : a = a) : a = a := by diff --git a/tests/lean/run/grind_try_trace.lean b/tests/lean/run/grind_try_trace.lean index 9ffe8a1adf14..6e9df71d0403 100644 --- a/tests/lean/run/grind_try_trace.lean +++ b/tests/lean/run/grind_try_trace.lean @@ -4,11 +4,11 @@ reset_grind_attrs% /-- info: Try these: - • simp - • simp only [ne_eq, reduceCtorEq, not_false_eq_true] - • grind - • grind only - • simp_all + [apply] simp + [apply] simp only [ne_eq, reduceCtorEq, not_false_eq_true] + [apply] grind + [apply] grind only + [apply] simp_all -/ #guard_msgs (info) in example : [1, 2] ≠ [] := by @@ -16,10 +16,10 @@ example : [1, 2] ≠ [] := by /-- info: Try these: - • simp +arith - • simp +arith only [ge_iff_le] - • grind - • grind only + [apply] simp +arith + [apply] simp +arith only [ge_iff_le] + [apply] grind + [apply] grind only -/ #guard_msgs (info) in example : 4 + x + y ≥ 1 + x := by @@ -27,8 +27,8 @@ example : 4 + x + y ≥ 1 + x := by /-- info: Try these: - • simp +arith - • grind + [apply] simp +arith + [apply] grind -/ #guard_msgs (info) in example : 4 + x + y ≥ 1 + x := by @@ -36,8 +36,8 @@ example : 4 + x + y ≥ 1 + x := by /-- info: Try these: - • grind - • grind only + [apply] grind + [apply] grind only -/ #guard_msgs (info) in example (f : Nat → Nat) : f a = b → a = c → f c = b := by @@ -49,8 +49,8 @@ public def f : Nat → Nat /-- info: Try these: - • grind [= f] - • grind only [f] + [apply] grind [= f] + [apply] grind only [f] -/ #guard_msgs (info) in example : f (f 0) > 0 := by @@ -58,8 +58,8 @@ example : f (f 0) > 0 := by /-- info: Try these: - • grind [= f.eq_def] - • grind only [= f.eq_def] + [apply] grind [= f.eq_def] + [apply] grind only [= f.eq_def] -/ #guard_msgs (info) in example : f x > 0 := by @@ -71,9 +71,9 @@ def app : List α → List α → List α /-- info: Try these: - • rfl - • grind [= app] - • grind only [app] + [apply] rfl + [apply] grind [= app] + [apply] grind only [app] -/ #guard_msgs (info) in example : app [a, b] [c] = [a, b, c] := by @@ -81,8 +81,8 @@ example : app [a, b] [c] = [a, b, c] := by /-- info: Try these: - • fun_induction app as bs <;> grind [= app] - • fun_induction app as bs <;> grind only [app] + [apply] (fun_induction app as bs) <;> grind [= app] + [apply] (fun_induction app as bs) <;> grind only [app] -/ #guard_msgs (info) in example : app (app as bs) cs = app as (app bs cs) := by @@ -90,7 +90,7 @@ example : app (app as bs) cs = app as (app bs cs) := by /-- info: Try this: - fun_induction app as bs <;> grind [= app] + [apply] (fun_induction app as bs) <;> grind [= app] -/ #guard_msgs (info) in example : app (app as bs) cs = app as (app bs cs) := by @@ -98,8 +98,8 @@ example : app (app as bs) cs = app as (app bs cs) := by /-- info: Try these: - • · expose_names; fun_induction app as bs_1 <;> grind [= app] - • · expose_names; fun_induction app as bs_1 <;> grind only [app] + [apply] · expose_names; fun_induction app as bs_1 <;> grind [= app] + [apply] · expose_names; fun_induction app as bs_1 <;> grind only [app] -/ #guard_msgs (info) in example : app (app as bs) cs = app as (app bs cs) := by @@ -108,8 +108,8 @@ example : app (app as bs) cs = app as (app bs cs) := by /-- info: Try these: - • · expose_names; fun_induction app as bs <;> grind [= app] - • · expose_names; fun_induction app as bs <;> grind only [app] + [apply] · expose_names; fun_induction app as bs <;> grind [= app] + [apply] · expose_names; fun_induction app as bs <;> grind only [app] -/ #guard_msgs (info) in example : app (app as bs) cs = app as (app bs cs) := by @@ -126,8 +126,8 @@ attribute [simp] concat /-- info: Try these: - • fun_induction concat <;> simp_all - • fun_induction concat <;> simp [*] + [apply] (fun_induction concat) <;> simp_all + [apply] (fun_induction concat) <;> simp [*] -/ #guard_msgs (info) in example (as : List α) (a : α) : concat as a = as ++ [a] := by @@ -135,10 +135,11 @@ example (as : List α) (a : α) : concat as a = as ++ [a] := by /-- info: Try these: - • fun_induction concat <;> simp_all - • · fun_induction concat - · simp - · simp [*] + [apply] (fun_induction concat) <;> simp_all + [apply] · + fun_induction concat + · simp + · simp [*] -/ #guard_msgs (info) in example (as : List α) (a : α) : concat as a = as ++ [a] := by @@ -150,8 +151,8 @@ def map (f : α → β) : List α → List β /-- info: Try these: - • fun_induction map f xs <;> grind [= map] - • fun_induction map f xs <;> grind only [map] + [apply] (fun_induction map f xs) <;> grind [= map] + [apply] (fun_induction map f xs) <;> grind only [map] -/ #guard_msgs (info) in theorem map_map (f : α → β) (g : β → γ) xs : @@ -165,15 +166,18 @@ def foo : Nat → Nat /-- info: Try these: - • · fun_induction foo - · simp - · sorry - • · fun_induction foo - · grind - · sorry - • · fun_induction foo - · simp_all - · sorry + [apply] · + fun_induction foo + · simp + · sorry + [apply] · + fun_induction foo + · grind + · sorry + [apply] · + fun_induction foo + · simp_all + · sorry -/ #guard_msgs (info) in example : foo x > 0 := by @@ -198,11 +202,11 @@ attribute [grind] List.length_reverse bla /-- info: Try these: - • fun_induction bla <;> grind - • fun_induction bla <;> simp_all - • fun_induction bla <;> simp [*] - • fun_induction bla <;> simp only [List.length_reverse, *] - • fun_induction bla <;> grind only [= List.length_reverse] + [apply] (fun_induction bla) <;> grind + [apply] (fun_induction bla) <;> simp_all + [apply] (fun_induction bla) <;> simp [*] + [apply] (fun_induction bla) <;> simp only [List.length_reverse, *] + [apply] (fun_induction bla) <;> grind only [= List.length_reverse] -/ #guard_msgs (info) in example : (bla xs ys).length = ys.length := by diff --git a/tests/lean/run/info_trees.lean b/tests/lean/run/info_trees.lean index f3621cd5d2e5..9d8b33053aef 100644 --- a/tests/lean/run/info_trees.lean +++ b/tests/lean/run/info_trees.lean @@ -70,7 +70,7 @@ info: • [Command] @ ⟨77, 0⟩-⟨77, 40⟩ @ Lean.Elab.Command.elabDeclarati • [Term] t (isBinder := true) : ∀ (n : Nat), 0 ≤ n @ ⟨77, 8⟩-⟨77, 9⟩ --- info: Try this: - exact Nat.zero_le n + [apply] exact Nat.zero_le n -/ #guard_msgs in #info_trees in diff --git a/tests/lean/run/issue8257.lean b/tests/lean/run/issue8257.lean index 273deeab14dc..2ebabd0888e4 100644 --- a/tests/lean/run/issue8257.lean +++ b/tests/lean/run/issue8257.lean @@ -1,6 +1,6 @@ /-- info: Try this: - termination_by xs.length / 2 - i + [apply] termination_by xs.length / 2 - i -/ #guard_msgs in def foo (xs : String) (i : Nat) (a b : String.Iterator) : Bool := @@ -14,7 +14,7 @@ termination_by? /-- info: Try this: - termination_by xs.length / 2 - i + [apply] termination_by xs.length / 2 - i -/ #guard_msgs in def bar (xs : String) (i : Nat) (a b : String.Iterator) : Bool := @@ -30,7 +30,7 @@ termination_by? /-- info: Try this: - termination_by xs.length / 2 - i + [apply] termination_by xs.length / 2 - i -/ #guard_msgs in def baz (xs : String) (i : Nat) (a b : String.Iterator) : Bool := diff --git a/tests/lean/run/issue8490.lean b/tests/lean/run/issue8490.lean index de88b2d26218..07599b586877 100644 --- a/tests/lean/run/issue8490.lean +++ b/tests/lean/run/issue8490.lean @@ -14,7 +14,7 @@ inductive_fixpoint /-- info: Try this: - simp_all only [asimp_const, reduceCtorEq] + [apply] simp_all only [asimp_const, reduceCtorEq] -/ #guard_msgs in example (x y : Aexp) (k : Int) (h : asimp_const (.val k) = x.plus y) : optimal' x := by diff --git a/tests/lean/run/issue8939.lean b/tests/lean/run/issue8939.lean index ec1c765af6c6..8af1ad63ea8b 100644 --- a/tests/lean/run/issue8939.lean +++ b/tests/lean/run/issue8939.lean @@ -11,7 +11,7 @@ public inductive AckFuel : (n m : Nat) → Type where namespace Test1 /-- info: Try this: - termination_by structural x _ x => x + [apply] termination_by structural x _ x => x -/ #guard_msgs in public def ackermann_fuel : (n m : Nat) → (fuel : AckFuel n m) → Nat @@ -25,7 +25,7 @@ end Test1 namespace Test2 /-- info: Try this: - termination_by structural x _ x => x + [apply] termination_by structural x _ x => x -/ #guard_msgs in public def ackermann_fuel : (n m : Nat) → (fuel : AckFuel n m) → Nat @@ -39,7 +39,7 @@ end Test2 namespace Test3 /-- info: Try this: - termination_by structural x _ x => x + [apply] termination_by structural x _ x => x -/ #guard_msgs in @[expose] public def ackermann_fuel : (n m : Nat) → (fuel : AckFuel n m) → Nat diff --git a/tests/lean/run/mvcgenInvariantsSuggestions.lean b/tests/lean/run/mvcgenInvariantsSuggestions.lean index d98dda2fb3a6..6e3dccf3a5df 100644 --- a/tests/lean/run/mvcgenInvariantsSuggestions.lean +++ b/tests/lean/run/mvcgenInvariantsSuggestions.lean @@ -16,8 +16,8 @@ def mySum (l : List Nat) : Nat := Id.run do /-- info: Try this: - invariants - · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = 0 ∨ xs.suffix = [] ∧ letMuts = l.sum⌝ + [apply] invariants + · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = 0 ∨ xs.suffix = [] ∧ letMuts = l.sum⌝ -/ #guard_msgs (info) in theorem mySum_suggest_invariant (l : List Nat) : mySum l = l.sum := by @@ -36,8 +36,8 @@ def mySum2 (l : List Nat) : Nat := Id.run do /-- info: Try this: - invariants - · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = ⟨0, 0⟩ ∨ xs.suffix = [] ∧ letMuts.fst = l.sum⌝ + [apply] invariants + · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = ⟨0, 0⟩ ∨ xs.suffix = [] ∧ letMuts.fst = l.sum⌝ -/ #guard_msgs (info) in theorem mySum2_suggest_invariant (l : List Nat) : mySum2 l = l.sum := by @@ -54,8 +54,8 @@ def copy (l : List Nat) : Id (Array Nat) := do /-- info: Try this: - invariants - · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = acc✝ ∨ xs.suffix = [] ∧ letMuts = l.toArray⌝ + [apply] invariants + · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = acc✝ ∨ xs.suffix = [] ∧ letMuts = l.toArray⌝ -/ #guard_msgs (info) in theorem copy_suggest_invariant (l : List Nat) : ⦃⌜True⌝⦄ copy l ⦃⇓ r => ⌜r = l.toArray⌝⦄ := by @@ -79,10 +79,10 @@ def nodup (l : List Int) : Bool := Id.run do /-- info: Try this: - invariants - · - Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜l.Nodup ∧ (r = true ↔ l.Nodup)⌝) - (onContinue := fun xs letMuts => ⌜xs.prefix = [] ∧ letMuts = ∅ ∨ xs.suffix = [] ∧ l.Nodup⌝) + [apply] invariants + · + Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜l.Nodup ∧ (r = true ↔ l.Nodup)⌝) (onContinue := + fun xs letMuts => ⌜xs.prefix = [] ∧ letMuts = ∅ ∨ xs.suffix = [] ∧ l.Nodup⌝) -/ #guard_msgs (info) in theorem nodup_suggest_invariant (l : List Int) : nodup l ↔ l.Nodup := by @@ -106,19 +106,18 @@ def nodup_twice (l : List Int) : Bool := Id.run do /-- info: Try this: - invariants - · - Invariant.withEarlyReturn (onReturn := fun r letMuts => - spred(Prod.fst ?inv2 ({ «prefix» := [], suffix := l, property := ⋯ }, ⟨none, ∅⟩) ∧ - { down := r = true ↔ l.Nodup })) - (onContinue := fun xs letMuts => - spred({ down := xs.prefix = [] ∧ letMuts = ∅ } ∨ - ⌜xs.suffix = []⌝ ∧ - Prod.fst ?inv2 ({ «prefix» := [], suffix := l, property := ⋯ }, ⟨none, ∅⟩) ∧ - { down := True })) - · - Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜l.Nodup ∧ (r = true ↔ l.Nodup)⌝) - (onContinue := fun xs letMuts => ⌜xs.prefix = [] ∧ letMuts = ∅ ∨ xs.suffix = [] ∧ l.Nodup⌝) + [apply] invariants + · + Invariant.withEarlyReturn (onReturn := fun r letMuts => + spred(Prod.fst ?inv2 ({ «prefix» := [], suffix := l, property := ⋯ }, ⟨none, ∅⟩) ∧ + { down := r = true ↔ l.Nodup })) + (onContinue := fun xs letMuts => + spred({ down := xs.prefix = [] ∧ letMuts = ∅ } ∨ + ⌜xs.suffix = []⌝ ∧ + Prod.fst ?inv2 ({ «prefix» := [], suffix := l, property := ⋯ }, ⟨none, ∅⟩) ∧ { down := True })) + · + Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜l.Nodup ∧ (r = true ↔ l.Nodup)⌝) (onContinue := + fun xs letMuts => ⌜xs.prefix = [] ∧ letMuts = ∅ ∨ xs.suffix = [] ∧ l.Nodup⌝) -/ #guard_msgs (info) in theorem nodup_twice_suggest_invariant (l : List Int) : nodup_twice l ↔ l.Nodup := by @@ -148,8 +147,8 @@ def mkFreshN (n : Nat) : AppM (List Nat) := do -- In the following, we suggest an inaccessible variable `acc` in the invariant. Unfortunate. /-- info: Try this: - invariants - · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = acc✝ ∨ xs.suffix = [] ∧ letMuts.toList.Nodup⌝ + [apply] invariants + · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = acc✝ ∨ xs.suffix = [] ∧ letMuts.toList.Nodup⌝ -/ #guard_msgs (info) in theorem mkFreshN_suggest_invariant (n : Nat) : ⦃⌜True⌝⦄ mkFreshN n ⦃⇓ r => ⌜r.Nodup⌝⦄ := by @@ -166,11 +165,10 @@ def mkFreshN_early_return (n : Nat) : AppM (List Nat) := do /-- info: Try this: - invariants - · - Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜letMuts.toList.Nodup ∧ r.Nodup⌝) - (onContinue := fun xs letMuts => - ⌜xs.prefix = [] ∧ letMuts = acc✝ ∨ xs.suffix = [] ∧ letMuts.toList.Nodup⌝) + [apply] invariants + · + Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜letMuts.toList.Nodup ∧ r.Nodup⌝) (onContinue := + fun xs letMuts => ⌜xs.prefix = [] ∧ letMuts = acc✝ ∨ xs.suffix = [] ∧ letMuts.toList.Nodup⌝) -/ #guard_msgs (info) in theorem mkFreshN_early_return_suggest_invariant (n : Nat) : ⦃⌜True⌝⦄ mkFreshN_early_return n ⦃⇓ r => ⌜r.Nodup⌝⦄ := by @@ -184,10 +182,10 @@ def earlyReturnWithoutLetMut (l : List Int) : Bool := Id.run do /-- info: Try this: - invariants - · - Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜r = true⌝) (onContinue := - fun xs letMuts => ⌜xs.prefix = [] ∨ xs.suffix = []⌝) + [apply] invariants + · + Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜r = true⌝) (onContinue := fun xs letMuts => + ⌜xs.prefix = [] ∨ xs.suffix = []⌝) -/ #guard_msgs (info) in theorem earlyReturnWithoutLetMut_suggest_invariant (l : List Int) : earlyReturnWithoutLetMut l := by @@ -208,9 +206,8 @@ def notQuiteEarlyReturn (l : List Nat) : Option Nat := Id.run do /-- info: Try this: - invariants - · ⇓⟨xs, letMuts⟩ => - ⌜xs.prefix = [] ∧ letMuts = ⟨none, ()⟩ ∨ xs.suffix = [] ∧ letMuts.fst = l.getLast?⌝ + [apply] invariants + · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = ⟨none, ()⟩ ∨ xs.suffix = [] ∧ letMuts.fst = l.getLast?⌝ -/ #guard_msgs (info) in theorem notQuiteEarlyReturn_suggest_invariant (l : List Nat) : notQuiteEarlyReturn l = l.getLast? := by @@ -229,8 +226,8 @@ def polySum [Monad m] (l : List Nat) : m Nat := do /-- info: Try this: - invariants - · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = ⟨0, 0⟩ ∨ xs.suffix = [] ∧ letMuts.fst = l.sum⌝ + [apply] invariants + · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = ⟨0, 0⟩ ∨ xs.suffix = [] ∧ letMuts.fst = l.sum⌝ -/ #guard_msgs (info) in theorem polySum_suggest_invariant [Monad m] [WPMonad m ps] (l : List Nat) : ⦃⌜True⌝⦄ @polySum m _ l ⦃⇓ r => ⌜r = l.sum⌝⦄ := by @@ -247,11 +244,10 @@ def polyNodup [Monad m] (l : List Int) : m Bool := do /-- info: Try this: - invariants - · - Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜l.Nodup ∧ (r = true ↔ l.Nodup)⌝) - (onContinue := fun xs letMuts => - ⌜xs.prefix = [] ∧ letMuts = seen✝ ∨ xs.suffix = [] ∧ l.Nodup⌝) + [apply] invariants + · + Invariant.withEarlyReturn (onReturn := fun r letMuts => ⌜l.Nodup ∧ (r = true ↔ l.Nodup)⌝) (onContinue := + fun xs letMuts => ⌜xs.prefix = [] ∧ letMuts = seen✝ ∨ xs.suffix = [] ∧ l.Nodup⌝) -/ #guard_msgs (info) in theorem polyNodup_suggest_invariant [Monad m] [WPMonad m ps] (l : List Int) : ⦃⌜True⌝⦄ @polyNodup m _ l ⦃⇓r => ⌜r ↔ l.Nodup⌝⦄ := by @@ -277,9 +273,8 @@ open Std.Do /-- info: Try this: - invariants - · ⇓⟨xs, letMuts⟩ => - ⌜xs.prefix = [] ∧ letMuts = ⟨n, x, 1⟩ ∨ xs.suffix = [] ∧ letMuts.2.snd = x ^ n⌝ + [apply] invariants + · ⇓⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = ⟨n, x, 1⟩ ∨ xs.suffix = [] ∧ letMuts.2.snd = x ^ n⌝ -/ #guard_msgs (info) in theorem fast_expo_correct (x n : Nat) : fast_expo x n = x^n := by @@ -321,11 +316,10 @@ theorem mkFresh_spec (c : Nat) : /-- info: Try this: - invariants - · - post⟨fun ⟨xs, letMuts⟩ => - ⌜xs.prefix = [] ∧ letMuts = acc✝ ∨ xs.suffix = [] ∧ letMuts.toList.Nodup⌝, fun x => ⌜False⌝, - fun _msg state => ⌜state.counter = state.limit⌝⟩ + [apply] invariants + · + post⟨fun ⟨xs, letMuts⟩ => ⌜xs.prefix = [] ∧ letMuts = acc✝ ∨ xs.suffix = [] ∧ letMuts.toList.Nodup⌝, fun x => + ⌜False⌝, fun _msg state => ⌜state.counter = state.limit⌝⟩ -/ #guard_msgs (info) in theorem mkFreshN_spec (n : Nat) : diff --git a/tests/lean/run/partial_fixpoint_explicit.lean b/tests/lean/run/partial_fixpoint_explicit.lean index 06bb7dc2ac0c..efb0f7547822 100644 --- a/tests/lean/run/partial_fixpoint_explicit.lean +++ b/tests/lean/run/partial_fixpoint_explicit.lean @@ -52,7 +52,7 @@ partial_fixpoint monotonicity /-- info: Try this: - exact fun x y a x_1 => a (x_1 + 1) + [apply] exact fun x y a x_1 => a (x_1 + 1) -/ #guard_msgs in def nullary6 (x : Nat) : Option Unit := nullary6 (x + 1) diff --git a/tests/lean/run/rewrites.lean b/tests/lean/run/rewrites.lean index a18211c9c442..bd48ab1e37be 100644 --- a/tests/lean/run/rewrites.lean +++ b/tests/lean/run/rewrites.lean @@ -5,7 +5,7 @@ private axiom test_sorry : ∀ {α}, α /-- info: Try this: - rw [@List.map_append] + [apply] rw [List.map_append] -- no goals -/ #guard_msgs in @@ -14,7 +14,7 @@ example (f : α → β) (L M : List α) : (L ++ M).map f = L.map f ++ M.map f := /-- info: Try this: - rw [Nat.one_mul] + [apply] rw [Nat.one_mul] -- no goals -/ #guard_msgs in @@ -100,7 +100,7 @@ example : zero = 0 := by -- Discharge side conditions from local hypotheses. /-- info: Try this: - rw [h p] + [apply] rw [h p] -- no goals -/ #guard_msgs in @@ -110,7 +110,7 @@ example {P : Prop} (p : P) (h : P → 1 = 2) : 2 = 1 := by -- Use `solve_by_elim` to discharge side conditions. /-- info: Try this: - rw [h (f p)] + [apply] rw [h (f p)] -- no goals -/ #guard_msgs in @@ -120,7 +120,7 @@ example {P Q : Prop} (p : P) (f : P → Q) (h : Q → 1 = 2) : 2 = 1 := by -- Rewrite in reverse, discharging side conditions from local hypotheses. /-- info: Try this: - rw [← h₁ p] + [apply] rw [← h₁ p] -- Q a -/ #guard_msgs in diff --git a/tests/lean/run/show_term.lean b/tests/lean/run/show_term.lean index f7557aa35c92..cf6ccbce0421 100644 --- a/tests/lean/run/show_term.lean +++ b/tests/lean/run/show_term.lean @@ -6,7 +6,7 @@ Authors: Kim Morrison /-- info: Try this: - exact (n, 37) + [apply] exact (n, 37) -/ #guard_msgs in example (n : Nat) : Nat × Nat := by show_term @@ -16,7 +16,7 @@ info: Try this: /-- info: Try this: - refine (?_, ?_) + [apply] refine (?_, ?_) -/ #guard_msgs in example : Nat × Nat := by show_term constructor @@ -24,7 +24,7 @@ info: Try this: /-- info: Try this: - fun {X} => X + [apply] fun {X} => X -/ #guard_msgs in example : {_a : Nat} → Nat := show_term by diff --git a/tests/lean/run/simpLoopProtection.lean b/tests/lean/run/simpLoopProtection.lean index 59f0a3c759a4..4dab7f30d1f4 100644 --- a/tests/lean/run/simpLoopProtection.lean +++ b/tests/lean/run/simpLoopProtection.lean @@ -362,7 +362,7 @@ example : d > 0 := by simp? only [dc, ca, ac]; exact testSorry /-- info: Try this: - simp only [dc, h, Nat.one_pos] + [apply] simp only [dc, h, Nat.one_pos] -/ #guard_msgs in example (h : c = 1) : d > 0 := by simp? only [dc, h, ca, ac, Nat.one_pos] diff --git a/tests/lean/run/simpTrace.lean b/tests/lean/run/simpTrace.lean index d103763860de..81207322d7b6 100644 --- a/tests/lean/run/simpTrace.lean +++ b/tests/lean/run/simpTrace.lean @@ -21,7 +21,7 @@ def x4 := 1 /-- info: Try this: - simp only [x1_eq_x2, x2_eq_x3, x3_eq_x4] + [apply] simp only [x1_eq_x2, x2_eq_x3, x3_eq_x4] -/ #guard_msgs in example : x1 = x4 := by @@ -33,7 +33,7 @@ example : x1 = x4 := by /-- info: Try this: - simp only [x1_eq_x2, x2_eq_x3, x3_eq_x4] + [apply] simp only [x1_eq_x2, x2_eq_x3, x3_eq_x4] -/ #guard_msgs in example : x1 * x1 = x4 * x4 := by @@ -54,7 +54,7 @@ def test : Nat → Nat /-- info: Try this: - simp only [test] + [apply] simp only [test] -/ #guard_msgs in example : [test 3, test 2, test 4, test 5, test 0] = [4, 9, 16, 0, 3] := by @@ -66,7 +66,7 @@ example : [test 3, test 2, test 4, test 5, test 0] = [4, 9, 16, 0, 3] := by /-- info: Try this: - simp only [a] + [apply] simp only [a] -/ #guard_msgs in example : let a := 5; a = 5 := by @@ -79,7 +79,7 @@ example : let a := 5; a = 5 := by /-- info: Try this: - simp only [h] + [apply] simp only [h] --- error: unsolved goals a b : Nat @@ -108,7 +108,7 @@ simproc dontRewriteY2 (y2) := fun _ => do info: was run --- info: Try this: - simp only [rewriteY1] + [apply] simp only [rewriteY1] -/ #guard_msgs in example : y1 = y2 := by @@ -116,7 +116,7 @@ example : y1 = y2 := by /-- info: Try this: - simp only [Nat.reduceAdd] + [apply] simp only [Nat.reduceAdd] -/ #guard_msgs in example : 1 + 1 = 2 := by @@ -135,8 +135,8 @@ theorem wait_this_is_rewritten_backwards_and_wow_it's_very_clear_and_obvious : z /-- info: Try this: - simp only [very_long_lemma_oh_no_can_you_please_stop_we're_getting_to_the_limit, - ← wait_this_is_rewritten_backwards_and_wow_it's_very_clear_and_obvious] + [apply] simp only [very_long_lemma_oh_no_can_you_please_stop_we're_getting_to_the_limit, + ← wait_this_is_rewritten_backwards_and_wow_it's_very_clear_and_obvious] -/ #guard_msgs in example : z1 = z3 := by diff --git a/tests/lean/run/simp_arith_deprecated.lean b/tests/lean/run/simp_arith_deprecated.lean index 93f90ea1e7a0..7ac37f7e5102 100644 --- a/tests/lean/run/simp_arith_deprecated.lean +++ b/tests/lean/run/simp_arith_deprecated.lean @@ -1,7 +1,7 @@ /-- info: Try these: - • simp +arith - • simp +arith +decide + [apply] simp +arith + [apply] simp +arith +decide --- error: `simp_arith` has been deprecated. It was a shorthand for `simp +arith +decide`, but most of the time, `+decide` was redundant since simprocs have been implemented. -/ @@ -11,8 +11,8 @@ example : x + 2 = 1 + 1 + x := by /-- info: Try these: - • simp +arith [h, id] at h₂ - • simp +arith +decide [h, id] at h₂ + [apply] simp +arith [h, id] at h₂ + [apply] simp +arith +decide [h, id] at h₂ --- error: `simp_arith` has been deprecated. It was a shorthand for `simp +arith +decide`, but most of the time, `+decide` was redundant since simprocs have been implemented. -/ @@ -22,8 +22,8 @@ example (h : x = y) (h₂ : y + 2 = 1 + 1 + x) : True := by /-- info: Try these: - • simp! +arith [h, id] at h₂ - • simp! +arith +decide [h, id] at h₂ + [apply] simp! +arith [h, id] at h₂ + [apply] simp! +arith +decide [h, id] at h₂ --- error: `simp_arith!` has been deprecated. It was a shorthand for `simp! +arith +decide`, but most of the time, `+decide` was redundant since simprocs have been implemented. -/ @@ -34,8 +34,8 @@ example (h : x = y) (h₂ : y + 2 = 1 + 1 + x) : True := by /-- info: Try these: - • simp_all +arith - • simp_all +arith +decide + [apply] simp_all +arith + [apply] simp_all +arith +decide --- error: `simp_all_arith` has been deprecated. It was a shorthand for `simp_all +arith +decide`, but most of the time, `+decide` was redundant since simprocs have been implemented. -/ @@ -45,8 +45,8 @@ example (h : x = y) (h₂ : y + 2 = 1 + 1 + x) : True := by /-- info: Try these: - • simp_all! +arith - • simp_all! +arith +decide + [apply] simp_all! +arith + [apply] simp_all! +arith +decide --- error: `simp_all_arith!` has been deprecated. It was a shorthand for `simp_all! +arith +decide`, but most of the time, `+decide` was redundant since simprocs have been implemented. -/ diff --git a/tests/lean/run/zetaDeltaTryThisIssue.lean b/tests/lean/run/zetaDeltaTryThisIssue.lean index ca0826884b04..4655e03a2ba6 100644 --- a/tests/lean/run/zetaDeltaTryThisIssue.lean +++ b/tests/lean/run/zetaDeltaTryThisIssue.lean @@ -2,7 +2,7 @@ opaque f : Nat → Nat /-- info: Try this: - simp only [h1, x] + [apply] simp only [h1, x] --- warning: declaration uses 'sorry' -/ @@ -16,7 +16,7 @@ example (a : Nat) : True := by /-- info: Try this: - simp only [this, x] + [apply] simp only [this, x] --- warning: declaration uses 'sorry' -/ diff --git a/tests/lean/termination_by.lean.expected.out b/tests/lean/termination_by.lean.expected.out index 2b18ba7129c5..9995993706fc 100644 --- a/tests/lean/termination_by.lean.expected.out +++ b/tests/lean/termination_by.lean.expected.out @@ -14,4 +14,4 @@ termination_by.lean:62:2-62:23: error: incomplete set of termination hints: This function is mutually recursive with isOdd, which does not have a termination hint. The present clause is ignored. Try this: - termination_by x1 => x1 + [apply] termination_by x1 => x1