Support for the linear arithmetic module for IntModule
in grind
@[reducible, inline]
Equations
Instances For
Equations
Equations
Equations
- Lean.Grind.Linarith.instReprExpr = { reprPrec := Lean.Grind.Linarith.reprExpr✝ }
@[reducible, inline]
Equations
Instances For
Equations
- Lean.Grind.Linarith.Var.denote ctx v = Lean.RArray.get ctx v
Instances For
Equations
- Lean.Grind.Linarith.Expr.denote ctx Lean.Grind.Linarith.Expr.zero = 0
- Lean.Grind.Linarith.Expr.denote ctx (Lean.Grind.Linarith.Expr.var v) = Lean.Grind.Linarith.Var.denote ctx v
- Lean.Grind.Linarith.Expr.denote ctx (a.add b) = Lean.Grind.Linarith.Expr.denote ctx a + Lean.Grind.Linarith.Expr.denote ctx b
- Lean.Grind.Linarith.Expr.denote ctx (a.sub b) = Lean.Grind.Linarith.Expr.denote ctx a - Lean.Grind.Linarith.Expr.denote ctx b
- Lean.Grind.Linarith.Expr.denote ctx (Lean.Grind.Linarith.Expr.natMul k a) = k * Lean.Grind.Linarith.Expr.denote ctx a
- Lean.Grind.Linarith.Expr.denote ctx (Lean.Grind.Linarith.Expr.intMul k a) = k * Lean.Grind.Linarith.Expr.denote ctx a
- Lean.Grind.Linarith.Expr.denote ctx a.neg = -Lean.Grind.Linarith.Expr.denote ctx a
Instances For
Equations
Equations
- Lean.Grind.Linarith.instReprPoly = { reprPrec := Lean.Grind.Linarith.reprPoly✝ }
Equations
- Lean.Grind.Linarith.Poly.denote ctx Lean.Grind.Linarith.Poly.nil = 0
- Lean.Grind.Linarith.Poly.denote ctx (Lean.Grind.Linarith.Poly.add a a_1 a_2) = a * Lean.Grind.Linarith.Var.denote ctx a_1 + Lean.Grind.Linarith.Poly.denote ctx a_2
Instances For
Similar to Poly.denote
, but produces a denotation better for normalization.
Equations
- Lean.Grind.Linarith.Poly.denote' ctx Lean.Grind.Linarith.Poly.nil = 0
- Lean.Grind.Linarith.Poly.denote' ctx (Lean.Grind.Linarith.Poly.add 1 v p_2) = Lean.Grind.Linarith.Poly.denote'.go ctx (Lean.Grind.Linarith.Var.denote ctx v) p_2
- Lean.Grind.Linarith.Poly.denote' ctx (Lean.Grind.Linarith.Poly.add k v p_2) = Lean.Grind.Linarith.Poly.denote'.go ctx (k * Lean.Grind.Linarith.Var.denote ctx v) p_2
Instances For
def
Lean.Grind.Linarith.Poly.denote'.go
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(r : α)
(p : Poly)
:
α
Equations
- Lean.Grind.Linarith.Poly.denote'.go ctx r Lean.Grind.Linarith.Poly.nil = r
- Lean.Grind.Linarith.Poly.denote'.go ctx r (Lean.Grind.Linarith.Poly.add 1 v p_2) = Lean.Grind.Linarith.Poly.denote'.go ctx (r + Lean.Grind.Linarith.Var.denote ctx v) p_2
- Lean.Grind.Linarith.Poly.denote'.go ctx r (Lean.Grind.Linarith.Poly.add k v p_2) = Lean.Grind.Linarith.Poly.denote'.go ctx (r + k * Lean.Grind.Linarith.Var.denote ctx v) p_2
Instances For
theorem
Lean.Grind.Linarith.instAssociativeHAdd
{α : Type u_1}
[IntModule α]
:
Std.Associative fun (x1 x2 : α) => x1 + x2
theorem
Lean.Grind.Linarith.instCommutativeHAdd
{α : Type u_1}
[IntModule α]
:
Std.Commutative fun (x1 x2 : α) => x1 + x2
theorem
Lean.Grind.Linarith.Poly.denote'_go_eq_denote
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(p : Poly)
(r : α)
:
Equations
- (Lean.Grind.Linarith.Poly.add a y p_2).coeff x = bif x == y then a else p_2.coeff x
- Lean.Grind.Linarith.Poly.nil.coeff x = 0
Instances For
Equations
- One or more equations did not get rendered due to their size.
- Lean.Grind.Linarith.Poly.insert k v Lean.Grind.Linarith.Poly.nil = Lean.Grind.Linarith.Poly.add k v Lean.Grind.Linarith.Poly.nil
Instances For
Normalizes the given polynomial by fusing monomial and constants.
Equations
Instances For
Equations
- Lean.Grind.Linarith.Poly.nil.append p₂ = p₂
- (Lean.Grind.Linarith.Poly.add a a_1 a_2).append p₂ = Lean.Grind.Linarith.Poly.add a a_1 (a_2.append p₂)
Instances For
Equations
- One or more equations did not get rendered due to their size.
- Lean.Grind.Linarith.Poly.combine' 0 p₁ p₂ = p₁.append p₂
- Lean.Grind.Linarith.Poly.combine' fuel_2.succ Lean.Grind.Linarith.Poly.nil p₂ = p₂
- Lean.Grind.Linarith.Poly.combine' fuel_2.succ p₁ Lean.Grind.Linarith.Poly.nil = p₁
Instances For
Equations
- p₁.combine p₂ = Lean.Grind.Linarith.Poly.combine' 100000000 p₁ p₂
Instances For
Converts the given expression into a polynomial.
Equations
Instances For
Equations
- Lean.Grind.Linarith.Expr.toPoly'.go coeff Lean.Grind.Linarith.Expr.zero = id
- Lean.Grind.Linarith.Expr.toPoly'.go coeff (Lean.Grind.Linarith.Expr.var v) = fun (x : Lean.Grind.Linarith.Poly) => Lean.Grind.Linarith.Poly.add coeff v x
- Lean.Grind.Linarith.Expr.toPoly'.go coeff (a.add b) = Lean.Grind.Linarith.Expr.toPoly'.go coeff a ∘ Lean.Grind.Linarith.Expr.toPoly'.go coeff b
- Lean.Grind.Linarith.Expr.toPoly'.go coeff (a.sub b) = Lean.Grind.Linarith.Expr.toPoly'.go coeff a ∘ Lean.Grind.Linarith.Expr.toPoly'.go (-coeff) b
- Lean.Grind.Linarith.Expr.toPoly'.go coeff (Lean.Grind.Linarith.Expr.natMul k a) = bif k == 0 then id else Lean.Grind.Linarith.Expr.toPoly'.go (coeff.mul ↑k) a
- Lean.Grind.Linarith.Expr.toPoly'.go coeff (Lean.Grind.Linarith.Expr.intMul k a) = bif k == 0 then id else Lean.Grind.Linarith.Expr.toPoly'.go (coeff.mul k) a
- Lean.Grind.Linarith.Expr.toPoly'.go coeff a.neg = Lean.Grind.Linarith.Expr.toPoly'.go (-coeff) a
Instances For
p.mul k
multiplies all coefficients and constant of the polynomial p
by k
.
Equations
- Lean.Grind.Linarith.Poly.nil.mul' k = Lean.Grind.Linarith.Poly.nil
- (Lean.Grind.Linarith.Poly.add a a_1 a_2).mul' k = Lean.Grind.Linarith.Poly.add (k * a) a_1 (a_2.mul' k)
Instances For
theorem
Lean.Grind.Linarith.Expr.denote_norm
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(e : Expr)
:
Equations
- (Lean.Grind.Linarith.Poly.add a v p_2).leadCoeff = a
- p.leadCoeff = 1
Instances For
Helper theorems for conflict resolution during model construction.
theorem
Lean.Grind.Linarith.le_le_combine
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(p₁ p₂ p₃ : Poly)
:
le_le_combine_cert p₁ p₂ p₃ = true → Poly.denote' ctx p₁ ≤ 0 → Poly.denote' ctx p₂ ≤ 0 → Poly.denote' ctx p₃ ≤ 0
theorem
Lean.Grind.Linarith.le_lt_combine
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(p₁ p₂ p₃ : Poly)
:
le_lt_combine_cert p₁ p₂ p₃ = true → Poly.denote' ctx p₁ ≤ 0 → Poly.denote' ctx p₂ < 0 → Poly.denote' ctx p₃ < 0
Equations
- One or more equations did not get rendered due to their size.
Instances For
theorem
Lean.Grind.Linarith.lt_lt_combine
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(p₁ p₂ p₃ : Poly)
:
lt_lt_combine_cert p₁ p₂ p₃ = true → Poly.denote' ctx p₁ < 0 → Poly.denote' ctx p₂ < 0 → Poly.denote' ctx p₃ < 0
Equations
- Lean.Grind.Linarith.diseq_split_cert p₁ p₂ = (p₂ == p₁.mul (-1))
Instances For
theorem
Lean.Grind.Linarith.diseq_split
{α : Type u_1}
[IntModule α]
[LinearOrder α]
[OrderedAdd α]
(ctx : Context α)
(p₁ p₂ : Poly)
:
diseq_split_cert p₁ p₂ = true → Poly.denote' ctx p₁ ≠ 0 → Poly.denote' ctx p₁ < 0 ∨ Poly.denote' ctx p₂ < 0
theorem
Lean.Grind.Linarith.diseq_split_resolve
{α : Type u_1}
[IntModule α]
[LinearOrder α]
[OrderedAdd α]
(ctx : Context α)
(p₁ p₂ : Poly)
:
diseq_split_cert p₁ p₂ = true → Poly.denote' ctx p₁ ≠ 0 → ¬Poly.denote' ctx p₁ < 0 → Poly.denote' ctx p₂ < 0
Helper theorems for internalizing facts into the linear arithmetic procedure
Equations
- Lean.Grind.Linarith.norm_cert lhs rhs p = (p == (lhs.sub rhs).norm)
Instances For
theorem
Lean.Grind.Linarith.eq_norm
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert lhs rhs p = true → Expr.denote ctx lhs = Expr.denote ctx rhs → Poly.denote' ctx p = 0
theorem
Lean.Grind.Linarith.le_of_eq
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert lhs rhs p = true → Expr.denote ctx lhs = Expr.denote ctx rhs → Poly.denote' ctx p ≤ 0
theorem
Lean.Grind.Linarith.diseq_norm
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert lhs rhs p = true → Expr.denote ctx lhs ≠ Expr.denote ctx rhs → Poly.denote' ctx p ≠ 0
theorem
Lean.Grind.Linarith.le_norm
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert lhs rhs p = true → Expr.denote ctx lhs ≤ Expr.denote ctx rhs → Poly.denote' ctx p ≤ 0
theorem
Lean.Grind.Linarith.lt_norm
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert lhs rhs p = true → Expr.denote ctx lhs < Expr.denote ctx rhs → Poly.denote' ctx p < 0
theorem
Lean.Grind.Linarith.not_le_norm
{α : Type u_1}
[IntModule α]
[LinearOrder α]
[OrderedAdd α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert rhs lhs p = true → ¬Expr.denote ctx lhs ≤ Expr.denote ctx rhs → Poly.denote' ctx p < 0
theorem
Lean.Grind.Linarith.not_lt_norm
{α : Type u_1}
[IntModule α]
[LinearOrder α]
[OrderedAdd α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert rhs lhs p = true → ¬Expr.denote ctx lhs < Expr.denote ctx rhs → Poly.denote' ctx p ≤ 0
theorem
Lean.Grind.Linarith.not_le_norm'
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert lhs rhs p = true → ¬Expr.denote ctx lhs ≤ Expr.denote ctx rhs → ¬Poly.denote' ctx p ≤ 0
theorem
Lean.Grind.Linarith.not_lt_norm'
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(lhs rhs : Expr)
(p : Poly)
:
norm_cert lhs rhs p = true → ¬Expr.denote ctx lhs < Expr.denote ctx rhs → ¬Poly.denote' ctx p < 0
Equality detection
Equations
- Lean.Grind.Linarith.eq_of_le_ge_cert p₁ p₂ = (p₂ == p₁.mul (-1))
Instances For
theorem
Lean.Grind.Linarith.eq_of_le_ge
{α : Type u_1}
[IntModule α]
[PartialOrder α]
[OrderedAdd α]
(ctx : Context α)
(p₁ p₂ : Poly)
:
eq_of_le_ge_cert p₁ p₂ = true → Poly.denote' ctx p₁ ≤ 0 → Poly.denote' ctx p₂ ≤ 0 → Poly.denote' ctx p₁ = 0
Helper theorems for closing the goal
theorem
Lean.Grind.Linarith.diseq_unsat
{α : Type u_1}
[IntModule α]
(ctx : Context α)
:
Poly.denote ctx Poly.nil ≠ 0 → False
theorem
Lean.Grind.Linarith.lt_unsat
{α : Type u_1}
[IntModule α]
[Preorder α]
(ctx : Context α)
:
Poly.denote ctx Poly.nil < 0 → False
Equations
Instances For
theorem
Lean.Grind.Linarith.zero_lt_one
{α : Type u_1}
[Ring α]
[Preorder α]
[OrderedRing α]
(ctx : Context α)
(p : Poly)
:
zero_lt_one_cert p = true → Var.denote ctx 0 = One.one → Poly.denote' ctx p < 0
Equations
Instances For
theorem
Lean.Grind.Linarith.zero_ne_one_of_ord_ring
{α : Type u_1}
[Ring α]
[Preorder α]
[OrderedRing α]
(ctx : Context α)
(p : Poly)
:
zero_ne_one_cert p = true → Var.denote ctx 0 = One.one → Poly.denote' ctx p ≠ 0
theorem
Lean.Grind.Linarith.zero_ne_one_of_field
{α : Type u_1}
[Field α]
(ctx : Context α)
(p : Poly)
:
zero_ne_one_cert p = true → Var.denote ctx 0 = One.one → Poly.denote' ctx p ≠ 0
theorem
Lean.Grind.Linarith.zero_ne_one_of_char0
{α : Type u_1}
[Ring α]
[IsCharP α 0]
(ctx : Context α)
(p : Poly)
:
zero_ne_one_cert p = true → Var.denote ctx 0 = One.one → Poly.denote' ctx p ≠ 0
Equations
Instances For
theorem
Lean.Grind.Linarith.zero_ne_one_of_charC
{α : Type u_1}
{c : Nat}
[Ring α]
[IsCharP α c]
(ctx : Context α)
(p : Poly)
:
zero_ne_one_of_charC_cert c p = true → Var.denote ctx 0 = One.one → Poly.denote' ctx p ≠ 0
Coefficient normalization
Equations
- Lean.Grind.Linarith.eq_neg_cert p₁ p₂ = (p₂ == p₁.mul (-1))
Instances For
theorem
Lean.Grind.Linarith.eq_neg
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(p₁ p₂ : Poly)
:
eq_neg_cert p₁ p₂ = true → Poly.denote' ctx p₁ = 0 → Poly.denote' ctx p₂ = 0
theorem
Lean.Grind.Linarith.eq_coeff
{α : Type u_1}
[IntModule α]
[NoNatZeroDivisors α]
(ctx : Context α)
(p₁ p₂ : Poly)
(k : Nat)
:
eq_coeff_cert p₁ p₂ k = true → Poly.denote' ctx p₁ = 0 → Poly.denote' ctx p₂ = 0
theorem
Lean.Grind.Linarith.le_coeff
{α : Type u_1}
[IntModule α]
[LinearOrder α]
[OrderedAdd α]
(ctx : Context α)
(p₁ p₂ : Poly)
(k : Nat)
:
coeff_cert p₁ p₂ k = true → Poly.denote' ctx p₁ ≤ 0 → Poly.denote' ctx p₂ ≤ 0
theorem
Lean.Grind.Linarith.lt_coeff
{α : Type u_1}
[IntModule α]
[LinearOrder α]
[OrderedAdd α]
(ctx : Context α)
(p₁ p₂ : Poly)
(k : Nat)
:
coeff_cert p₁ p₂ k = true → Poly.denote' ctx p₁ < 0 → Poly.denote' ctx p₂ < 0
theorem
Lean.Grind.Linarith.diseq_neg
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(p p' : Poly)
:
(p' == p.mul (-1)) = true → Poly.denote' ctx p ≠ 0 → Poly.denote' ctx p' ≠ 0
Substitution
theorem
Lean.Grind.Linarith.eq_diseq_subst
{α : Type u_1}
[IntModule α]
[NoNatZeroDivisors α]
(ctx : Context α)
(k₁ k₂ : Int)
(p₁ p₂ p₃ : Poly)
:
eq_diseq_subst_cert k₁ k₂ p₁ p₂ p₃ = true → Poly.denote' ctx p₁ = 0 → Poly.denote' ctx p₂ ≠ 0 → Poly.denote' ctx p₃ ≠ 0
Equations
- Lean.Grind.Linarith.eq_diseq_subst1_cert k p₁ p₂ p₃ = (p₃ == (p₁.mul k).combine p₂)
Instances For
theorem
Lean.Grind.Linarith.eq_diseq_subst1
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(k : Int)
(p₁ p₂ p₃ : Poly)
:
eq_diseq_subst1_cert k p₁ p₂ p₃ = true → Poly.denote' ctx p₁ = 0 → Poly.denote' ctx p₂ ≠ 0 → Poly.denote' ctx p₃ ≠ 0
theorem
Lean.Grind.Linarith.eq_le_subst
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(x : Var)
(p₁ p₂ p₃ : Poly)
:
eq_le_subst_cert x p₁ p₂ p₃ = true → Poly.denote' ctx p₁ = 0 → Poly.denote' ctx p₂ ≤ 0 → Poly.denote' ctx p₃ ≤ 0
theorem
Lean.Grind.Linarith.eq_lt_subst
{α : Type u_1}
[IntModule α]
[Preorder α]
[OrderedAdd α]
(ctx : Context α)
(x : Var)
(p₁ p₂ p₃ : Poly)
:
eq_lt_subst_cert x p₁ p₂ p₃ = true → Poly.denote' ctx p₁ = 0 → Poly.denote' ctx p₂ < 0 → Poly.denote' ctx p₃ < 0
theorem
Lean.Grind.Linarith.eq_eq_subst
{α : Type u_1}
[IntModule α]
(ctx : Context α)
(x : Var)
(p₁ p₂ p₃ : Poly)
:
eq_eq_subst_cert x p₁ p₂ p₃ = true → Poly.denote' ctx p₁ = 0 → Poly.denote' ctx p₂ = 0 → Poly.denote' ctx p₃ = 0