Skip to content

Commit

Permalink
feat: display diagnostic information at term and tactic `set_option d…
Browse files Browse the repository at this point in the history
…iagnostics true` (#4048)

We don't need to include reduction info at `simp` diagnostic
information.
  • Loading branch information
leodemoura authored May 1, 2024
1 parent 5f72769 commit 5f1c4df
Show file tree
Hide file tree
Showing 14 changed files with 103 additions and 37 deletions.
2 changes: 1 addition & 1 deletion src/Lean/Compiler/ExternAttr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -156,7 +156,7 @@ private def getExternConstArity (declName : Name) : CoreM Nat := do
@[export lean_get_extern_const_arity]
def getExternConstArityExport (env : Environment) (declName : Name) : IO (Option Nat) := do
try
let (arity, _) ← (getExternConstArity declName).toIO { fileName := "<compiler>", fileMap := default, diag := false } { env := env }
let (arity, _) ← (getExternConstArity declName).toIO { fileName := "<compiler>", fileMap := default } { env := env }
return some arity
catch
| IO.Error.userError _ => return none
Expand Down
15 changes: 11 additions & 4 deletions src/Lean/CoreM.lean
Original file line number Diff line number Diff line change
Expand Up @@ -121,7 +121,15 @@ instance : MonadOptions CoreM where
getOptions := return (← read).options

instance : MonadWithOptions CoreM where
withOptions f x := withReader (fun ctx => { ctx with options := f ctx.options }) x
withOptions f x :=
withReader
(fun ctx =>
let options := f ctx.options
{ ctx with
options
diag := diagnostics.get options
maxRecDepth := maxRecDepth.get options })
x

instance : AddMessageContext CoreM where
addMessageContext := addMessageContextPartial
Expand Down Expand Up @@ -223,7 +231,7 @@ def mkFreshUserName (n : Name) : CoreM Name :=
instance [MetaEval α] : MetaEval (CoreM α) where
eval env opts x _ := do
let x : CoreM α := do try x finally printTraces
let (a, s) ← x.toIO { maxRecDepth := maxRecDepth.get opts, options := opts, fileName := "<CoreM>", fileMap := default, diag := diagnostics.get opts } { env := env }
let (a, s) ← (withOptions (fun _ => opts) x).toIO { fileName := "<CoreM>", fileMap := default } { env := env }
MetaEval.eval s.env opts a (hideUnit := true)

-- withIncRecDepth for a monad `m` such that `[MonadControlT CoreM n]`
Expand Down Expand Up @@ -398,8 +406,7 @@ def isDiagnosticsEnabled : CoreM Bool :=

def ImportM.runCoreM (x : CoreM α) : ImportM α := do
let ctx ← read
let opts := ctx.opts
let (a, _) ← x.toIO { options := opts, fileName := "<ImportM>", fileMap := default, diag := getDiag opts } { env := ctx.env }
let (a, _) ← (withOptions (fun _ => ctx.opts) x).toIO { fileName := "<ImportM>", fileMap := default } { env := ctx.env }
return a

/-- Return `true` if the exception was generated by one our resource limits. -/
Expand Down
9 changes: 7 additions & 2 deletions src/Lean/Elab/BuiltinTerm.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
-/
prelude
import Lean.Meta.Diagnostics
import Lean.Elab.Open
import Lean.Elab.SetOption
import Lean.Elab.Eval
Expand Down Expand Up @@ -313,8 +314,12 @@ private def mkSilentAnnotationIfHole (e : Expr) : TermElabM Expr := do

@[builtin_term_elab «set_option»] def elabSetOption : TermElab := fun stx expectedType? => do
let options ← Elab.elabSetOption stx[1] stx[3]
withTheReader Core.Context (fun ctx => { ctx with maxRecDepth := maxRecDepth.get options, options := options, diag := getDiag options }) do
elabTerm stx[5] expectedType?
withOptions (fun _ => options) do
try
elabTerm stx[5] expectedType?
finally
if stx[1].getId == `diagnostics then
reportDiag

@[builtin_term_elab withAnnotateTerm] def elabWithAnnotateTerm : TermElab := fun stx expectedType? => do
match stx with
Expand Down
7 changes: 4 additions & 3 deletions src/Lean/Elab/InfoTree/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -102,9 +102,10 @@ def ContextInfo.runCoreM (info : ContextInfo) (x : CoreM α) : IO α := do
have been used in `lctx` and `info.mctx`.
-/
(·.1) <$>
x.toIO { options := info.options, currNamespace := info.currNamespace, openDecls := info.openDecls
fileName := "<InfoTree>", fileMap := default, diag := getDiag info.options }
{ env := info.env, ngen := info.ngen }
(withOptions (fun _ => info.options) x).toIO
{ currNamespace := info.currNamespace, openDecls := info.openDecls
fileName := "<InfoTree>", fileMap := default }
{ env := info.env, ngen := info.ngen }

def ContextInfo.runMetaM (info : ContextInfo) (lctx : LocalContext) (x : MetaM α) : IO α := do
(·.1) <$> info.runCoreM (x.run { lctx := lctx } { mctx := info.mctx })
Expand Down
9 changes: 7 additions & 2 deletions src/Lean/Elab/Tactic/BuiltinTactic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
-/
prelude
import Lean.Meta.Diagnostics
import Lean.Meta.Tactic.Apply
import Lean.Meta.Tactic.Assumption
import Lean.Meta.Tactic.Contradiction
Expand Down Expand Up @@ -163,8 +164,12 @@ private def getOptRotation (stx : Syntax) : Nat :=

@[builtin_tactic Parser.Tactic.set_option] def elabSetOption : Tactic := fun stx => do
let options ← Elab.elabSetOption stx[1] stx[3]
withTheReader Core.Context (fun ctx => { ctx with maxRecDepth := maxRecDepth.get options, options := options, diag := getDiag options }) do
evalTactic stx[5]
withOptions (fun _ => options) do
try
evalTactic stx[5]
finally
if stx[1].getId == `diagnostics then
reportDiag

@[builtin_tactic Parser.Tactic.allGoals] def evalAllGoals : Tactic := fun stx => do
let mvarIds ← getGoals
Expand Down
4 changes: 1 addition & 3 deletions src/Lean/Elab/Tactic/Simp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -414,10 +414,8 @@ where
return stats

def withSimpDiagnostics (x : TacticM Simp.Diagnostics) : TacticM Unit := do
let origDiag := (← getThe Meta.State).diag
let stats ← x
Simp.reportDiag stats origDiag
return ()
Simp.reportDiag stats

/-
"simp" (config)? (discharger)? (" only")? (" [" ((simpStar <|> simpErase <|> simpLemma),*,?) "]")?
Expand Down
11 changes: 2 additions & 9 deletions src/Lean/Meta/Tactic/Simp/Diagnostics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -32,23 +32,16 @@ def mkSimpDiagSummary (counters : PHashMap Origin Nat) (usedCounters? : Option (
data := data.push m!"{if data.isEmpty then " " else "\n"}{key} ↦ {counter}{usedMsg}"
return { data, max := entries[0]!.2 }

def reportDiag (diag : Simp.Diagnostics) (diagOrig : Meta.Diagnostics) : MetaM Unit := do
def reportDiag (diag : Simp.Diagnostics) : MetaM Unit := do
if (← isDiagnosticsEnabled) then
let used ← mkSimpDiagSummary diag.usedThmCounter
let tried ← mkSimpDiagSummary diag.triedThmCounter diag.usedThmCounter
let congr ← mkDiagSummary diag.congrThmCounter
let unfoldCounter := subCounters (← get).diag.unfoldCounter diagOrig.unfoldCounter
let unfoldDefault ← mkDiagSummaryForUnfolded unfoldCounter
let unfoldInstance ← mkDiagSummaryForUnfolded unfoldCounter (instances := true)
let unfoldReducible ← mkDiagSummaryForUnfoldedReducible unfoldCounter
unless used.isEmpty && tried.isEmpty && unfoldDefault.isEmpty && unfoldInstance.isEmpty && unfoldReducible.isEmpty do
unless used.isEmpty && tried.isEmpty && congr.isEmpty do
let m := MessageData.nil
let m := appendSection m `simp "used theorems" used
let m := appendSection m `simp "tried theorems" tried
let m := appendSection m `simp "tried congruence theorems" congr
let m := appendSection m `reduction "unfolded declarations" unfoldDefault
let m := appendSection m `reduction "unfolded instances" unfoldInstance
let m := appendSection m `reduction "unfolded reducible declarations" unfoldReducible
let m := m ++ "use `set_option diagnostics.threshold <num>` to control threshold for reporting counters"
logInfo m

Expand Down
6 changes: 2 additions & 4 deletions src/Lean/Meta/Tactic/Simp/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -663,11 +663,10 @@ def main (e : Expr) (ctx : Context) (stats : Stats := {}) (methods : Methods :=
return (r, { s with })
where
simpMain (e : Expr) : SimpM Result := withCatchingRuntimeEx do
let origDiag := (← getThe Meta.State).diag
try
withoutCatchingRuntimeEx <| simp e
catch ex =>
reportDiag (← get).diag origDiag
reportDiag (← get).diag
if ex.isRuntime then
throwNestedTacticEx `simp ex
else
Expand All @@ -679,11 +678,10 @@ def dsimpMain (e : Expr) (ctx : Context) (stats : Stats := {}) (methods : Method
pure (r, { s with })
where
dsimpMain (e : Expr) : SimpM Expr := withCatchingRuntimeEx do
let origDiag := (← getThe Meta.State).diag
try
withoutCatchingRuntimeEx <| dsimp e
catch ex =>
reportDiag (← get).diag origDiag
reportDiag (← get).diag
if ex.isRuntime then
throwNestedTacticEx `simp ex
else
Expand Down
6 changes: 3 additions & 3 deletions src/Lean/PrettyPrinter.lean
Original file line number Diff line number Diff line change
Expand Up @@ -51,9 +51,9 @@ def ppExprWithInfos (e : Expr) (optsPerPos : Delaborator.OptionsPerPos := {}) (d

@[export lean_pp_expr]
def ppExprLegacy (env : Environment) (mctx : MetavarContext) (lctx : LocalContext) (opts : Options) (e : Expr) : IO Format :=
Prod.fst <$> ((ppExpr e).run' { lctx := lctx } { mctx := mctx }).toIO {
options := opts, fileName := "<PrettyPrinter>", fileMap := default, diag := getDiag opts
} { env := env }
Prod.fst <$> ((withOptions (fun _ => opts) <| ppExpr e).run' { lctx := lctx } { mctx := mctx }).toIO
{ fileName := "<PrettyPrinter>", fileMap := default }
{ env := env }

def ppTactic (stx : TSyntax `tactic) : CoreM Format := ppCategory `tactic stx

Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Replay.lean
Original file line number Diff line number Diff line change
Expand Up @@ -50,7 +50,7 @@ def isTodo (name : Name) : M Bool := do

/-- Use the current `Environment` to throw a `KernelException`. -/
def throwKernelException (ex : KernelException) : M Unit := do
let ctx := { fileName := "", options := ({} : KVMap), fileMap := default, diag := false }
let ctx := { fileName := "", options := ({} : KVMap), fileMap := default }
let state := { env := (← get).env }
Prod.fst <$> (Lean.Core.CoreM.toIO · ctx state) do Lean.throwKernelException ex

Expand Down
2 changes: 1 addition & 1 deletion src/lake/Lake/CLI/Translate.lean
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ def Package.mkConfigString (pkg : Package) (lang : ConfigLang) : LogIO String :=
| .lean => do
let env ← importModulesUsingCache #[`Lake] {} 1024
let pp := ppModule <| descopeTSyntax <| pkg.mkLeanConfig
match (← pp.toIO {fileName := "", fileMap := default, diag := false} {env} |>.toBaseIO) with
match (← pp.toIO {fileName := "", fileMap := default} {env} |>.toBaseIO) with
| .ok (fmt, _) => pure <| (toString fmt).trim ++ "\n"
| .error ex =>
error s!"(internal) failed to pretty print Lean configuration: {ex.toString}"
2 changes: 1 addition & 1 deletion src/lake/Lake/Toml/Load.lean
Original file line number Diff line number Diff line change
Expand Up @@ -22,7 +22,7 @@ def loadToml (ictx : InputContext) : EIO MessageLog Table := do
throw <| MessageLog.empty.add <| mkParserErrorMessage ictx s errorMsg
else if ictx.input.atEnd s.pos then
let act := elabToml ⟨s.stxStack.back⟩
match (← act.run {fileName := ictx.fileName, fileMap := ictx.fileMap, diag := false} {env} |>.toBaseIO) with
match (← act.run {fileName := ictx.fileName, fileMap := ictx.fileMap} {env} |>.toBaseIO) with
| .ok (t, s) =>
if s.messages.hasErrors then
throw s.messages
Expand Down
53 changes: 53 additions & 0 deletions tests/lean/run/diagnostics.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,53 @@
opaque q : Nat → Nat
def f (x : Nat) : Nat :=
match x with
| 0 => 1
| x+1 => q (f x)

theorem f_eq : f (x + 1) = q (f x) := rfl

set_option trace.Meta.debug true

/--
info: [reduction] unfolded declarations (max: 15, num: 6):
Nat.rec ↦ 15
Add.add ↦ 10
HAdd.hAdd ↦ 10
Nat.add ↦ 10
f ↦ 5
OfNat.ofNat ↦ 5[reduction] unfolded instances (max: 5, num: 1):
instOfNatNat ↦ 5[reduction] unfolded reducible declarations (max: 15, num: 1):
Nat.casesOn ↦ 15use `set_option diagnostics.threshold <num>` to control threshold for reporting counters
-/
#guard_msgs in
example : f (x + 5) = q (q (q (q (q (f x))))) :=
set_option diagnostics.threshold 4 in
set_option diagnostics true in
rfl

/--
info: [reduction] unfolded declarations (max: 15, num: 6):
Nat.rec ↦ 15
Add.add ↦ 10
HAdd.hAdd ↦ 10
Nat.add ↦ 10
f ↦ 5
OfNat.ofNat ↦ 5[reduction] unfolded instances (max: 5, num: 1):
instOfNatNat ↦ 5[reduction] unfolded reducible declarations (max: 15, num: 1):
Nat.casesOn ↦ 15use `set_option diagnostics.threshold <num>` to control threshold for reporting counters
-/
#guard_msgs in
example : f (x + 5) = q (q (q (q (q (f x))))) := by
set_option diagnostics.threshold 4 in
set_option diagnostics true in
rfl
12 changes: 9 additions & 3 deletions tests/lean/run/simpDiag.lean
Original file line number Diff line number Diff line change
Expand Up @@ -98,11 +98,12 @@ opaque q1 : Nat → Nat → Prop
/--
info: [simp] used theorems (max: 1, num: 1):
q1_ax ↦ 1[simp] tried theorems (max: 1, num: 1):
q1_ax ↦ 1, succeeded: 1[reduction] unfolded declarations (max: 246, num: 2):
q1_ax ↦ 1, succeeded: 1use `set_option diagnostics.threshold <num>` to control threshold for reporting counters
---
info: [reduction] unfolded declarations (max: 246, num: 2):
Nat.rec ↦ 246
OfNat.ofNat ↦ 24[reduction] unfolded instances (max: 12, num: 1):
instOfNatNat ↦ 12[reduction] unfolded reducible declarations (max: 246, num: 2):
OfNat.ofNat ↦ 24[reduction] unfolded reducible declarations (max: 246, num: 2):
h ↦ 246
Nat.casesOn ↦ 246use `set_option diagnostics.threshold <num>` to control threshold for reporting counters
-/
Expand All @@ -111,3 +112,8 @@ example : q1 x (h 40) := by
set_option diagnostics true in
set_option diagnostics.threshold 0 in
simp

example : q1 x (h 40) := by
set_option diagnostics true in
set_option diagnostics.threshold 0 in
simp

0 comments on commit 5f1c4df

Please sign in to comment.