Skip to content

Commit

Permalink
feat: add Grind.Config.verbose and reportIssue! macro (#6904)
Browse files Browse the repository at this point in the history
This PR adds the `grind` configuration option `verbose`. For example,
`grind -verbose` disables all diagnostics. We are going to use this flag
to implement `try?`.
  • Loading branch information
leodemoura authored Feb 1, 2025
1 parent deb3299 commit 38086a8
Show file tree
Hide file tree
Showing 14 changed files with 58 additions and 30 deletions.
2 changes: 2 additions & 0 deletions src/Init/Grind/Tactics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -65,6 +65,8 @@ structure Config where
canonHeartbeats : Nat := 1000
/-- If `ext` is `true`, `grind` uses extensionality theorems available in the environment. -/
ext : Bool := true
/-- If `verbose` is `false`, additional diagnostics information is not collected. -/
verbose : Bool := true
deriving Inhabited, BEq

end Lean.Grind
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind/Canon.lean
Original file line number Diff line number Diff line change
Expand Up @@ -60,7 +60,7 @@ private def isDefEqBounded (a b : Expr) (parent : Expr) : GoalM Bool := do
withDefault <| isDefEq a b)
fun ex => do
if ex.isRuntime then
reportIssue m!"failed to show that{indentExpr a}\nis definitionally equal to{indentExpr b}\nwhile canonicalizing{indentExpr parent}\nusing `{curr}*1000` heartbeats, `(canonHeartbeats := {curr})`"
reportIssue! "failed to show that{indentExpr a}\nis definitionally equal to{indentExpr b}\nwhile canonicalizing{indentExpr parent}\nusing `{curr}*1000` heartbeats, `(canonHeartbeats := {curr})`"
return false
else
throw ex
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
| _ =>
reportIssue m!"unexpected injectivity theorem result type{indentExpr eqs}"
reportIssue! "unexpected injectivity theorem result type{indentExpr eqs}"
return ()

/--
Expand Down
8 changes: 4 additions & 4 deletions src/Lean/Meta/Tactic/Grind/EMatch.lean
Original file line number Diff line number Diff line change
Expand Up @@ -268,7 +268,7 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w
assert! c.assignment.size == numParams
let (mvars, bis, _) ← forallMetaBoundedTelescope (← inferType proof) numParams
if mvars.size != thm.numParams then
reportIssue m!"unexpected number of parameters at {← thm.origin.pp}"
reportIssue! "unexpected number of parameters at {← thm.origin.pp}"
return ()
-- Apply assignment
for h : i in [:mvars.size] do
Expand All @@ -278,7 +278,7 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w
let mvarIdType ← mvarId.getType
let vType ← inferType v
let report : M Unit := do
reportIssue m!"type error constructing proof for {← thm.origin.pp}\nwhen assigning metavariable {mvars[i]} with {indentExpr v}\n{← mkHasTypeButIsExpectedMsg vType mvarIdType}"
reportIssue! "type error constructing proof for {← thm.origin.pp}\nwhen assigning metavariable {mvars[i]} with {indentExpr v}\n{← mkHasTypeButIsExpectedMsg vType mvarIdType}"
unless (← withDefault <| isDefEq mvarIdType vType) do
let some heq ← proveEq? vType mvarIdType
| report
Expand All @@ -297,15 +297,15 @@ private partial def instantiateTheorem (c : Choice) : M Unit := withDefault do w
if bi.isInstImplicit && !(← mvar.mvarId!.isAssigned) then
let type ← inferType mvar
unless (← synthesizeInstanceAndAssign mvar type) do
reportIssue m!"failed to synthesize instance when instantiating {← thm.origin.pp}{indentExpr type}"
reportIssue! "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 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
reportIssue m!"failed to instantiate {← thm.origin.pp}, failed to instantiate non propositional argument with type{indentExpr (← inferType mvarBad)}"
reportIssue! "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 proof c.gen

Expand Down
6 changes: 3 additions & 3 deletions src/Lean/Meta/Tactic/Grind/Ext.lean
Original file line number Diff line number Diff line change
Expand Up @@ -14,14 +14,14 @@ def instantiateExtTheorem (thm : Ext.ExtTheorem) (e : Expr) : GoalM Unit := with
let c ← mkConstWithFreshMVarLevels thm.declName
let (mvars, bis, type) ← withDefault <| forallMetaTelescopeReducing (← inferType c)
unless (← isDefEq e type) do
reportIssue m!"failed to apply extensionality theorem `{thm.declName}` for {indentExpr e}\nis not definitionally equal to{indentExpr type}"
reportIssue! "failed to apply extensionality theorem `{thm.declName}` for {indentExpr e}\nis not definitionally equal to{indentExpr type}"
return ()
-- Instantiate type class instances
for mvar in mvars, bi in bis do
if bi.isInstImplicit && !(← mvar.mvarId!.isAssigned) then
let type ← inferType mvar
unless (← synthesizeInstanceAndAssign mvar type) do
reportIssue m!"failed to synthesize instance when instantiating extensionality theorem `{thm.declName}` for {indentExpr e}"
reportIssue! "failed to synthesize instance when instantiating extensionality theorem `{thm.declName}` for {indentExpr e}"
return ()
-- Remark: `proof c mvars` has type `e`
let proof ← instantiateMVars (mkAppN c mvars)
Expand All @@ -33,7 +33,7 @@ def instantiateExtTheorem (thm : Ext.ExtTheorem) (e : Expr) : GoalM Unit := with
let proof' ← instantiateMVars (← mkLambdaFVars mvars proof)
let prop' ← inferType proof'
if proof'.hasMVar || prop'.hasMVar then
reportIssue m!"failed to apply extensionality theorem `{thm.declName}` for {indentExpr e}\nresulting terms contain metavariables"
reportIssue! "failed to apply extensionality theorem `{thm.declName}` for {indentExpr e}\nresulting terms contain metavariables"
return ()
addNewFact proof' prop' ((← getGeneration e) + 1)

Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind/ForallProp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -78,7 +78,7 @@ private def addLocalEMatchTheorems (e : Expr) : GoalM Unit := do
if let some thm ← mkEMatchTheoremWithKind'? origin proof .default then
activateTheorem thm gen
if (← get).newThms.size == size then
reportIssue m!"failed to create E-match local theorem for{indentExpr e}"
reportIssue! "failed to create E-match local theorem for{indentExpr e}"

def propagateForallPropDown (e : Expr) : GoalM Unit := do
let .forallE n a b bi := e | return ()
Expand Down
8 changes: 4 additions & 4 deletions src/Lean/Meta/Tactic/Grind/Internalize.lean
Original file line number Diff line number Diff line change
Expand Up @@ -27,7 +27,7 @@ def addCongrTable (e : Expr) : GoalM Unit := do
let g := e'.getAppFn
unless isSameExpr f g do
unless (← hasSameType f g) do
reportIssue m!"found congruence between{indentExpr e}\nand{indentExpr e'}\nbut functions have different types"
reportIssue! "found congruence between{indentExpr e}\nand{indentExpr e'}\nbut functions have different types"
return ()
trace_goal[grind.debug.congr] "{e} = {e'}"
pushEqHEq e e' congrPlaceholderProof
Expand Down Expand Up @@ -231,13 +231,13 @@ private partial def internalizeImpl (e : Expr) (generation : Nat) (parent? : Opt
| .lit .. | .const .. =>
mkENode e generation
| .mvar .. =>
reportIssue m!"unexpected metavariable during internalization{indentExpr e}\n`grind` is not supposed to be used in goals containing metavariables."
reportIssue! "unexpected metavariable during internalization{indentExpr e}\n`grind` is not supposed to be used in goals containing metavariables."
mkENode' e generation
| .mdata .. =>
reportIssue m!"unexpected metadata found during internalization{indentExpr e}\n`grind` uses a pre-processing step that eliminates metadata"
reportIssue! "unexpected metadata found during internalization{indentExpr e}\n`grind` uses a pre-processing step that eliminates metadata"
mkENode' e generation
| .proj .. =>
reportIssue m!"unexpected kernel projection term during internalization{indentExpr e}\n`grind` uses a pre-processing step that folds them as projection applications, the pre-processor should have failed to fold this term"
reportIssue! "unexpected kernel projection term during internalization{indentExpr e}\n`grind` uses a pre-processing step that folds them as projection applications, the pre-processor should have failed to fold this term"
mkENode' e generation
| .app .. =>
if (← isLitValue e) then
Expand Down
17 changes: 9 additions & 8 deletions src/Lean/Meta/Tactic/Grind/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -131,14 +131,15 @@ def Result.hasFailures (r : Result) : Bool :=

def Result.toMessageData (result : Result) : MetaM MessageData := do
let mut msgs ← result.failures.mapM (goalToMessageData · result.config)
let mut issues := result.issues
unless result.skipped.isEmpty do
let m := m!"#{result.skipped.length} other goal(s) were not fully processed due to previous failures, threshold: `(failures := {result.config.failures})`"
issues := .trace { cls := `issue } m #[] :: issues
unless issues.isEmpty do
msgs := msgs ++ [.trace { cls := `grind } "Issues" issues.reverse.toArray]
if let some msg ← result.counters.toMessageData? then
msgs := msgs ++ [msg]
if result.config.verbose then
let mut issues := result.issues
unless result.skipped.isEmpty do
let m := m!"#{result.skipped.length} other goal(s) were not fully processed due to previous failures, threshold: `(failures := {result.config.failures})`"
issues := .trace { cls := `issue } m #[] :: issues
unless issues.isEmpty do
msgs := msgs ++ [.trace { cls := `grind } "Issues" issues.reverse.toArray]
if let some msg ← result.counters.toMessageData? then
msgs := msgs ++ [msg]
return MessageData.joinSep msgs m!"\n"

def main (mvarId : MVarId) (params : Params) (mainDeclName : Name) (fallback : Fallback) : MetaM Result := do profileitM Exception "grind" (← getOptions) do
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Meta/Tactic/Grind/MatchCond.lean
Original file line number Diff line number Diff line change
Expand Up @@ -301,7 +301,7 @@ where
-- created when we infer the type of the `noConfusion` term below
let lhs ← shareCommon lhs
let some root ← getRootENode? lhs
| reportIssue "found term that has not been internalized{indentExpr lhs}\nwhile trying to construct a proof for `MatchCond`{indentExpr e}"
| reportIssue! "found term that has not been internalized{indentExpr lhs}\nwhile trying to construct a proof for `MatchCond`{indentExpr e}"
return none
let isHEq := α?.isSome
let h ← if isHEq then
Expand Down Expand Up @@ -430,7 +430,7 @@ builtin_grind_propagator propagateMatchCondUp ↑Grind.MatchCond := fun e => do
else
if !(← isStatisfied e) then return ()
let some h ← mkMatchCondProof? e
| reportIssue m!"failed to construct proof for{indentExpr e}"; return ()
| reportIssue! "failed to construct proof for{indentExpr e}"; return ()
trace_goal[grind.debug.matchCond] "{← inferType h}"
pushEqTrue e <| mkEqTrueCore e h

Expand Down
11 changes: 7 additions & 4 deletions src/Lean/Meta/Tactic/Grind/PP.lean
Original file line number Diff line number Diff line change
Expand Up @@ -151,10 +151,13 @@ private def ppCasesTrace : M Unit := do
pushMsg <| .trace { cls := `cases } "Case analyses" msgs

def goalToMessageData (goal : Goal) (config : Grind.Config) : MetaM MessageData := goal.mvarId.withContext do
let (_, m) ← go goal |>.run #[]
let gm := MessageData.trace { cls := `grind, collapsed := false } "Diagnostics" m
let r := m!"{.ofGoal goal.mvarId}\n{gm}"
addMessageContextFull r
if config.verbose then
let (_, m) ← go goal |>.run #[]
let gm := MessageData.trace { cls := `grind, collapsed := false } "Diagnostics" m
let r := m!"{.ofGoal goal.mvarId}\n{gm}"
addMessageContextFull r
else
return .ofGoal goal.mvarId
where
go : M Unit := do
pushMsg <| ppExprArray `facts "Asserted facts" goal.facts.toArray `prop
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind/Solve.lean
Original file line number Diff line number Diff line change
Expand Up @@ -40,7 +40,7 @@ def pushFailure (goal : Goal) : M Unit := do
x goal
catch ex =>
if ex.isMaxHeartbeat || ex.isMaxRecDepth then
reportIssue ex.toMessageData
reportIssue! ex.toMessageData
pushFailure goal
return true
else
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 @@ -104,7 +104,7 @@ private def checkCaseSplitStatus (e : Expr) : GoalM CaseSplitStatus := do
if e.isFVar then
let type ← whnfD (← inferType e)
let report : GoalM Unit := do
reportIssue "cannot perform case-split on {e}, unexpected type{indentExpr type}"
reportIssue! "cannot perform case-split on {e}, unexpected type{indentExpr type}"
let .const declName _ := type.getAppFn | report; return .resolved
let .inductInfo info ← getConstInfo declName | report; return .resolved
return .ready info.ctors.length info.isRec
Expand Down
9 changes: 9 additions & 0 deletions src/Lean/Meta/Tactic/Grind/Types.lean
Original file line number Diff line number Diff line change
Expand Up @@ -244,6 +244,15 @@ def reportIssue (msg : MessageData) : GrindM Unit := do
-/
trace[grind.issues] msg

private def expandReportIssueMacro (s : Syntax) : MacroM (TSyntax `doElem) := do
let msg ← if s.getKind == interpolatedStrKind then `(m! $(⟨s⟩)) else `(($(⟨s⟩) : MessageData))
`(doElem| do
if (← getConfig).verbose then
reportIssue $msg)

macro "reportIssue!" s:(interpolatedStr(term) <|> term) : doElem => do
expandReportIssueMacro s.raw

/--
Stores information for a node in the egraph.
Each internalized expression `e` has an `ENode` associated with it.
Expand Down
13 changes: 13 additions & 0 deletions tests/lean/run/grind_t1.lean
Original file line number Diff line number Diff line change
Expand Up @@ -259,6 +259,19 @@ set_option trace.grind.split true in
example (p q : Prop) : (p ↔ q) → p → False := by
grind -- should not split on (p ↔ q)

/--
error: `grind` failed
case grind
p q : Prop
a✝¹ : p = q
a✝ : p
⊢ False
-/
#guard_msgs (error) in
example (p q : Prop) : (p ↔ q) → p → False := by
grind -verbose -- We should not get any diagnostics


/--
error: `grind` failed
case grind
Expand Down

0 comments on commit 38086a8

Please sign in to comment.