Skip to content

Commit

Permalink
fix: grind canonicalizer
Browse files Browse the repository at this point in the history
This PR fixes a bug in the `grind` canonicalizer.
  • Loading branch information
leodemoura committed Jan 9, 2025
1 parent 0afa1d1 commit 84cdab1
Show file tree
Hide file tree
Showing 3 changed files with 28 additions and 4 deletions.
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind.lean
Original file line number Diff line number Diff line change
Expand Up @@ -53,5 +53,5 @@ builtin_initialize registerTraceClass `grind.debug.parent
builtin_initialize registerTraceClass `grind.debug.final
builtin_initialize registerTraceClass `grind.debug.forallPropagator
builtin_initialize registerTraceClass `grind.debug.split

builtin_initialize registerTraceClass `grind.debug.canon
end Lean
17 changes: 14 additions & 3 deletions src/Lean/Meta/Tactic/Grind/Canon.lean
Original file line number Diff line number Diff line change
Expand Up @@ -66,11 +66,13 @@ def canonElemCore (f : Expr) (i : Nat) (e : Expr) (isInst : Bool) : StateT State
-- in general safe to replace `e` with `c` if `c` has more free variables than `e`.
-- However, we don't revert previously canonicalized elements in the `grind` tactic.
modify fun s => { s with canon := s.canon.insert e c }
trace[grind.debug.canon] "found {e} ===> {c}"
return c
if isInst then
if (← isDiagnosticsEnabled <&&> pure (c.fvarsSubset e) <&&> (withDefault <| isDefEq e c)) then
if (← isDiagnosticsEnabled <&&> (withDefault <| isDefEq e c)) then
-- TODO: consider storing this information in some structure that can be browsed later.
trace[grind.issues] "the following `grind` static elements are definitionally equal with `default` transparency, but not with `instances` transparency{indentExpr e}\nand{indentExpr c}"
trace[grind.debug.canon] "({f}, {i}) ↦ {e}"
modify fun s => { s with canon := s.canon.insert e e, argMap := s.argMap.insert key (e::cs) }
return e

Expand All @@ -92,6 +94,12 @@ private inductive ShouldCanonResult where
visit
deriving Inhabited

instance : Repr ShouldCanonResult where
reprPrec r _ := match r with
| .canonType => "canonType"
| .canonInst => "canonInst"
| .visit => "visit"

/--
See comments at `ShouldCanonResult`.
-/
Expand All @@ -102,7 +110,9 @@ def shouldCanon (pinfos : Array ParamInfo) (i : Nat) (arg : Expr) : MetaM Should
return .canonInst
else if pinfo.isProp then
return .visit
if (← isTypeFormer arg) then
if (← isProp arg) then
return .visit
else if (← isTypeFormer arg) then
return .canonType
else
return .visit
Expand Down Expand Up @@ -151,7 +161,8 @@ where
return e'

/-- Canonicalizes nested types, type formers, and instances in `e`. -/
def canon (e : Expr) : StateT State MetaM Expr :=
def canon (e : Expr) : StateT State MetaM Expr := do
trace[grind.debug.canon] "{e}"
unsafe canonImpl e

end Canon
Expand Down
13 changes: 13 additions & 0 deletions tests/lean/run/grind_canon_bug.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
open List

attribute [grind =] getElem_cons_zero
attribute [grind =] getElem?_cons_zero

example (h : (a :: t)[0]? = some b) : (a :: t)[0] = b := by
grind --

example [Inhabited α] : ∀ {l : List α} {i : Nat}, l[i]? = some b → l[i]! = b
| a::t, 0, _ => by
rw [getElem!_pos _ _ sorry]
grind
| _::l, _+1, e => sorry

0 comments on commit 84cdab1

Please sign in to comment.