Skip to content

Commit 97d63db

Browse files
authored
fix: mbtc for nonlinear terms in grind cutsat (#10965)
This PR ensures that model-based theory combination in `grind cutsat` considers nonlinear terms. Nonlinear multiplications such as `x * y` are treated as uninterpreted symbols in `cutsat`. Closes #10885
1 parent f8b0bee commit 97d63db

File tree

2 files changed

+39
-1
lines changed

2 files changed

+39
-1
lines changed

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

Lines changed: 28 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -46,8 +46,35 @@ private def getAssignmentExt? (e : Expr) : GoalM (Option Rat) := do
4646
private def hasTheoryVar (e : Expr) : GoalM Bool := do
4747
cutsatExt.hasTermAtRoot e
4848

49+
/-
50+
**Note**: cutsat is a procedure for linear integer arithmetic. Thus, morally a
51+
nonlinear multiplication, division, and modulo are **not** interpreted by cutsat.
52+
Thus, we enable model-theory combination for them. This is necessary for examples
53+
such as:
54+
```
55+
example {a b : Nat} (ha : 1 ≤ a) : (a - 1 + 1) * b = a * b := by grind
56+
```
57+
Note that we currently use a restrictive/cheaper version of mbtc. We only case-split
58+
on `a = b`, if they have the same assignment **and** occur as the `i`-th argument of
59+
the same function symbol. The latter reduces the number of case-splits we have to
60+
perform, but misses the following variant of the problem above.
61+
```
62+
example {a b : Nat} (ha : 1 ≤ a) : (a - 1 + 1) * b = b * a := by grind
63+
```
64+
If this becomes an issue in practice, we can add a flag for enabling the more
65+
expensive version of `mbtc`.
66+
-/
67+
68+
private def isNonlinearTerm (e : Expr) : GoalM Bool :=
69+
match_expr e with
70+
| HMul.hMul _ _ _ _ a b => return (← getIntValue? a <|> getIntValue? b).isNone
71+
| HDiv.hDiv _ _ _ _ _ b => return (← getIntValue? b).isNone
72+
| HMod.hMod _ _ _ _ _ b => return (← getIntValue? b).isNone
73+
| _ => return false
74+
4975
private def isInterpreted (e : Expr) : GoalM Bool := do
50-
if isInterpretedTerm e then return true
76+
if isInterpretedTerm e then
77+
return !(← isNonlinearTerm e)
5178
let f := e.getAppFn
5279
/-
5380
**Note**: `grind` normalizes terms, but some of them cannot be rewritten by `simp` because

tests/lean/run/grind_10885.lean

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,11 @@
1+
example {a b : Nat} (ha : 1 ≤ a) : (a - 1 + 1) * b = a * b := by grind
2+
3+
/--
4+
info: Try this:
5+
[apply]
6+
mbtc
7+
cases #9501
8+
-/
9+
#guard_msgs in
10+
example {a b : Nat} (ha : 1 ≤ a) : (a - 1 + 1) * b = a * b := by
11+
grind => finish? -- mbtc was applied consider nonlinear `*`

0 commit comments

Comments
 (0)