Skip to content

Commit 00d41d6

Browse files
authored
feat: remove +premises from grind? output (#11028)
This PR ensures that `grind? +premises` removes `+premises` from the "Try this" suggestion.
1 parent 40e1e09 commit 00d41d6

File tree

2 files changed

+29
-1
lines changed

2 files changed

+29
-1
lines changed

src/Lean/Elab/Tactic/Grind/Main.lean

Lines changed: 17 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -201,6 +201,21 @@ def setGrindParams (stx : TSyntax `tactic) (params : Array Syntax) : TSyntax `ta
201201
def getGrindParams (stx : TSyntax `tactic) : Array Syntax :=
202202
stx.raw[grindParamsPos][1].getSepArgs
203203

204+
/-- Filter out `+premises` from the config syntax -/
205+
def filterPremisesFromConfig (config : TSyntax ``Lean.Parser.Tactic.optConfig) : TSyntax ``Lean.Parser.Tactic.optConfig :=
206+
let configItems := config.raw.getArgs
207+
let filteredItems := configItems.filter fun item =>
208+
-- Keep all items except +premises
209+
-- Structure: null node -> configItem -> posConfigItem -> ["+", ident]
210+
match item[0]? with
211+
| some configItem => match configItem[0]? with
212+
| some posConfigItem => match posConfigItem[1]? with
213+
| some ident => !(posConfigItem.getKind == ``Lean.Parser.Tactic.posConfigItem && ident.getId == `premises)
214+
| none => true
215+
| none => true
216+
| none => true
217+
⟨config.raw.setArgs filteredItems⟩
218+
204219
def mkGrindOnly
205220
(config : TSyntax ``Lean.Parser.Tactic.optConfig)
206221
(trace : Grind.Trace)
@@ -217,7 +232,8 @@ def mkGrindOnly
217232
else
218233
let param ← Grind.globalDeclToGrindParamSyntax declName kind minIndexable
219234
params := params.push param
220-
let result ← `(tactic| grind $config:optConfig only)
235+
let filteredConfig := filterPremisesFromConfig config
236+
let result ← `(tactic| grind $filteredConfig:optConfig only)
221237
return setGrindParams result params
222238

223239
private def elabGrindConfig' (config : TSyntax ``Lean.Parser.Tactic.optConfig) (interactive : Bool) : TacticM Grind.Config := do
Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,12 @@
1+
import Lean
2+
3+
set_premise_selector Lean.PremiseSelection.sineQuaNonSelector
4+
5+
-- Test that grind? +premises does NOT include +premises in its output
6+
/--
7+
info: Try this:
8+
[apply] grind only
9+
-/
10+
#guard_msgs in
11+
example {x y : Nat} (h : x = y) : y = x := by
12+
grind? +premises

0 commit comments

Comments
 (0)