Skip to content

Commit

Permalink
fix: redundant information in the offset constrainst module
Browse files Browse the repository at this point in the history
This PR ensures there are no redundant entries in the offset
constraint model produced by `grind`
  • Loading branch information
leodemoura committed Jan 21, 2025
1 parent 0c2fb34 commit e14474d
Show file tree
Hide file tree
Showing 3 changed files with 27 additions and 1 deletion.
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind/Arith/Model.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,7 +39,7 @@ def mkModel (goal : Goal) : MetaM (Array (Expr × Nat)) := do
We should not include the assignment for auxiliary offset terms since
they do not provide any additional information.
-/
if (isNatOffset? e).isNone && isNatNum? e != some 0 then
if !(← isLitValue e) && (isNatOffset? e).isNone && isNatNum? e != some 0 then
r := r.push (e, val)
return r

Expand Down
1 change: 1 addition & 0 deletions src/Lean/Meta/Tactic/Grind/PP.lean
Original file line number Diff line number Diff line change
Expand Up @@ -121,6 +121,7 @@ private def ppOffset : M Unit := do
let nodes := s.nodes
if nodes.isEmpty then return ()
let model ← Arith.Offset.mkModel goal
if model.isEmpty then return ()
let mut ms := #[]
for (e, val) in model do
ms := ms.push <| .trace { cls := `assign } m!"{e} := {val}" #[]
Expand Down
25 changes: 25 additions & 0 deletions tests/lean/run/grind_t1.lean
Original file line number Diff line number Diff line change
Expand Up @@ -325,3 +325,28 @@ example (x y : S) : x.a = y.a → y.b = x.b → x = y := by

example (x : S) : x.a = 10 → false ≠ x.b → x = { a := 10, b := true } := by
grind


-- In the following test, we should not display `10 := 10` and `20 := 20` in the
-- assignment produced by the offset module
/--
error: `grind` failed
case grind
a : Nat
b : Bool
a✝¹ : (if b = true then 10 else 20) = a
a✝ : b = true
⊢ False
[grind] Diagnostics
[facts] Asserted facts
[prop] (if b = true then 10 else 20) = a
[prop] b = true
[eqc] True propositions
[prop] b = true
[eqc] Equivalence classes
[eqc] {a, if b = true then 10 else 20, 10}
[eqc] {b, true}
-/
#guard_msgs (error) in
example (b : Bool) : (if b then 10 else 20) = a → b = true → False := by
grind

0 comments on commit e14474d

Please sign in to comment.