Skip to content

Commit

Permalink
feat: trace messages for working and closing goals in the grind tac…
Browse files Browse the repository at this point in the history
…tic (#6567)

This PR adds support for erasing the `[grind]` attribute used to mark
theorems for heuristic instantiation in the `grind` tactic.
  • Loading branch information
leodemoura authored Jan 7, 2025
1 parent 0da5be1 commit 5decd2c
Show file tree
Hide file tree
Showing 15 changed files with 97 additions and 37 deletions.
5 changes: 4 additions & 1 deletion src/Lean/Meta/Tactic/Grind/Cases.lean
Original file line number Diff line number Diff line change
Expand Up @@ -36,14 +36,17 @@ def cases (mvarId : MVarId) (e : Expr) : MetaM (List MVarId) := mvarId.withConte
let mut recursor := mkApp (mkAppN recursor indicesExpr) (mkFVar fvarId)
let mut recursorType ← inferType recursor
let mut mvarIdsNew := #[]
let mut idx := 1
for _ in [:recursorInfo.numMinors] do
let .forallE _ targetNew recursorTypeNew _ ← whnf recursorType
| throwTacticEx `grind.cases mvarId "unexpected recursor type"
recursorType := recursorTypeNew
let mvar ← mkFreshExprSyntheticOpaqueMVar targetNew tag
let tagNew := if recursorInfo.numMinors > 1 then Name.num tag idx else tag
let mvar ← mkFreshExprSyntheticOpaqueMVar targetNew tagNew
recursor := mkApp recursor mvar
let mvarIdNew ← mvar.mvarId!.tryClearMany (indices.push fvarId)
mvarIdsNew := mvarIdsNew.push mvarIdNew
idx := idx + 1
mvarId.assign recursor
return mvarIdsNew.toList
if recursorInfo.numIndices > 0 then
Expand Down
16 changes: 8 additions & 8 deletions src/Lean/Meta/Tactic/Grind/Core.lean
Original file line number Diff line number Diff line change
Expand Up @@ -43,7 +43,7 @@ private def removeParents (root : Expr) : GoalM ParentSet := do
for parent in parents do
-- Recall that we may have `Expr.forallE` in `parents` because of `ForallProp.lean`
if (← pure parent.isApp <&&> isCongrRoot parent) then
trace[grind.debug.parent] "remove: {parent}"
trace_goal[grind.debug.parent] "remove: {parent}"
modify fun s => { s with congrTable := s.congrTable.erase { e := parent } }
return parents

Expand All @@ -54,7 +54,7 @@ This is an auxiliary function performed while merging equivalence classes.
private def reinsertParents (parents : ParentSet) : GoalM Unit := do
for parent in parents do
if (← pure parent.isApp <&&> isCongrRoot parent) then
trace[grind.debug.parent] "reinsert: {parent}"
trace_goal[grind.debug.parent] "reinsert: {parent}"
addCongrTable parent

/-- Closes the goal when `True` and `False` are in the same equivalence class. -/
Expand Down Expand Up @@ -91,9 +91,9 @@ private partial def addEqStep (lhs rhs proof : Expr) (isHEq : Bool) : GoalM Unit
let rhsNode ← getENode rhs
if isSameExpr lhsNode.root rhsNode.root then
-- `lhs` and `rhs` are already in the same equivalence class.
trace[grind.debug] "{← ppENodeRef lhs} and {← ppENodeRef rhs} are already in the same equivalence class"
trace_goal[grind.debug] "{← ppENodeRef lhs} and {← ppENodeRef rhs} are already in the same equivalence class"
return ()
trace[grind.eqc] "{← if isHEq then mkHEq lhs rhs else mkEq lhs rhs}"
trace_goal[grind.eqc] "{← if isHEq then mkHEq lhs rhs else mkEq lhs rhs}"
let lhsRoot ← getENode lhsNode.root
let rhsRoot ← getENode rhsNode.root
let mut valueInconsistency := false
Expand All @@ -118,11 +118,11 @@ private partial def addEqStep (lhs rhs proof : Expr) (isHEq : Bool) : GoalM Unit
unless (← isInconsistent) do
if valueInconsistency then
closeGoalWithValuesEq lhsRoot.self rhsRoot.self
trace[grind.debug] "after addEqStep, {← ppState}"
trace_goal[grind.debug] "after addEqStep, {← ppState}"
checkInvariants
where
go (lhs rhs : Expr) (lhsNode rhsNode lhsRoot rhsRoot : ENode) (flipped : Bool) : GoalM Unit := do
trace[grind.debug] "adding {← ppENodeRef lhs} ↦ {← ppENodeRef rhs}"
trace_goal[grind.debug] "adding {← ppENodeRef lhs} ↦ {← ppENodeRef rhs}"
/-
We have the following `target?/proof?`
`lhs -> ... -> lhsNode.root`
Expand All @@ -139,7 +139,7 @@ where
}
let parents ← removeParents lhsRoot.self
updateRoots lhs rhsNode.root
trace[grind.debug] "{← ppENodeRef lhs} new root {← ppENodeRef rhsNode.root}, {← ppENodeRef (← getRoot lhs)}"
trace_goal[grind.debug] "{← ppENodeRef lhs} new root {← ppENodeRef rhsNode.root}, {← ppENodeRef (← getRoot lhs)}"
reinsertParents parents
setENode lhsNode.root { (← getENode lhsRoot.self) with -- We must retrieve `lhsRoot` since it was updated.
next := rhsRoot.next
Expand Down Expand Up @@ -208,7 +208,7 @@ def addNewEq (lhs rhs proof : Expr) (generation : Nat) : GoalM Unit := do

/-- Adds a new `fact` justified by the given proof and using the given generation. -/
def add (fact : Expr) (proof : Expr) (generation := 0) : GoalM Unit := do
trace[grind.assert] "{fact}"
trace_goal[grind.assert] "{fact}"
if (← isInconsistent) then return ()
resetNewEqs
let_expr Not p := fact
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind/Ctor.lean
Original file line number Diff line number Diff line change
Expand Up @@ -20,7 +20,7 @@ private partial def propagateInjEqs (eqs : Expr) (proof : Expr) : GoalM Unit :=
| HEq _ lhs _ rhs =>
pushHEq (← shareCommon lhs) (← shareCommon rhs) proof
| _ =>
trace[grind.issues] "unexpected injectivity theorem result type{indentExpr eqs}"
trace_goal[grind.issues] "unexpected injectivity theorem result type{indentExpr eqs}"
return ()

/--
Expand Down
12 changes: 6 additions & 6 deletions src/Lean/Meta/Tactic/Grind/EMatch.lean
Original file line number Diff line number Diff line change
Expand Up @@ -221,7 +221,7 @@ private def addNewInstance (origin : Origin) (proof : Expr) (generation : Nat) :
-- `initApp` is a match-application that we don't need to split at anymore.
markCaseSplitAsResolved (← read).initApp
prop ← annotateMatchEqnType prop
trace[grind.ematch.instance] "{← origin.pp}: {prop}"
trace_goal[grind.ematch.instance] "{← origin.pp}: {prop}"
addTheoremInstance proof prop (generation+1)

/--
Expand All @@ -232,36 +232,36 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w
let thm := (← read).thm
unless (← markTheoremInstance thm.proof c.assignment) do
return ()
trace[grind.ematch.instance.assignment] "{← thm.origin.pp}: {assignmentToMessageData c.assignment}"
trace_goal[grind.ematch.instance.assignment] "{← thm.origin.pp}: {assignmentToMessageData c.assignment}"
let proof ← thm.getProofWithFreshMVarLevels
let numParams := thm.numParams
assert! c.assignment.size == numParams
let (mvars, bis, _) ← forallMetaBoundedTelescope (← inferType proof) numParams
if mvars.size != thm.numParams then
trace[grind.issues] "unexpected number of parameters at {← thm.origin.pp}"
trace_goal[grind.issues] "unexpected number of parameters at {← thm.origin.pp}"
return ()
-- Apply assignment
for h : i in [:mvars.size] do
let v := c.assignment[numParams - i - 1]!
unless isSameExpr v unassigned do
let mvarId := mvars[i].mvarId!
unless (← isDefEq (← mvarId.getType) (← inferType v) <&&> mvarId.checkedAssign v) do
trace[grind.issues] "type error constructing proof for {← thm.origin.pp}\nwhen assigning metavariable {mvars[i]} with {indentExpr v}"
trace_goal[grind.issues] "type error constructing proof for {← thm.origin.pp}\nwhen assigning metavariable {mvars[i]} with {indentExpr v}"
return ()
-- Synthesize instances
for mvar in mvars, bi in bis do
if bi.isInstImplicit && !(← mvar.mvarId!.isAssigned) then
let type ← inferType mvar
unless (← synthesizeInstance mvar type) do
trace[grind.issues] "failed to synthesize instance when instantiating {← thm.origin.pp}{indentExpr type}"
trace_goal[grind.issues] "failed to synthesize instance when instantiating {← thm.origin.pp}{indentExpr type}"
return ()
let proof := mkAppN proof mvars
if (← mvars.allM (·.mvarId!.isAssigned)) then
addNewInstance thm.origin proof c.gen
else
let mvars ← mvars.filterM fun mvar => return !(← mvar.mvarId!.isAssigned)
if let some mvarBad ← mvars.findM? fun mvar => return !(← isProof mvar) then
trace[grind.issues] "failed to instantiate {← thm.origin.pp}, failed to instantiate non propositional argument with type{indentExpr (← inferType mvarBad)}"
trace_goal[grind.issues] "failed to instantiate {← thm.origin.pp}, failed to instantiate non propositional argument with type{indentExpr (← inferType mvarBad)}"
let proof ← mkLambdaFVars (binderInfoForMVars := .default) mvars (← instantiateMVars proof)
addNewInstance thm.origin proof c.gen
where
Expand Down
6 changes: 3 additions & 3 deletions src/Lean/Meta/Tactic/Grind/ForallProp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -17,19 +17,19 @@ If so, internalize the term `proj_i (ctor ... a_i ...)` and add the equality `pr
-/
def propagateForallPropUp (e : Expr) : GoalM Unit := do
let .forallE n p q bi := e | return ()
trace[grind.debug.forallPropagator] "{e}"
trace_goal[grind.debug.forallPropagator] "{e}"
if !q.hasLooseBVars then
propagateImpliesUp p q
else
unless (← isEqTrue p) do return
trace[grind.debug.forallPropagator] "isEqTrue, {e}"
trace_goal[grind.debug.forallPropagator] "isEqTrue, {e}"
let h₁ ← mkEqTrueProof p
let qh₁ := q.instantiate1 (mkApp2 (mkConst ``of_eq_true) p h₁)
let r ← simp qh₁
let q := mkLambda n bi p q
let q' := r.expr
internalize q' (← getGeneration e)
trace[grind.debug.forallPropagator] "q': {q'} for{indentExpr e}"
trace_goal[grind.debug.forallPropagator] "q': {q'} for{indentExpr e}"
let h₂ ← r.getProof
let h := mkApp5 (mkConst ``Lean.Grind.forall_propagator) p q q' h₁ h₂
pushEq e q' h
Expand Down
14 changes: 7 additions & 7 deletions src/Lean/Meta/Tactic/Grind/Internalize.lean
Original file line number Diff line number Diff line change
Expand Up @@ -22,9 +22,9 @@ def addCongrTable (e : Expr) : GoalM Unit := do
let g := e'.getAppFn
unless isSameExpr f g do
unless (← hasSameType f g) do
trace[grind.issues] "found congruence between{indentExpr e}\nand{indentExpr e'}\nbut functions have different types"
trace_goal[grind.issues] "found congruence between{indentExpr e}\nand{indentExpr e'}\nbut functions have different types"
return ()
trace[grind.debug.congr] "{e} = {e'}"
trace_goal[grind.debug.congr] "{e} = {e'}"
pushEqHEq e e' congrPlaceholderProof
let node ← getENode e
setENode e { node with congr := e' }
Expand All @@ -46,7 +46,7 @@ private def updateAppMap (e : Expr) : GoalM Unit := do

/-- Inserts `e` into the list of case-split candidates. -/
private def addSplitCandidate (e : Expr) : GoalM Unit := do
trace[grind.split.candidate] "{e}"
trace_goal[grind.split.candidate] "{e}"
modify fun s => { s with splitCandidates := e :: s.splitCandidates }

-- TODO: add attribute to make this extensible
Expand Down Expand Up @@ -89,7 +89,7 @@ private partial def activateTheorem (thm : EMatchTheorem) (generation : Nat) : G
-- We don't want to use structural equality when comparing keys.
let proof ← shareCommon thm.proof
let thm := { thm with proof, patterns := (← thm.patterns.mapM (internalizePattern · generation)) }
trace[grind.ematch] "activated `{thm.origin.key}`, {thm.patterns.map ppPattern}"
trace_goal[grind.ematch] "activated `{thm.origin.key}`, {thm.patterns.map ppPattern}"
modify fun s => { s with newThms := s.newThms.push thm }

/--
Expand Down Expand Up @@ -117,12 +117,12 @@ private partial def activateTheoremPatterns (fName : Name) (generation : Nat) :
match symbols with
| [] => activateTheorem thm generation
| _ =>
trace[grind.ematch] "reinsert `{thm.origin.key}`"
trace_goal[grind.ematch] "reinsert `{thm.origin.key}`"
modify fun s => { s with thmMap := s.thmMap.insert thm }

partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do
if (← alreadyInternalized e) then return ()
trace[grind.internalize] "{e}"
trace_goal[grind.internalize] "{e}"
match e with
| .bvar .. => unreachable!
| .sort .. => return ()
Expand All @@ -142,7 +142,7 @@ partial def internalize (e : Expr) (generation : Nat) : GoalM Unit := do
| .mvar ..
| .mdata ..
| .proj .. =>
trace[grind.issues] "unexpected term during internalization{indentExpr e}"
trace_goal[grind.issues] "unexpected term during internalization{indentExpr e}"
mkENodeCore e (ctor := false) (interpreted := false) (generation := generation)
| .app .. =>
if (← isLitValue e) then
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Meta/Tactic/Grind/Inv.lean
Original file line number Diff line number Diff line change
Expand Up @@ -85,9 +85,9 @@ private def checkProofs : GoalM Unit := do
for b in eqc do
unless isSameExpr a b do
let p ← mkEqHEqProof a b
trace[grind.debug.proofs] "{a} = {b}"
trace_goal[grind.debug.proofs] "{a} = {b}"
check p
trace[grind.debug.proofs] "checked: {← inferType p}"
trace_goal[grind.debug.proofs] "checked: {← inferType p}"

/--
Checks basic invariants if `grind.debug` is enabled.
Expand Down
1 change: 1 addition & 0 deletions src/Lean/Meta/Tactic/Grind/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -60,6 +60,7 @@ private def initCore (mvarId : MVarId) : GrindM (List Goal) := do
let mvarId ← mvarId.revertAll
let mvarId ← mvarId.unfoldReducible
let mvarId ← mvarId.betaReduce
appendTagSuffix mvarId `grind
let goals ← intros (← mkGoal mvarId) (generation := 0)
goals.forM (·.checkInvariants (expensive := true))
return goals.filter fun goal => !goal.inconsistent
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind/Proj.lean
Original file line number Diff line number Diff line change
Expand Up @@ -30,7 +30,7 @@ def propagateProjEq (parent : Expr) : GoalM Unit := do
let parentNew ← shareCommon (mkApp parent.appFn! ctor)
internalize parentNew (← getGeneration parent)
pure parentNew
trace[grind.debug.proj] "{parentNew}"
trace_goal[grind.debug.proj] "{parentNew}"
let idx := info.numParams + info.i
unless idx < ctor.getAppNumArgs do return ()
let v := ctor.getArg! idx
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind/Split.lean
Original file line number Diff line number Diff line change
Expand Up @@ -110,7 +110,7 @@ def splitNext : GrindTactic := fun goal => do
let some c ← selectNextSplit?
| return none
let gen ← getGeneration c
trace[grind.split] "{c}, generation: {gen}"
trace_goal[grind.split] "{c}, generation: {gen}"
-- TODO: `match`
let major ← mkCasesMajor c
let mvarIds ← cases (← get).mvarId major
Expand Down
34 changes: 30 additions & 4 deletions src/Lean/Meta/Tactic/Grind/Types.lean
Original file line number Diff line number Diff line change
Expand Up @@ -83,6 +83,11 @@ structure State where
simpStats : Simp.Stats := {}
trueExpr : Expr
falseExpr : Expr
/--
Used to generate trace messages of the for `[grind] working on <tag>`,
and implement the macro `trace_goal`.
-/
lastTag : Name := .anonymous

private opaque MethodsRefPointed : NonemptyType.{0}
private def MethodsRef : Type := MethodsRefPointed.type
Expand Down Expand Up @@ -126,9 +131,9 @@ Applies hash-consing to `e`. Recall that all expressions in a `grind` goal have
been hash-consed. We perform this step before we internalize expressions.
-/
def shareCommon (e : Expr) : GrindM Expr := do
modifyGet fun { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats } =>
modifyGet fun { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats, lastTag } =>
let (e, scState) := ShareCommon.State.shareCommon scState e
(e, { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats })
(e, { canon, scState, nextThmIdx, congrThms, trueExpr, falseExpr, simpStats, lastTag })

/--
Canonicalizes nested types, type formers, and instances in `e`.
Expand Down Expand Up @@ -401,6 +406,25 @@ abbrev GoalM := StateRefT Goal GrindM
@[inline] def GoalM.run' (goal : Goal) (x : GoalM Unit) : GrindM Goal :=
goal.mvarId.withContext do StateRefT'.run' (x *> get) goal

def updateLastTag : GoalM Unit := do
if (← isTracingEnabledFor `grind) then
let currTag ← (← get).mvarId.getTag
if currTag != (← getThe Grind.State).lastTag then
trace[grind] "working on goal `{currTag}`"
modifyThe Grind.State fun s => { s with lastTag := currTag }

/--
Macro similar to `trace[...]`, but it includes the trace message `trace[grind] "working on <current goal>"`
if the tag has changed since the last trace message.
-/
macro "trace_goal[" id:ident "]" s:(interpolatedStr(term) <|> term) : doElem => do
let msg ← if s.raw.getKind == interpolatedStrKind then `(m! $(⟨s⟩)) else `(($(⟨s⟩) : MessageData))
`(doElem| do
let cls := $(quote id.getId.eraseMacroScopes)
if (← Lean.isTracingEnabledFor cls) then
updateLastTag
Lean.addTrace cls $msg)

/--
A helper function used to mark a theorem instance found by the E-matching module.
It returns `true` if it is a new instance and `false` otherwise.
Expand Down Expand Up @@ -652,7 +676,9 @@ def mkEqFalseProof (a : Expr) : GoalM Expr := do

/-- Marks current goal as inconsistent without assigning `mvarId`. -/
def markAsInconsistent : GoalM Unit := do
modify fun s => { s with inconsistent := true }
unless (← get).inconsistent do
trace[grind] "closed `{← (← get).mvarId.getTag}`"
modify fun s => { s with inconsistent := true }

/--
Closes the current goal using the given proof of `False` and
Expand Down Expand Up @@ -751,7 +777,7 @@ Remark: we currently use this feature to disable `match`-case-splits
-/
def markCaseSplitAsResolved (e : Expr) : GoalM Unit := do
unless (← isResolvedCaseSplit e) do
trace[grind.split.resolved] "{e}"
trace_goal[grind.split.resolved] "{e}"
modify fun s => { s with resolvedSplits := s.resolvedSplits.insert { expr := e } }

end Lean.Meta.Grind
7 changes: 5 additions & 2 deletions src/Lean/Util/Trace.lean
Original file line number Diff line number Diff line change
Expand Up @@ -293,13 +293,16 @@ def registerTraceClass (traceClassName : Name) (inherited := false) (ref : Name
if inherited then
inheritedTraceOptions.modify (·.insert optionName)

macro "trace[" id:ident "]" s:(interpolatedStr(term) <|> term) : doElem => do
let msg ← if s.raw.getKind == interpolatedStrKind then `(m! $(⟨s⟩)) else `(($(⟨s⟩) : MessageData))
def expandTraceMacro (id : Syntax) (s : Syntax) : MacroM (TSyntax `doElem) := do
let msg ← if s.getKind == interpolatedStrKind then `(m! $(⟨s⟩)) else `(($(⟨s⟩) : MessageData))
`(doElem| do
let cls := $(quote id.getId.eraseMacroScopes)
if (← Lean.isTracingEnabledFor cls) then
Lean.addTrace cls $msg)

macro "trace[" id:ident "]" s:(interpolatedStr(term) <|> term) : doElem => do
expandTraceMacro id s.raw

def bombEmoji := "💥️"
def checkEmoji := "✅️"
def crossEmoji := "❌️"
Expand Down
2 changes: 2 additions & 0 deletions tests/lean/run/grind_erase_attr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -25,6 +25,7 @@ attribute [-grind] fthm

/--
error: `grind` failed
case grind
a : Nat
x✝ : ¬f (f (f a)) = f a
⊢ False
Expand Down Expand Up @@ -60,6 +61,7 @@ attribute [-grind] g

/--
error: `grind` failed
case grind
a b : Nat
a✝¹ : g a = b
a✝ : a = 0
Expand Down
Loading

0 comments on commit 5decd2c

Please sign in to comment.