Skip to content

Commit

Permalink
fix: Int.reduceNeg simproc (#6468)
Browse files Browse the repository at this point in the history
This PR fixes issue #6467 


closes #6467
  • Loading branch information
leodemoura authored Dec 28, 2024
1 parent 3fc7485 commit 5930db9
Show file tree
Hide file tree
Showing 2 changed files with 18 additions and 6 deletions.
8 changes: 2 additions & 6 deletions src/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Int.lean
Original file line number Diff line number Diff line change
Expand Up @@ -49,17 +49,13 @@ If they do, they must disable the following `simprocs`.
-/

builtin_dsimproc [simp, seval] reduceNeg ((- _ : Int)) := fun e => do
unless e.isAppOfArity ``Neg.neg 3 do return .continue
let arg := e.appArg!
let_expr Neg.neg _ _ arg ← e | return .continue
if arg.isAppOfArity ``OfNat.ofNat 3 then
-- We return .done to ensure `Neg.neg` is not unfolded even when `ground := true`.
return .done e
else
let some v ← fromExpr? arg | return .continue
if v < 0 then
return .done <| toExpr (- v)
else
return .done <| toExpr v
return .done <| toExpr (- v)

/-- Return `.done` for positive Int values. We don't want to unfold in the symbolic evaluator. -/
builtin_dsimproc [seval] isPosValue ((OfNat.ofNat _ : Int)) := fun e => do
Expand Down
16 changes: 16 additions & 0 deletions tests/lean/run/6467.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,16 @@
theorem confuses_the_user : -(no_index (OfNat.ofNat <| nat_lit 2)) = -2 := by
simp

theorem ex : -(no_index (OfNat.ofNat <| nat_lit 2)) = (2 : Int) := by
simp only [Int.reduceNeg]
guard_target =ₛ -2 = 2
sorry

theorem its_true_really : -(no_index (OfNat.ofNat <| nat_lit 2)) = -2 := by
rfl

example : -(-(no_index (OfNat.ofNat <| nat_lit 2))) = 2 := by
simp

example : -(no_index (-(no_index (OfNat.ofNat <| nat_lit 2)))) = -(-(-(-3+1))) := by
simp

0 comments on commit 5930db9

Please sign in to comment.