Skip to content

Commit

Permalink
feat: pretty print Array DiscrTree.Key (#4208)
Browse files Browse the repository at this point in the history
  • Loading branch information
leodemoura authored May 17, 2024
1 parent ee0bcc8 commit 4d2ff6f
Show file tree
Hide file tree
Showing 4 changed files with 69 additions and 29 deletions.
45 changes: 45 additions & 0 deletions src/Lean/Meta/DiscrTree.lean
Original file line number Diff line number Diff line change
Expand Up @@ -80,6 +80,51 @@ def Key.format : Key → Format

instance : ToFormat Key := ⟨Key.format⟩

/--
Helper function for converting an entry (i.e., `Array Key`) to the discrimination tree into
`MessageData` that is more user-friendly. We use this function to implement diagnostic information.
-/
partial def keysAsPattern (keys : Array Key) : CoreM MessageData := do
go (parenIfNonAtomic := false) |>.run' keys.toList
where
next? : StateRefT (List Key) CoreM (Option Key) := do
let key :: keys ← get | return none
set keys
return some key

mkApp (f : MessageData) (args : Array MessageData) (parenIfNonAtomic : Bool) : CoreM MessageData := do
if args.isEmpty then
return f
else
let mut r := f
for arg in args do
r := r ++ m!" {arg}"
if parenIfNonAtomic then
return m!"({r})"
else
return r

go (parenIfNonAtomic := true) : StateRefT (List Key) CoreM MessageData := do
let some key ← next? | return .nil
match key with
| .const declName nargs =>
mkApp m!"{← mkConstWithLevelParams declName}" (← goN nargs) parenIfNonAtomic
| .fvar fvarId nargs =>
mkApp m!"{mkFVar fvarId}" (← goN nargs) parenIfNonAtomic
| .proj _ i nargs =>
mkApp m!"{← go}.{i+1}" (← goN nargs) parenIfNonAtomic
| .arrow => return "<arrow>"
| .star => return "_"
| .other => return "<other>"
| .lit (.natVal v) => return m!"{v}"
| .lit (.strVal v) => return m!"{v}"

goN (num : Nat) : StateRefT (List Key) CoreM (Array MessageData) := do
let mut r := #[]
for _ in [: num] do
r := r.push (← go)
return r

def Key.arity : Key → Nat
| .const _ a => a
| .fvar _ a => a
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Simp/Diagnostics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -41,7 +41,7 @@ private def mkTheoremsWithBadKeySummary (thms : PArray SimpTheorem) : MetaM Diag
else
let mut data := #[]
for thm in thms do
data := data.push m!"{if data.isEmpty then " " else "\n"}{← originToKey thm.origin}, key: {thm.keys.map (·.format)}"
data := data.push m!"{if data.isEmpty then " " else "\n"}{← originToKey thm.origin}, key: {← DiscrTree.keysAsPattern thm.keys}"
pure ()
return { data }

Expand Down
30 changes: 2 additions & 28 deletions tests/lean/run/4171.lean
Original file line number Diff line number Diff line change
Expand Up @@ -717,20 +717,7 @@ example (M : Comon_ (Mon_ C)) : Mon_ (Comon_ C) where

/--
info: [simp] theorems with bad keys
foo, key: [Quiver.Hom.unop,
*,
*,
*,
*,
Opposite.op,
Quiver.Hom,
*,
*,
Opposite.0,
*,
Opposite.0,
*,
*]use `set_option diagnostics.threshold <num>` to control threshold for reporting counters
foo, key: @Quiver.Hom.unop _ _ _ _ (@Opposite.op (@Quiver.Hom _ _ _.1 _.1) _)use `set_option diagnostics.threshold <num>` to control threshold for reporting counters
-/
#guard_msgs in
example (M : Comon_ (Mon_ C)) : Mon_ (Comon_ C) where
Expand All @@ -749,20 +736,7 @@ attribute [simp] foo

/--
info: [simp] theorems with bad keys
foo, key: [Quiver.Hom.unop,
*,
*,
*,
*,
Opposite.op,
Quiver.Hom,
*,
*,
Opposite.0,
*,
Opposite.0,
*,
*]use `set_option diagnostics.threshold <num>` to control threshold for reporting counters
foo, key: @Quiver.Hom.unop _ _ _ _ (@Opposite.op (@Quiver.Hom _ _ _.1 _.1) _)use `set_option diagnostics.threshold <num>` to control threshold for reporting counters
-/
#guard_msgs in
example (M : Comon_ (Mon_ C)) : Mon_ (Comon_ C) where
Expand Down
21 changes: 21 additions & 0 deletions tests/lean/run/simpIndexDiag.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,21 @@
namespace Ex1
opaque f : Nat → Nat → Nat

@[simp] theorem foo : f x (x, y).2 = y := by sorry

example : f a b ≤ b := by
fail_if_success simp -- should fail
set_option diagnostics true in
simp (config := { index := false })

end Ex1

namespace Ex2
opaque f : Nat → Nat → Nat

@[simp] theorem foo : f x (no_index (x, y).2) = y := by sorry

example : f a b ≤ b := by
simp -- should work since `foo` is using `no_index`

end Ex2

0 comments on commit 4d2ff6f

Please sign in to comment.