Skip to content

Commit 32eedc2

Browse files
authored
feat: grind -cutsat (#8774)
This PR adds an option for disabling the cutsat procedure in `grind`. The linarith module takes over linear integer/nat constraints. Example: ```lean set_option trace.grind.cutsat.assert true in -- cutsat should **not** process the following constraints example (x y z : Int) (h1 : 2 * x < 3 * y) (h2 : -4 * x + 2 * z < 0) : ¬ 12*y - 4* z < 0 := by grind -cutsat -- `linarith` module solves it ```
1 parent 95e532a commit 32eedc2

File tree

7 files changed

+29
-1
lines changed

7 files changed

+29
-1
lines changed

src/Init/Grind/Tactics.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -189,6 +189,10 @@ structure Config where
189189
`CommRing`.
190190
-/
191191
linarith := true
192+
/--
193+
When `true` (default: `true`), uses procedure for handling linear integer arithmetic for `Int` and `Nat`.
194+
-/
195+
cutsat := true
192196
deriving Inhabited, BEq
193197

194198
end Lean.Grind

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

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -112,6 +112,7 @@ def propagateNatDvd (e : Expr) : GoalM Unit := do
112112
pushNewFact <| mkApp3 (mkConst ``Nat.emod_pos_of_not_dvd) a b (mkOfEqFalseCore e (← mkEqFalseProof e))
113113

114114
builtin_grind_propagator propagateDvd ↓Dvd.dvd := fun e => do
115+
unless (← getConfig).cutsat do return ()
115116
let_expr Dvd.dvd α _ _ _ ← e | return ()
116117
if α.isConstOf ``Nat then
117118
propagateNatDvd e

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

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -254,6 +254,7 @@ private def processNewNatEq (a b : Expr) : GoalM Unit := do
254254

255255
@[export lean_process_cutsat_eq]
256256
def processNewEqImpl (a b : Expr) : GoalM Unit := do
257+
unless (← getConfig).cutsat do return ()
257258
match (← foreignTerm? a), (← foreignTerm? b) with
258259
| none, none => processNewIntEq a b
259260
| some .nat, some .nat => processNewNatEq a b
@@ -281,6 +282,7 @@ private def processNewNatDiseq (a b : Expr) : GoalM Unit := do
281282

282283
@[export lean_process_cutsat_diseq]
283284
def processNewDiseqImpl (a b : Expr) : GoalM Unit := do
285+
unless (← getConfig).cutsat do return ()
284286
match (← foreignTerm? a), (← foreignTermOrLit? b) with
285287
| none, none => processNewIntDiseq a b
286288
| some .nat, some .nat => processNewNatDiseq a b
@@ -404,6 +406,7 @@ Internalizes an integer (and `Nat`) expression. Here are the different cases tha
404406
back to the congruence closure module. Example: we have `f 5`, `f x`, `x - y = 3`, `y = 2`.
405407
-/
406408
def internalize (e : Expr) (parent? : Option Expr) : GoalM Unit := do
409+
unless (← getConfig).cutsat do return ()
407410
let some (k, type) := getKindAndType? e | return ()
408411
if isForbiddenParent parent? k then return ()
409412
trace[grind.debug.cutsat.internalize] "{e} : {type}"

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

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -162,6 +162,7 @@ def propagateNatLe (e : Expr) (eqTrue : Bool) : GoalM Unit := do
162162
c.assert
163163

164164
def propagateIfSupportedLe (e : Expr) (eqTrue : Bool) : GoalM Unit := do
165+
unless (← getConfig).cutsat do return ()
165166
let_expr LE.le α _ _ _ := e | return ()
166167
if α.isConstOf ``Nat then
167168
propagateNatLe e eqTrue

src/Lean/Meta/Tactic/Grind/Arith/Linear/StructId.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -39,7 +39,7 @@ private def ensureDefEq (a b : Expr) : MetaM Unit := do
3939
throwError (← mkExpectedDefEqMsg a b)
4040

4141
def getStructId? (type : Expr) : GoalM (Option Nat) := do
42-
if Cutsat.isSupportedType type then
42+
if (← getConfig).cutsat && Cutsat.isSupportedType type then
4343
-- If `type` is supported by cutsat, let it handle
4444
return none
4545
if let some id? := (← get').typeIdOf.find? { expr := type } then

tests/lean/run/grind_cutsat_div_1.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -79,3 +79,7 @@ example (x : Int) (_ : 10 ∣ x) (_ : ¬ 5 ∣ x) : False := by
7979

8080
example (x : Nat) (_ : 10 ∣ x) (_ : ¬ 5 ∣ x) : False := by
8181
grind
82+
83+
example (a b : Int) (h₀ : 2 ∣ a + 1) (h₁ : 2 ∣ b + a) (h₂ : 2 ∣ b + 2*a) : False := by
84+
fail_if_success grind -cutsat
85+
sorry

tests/lean/run/grind_linarith_1.lean

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -148,3 +148,18 @@ set_option trace.grind.linarith.internalize true
148148
example [CommRing α] [LinearOrder α] [Ring.IsOrdered α] (a b c d : α)
149149
: a < b + c + d → c = b → a < b + b + d := by
150150
grind
151+
152+
/--
153+
trace: [grind.linarith.assert] -3 * y + 2 * x + One.one ≤ 0
154+
[grind.linarith.assert] 2 * z + -4 * x + One.one ≤ 0
155+
[grind.linarith.assert] -1 * z + 3 * y + One.one ≤ 0
156+
[grind.linarith.assert] 6 * y + -4 * x + 3 * One.one ≤ 0
157+
[grind.linarith.assert] 15 * One.one ≤ 0
158+
[grind.linarith.assert] Zero.zero < 0
159+
-/
160+
#guard_msgs (trace) in
161+
set_option trace.grind.cutsat.assert true in -- cutsat should **not** process the following constraints
162+
set_option trace.grind.linarith.assert true in -- linarith should take over
163+
set_option trace.grind.linarith.assert.store false in
164+
example (x y z : Int) (h1 : 2 * x < 3 * y) (h2 : -4 * x + 2 * z < 0) : ¬ 12*y - 4* z < 0 := by
165+
grind -cutsat

0 commit comments

Comments
 (0)