Skip to content

Commit

Permalink
fix: ofScientific at simp (#3628)
Browse files Browse the repository at this point in the history
closes #2159
  • Loading branch information
leodemoura authored Mar 7, 2024
1 parent d731854 commit 611b174
Show file tree
Hide file tree
Showing 2 changed files with 36 additions and 11 deletions.
34 changes: 23 additions & 11 deletions src/Lean/Meta/Tactic/Simp/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -50,6 +50,10 @@ def foldRawNatLit (e : Expr) : SimpM Expr := do
return toExpr n
| none => return e

/-- Return true if `e` is of the form `ofScientific n b m` where `n` and `m` are kernel Nat literals. -/
def isOfScientificLit (e : Expr) : Bool :=
e.isAppOfArity ``OfScientific.ofScientific 5 && (e.getArg! 4).isRawNatLit && (e.getArg! 2).isRawNatLit

private def reduceProjFn? (e : Expr) : SimpM (Option Expr) := do
matchConst e.getAppFn (fun _ => pure none) fun cinfo _ => do
match (← getProjectionFnInfo? cinfo.name) with
Expand Down Expand Up @@ -410,27 +414,35 @@ private def dsimpReduce : DSimproc := fun e => do
eNew ← reduceFVar (← getConfig) (← getSimpTheorems) eNew
if eNew != e then return .visit eNew else return .done e

/--
Auliliary `dsimproc` for not visiting `OfNat.ofNat` application subterms.
This is the `dsimp` equivalent of the approach used at `visitApp`.
Recall that we fold orphan raw Nat literals.
-/
private def doNotVisitOfNat : DSimproc := fun e => do
if isOfNatNatLit e then
if (← readThe Simp.Context).isDeclToUnfold ``OfNat.ofNat then
/-- Helper `dsimproc` for `doNotVisitOfNat` and `doNotVisitOfScientific` -/
private def doNotVisit (pred : Expr → Bool) (declName : Name) : DSimproc := fun e => do
if pred e then
if (← readThe Simp.Context).isDeclToUnfold declName then
return .continue e
else
return .done e
else
return .continue e

/--
Auliliary `dsimproc` for not visiting `OfNat.ofNat` application subterms.
This is the `dsimp` equivalent of the approach used at `visitApp`.
Recall that we fold orphan raw Nat literals.
-/
private def doNotVisitOfNat : DSimproc := doNotVisit isOfNatNatLit ``OfNat.ofNat

/--
Auliliary `dsimproc` for not visiting `OfScientific.ofScientific` application subterms.
-/
private def doNotVisitOfScientific : DSimproc := doNotVisit isOfScientificLit ``OfScientific.ofScientific

@[export lean_dsimp]
private partial def dsimpImpl (e : Expr) : SimpM Expr := do
let cfg ← getConfig
unless cfg.dsimp do
return e
let m ← getMethods
let pre := m.dpre >> doNotVisitOfNat
let pre := m.dpre >> doNotVisitOfNat >> doNotVisitOfScientific
let post := m.dpost >> dsimpReduce
transform (usedLetOnly := cfg.zeta) e (pre := pre) (post := post)

Expand Down Expand Up @@ -550,8 +562,8 @@ def congr (e : Expr) : SimpM Result := do
congrDefault e

def simpApp (e : Expr) : SimpM Result := do
if isOfNatNatLit e then
-- Recall that we expand "orphan" kernel Nat literals `n` into `OfNat.ofNat n`
if isOfNatNatLit e || isOfScientificLit e then
-- Recall that we fold "orphan" kernel Nat literals `n` into `OfNat.ofNat n`
return { expr := e }
else
congr e
Expand Down
13 changes: 13 additions & 0 deletions tests/lean/run/2159.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
/--
warning: declaration uses 'sorry'
---
info: ⊢ 1.2 < 2
---
info: ⊢ 1.2 < 2
-/
#guard_msgs in
example : 1.2 < 2 := by
trace_state
fail_if_success simp only
trace_state
sorry

0 comments on commit 611b174

Please sign in to comment.