@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44Authors: Leonardo de Moura
55-/
66prelude
7+ import Init.Grind.CommRing.Field
78import Lean.Meta.Tactic.Grind.Simp
89import Lean.Meta.Tactic.Grind.Arith.CommRing.Util
910
@@ -12,41 +13,38 @@ namespace Lean.Meta.Grind.Arith.CommRing
1213private 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}\n is 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}\n is 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}\n is 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}\n is 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
5149private 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
0 commit comments