Skip to content

Commit

Permalink
fix: proof canonicalizer in grind (#6499)
Browse files Browse the repository at this point in the history
This PR fixes the proof canonicalizer for `grind`.
  • Loading branch information
leodemoura authored Jan 2, 2025
1 parent f7c4edc commit a8d09da
Showing 1 changed file with 14 additions and 8 deletions.
22 changes: 14 additions & 8 deletions src/Lean/Meta/Tactic/Grind/Canon.lean
Original file line number Diff line number Diff line change
Expand Up @@ -41,8 +41,9 @@ Furthermore, `grind` will not be able to infer that `HEq (a + a) (b + b)` even
-/

structure State where
argMap : PHashMap (Expr × Nat) (List Expr) := {}
canon : PHashMap Expr Expr := {}
argMap : PHashMap (Expr × Nat) (List Expr) := {}
canon : PHashMap Expr Expr := {}
proofCanon : PHashMap Expr Expr := {}
deriving Inhabited

/--
Expand All @@ -61,10 +62,11 @@ def canonElemCore (f : Expr) (i : Nat) (e : Expr) (isInst : Bool) : StateT State
let cs := s.argMap.find? key |>.getD []
for c in cs do
if (← isDefEq e c) then
if c.fvarsSubset e then
-- It is not in general safe to replace `e` with `c` if `c` has more free variables than `e`.
modify fun s => { s with canon := s.canon.insert e c }
return c
-- We used to check `c.fvarsSubset e` because it is not
-- 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 }
return c
if isInst then
if (← isDiagnosticsEnabled <&&> pure (c.fvarsSubset e) <&&> (withDefault <| isDefEq e c)) then
-- TODO: consider storing this information in some structure that can be browsed later.
Expand Down Expand Up @@ -125,10 +127,14 @@ where
return r
e.withApp fun f args => do
let e' ← if f.isConstOf ``Lean.Grind.nestedProof && args.size == 2 then
-- We just canonize the proposition
let prop := args[0]!
let prop' ← visit prop
pure <| if ptrEq prop prop' then mkAppN f (args.set! 0 prop') else e
if let some r := (← getThe State).proofCanon.find? prop' then
pure r
else
let e' := if ptrEq prop prop' then e else mkAppN f (args.set! 0 prop')
modifyThe State fun s => { s with proofCanon := s.proofCanon.insert prop' e' }
pure e'
else
let pinfos := (← getFunInfo f).paramInfo
let mut modified := false
Expand Down

0 comments on commit a8d09da

Please sign in to comment.