Skip to content

Commit aef4a29

Browse files
authored
feat: Field support in grind ring (#8777)
This PR implements basic `Field` support in the commutative ring module in `grind`. It is just division by numerals for now. Examples: ```lean open Lean Grind example [Field α] [IsCharP α 0] (a b c : α) : a/3 = b → c = a/3 → a/2 + a/2 = b + 2*c := by grind example [Field α] (a b : α) : b = 0 → (a + a) / 0 = b := by grind example [Field α] [IsCharP α 3] (a b : α) : a/3 = b → b = 0 := by grind example [Field α] [IsCharP α 7] (a b c : α) : a/3 = b → c = a/3 → a/2 + a/2 = b + 2*c + 7 := by grind example [Field R] [IsCharP R 0] (x : R) (cos : R → R) : (cos x ^ 2 + (2 * cos x ^ 2 - 1) ^ 2 + (4 * cos x ^ 3 - 3 * cos x) ^ 2 - 1) / 4 = cos x * (cos x ^ 2 - 1 / 2) * (4 * cos x ^ 3 - 3 * cos x) := by grind ```
1 parent 5d50433 commit aef4a29

File tree

6 files changed

+172
-43
lines changed

6 files changed

+172
-43
lines changed

src/Init/Grind/CommRing/Poly.lean

Lines changed: 26 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -11,14 +11,13 @@ import Init.Data.Hashable
1111
import all Init.Data.Ord
1212
import Init.Data.RArray
1313
import Init.Grind.CommRing.Basic
14+
import Init.Grind.CommRing.Field
1415
import Init.Grind.Ordered.Ring
1516

1617
namespace Lean.Grind
17-
namespace CommRing
18-
1918
-- These are no longer global instances, so we need to turn them on here.
2019
attribute [local instance] Semiring.natCast Ring.intCast
21-
20+
namespace CommRing
2221
abbrev Var := Nat
2322

2423
inductive Expr where
@@ -1225,5 +1224,29 @@ theorem not_lt_norm' {α} [CommRing α] [Preorder α] [Ring.IsOrdered α] (ctx :
12251224
rw [sub_eq_add_neg, add_left_comm, ← sub_eq_add_neg, sub_self] at h; simp [add_zero] at h
12261225
contradiction
12271226

1227+
theorem div_int_eq {α} [Field α] [IsCharP α 0] (a : α) (b : Int) : b != 0 → a = denoteInt b * (a / denoteInt b) := by
1228+
simp [Field.div_eq_mul_inv]; intro h
1229+
have : (denoteInt b : α) ≠ 0 := by
1230+
simp [denoteInt_eq]; intro h
1231+
have := IsCharP.intCast_eq_zero_iff (α := α) 0 b; simp [*] at this
1232+
rw [CommRing.mul_comm, Semiring.mul_assoc, CommRing.mul_comm _ (denoteInt b), Field.mul_inv_cancel this, Semiring.mul_one]
1233+
1234+
theorem div_int_eqC {α c} [Field α] [IsCharP α c] (a : α) (b : Int) : b % c != 0 → a = (denoteInt b) * (a / denoteInt b) := by
1235+
simp [Field.div_eq_mul_inv]; intro h
1236+
have : (denoteInt b : α) ≠ 0 := by
1237+
simp [denoteInt_eq]; intro h
1238+
have := IsCharP.intCast_eq_zero_iff (α := α) c b; simp [*] at this
1239+
rw [CommRing.mul_comm, Semiring.mul_assoc, CommRing.mul_comm _ (denoteInt b), Field.mul_inv_cancel this, Semiring.mul_one]
1240+
1241+
theorem div_zero_eq {α} [Field α] (a : α) : a / 0 = 0 := by
1242+
simp [Field.div_eq_mul_inv, Field.inv_zero, Semiring.mul_zero]
1243+
1244+
theorem div_zero_eqC {α c} [Field α] [IsCharP α c] (a : α) (b : Int) : b % c == 0 → (a / denoteInt b) = 0 := by
1245+
simp [Field.div_eq_mul_inv, denoteInt_eq]; intro h
1246+
have : (b : α) = 0 := by
1247+
have := IsCharP.intCast_eq_zero_iff (α := α) c b
1248+
simp [*]
1249+
simp [this, Field.div_eq_mul_inv, Field.inv_zero, Semiring.mul_zero]
12281250
end CommRing
1251+
12291252
end Lean.Grind

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

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ Helper functions for converting reified terms back into their denotations.
1414

1515
variable [Monad M] [MonadGetRing M]
1616

17-
private def denoteNum (k : Int) : M Expr := do
17+
def denoteNum (k : Int) : M Expr := do
1818
let ring ← getRing
1919
let n := mkRawNatLit k.natAbs
2020
let ofNatInst := mkApp3 (mkConst ``Grind.Semiring.ofNat [ring.u]) ring.type ring.semiringInst n

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

Lines changed: 72 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,7 @@ prelude
77
import Lean.Meta.Tactic.Grind.Simp
88
import Lean.Meta.Tactic.Grind.Arith.CommRing.RingId
99
import Lean.Meta.Tactic.Grind.Arith.CommRing.Reify
10+
import Lean.Meta.Tactic.Grind.Arith.CommRing.DenoteExpr
1011

1112
namespace Lean.Meta.Grind.Arith.CommRing
1213

@@ -41,8 +42,79 @@ private def isForbiddenParent (parent? : Option Expr) : Bool :=
4142
else
4243
true
4344

45+
private partial def toInt? (e : Expr) : RingM (Option Int) := do
46+
match_expr e with
47+
| Neg.neg _ i a =>
48+
if isNegInst (← getRing) i then return (- .) <$> (← toInt? a) else return none
49+
| IntCast.intCast _ i a =>
50+
if isIntCastInst (← getRing) i then getIntValue? a else return none
51+
| NatCast.natCast _ i a =>
52+
if isNatCastInst (← getRing) i then
53+
let some v ← getNatValue? a | return none
54+
return some (Int.ofNat v)
55+
else
56+
return none
57+
| OfNat.ofNat _ n _ =>
58+
let some v ← getNatValue? n | return none
59+
return some (Int.ofNat v)
60+
| _ => return none
61+
62+
private def isDivInst (inst : Expr) : RingM Bool := do
63+
let some fn := (← getRing).divFn? | return false
64+
return isSameExpr fn.appArg! inst
65+
66+
/--
67+
Given `e` of the form `@HDiv.hDiv _ _ _ inst a b`,
68+
asserts `a = b * e` if `b` is a numeral.
69+
Otherwise, asserts `b = 0 ∨ a = b * e`
70+
-/
71+
private def processDiv (e inst a b : Expr) : RingM Unit := do
72+
unless (← isDivInst inst) do return ()
73+
if (← getRing).divSet.contains (a, b) then return ()
74+
modifyRing fun s => { s with divSet := s.divSet.insert (a, b) }
75+
if let some k ← toInt? b then
76+
let ring ← getRing
77+
let some fieldInst := ring.fieldInst? | return ()
78+
if k == 0 then
79+
pushNewFact <| mkApp3 (mkConst ``Grind.CommRing.div_zero_eq [ring.u]) ring.type fieldInst a
80+
else if (← hasChar) then
81+
let (charInst, c) ← getCharInst
82+
if c == 0 then
83+
let expected ← mkEq a (mkApp2 ring.mulFn b e)
84+
pushNewFact <| mkExpectedPropHint
85+
(mkApp6 (mkConst ``Grind.CommRing.div_int_eq [ring.u]) ring.type fieldInst charInst a (mkIntLit k) reflBoolTrue)
86+
expected
87+
else if k % c == 0 then
88+
let expected ← mkEq e (← denoteNum 0)
89+
pushNewFact <| mkExpectedPropHint
90+
(mkApp7 (mkConst ``Grind.CommRing.div_zero_eqC [ring.u]) ring.type (mkNatLit c) fieldInst charInst a (mkIntLit k) reflBoolTrue)
91+
expected
92+
else
93+
let expected ← mkEq a (mkApp2 ring.mulFn b e)
94+
pushNewFact <| mkExpectedPropHint
95+
(mkApp7 (mkConst ``Grind.CommRing.div_int_eqC [ring.u]) ring.type (mkNatLit c) fieldInst charInst a (mkIntLit k) reflBoolTrue)
96+
expected
97+
else
98+
-- TODO
99+
return ()
100+
101+
/--
102+
Returns `true` if `e` is a term `a/b` or `a⁻¹`.
103+
-/
104+
private def internalizeDivInv (e : Expr) : GoalM Bool := do
105+
match_expr e with
106+
| HDiv.hDiv α _ _ inst a b =>
107+
let some ringId ← getRingId? α | return true
108+
RingM.run ringId do processDiv e inst a b
109+
return true
110+
| Inv.inv _α _inst _a =>
111+
-- TODO
112+
return true
113+
| _ => return false
114+
44115
def internalize (e : Expr) (parent? : Option Expr) : GoalM Unit := do
45116
if !(← getConfig).ring && !(← getConfig).ringNull then return ()
117+
if (← internalizeDivInv e) then return ()
46118
let some type := getType? e | return ()
47119
if isForbiddenParent parent? then return ()
48120
let some ringId ← getRingId? type | return ()

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

Lines changed: 43 additions & 37 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Leonardo de Moura
55
-/
66
prelude
7+
import Init.Grind.CommRing.Field
78
import Lean.Meta.Tactic.Grind.Simp
89
import Lean.Meta.Tactic.Grind.Arith.CommRing.Util
910

@@ -12,41 +13,38 @@ namespace Lean.Meta.Grind.Arith.CommRing
1213
private def internalizeFn (fn : Expr) : GoalM Expr := do
1314
shareCommon (← canon fn)
1415

15-
private def getAddFn (type : Expr) (u : Level) (semiringInst : Expr) : GoalM Expr := do
16-
let instType := mkApp3 (mkConst ``HAdd [u, u, u]) type type type
17-
let .some inst ← trySynthInstance instType |
18-
throwError "failed to find instance for ring addition{indentExpr instType}"
19-
let inst' := mkApp2 (mkConst ``instHAdd [u]) type <| mkApp2 (mkConst ``Grind.Semiring.toAdd [u]) type semiringInst
20-
unless (← withDefault <| isDefEq inst inst') do
21-
throwError "instance for addition{indentExpr inst}\nis not definitionally equal to the `Grind.Semiring` one{indentExpr inst'}"
22-
internalizeFn <| mkApp4 (mkConst ``HAdd.hAdd [u, u, u]) type type type inst
16+
private def getUnaryFn (type : Expr)(u : Level) (instDeclName : Name) (declName : Name) : GoalM Expr := do
17+
let instType := mkApp (mkConst instDeclName [u]) type
18+
let .some inst ← trySynthInstance instType
19+
| throwError "`grind ring` failed to find instance{indentExpr instType}"
20+
internalizeFn <| mkApp2 (mkConst declName [u]) type inst
2321

24-
private def getMulFn (type : Expr) (u : Level) (semiringInst : Expr) : GoalM Expr := do
25-
let instType := mkApp3 (mkConst ``HMul [u, u, u]) type type type
26-
let .some inst ← trySynthInstance instType |
27-
throwError "failed to find instance for ring multiplication{indentExpr instType}"
28-
let inst' := mkApp2 (mkConst ``instHMul [u]) type <| mkApp2 (mkConst ``Grind.Semiring.toMul [u]) type semiringInst
29-
unless (← withDefault <| isDefEq inst inst') do
30-
throwError "instance for multiplication{indentExpr inst}\nis not definitionally equal to the `Grind.Semiring` one{indentExpr inst'}"
31-
internalizeFn <| mkApp4 (mkConst ``HMul.hMul [u, u, u]) type type type inst
22+
private def getBinHomoFn (type : Expr)(u : Level) (instDeclName : Name) (declName : Name) : GoalM Expr := do
23+
let instType := mkApp3 (mkConst instDeclName [u, u, u]) type type type
24+
let .some inst ← trySynthInstance instType
25+
| throwError "`grind ring` failed to find instance{indentExpr instType}"
26+
internalizeFn <| mkApp4 (mkConst declName [u, u, u]) type type type inst
3227

33-
private def getSubFn (type : Expr) (u : Level) (ringInst : Expr) : GoalM Expr := do
34-
let instType := mkApp3 (mkConst ``HSub [u, u, u]) type type type
35-
let .some inst ← trySynthInstance instType |
36-
throwError "failed to find instance for ring subtraction{indentExpr instType}"
37-
let inst' := mkApp2 (mkConst ``instHSub [u]) type <| mkApp2 (mkConst ``Grind.Ring.toSub [u]) type ringInst
38-
unless (← withDefault <| isDefEq inst inst') do
39-
throwError "instance for subtraction{indentExpr inst}\nis not definitionally equal to the `Grind.Ring` one{indentExpr inst'}"
40-
internalizeFn <| mkApp4 (mkConst ``HSub.hSub [u, u, u]) type type type inst
28+
-- Remark: we removed consistency checks such as the one that ensures `HAdd` instance matches `Semiring.toAdd`
29+
-- That is, we are assuming the type classes were properly setup.
4130

42-
private def getNegFn (type : Expr) (u : Level) (ringInst : Expr) : GoalM Expr := do
43-
let instType := mkApp (mkConst ``Neg [u]) type
44-
let .some inst ← trySynthInstance instType |
45-
throwError "failed to find instance for ring negation{indentExpr instType}"
46-
let inst' := mkApp2 (mkConst ``Grind.Ring.toNeg [u]) type ringInst
47-
unless (← withDefault <| isDefEq inst inst') do
48-
throwError "instance for negation{indentExpr inst}\nis not definitionally equal to the `Grind.Ring` one{indentExpr inst'}"
49-
internalizeFn <| mkApp2 (mkConst ``Neg.neg [u]) type inst
31+
private def getAddFn (type : Expr) (u : Level) : GoalM Expr := do
32+
getBinHomoFn type u ``HAdd ``HAdd.hAdd
33+
34+
private def getMulFn (type : Expr) (u : Level) : GoalM Expr := do
35+
getBinHomoFn type u ``HMul ``HMul.hMul
36+
37+
private def getSubFn (type : Expr) (u : Level) : GoalM Expr := do
38+
getBinHomoFn type u ``HSub ``HSub.hSub
39+
40+
private def getDivFn (type : Expr) (u : Level) : GoalM Expr := do
41+
getBinHomoFn type u ``HDiv ``HDiv.hDiv
42+
43+
private def getNegFn (type : Expr) (u : Level) : GoalM Expr := do
44+
getUnaryFn type u ``Neg ``Neg.neg
45+
46+
private def getInvFn (type : Expr) (u : Level) : GoalM Expr := do
47+
getUnaryFn type u ``Inv ``Inv.inv
5048

5149
private def getPowFn (type : Expr) (u : Level) (semiringInst : Expr) : GoalM Expr := do
5250
let instType := mkApp3 (mkConst ``HPow [u, 0, u]) type Nat.mkType type
@@ -135,15 +133,23 @@ where
135133
let noZeroDivType := mkApp3 (mkConst ``Grind.NoNatZeroDivisors [u]) type zeroInst hmulInst
136134
LOption.toOption <$> trySynthInstance noZeroDivType
137135
trace_goal[grind.ring] "NoNatZeroDivisors available: {noZeroDivInst?.isSome}"
138-
let addFn ← getAddFn type u semiringInst
139-
let mulFn ← getMulFn type u semiringInst
140-
let subFn ← getSubFn type u ringInst
141-
let negFn ← getNegFn type u ringInst
136+
let field := mkApp (mkConst ``Grind.Field [u]) type
137+
let fieldInst? : Option Expr ← LOption.toOption <$> trySynthInstance field
138+
let addFn ← getAddFn type u
139+
let mulFn ← getMulFn type u
140+
let subFn ← getSubFn type u
141+
let negFn ← getNegFn type u
142142
let powFn ← getPowFn type u semiringInst
143143
let intCastFn ← getIntCastFn type u ringInst
144144
let natCastFn ← getNatCastFn type u semiringInst
145+
let (invFn?, divFn?) ← if fieldInst?.isSome then
146+
pure (some (← getInvFn type u), some (← getDivFn type u))
147+
else
148+
pure (none, none)
145149
let id := (← get').rings.size
146-
let ring : Ring := { id, type, u, semiringInst, ringInst, commSemiringInst, commRingInst, charInst?, noZeroDivInst?, addFn, mulFn, subFn, negFn, powFn, intCastFn, natCastFn }
150+
let ring : Ring := {
151+
id, type, u, semiringInst, ringInst, commSemiringInst, commRingInst, charInst?, noZeroDivInst?, fieldInst?,
152+
addFn, mulFn, subFn, negFn, powFn, intCastFn, natCastFn, invFn?, divFn? }
147153
modify' fun s => { s with rings := s.rings.push ring }
148154
return some id
149155

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

Lines changed: 12 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -139,16 +139,22 @@ structure Ring where
139139
/-- `CommRing` instance for `type` -/
140140
commRingInst : Expr
141141
/-- `IsCharP` instance for `type` if available. -/
142-
charInst? : Option (Expr × Nat) := .none
142+
charInst? : Option (Expr × Nat)
143143
/-- `NoNatZeroDivisors` instance for `type` if available. -/
144-
noZeroDivInst? : Option Expr := .none
144+
noZeroDivInst? : Option Expr
145+
/-- `Field` instance for `type` if available. -/
146+
fieldInst? : Option Expr
145147
addFn : Expr
146148
mulFn : Expr
147149
subFn : Expr
148150
negFn : Expr
149151
powFn : Expr
150152
intCastFn : Expr
151153
natCastFn : Expr
154+
/-- Inverse if `fieldInst?` is `some inst` -/
155+
invFn? : Option Expr
156+
/-- Division if `fieldInst?` is `some inst` -/
157+
divFn? : Option Expr
152158
/--
153159
Mapping from variables to their denotations.
154160
Remark each variable can be in only one ring.
@@ -178,6 +184,10 @@ structure Ring where
178184
disequalities and implied equalities.
179185
-/
180186
recheck : Bool := false
187+
/-- Division theorems that have been already asserted. -/
188+
divSet : PHashSet (Expr × Expr) := {}
189+
/-- Inverse theorems that have been already asserted. -/
190+
invSet : PHashSet Expr := {}
181191
deriving Inhabited
182192

183193
/-- State for all `CommRing` types detected by `grind`. -/
Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,18 @@
1+
open Lean Grind
2+
3+
example [Field α] [IsCharP α 0] (a b c : α) : a/3 = b → c = a/3 → a/2 + a/2 = b + 2*c := by
4+
grind
5+
6+
example [Field α] (a b : α) : b = 0 → (a + a) / 0 = b := by
7+
grind
8+
9+
example [Field α] [IsCharP α 3] (a b : α) : a/3 = b → b = 0 := by
10+
grind
11+
12+
example [Field α] [IsCharP α 7] (a b c : α) : a/3 = b → c = a/3 → a/2 + a/2 = b + 2*c + 7 := by
13+
grind
14+
15+
example [Field R] [IsCharP R 0] (x : R) (cos : R → R) :
16+
(cos x ^ 2 + (2 * cos x ^ 2 - 1) ^ 2 + (4 * cos x ^ 3 - 3 * cos x) ^ 2 - 1) / 4 =
17+
cos x * (cos x ^ 2 - 1 / 2) * (4 * cos x ^ 3 - 3 * cos x) := by
18+
grind

0 commit comments

Comments
 (0)