Skip to content

Commit d5ca0c7

Browse files
authored
fix: bug in cutsat model construction (#10951)
This PR fixes a bug in the `cutsat` incremental model construction. The model was not being reset when new (unsatisfied) equalities were asserted.
1 parent 3c2ab0f commit d5ca0c7

File tree

3 files changed

+28
-0
lines changed

3 files changed

+28
-0
lines changed

src/Lean/Meta/Tactic/Grind/Arith/Cutsat/EqCnstr.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -359,6 +359,8 @@ def EqCnstr.assertImpl (c : EqCnstr) : GoalM Unit := do
359359
trace[grind.debug.cutsat.subst] ">> {← getVar x}, {← c.pp}"
360360
trace[grind.cutsat.assert.store] "{← c.pp}"
361361
trace[grind.debug.cutsat.elimEq] "{← getVar x}, {← c.pp}"
362+
if (← c.satisfied) == .false then
363+
resetAssignmentFrom x
362364
modify' fun s => { s with
363365
elimEqs := s.elimEqs.set x (some c)
364366
elimStack := x :: s.elimStack

src/Lean/Meta/Tactic/Grind/Arith/Cutsat/Util.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -218,6 +218,14 @@ def DiseqCnstr.satisfied (c : DiseqCnstr) : GoalM LBool := do
218218
let some v ← c.p.eval? | return .undef
219219
return v != 0 |>.toLBool
220220

221+
/--
222+
Returns `.true` if `c` is satisfied by the current partial model,
223+
`.undef` if `c` contains unassigned variables, and `.false` otherwise.
224+
-/
225+
def EqCnstr.satisfied (c : EqCnstr) : GoalM LBool := do
226+
let some v ← c.p.eval? | return .undef
227+
return v == 0 |>.toLBool
228+
221229
/--
222230
Given a polynomial `p`, returns `some (x, k, c)` if `p` contains the monomial `k*x`,
223231
and `x` has been eliminated using the equality `c`.

tests/lean/run/grind_finish_trace.lean

Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -162,3 +162,21 @@ example (f : Int → Int → Int) (x y : Int)
162162
grind =>
163163
-- We can use `have` to golf proofs using `mbtc` and `cases`
164164
have : x = 1
165+
166+
example (f : Int → Int) (x y : Int)
167+
: 0 ≤ x → x ≤ 2 → f 0 = y → f 1 = y → f 2 = y → f x = y := by
168+
grind
169+
170+
example (f : Int → Int) (x y : Int)
171+
: 0 ≤ x → x ≤ 2 → f 0 = y → f 1 = y → f 2 = y → f x = y := by
172+
grind =>
173+
mbtc
174+
cases #23ad <;> mbtc <;> cases #beb4 <;> mbtc <;> cases #beed
175+
176+
example (f : Int → Int) (x y : Int)
177+
: 0 ≤ x → x ≤ 2 → f 0 = y → f 1 = y → f 2 = y → f x = y := by
178+
grind =>
179+
-- Again, we can use `have` to golf the proof with `mbtc`
180+
have : x ≠ 0
181+
have : x ≠ 1
182+
have : x ≠ 2

0 commit comments

Comments
 (0)