Skip to content

Commit 41c41e4

Browse files
authored
feat: One.one support in linarith (#8694)
This PR implements special support for `One.one` in linarith when the structure is a ordered ring. It also fixes bugs during initialization.
1 parent f61a412 commit 41c41e4

File tree

7 files changed

+52
-4
lines changed

7 files changed

+52
-4
lines changed

src/Init/Grind/Ordered/Linarith.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Leonardo de Moura
66
module
77
prelude
88
import Init.Grind.Ordered.Module
9+
import Init.Grind.Ordered.Ring
910
import all Init.Data.Ord
1011
import all Init.Data.AC
1112
import Init.Data.RArray
@@ -450,6 +451,14 @@ theorem lt_unsat {α} [IntModule α] [Preorder α] (ctx : Context α) : (Poly.ni
450451
have := Preorder.lt_iff_le_not_le.mp h
451452
simp at this
452453

454+
def zero_lt_one_cert (p : Poly) : Bool :=
455+
p == .add (-1) 0 .nil
456+
457+
theorem zero_lt_one {α} [Ring α] [Preorder α] [Ring.IsOrdered α] (ctx : Context α) (p : Poly)
458+
: zero_lt_one_cert p → (0 : Var).denote ctx = One.one → p.denote' ctx < 0 := by
459+
simp [zero_lt_one_cert]; intro _ h; subst p; simp [Poly.denote, h, One.one, neg_hmul]
460+
rw [neg_lt_iff, neg_zero]; apply Ring.IsOrdered.zero_lt_one
461+
453462
/-!
454463
Coefficient normalization
455464
-/

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

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -102,7 +102,7 @@ def propagateIneq (e : Expr) (eqTrue : Bool) : GoalM Unit := do
102102
return ()
103103
let lhs := e.getArg! 2 numArgs
104104
let rhs := e.getArg! 3 numArgs
105-
if (← isCommRing) then
105+
if (← isOrderedCommRing) then
106106
propagateCommRingIneq e lhs rhs strict eqTrue
107107
-- TODO: non-commutative ring normalizer
108108
else

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

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -196,6 +196,10 @@ partial def IneqCnstr.toExprProof (c' : IneqCnstr) : ProofM Expr := caching c' d
196196
| false, false => (``Grind.Linarith.le_le_combine, c₁, c₂)
197197
return mkApp6 (← mkIntModPreOrdThmPrefix declName) (← mkPolyDecl c₁.p) (← mkPolyDecl c₂.p) (← mkPolyDecl c'.p) reflBoolTrue
198198
(← c₁.toExprProof) (← c₂.toExprProof)
199+
| .oneGtZero =>
200+
let s ← getStruct
201+
let h := mkApp5 (mkConst ``Grind.Linarith.zero_lt_one [s.u]) s.type (← getRingInst) s.preorderInst (← getRingIsOrdInst) (← getContext)
202+
return mkApp3 h (← mkPolyDecl c'.p) reflBoolTrue (← mkEqRefl (← getOne))
199203
| _ => throwError "NIY"
200204

201205
partial def DiseqCnstr.toExprProof (c' : DiseqCnstr) : ProofM Expr := caching c' do

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

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -96,6 +96,6 @@ partial def reify? (e : Expr) (skipVar : Bool) : LinearM (Option LinExpr) := do
9696
if skipVar then
9797
return none
9898
else
99-
return some (← asVar e)
99+
return some (← toVar e)
100100

101101
end Lean.Meta.Grind.Arith.Linear

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

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -139,7 +139,10 @@ where
139139
let getRingIsOrdInst? : GoalM (Option Expr) := do
140140
let some ringInst := ringInst? | return none
141141
let isOrdType := mkApp3 (mkConst ``Grind.Ring.IsOrdered [u]) type ringInst preorderInst
142-
return LOption.toOption (← trySynthInstance isOrdType)
142+
let .some inst ← trySynthInstance isOrdType
143+
| reportIssue! "type is an ordered `IntModule` and a `Ring`, but is not an ordered ring, failed to synthesize{indentExpr isOrdType}"
144+
return none
145+
return some inst
143146
let ringIsOrdInst? ← getRingIsOrdInst?
144147
let getNoNatZeroDivInst? : GoalM (Option Expr) := do
145148
let hmulNat := mkApp3 (mkConst ``HMul [0, u, u]) Nat.mkType type type

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

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -86,6 +86,9 @@ def withRingM (x : RingM α) : LinearM α := do
8686
def isCommRing : LinearM Bool :=
8787
return (← getStruct).ringId?.isSome
8888

89+
def isOrderedCommRing : LinearM Bool := do
90+
return (← isCommRing) && (← getStruct).ringIsOrdInst?.isSome
91+
8992
def isLinearOrder : LinearM Bool :=
9093
return (← getStruct).linearInst?.isSome
9194

@@ -109,6 +112,11 @@ def getLinearOrderInst : LinearM Expr := do
109112
| throwError "`grind linarith` internal error, structure is not a linear order"
110113
return inst
111114

115+
def getRingInst : LinearM Expr := do
116+
let some inst := (← getStruct).ringInst?
117+
| throwError "`grind linarith` internal error, structure is not a ring"
118+
return inst
119+
112120
def getCommRingInst : LinearM Expr := do
113121
let some inst := (← getStruct).commRingInst?
114122
| throwError "`grind linarith` internal error, structure is not a commutative ring"

tests/lean/run/grind_linarith_1.lean

Lines changed: 25 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -61,5 +61,29 @@ example [CommRing α] [LinearOrder α] [Ring.IsOrdered α] (a b c d : α)
6161
grind
6262

6363
example [CommRing α] [Preorder α] [Ring.IsOrdered α] (a b c : α)
64-
: a < b → b < c → c < a → False := by
64+
: a < b → 2*b < c → c < 2*a → False := by
65+
grind
66+
67+
-- Test misconfigured instances
68+
/--
69+
trace: [grind.issues] type is an ordered `IntModule` and a `Ring`, but is not an ordered ring, failed to synthesize
70+
Ring.IsOrdered α
71+
-/
72+
#guard_msgs (drop error, trace) in
73+
set_option trace.grind.issues true in
74+
example [CommRing α] [Preorder α] [IntModule.IsOrdered α] (a b c : α)
75+
: a < b → b + b < c → c < a → False := by
76+
grind
77+
78+
example [CommRing α] [Preorder α] [Ring.IsOrdered α] (a b : α)
79+
: a < 2 → b < a → b > 5 → False := by
6580
grind
81+
82+
example [CommRing α] [Preorder α] [Ring.IsOrdered α] (a b : α)
83+
: a < One.one + 4 → b < a → b ≥ 5 → False := by
84+
grind
85+
86+
example [CommRing α] [Preorder α] [Ring.IsOrdered α] (a b : α)
87+
: a < One.one + 5 → b < a → b ≥ 5 → False := by
88+
fail_if_success grind
89+
sorry

0 commit comments

Comments
 (0)