Skip to content

Commit 114f7e4

Browse files
authored
feat: lazy message with grind state (#10791)
This PR adds a silent info message with the `grind` state in its interactive mode. The message is shown only when there is exactly one goal in the grind interactive mode. The condition is a workaround for current limitations of our `InfoTree`.
1 parent 419982b commit 114f7e4

File tree

2 files changed

+33
-13
lines changed

2 files changed

+33
-13
lines changed

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

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,9 @@ public import Lean.Elab.Tactic.Basic
1010
public import Lean.Meta.Tactic.Grind.Types
1111
public import Lean.Meta.Tactic.Grind.Main
1212
public import Lean.Meta.Tactic.Grind.SearchM
13+
import Lean.CoreM
1314
import Lean.Meta.Tactic.Grind.Intro
15+
import Lean.Meta.Tactic.Grind.PP
1416
public section
1517
namespace Lean.Elab.Tactic.Grind
1618
open Meta
@@ -87,6 +89,19 @@ def mkTacticInfo (mctxBefore : MetavarContext) (goalsBefore : List MVarId) (stx
8789
def mkInitialTacticInfo (stx : Syntax) : GrindTacticM (GrindTacticM Info) := do
8890
let mctxBefore ← getMCtx
8991
let goalsBefore ← getUnsolvedGoalMVarIds
92+
/-
93+
**Note**: We only display the grind state if there is exactly one goal.
94+
This is a hack because we currently use a silent info to display the grind state, and we cannot attach it after each goal.
95+
We claim this is not a big deal since the user will probably use `next =>` to focus on subgoals.
96+
-/
97+
if let [goal] ← getGoals then goal.withContext do
98+
let config := (← read).params.config
99+
let msg := MessageData.lazy fun ctx => do
100+
let .ok msg ← EIO.toBaseIO <| ctx.runMetaM
101+
<| Grind.goalDiagToMessageData goal config (header := "Grind state") (collapsedMain := false)
102+
| return "Grind state could not be generated"
103+
return msg
104+
logAt (severity := .information) (isSilent := true) stx msg
90105
return mkTacticInfo mctxBefore goalsBefore stx
91106

92107
@[inline] def withTacticInfoContext (stx : Syntax) (x : GrindTacticM α) : GrindTacticM α := do

src/Lean/Meta/Tactic/Grind/PP.lean

Lines changed: 18 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -153,7 +153,7 @@ end EqcFilter
153153
def ppEqc (eqc : List Expr) (children : Array MessageData := #[]) : MessageData :=
154154
.trace { cls := `eqc } (.group ("{" ++ (MessageData.joinSep (eqc.map toMessageData) ("," ++ Format.line)) ++ "}")) children
155155

156-
private def ppEqcs : M Unit := do
156+
private def ppEqcs (collapsedProps := true) : M Unit := do
157157
let mut trueEqc? : Option MessageData := none
158158
let mut falseEqc? : Option MessageData := none
159159
let mut regularEqcs : Array MessageData := #[]
@@ -163,11 +163,11 @@ private def ppEqcs : M Unit := do
163163
if Option.isSome <| eqc.find? (·.isTrue) then
164164
let eqc := eqc.filter fun e => !e.isTrue
165165
unless eqc.isEmpty do
166-
trueEqc? := ppExprArray `eqc "True propositions" eqc.toArray `prop
166+
trueEqc? := ppExprArray `eqc "True propositions" eqc.toArray `prop (collapsed := collapsedProps)
167167
else if Option.isSome <| eqc.find? (·.isFalse) then
168168
let eqc := eqc.filter fun e => !e.isFalse
169169
unless eqc.isEmpty do
170-
falseEqc? := ppExprArray `eqc "False propositions" eqc.toArray `prop
170+
falseEqc? := ppExprArray `eqc "False propositions" eqc.toArray `prop (collapsed := collapsedProps)
171171
else if let e :: _ :: _ := eqc then
172172
-- We may want to add a flag to pretty print equivalence classes of nested proofs
173173
unless (← isProof e) do
@@ -270,18 +270,15 @@ private def ppCasesTrace : M Unit := do
270270
]
271271
pushMsg <| .trace { cls := `cases } "Case analyses" msgs
272272

273-
def goalToMessageData (goal : Goal) (config : Grind.Config) : MetaM MessageData := goal.mvarId.withContext do
274-
if config.verbose then
275-
let (_, m) ← go goal |>.run #[]
276-
let gm := MessageData.trace { cls := `grind, collapsed := false } "Goal diagnostics" m
277-
let r := m!"{.ofGoal goal.mvarId}\n{gm}"
278-
addMessageContextFull r
279-
else
280-
return .ofGoal goal.mvarId
273+
def goalDiagToMessageData (goal : Goal) (config : Grind.Config) (header := "Goal diagnostics") (collapsedMain := true)
274+
: MetaM MessageData := do
275+
let (_, m) ← go goal |>.run #[]
276+
let gm := MessageData.trace { cls := `grind, collapsed := false } header m
277+
return gm
281278
where
282279
go : M Unit := do
283-
pushMsg <| ppExprArray `facts "Asserted facts" goal.facts.toArray `prop
284-
ppEqcs
280+
pushMsg <| ppExprArray `facts "Asserted facts" goal.facts.toArray `prop (collapsed := collapsedMain)
281+
ppEqcs (collapsedProps := collapsedMain)
285282
ppCasesTrace
286283
ppActiveTheoremPatterns
287284
ppOffset
@@ -291,4 +288,12 @@ where
291288
ppAC
292289
ppThresholds config
293290

291+
def goalToMessageData (goal : Goal) (config : Grind.Config) : MetaM MessageData := goal.withContext do
292+
if config.verbose then
293+
let gm ← goalDiagToMessageData goal config
294+
let r := m!"{.ofGoal goal.mvarId}\n{gm}"
295+
addMessageContextFull r
296+
else
297+
return .ofGoal goal.mvarId
298+
294299
end Lean.Meta.Grind

0 commit comments

Comments
 (0)