Skip to content

Commit efad5a0

Browse files
committed
fix: grind order regression
This PR fixes a regression in the `grind order` module introduced by Closes #10953
1 parent b3ef7c9 commit efad5a0

File tree

2 files changed

+10
-19
lines changed

2 files changed

+10
-19
lines changed

src/Lean/Meta/Tactic/Grind/Order/Assert.lean

Lines changed: 9 additions & 19 deletions
Original file line numberDiff line numberDiff line change
@@ -126,15 +126,12 @@ Given `e` represented by constraint `c` (from `u` to `v`).
126126
Checks whether `e = True` can be propagated using the path `u --(k)--> v`.
127127
If it can, adds a new entry to propagation list.
128128
-/
129-
def checkEqTrue (u v : NodeId) (k : Weight) (c : Cnstr NodeId) (e : Expr) : OrderM Bool := do
130-
if (← isAlreadyTrue e) then return true
129+
def checkEqTrue (u v : NodeId) (k : Weight) (c : Cnstr NodeId) (e : Expr) : OrderM Unit := do
130+
if (← isAlreadyTrue e) then return ()
131131
let k' := c.getWeight
132132
trace[grind.debug.order.check_eq_true] "{← getExpr u}, {← getExpr v}, {k}, {k'}, {← c.pp}"
133133
if k ≤ k' then
134134
pushToPropagate <| .eqTrue c e u v k k'
135-
return true
136-
else
137-
return false
138135

139136
/--
140137
Returns `true` if `e` is already `False` in the `grind` core.
@@ -151,28 +148,21 @@ Given `e` represented by constraint `c` (from `v` to `u`).
151148
Checks whether `e = False` can be propagated using the path `u --(k)--> v`.
152149
If it can, adds a new entry to propagation list.
153150
-/
154-
def checkEqFalse (u v : NodeId) (k : Weight) (c : Cnstr NodeId) (e : Expr) : OrderM Bool := do
155-
if (← isAlreadyFalse e) then return true
151+
def checkEqFalse (u v : NodeId) (k : Weight) (c : Cnstr NodeId) (e : Expr) : OrderM Unit := do
152+
if (← isAlreadyFalse e) then return ()
156153
let k' := c.getWeight
157154
trace[grind.debug.order.check_eq_false] "{← getExpr u}, {← getExpr v}, {k}, {k'} {← c.pp}"
158155
if (k + k').isNeg then
159156
pushToPropagate <| .eqFalse c e u v k k'
160-
return true
161-
return false
162157

163158
/--
164159
Auxiliary function for implementing theory propagation.
165160
Traverses the constraints `c` (representing an expression `e`) s.t.
166-
`c.u = u` and `c.v = v`, it removes `c` from the list of constraints
167-
associated with `(u, v)` IF
168-
- `e` is already assigned, or
169-
- `f c e` returns true
161+
`c.u = u` and `c.v = v`.
170162
-/
171-
@[inline] def updateCnstrsOf (u v : NodeId) (f : Cnstr NodeId → Expr → OrderM Bool) : OrderM Unit := do
163+
@[inline] def forEachCnstrsOf (u v : NodeId) (f : Cnstr NodeId → Expr → OrderM Unit) : OrderM Unit := do
172164
if let some cs := (← getStruct).cnstrsOf.find? (u, v) then
173-
let cs' ← cs.filterM fun (c, e) => do
174-
return !(← f c e)
175-
modifyStruct fun s => { s with cnstrsOf := s.cnstrsOf.insert (u, v) cs' }
165+
cs.forM fun (c, e) => f c e
176166

177167
/-- Equality propagation. -/
178168
def checkEq (u v : NodeId) (k : Weight) : OrderM Unit := do
@@ -188,8 +178,8 @@ def checkEq (u v : NodeId) (k : Weight) : OrderM Unit := do
188178

189179
/-- Finds constrains and equalities to be propagated. -/
190180
def checkToPropagate (u v : NodeId) (k : Weight) : OrderM Unit := do
191-
updateCnstrsOf u v fun c e => checkEqTrue u v k c e
192-
updateCnstrsOf v u fun c e => checkEqFalse u v k c e
181+
forEachCnstrsOf u v fun c e => checkEqTrue u v k c e
182+
forEachCnstrsOf v u fun c e => checkEqFalse u v k c e
193183
checkEq u v k
194184

195185
/--

tests/lean/run/grind_10953.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1 @@
1+
example [LT α] [LE α] [Std.IsLinearOrder α] [Std.LawfulOrderLT α] {b : α} (h : b < b) : False := by grind

0 commit comments

Comments
 (0)