Skip to content
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
4 changes: 4 additions & 0 deletions src/Init/Grind/Tactics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -189,6 +189,10 @@ structure Config where
`CommRing`.
-/
linarith := true
/--
When `true` (default: `true`), uses procedure for handling linear integer arithmetic for `Int` and `Nat`.
-/
cutsat := true
deriving Inhabited, BEq

end Lean.Grind
Expand Down
1 change: 1 addition & 0 deletions src/Lean/Meta/Tactic/Grind/Arith/Cutsat/DvdCnstr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -112,6 +112,7 @@ def propagateNatDvd (e : Expr) : GoalM Unit := do
pushNewFact <| mkApp3 (mkConst ``Nat.emod_pos_of_not_dvd) a b (mkOfEqFalseCore e (← mkEqFalseProof e))

builtin_grind_propagator propagateDvd ↓Dvd.dvd := fun e => do
unless (← getConfig).cutsat do return ()
let_expr Dvd.dvd α _ _ _ ← e | return ()
if α.isConstOf ``Nat then
propagateNatDvd e
Expand Down
3 changes: 3 additions & 0 deletions src/Lean/Meta/Tactic/Grind/Arith/Cutsat/EqCnstr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -254,6 +254,7 @@ private def processNewNatEq (a b : Expr) : GoalM Unit := do

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

@[export lean_process_cutsat_diseq]
def processNewDiseqImpl (a b : Expr) : GoalM Unit := do
unless (← getConfig).cutsat do return ()
match (← foreignTerm? a), (← foreignTermOrLit? b) with
| none, none => processNewIntDiseq a b
| some .nat, some .nat => processNewNatDiseq a b
Expand Down Expand Up @@ -404,6 +406,7 @@ Internalizes an integer (and `Nat`) expression. Here are the different cases tha
back to the congruence closure module. Example: we have `f 5`, `f x`, `x - y = 3`, `y = 2`.
-/
def internalize (e : Expr) (parent? : Option Expr) : GoalM Unit := do
unless (← getConfig).cutsat do return ()
let some (k, type) := getKindAndType? e | return ()
if isForbiddenParent parent? k then return ()
trace[grind.debug.cutsat.internalize] "{e} : {type}"
Expand Down
1 change: 1 addition & 0 deletions src/Lean/Meta/Tactic/Grind/Arith/Cutsat/LeCnstr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -162,6 +162,7 @@ def propagateNatLe (e : Expr) (eqTrue : Bool) : GoalM Unit := do
c.assert

def propagateIfSupportedLe (e : Expr) (eqTrue : Bool) : GoalM Unit := do
unless (← getConfig).cutsat do return ()
let_expr LE.le α _ _ _ := e | return ()
if α.isConstOf ``Nat then
propagateNatLe e eqTrue
Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Meta/Tactic/Grind/Arith/Linear/StructId.lean
Original file line number Diff line number Diff line change
Expand Up @@ -39,7 +39,7 @@ private def ensureDefEq (a b : Expr) : MetaM Unit := do
throwError (← mkExpectedDefEqMsg a b)

def getStructId? (type : Expr) : GoalM (Option Nat) := do
if Cutsat.isSupportedType type then
if (← getConfig).cutsat && Cutsat.isSupportedType type then
-- If `type` is supported by cutsat, let it handle
return none
if let some id? := (← get').typeIdOf.find? { expr := type } then
Expand Down
4 changes: 4 additions & 0 deletions tests/lean/run/grind_cutsat_div_1.lean
Original file line number Diff line number Diff line change
Expand Up @@ -79,3 +79,7 @@ example (x : Int) (_ : 10 ∣ x) (_ : ¬ 5 ∣ x) : False := by

example (x : Nat) (_ : 10 ∣ x) (_ : ¬ 5 ∣ x) : False := by
grind

example (a b : Int) (h₀ : 2 ∣ a + 1) (h₁ : 2 ∣ b + a) (h₂ : 2 ∣ b + 2*a) : False := by
fail_if_success grind -cutsat
sorry
15 changes: 15 additions & 0 deletions tests/lean/run/grind_linarith_1.lean
Original file line number Diff line number Diff line change
Expand Up @@ -148,3 +148,18 @@ set_option trace.grind.linarith.internalize true
example [CommRing α] [LinearOrder α] [Ring.IsOrdered α] (a b c d : α)
: a < b + c + d → c = b → a < b + b + d := by
grind

/--
trace: [grind.linarith.assert] -3 * y + 2 * x + One.one ≤ 0
[grind.linarith.assert] 2 * z + -4 * x + One.one ≤ 0
[grind.linarith.assert] -1 * z + 3 * y + One.one ≤ 0
[grind.linarith.assert] 6 * y + -4 * x + 3 * One.one ≤ 0
[grind.linarith.assert] 15 * One.one ≤ 0
[grind.linarith.assert] Zero.zero < 0
-/
#guard_msgs (trace) in
set_option trace.grind.cutsat.assert true in -- cutsat should **not** process the following constraints
set_option trace.grind.linarith.assert true in -- linarith should take over
set_option trace.grind.linarith.assert.store false in
example (x y z : Int) (h1 : 2 * x < 3 * y) (h2 : -4 * x + 2 * z < 0) : ¬ 12*y - 4* z < 0 := by
grind -cutsat
Loading