Skip to content

Commit b3753ba

Browse files
authored
feat: grind propagators for Nat operators (#11522)
This PR implements `grind` propagators for `Nat` operators that have a simproc associated with them, but do not have any theory solver support. Examples: ```lean example (a b : Nat) : a = 3 → b = 6 → a &&& b = 2 := by grind example (a b : Nat) : a = 3 → b = 6 → a ||| b = 7 := by grind example (a b : Nat) : a = 3 → b = 6 → a ^^^ b = 5 := by grind example (a b : Nat) : a = 3 → b = 6 → a <<< b = 192 := by grind example (a b : Nat) : a = 1135 → b = 6 → a >>> b = 17 := by grind ``` Closes #11498
1 parent 8afaa1b commit b3753ba

File tree

4 files changed

+90
-0
lines changed

4 files changed

+90
-0
lines changed

src/Init/Grind/Lemmas.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -185,4 +185,12 @@ theorem decide_eq_false {p : Prop} {_ : Decidable p} : p = False → decide p =
185185
theorem of_lookahead (p : Prop) (h : (¬ p) → False) : p = True := by
186186
simp at h; simp [h]
187187

188+
/-! Nat propagators -/
189+
190+
theorem Nat.and_congr {a b : Nat} {k₁ k₂ k : Nat} (h₁ : a = k₁) (h₂ : b = k₂) : k == k₁ &&& k₂ → a &&& b = k := by simp_all
191+
theorem Nat.xor_congr {a b : Nat} {k₁ k₂ k : Nat} (h₁ : a = k₁) (h₂ : b = k₂) : k == k₁ ^^^ k₂ → a ^^^ b = k := by simp_all
192+
theorem Nat.or_congr {a b : Nat} {k₁ k₂ k : Nat} (h₁ : a = k₁) (h₂ : b = k₂) : k == k₁ ||| k₂ → a ||| b = k := by simp_all
193+
theorem Nat.shiftLeft_congr {a b : Nat} {k₁ k₂ k : Nat} (h₁ : a = k₁) (h₂ : b = k₂) : k == k₁ <<< k₂ → a <<< b = k := by simp_all
194+
theorem Nat.shiftRight_congr {a b : Nat} {k₁ k₂ k : Nat} (h₁ : a = k₁) (h₂ : b = k₂) : k == k₁ >>> k₂ → a >>> b = k := by simp_all
195+
188196
end Lean.Grind

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

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -12,3 +12,4 @@ public import Lean.Meta.Tactic.Grind.Arith.CommRing
1212
public import Lean.Meta.Tactic.Grind.Arith.Linear
1313
public import Lean.Meta.Tactic.Grind.Arith.Cutsat
1414
public import Lean.Meta.Tactic.Grind.Arith.Simproc
15+
public import Lean.Meta.Tactic.Grind.Arith.Propagate
Lines changed: 65 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,65 @@
1+
/-
2+
Copyright (c) 2025 Amazon.com, Inc. or its affiliates. All Rights Reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Leonardo de Moura
5+
-/
6+
module
7+
prelude
8+
public import Lean.Meta.Tactic.Grind.Types
9+
import Init.Grind
10+
import Lean.Meta.Tactic.Grind.PropagatorAttr
11+
public section
12+
namespace Lean.Meta.Grind
13+
14+
/-!
15+
This file defines propagators for `Nat` operators that have simprocs associated with them, but do not
16+
have support in satellite solvers. The goal is to workaround a nasty interaction between
17+
E-matching and normalization. Consider the following example:
18+
```
19+
@[grind] def mask := 15
20+
21+
@[grind =] theorem foo (x : Nat) : x &&& 15 = x % 16 := by sorry
22+
23+
example (x : Nat) : 3 &&& mask = 3 := by
24+
grind only [foo, mask]
25+
```
26+
This is a very simple version of issue #11498.
27+
The e-graph contains `3 &&& mask`. E-matching finds that `3 &&& mask` matches pattern `x &&& 15`
28+
modulo the equality `mask = 15`, binding `x := 3`.
29+
Thus is produces the theorem instance `3 &&& 15 = 3 % 16`, which simplifies to `True` using the
30+
builtin simprocs for `&&&` and `%`. So, nothing is learned.
31+
**Tension**: The invariant "all terms in e-graph are normalized" conflicts with congruence needing
32+
the intermediate term `3 &&& 15` to make the connection.
33+
This is yet another example that shows how tricky is to combine e-matching with a normalizer.
34+
It is yet another reason for not letting users to extend the normalizer.
35+
If we do, we should be able to mark which ones should be used not normalize theorem instances.
36+
The following propagator ensure that `3 &&& mask` is merged with the equivalence class `3` if
37+
`mask = 15`.
38+
-/
39+
40+
def propagateNatBinOp (declName : Name) (congrThmName : Name) (op : Nat → Nat → Nat) (e : Expr) : GoalM Unit := do
41+
let arity := 6
42+
unless e.isAppOfArity declName arity do return ()
43+
unless e.getArg! 0 |>.isConstOf ``Nat do return ()
44+
unless e.getArg! 1 |>.isConstOf ``Nat do return ()
45+
let a := e.getArg! (arity - 2) arity
46+
let aRoot ← getRoot a
47+
let some k₁ ← getNatValue? aRoot | return ()
48+
let b := e.getArg! (arity - 1) arity
49+
let bRoot ← getRoot b
50+
let some k₂ ← getNatValue? bRoot | return ()
51+
let k := op k₁ k₂
52+
let r ← shareCommon (mkNatLit k)
53+
internalize r 0
54+
let h := mkApp8 (mkConst congrThmName) a b aRoot bRoot r (← mkEqProof a aRoot) (← mkEqProof b bRoot) eagerReflBoolTrue
55+
pushEq e r h
56+
57+
builtin_grind_propagator propagateNatAnd ↑HAnd.hAnd := propagateNatBinOp ``HAnd.hAnd ``Grind.Nat.and_congr (· &&& ·)
58+
builtin_grind_propagator propagateNatOr ↑HOr.hOr := propagateNatBinOp ``HOr.hOr ``Grind.Nat.or_congr (· ||| ·)
59+
builtin_grind_propagator propagateNatXOr ↑HXor.hXor := propagateNatBinOp ``HXor.hXor ``Grind.Nat.xor_congr (· ^^^ ·)
60+
builtin_grind_propagator propagateNatShiftLeft ↑HShiftLeft.hShiftLeft :=
61+
propagateNatBinOp ``HShiftLeft.hShiftLeft ``Grind.Nat.shiftLeft_congr (· <<< ·)
62+
builtin_grind_propagator propagateNatShiftRight ↑HShiftRight.hShiftRight :=
63+
propagateNatBinOp ``HShiftRight.hShiftRight ``Grind.Nat.shiftRight_congr (· >>> ·)
64+
65+
end Lean.Meta.Grind

tests/lean/run/grind_11498.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,16 @@
1+
@[grind]
2+
def RMASK : Nat := 65535
3+
4+
@[grind =] theorem and_65535_eq_mod (x : Nat) : x &&& 65535 = x % 65536 := by sorry
5+
6+
example : 3327 = 3327 &&& RMASK := by
7+
grind
8+
9+
example : 3327 = 3327 &&& RMASK := by
10+
grind only [RMASK] -- and_... is not needed
11+
12+
example (a b : Nat) : a = 3 → b = 6 → a &&& b = 2 := by grind
13+
example (a b : Nat) : a = 3 → b = 6 → a ||| b = 7 := by grind
14+
example (a b : Nat) : a = 3 → b = 6 → a ^^^ b = 5 := by grind
15+
example (a b : Nat) : a = 3 → b = 6 → a <<< b = 192 := by grind
16+
example (a b : Nat) : a = 1135 → b = 6 → a >>> b = 17 := by grind

0 commit comments

Comments
 (0)