Documentation

Init.Data.Int.Linear

Helper definitions and theorems for constructing linear arithmetic proofs.

@[reducible, inline]
Equations
@[reducible, inline]
Equations

Similar to Poly.denote, but produces a denotation better for simp +arith. Remark: we used to convert Poly back into Expr to achieve that.

Equations
theorem Int.Linear.Poly.denote'_go_eq_denote (ctx : Context) (p : Poly) (r : Int) :
denote'.go ctx r p = denote ctx p + r
theorem Int.Linear.Poly.denote'_add (ctx : Context) (a : Int) (x : Var) (p : Poly) :
denote' ctx (add a x p) = a * Var.denote ctx x + denote ctx p
def Int.Linear.Poly.insert (k : Int) (v : Var) (p : Poly) :
Equations

Normalizes the given polynomial by fusing monomial and constants.

Equations
def Int.Linear.Poly.append (p₁ p₂ : Poly) :
Equations
def Int.Linear.Poly.combine (p₁ p₂ : Poly) :
Equations

Converts the given expression into a polynomial.

Equations

Converts the given expression into a polynomial, and then normalizes it.

Equations
def Int.Linear.cdiv (a b : Int) :

Returns the ceiling of the division a / b. That is, the result is equivalent to ⌈a / b⌉. Examples:

  • cdiv 7 3 returns 3
  • cdiv (-7) 3 returns -2.
Equations
def Int.Linear.cmod (a b : Int) :

Returns the ceiling-compatible remainder of the division a / b. This function ensures that the remainder is consistent with cdiv, meaning:

a = b * cdiv a b + cmod a b

See theorem cdiv_add_cmod. We also have

-b < cmod a b ≤ 0
Equations
theorem Int.Linear.cdiv_add_cmod (a b : Int) :
b * cdiv a b + cmod a b = a
theorem Int.Linear.cmod_gt_of_pos (a : Int) {b : Int} (h : 0 < b) :
cmod a b > -b
theorem Int.Linear.cmod_nonpos (a : Int) {b : Int} (h : b 0) :
cmod a b 0
theorem Int.Linear.cdiv_eq_div_of_divides {a b : Int} (h : a % b = 0) :
a / b = cdiv a b

Returns the constant of the given linear polynomial.

Equations

p.div k divides all coefficients of the polynomial p by k, but rounds up the constant using cdiv. Notes:

  • We only use this function with ks that divides all coefficients.
  • We use cdiv for the constant to implement the inequality tightening rule.
Equations

Returns true if k divides all coefficients and the constant of the given linear polynomial.

Equations

Returns true if k divides all coefficients of the given linear polynomial.

Equations

p.mul k multiplies all coefficients and constant of the polynomial p by k.

Equations
def Int.Linear.Poly.mul (p : Poly) (k : Int) :
Equations
@[simp]
theorem Int.Linear.Poly.denote_mul (ctx : Context) (p : Poly) (k : Int) :
denote ctx (p.mul k) = k * denote ctx p
theorem Int.Linear.Poly.denote_addConst (ctx : Context) (p : Poly) (k : Int) :
denote ctx (p.addConst k) = denote ctx p + k
theorem Int.Linear.Poly.denote_insert (ctx : Context) (k : Int) (v : Var) (p : Poly) :
denote ctx (insert k v p) = denote ctx p + k * Var.denote ctx v
theorem Int.Linear.Poly.denote_norm (ctx : Context) (p : Poly) :
denote ctx p.norm = denote ctx p
theorem Int.Linear.Poly.denote_append (ctx : Context) (p₁ p₂ : Poly) :
denote ctx (p₁.append p₂) = denote ctx p₁ + denote ctx p₂
theorem Int.Linear.Poly.denote_combine' (ctx : Context) (fuel : Nat) (p₁ p₂ : Poly) :
denote ctx (combine' fuel p₁ p₂) = denote ctx p₁ + denote ctx p₂
theorem Int.Linear.Poly.denote_combine (ctx : Context) (p₁ p₂ : Poly) :
denote ctx (p₁.combine p₂) = denote ctx p₁ + denote ctx p₂
theorem Int.Linear.sub_fold (a b : Int) :
a.sub b = a - b
theorem Int.Linear.neg_fold (a : Int) :
a.neg = -a
theorem Int.Linear.Poly.denote_div_eq_of_divAll (ctx : Context) (p : Poly) (k : Int) :
divAll k p = truedenote ctx (div k p) * k = denote ctx p
theorem Int.Linear.Poly.denote_div_eq_of_divCoeffs (ctx : Context) (p : Poly) (k : Int) :
divCoeffs k p = truedenote ctx (div k p) * k + cmod p.getConst k = denote ctx p
theorem Int.Linear.Expr.denote_toPoly'_go {k : Int} {p : Poly} (ctx : Context) (e : Expr) :
Poly.denote ctx (toPoly'.go k e p) = k * denote ctx e + Poly.denote ctx p
theorem Int.Linear.Expr.eq_of_norm_eq (ctx : Context) (e : Expr) (p : Poly) (h : (e.norm == p) = true) :
denote ctx e = Poly.denote' ctx p
def Int.Linear.norm_eq_cert (lhs rhs : Expr) (p : Poly) :
Equations
theorem Int.Linear.norm_eq (ctx : Context) (lhs rhs : Expr) (p : Poly) (h : norm_eq_cert lhs rhs p = true) :
(Expr.denote ctx lhs = Expr.denote ctx rhs) = (Poly.denote' ctx p = 0)
theorem Int.Linear.norm_le (ctx : Context) (lhs rhs : Expr) (p : Poly) (h : norm_eq_cert lhs rhs p = true) :
(Expr.denote ctx lhs Expr.denote ctx rhs) = (Poly.denote' ctx p 0)
theorem Int.Linear.norm_eq_var (ctx : Context) (lhs rhs : Expr) (x y : Var) (h : norm_eq_var_cert lhs rhs x y = true) :
(Expr.denote ctx lhs = Expr.denote ctx rhs) = (Var.denote ctx x = Var.denote ctx y)
theorem Int.Linear.norm_eq_var_const (ctx : Context) (lhs rhs : Expr) (x : Var) (k : Int) (h : norm_eq_var_const_cert lhs rhs x k = true) :
(Expr.denote ctx lhs = Expr.denote ctx rhs) = (Var.denote ctx x = k)
theorem Int.Linear.norm_eq_coeff' (ctx : Context) (p p' : Poly) (k : Int) :
p = p'.mul kk > 0 → (Poly.denote ctx p = 0 Poly.denote ctx p' = 0)
def Int.Linear.norm_eq_coeff_cert (lhs rhs : Expr) (p : Poly) (k : Int) :
Equations
theorem Int.Linear.norm_eq_coeff (ctx : Context) (lhs rhs : Expr) (p : Poly) (k : Int) :
norm_eq_coeff_cert lhs rhs p k = true → (Expr.denote ctx lhs = Expr.denote ctx rhs) = (Poly.denote' ctx p = 0)
theorem Int.Linear.norm_le_coeff (ctx : Context) (lhs rhs : Expr) (p : Poly) (k : Int) :
norm_eq_coeff_cert lhs rhs p k = true → (Expr.denote ctx lhs Expr.denote ctx rhs) = (Poly.denote' ctx p 0)
theorem Int.Linear.norm_le_coeff_tight (ctx : Context) (lhs rhs : Expr) (p : Poly) (k : Int) :
norm_le_coeff_tight_cert lhs rhs p k = true → (Expr.denote ctx lhs Expr.denote ctx rhs) = (Poly.denote' ctx p 0)
theorem Int.Linear.eq_eq_false (ctx : Context) (lhs rhs : Expr) :
(lhs.sub rhs).norm.isUnsatEq = true → (Expr.denote ctx lhs = Expr.denote ctx rhs) = False
theorem Int.Linear.eq_eq_true (ctx : Context) (lhs rhs : Expr) :
(lhs.sub rhs).norm.isValidEq = true → (Expr.denote ctx lhs = Expr.denote ctx rhs) = True
theorem Int.Linear.le_eq_false (ctx : Context) (lhs rhs : Expr) :
(lhs.sub rhs).norm.isUnsatLe = true → (Expr.denote ctx lhs Expr.denote ctx rhs) = False
theorem Int.Linear.le_eq_true (ctx : Context) (lhs rhs : Expr) :
(lhs.sub rhs).norm.isValidLe = true → (Expr.denote ctx lhs Expr.denote ctx rhs) = True
theorem Int.Linear.eq_eq_false_of_divCoeff (ctx : Context) (lhs rhs : Expr) (k : Int) :
unsatEqDivCoeffCert lhs rhs k = true → (Expr.denote ctx lhs = Expr.denote ctx rhs) = False
theorem Int.Linear.Poly.gcd_dvd_const {ctx : Context} {p : Poly} {k : Int} (h : k denote ctx p) :
theorem Int.Linear.dvd_unsat (ctx : Context) (k : Int) (p : Poly) :
theorem Int.Linear.norm_dvd (ctx : Context) (k : Int) (e : Expr) (p : Poly) :
(e.norm == p) = true → (k Expr.denote ctx e) = (k Poly.denote' ctx p)
theorem Int.Linear.dvd_eq_false (ctx : Context) (k : Int) (e : Expr) (h : Poly.isUnsatDvd k e.norm = true) :
def Int.Linear.dvd_coeff_cert (k₁ : Int) (p₁ : Poly) (k₂ : Int) (p₂ : Poly) (k : Int) :
Equations
def Int.Linear.norm_dvd_gcd_cert (k₁ : Int) (e₁ : Expr) (k₂ : Int) (p₂ : Poly) (k : Int) :
Equations
theorem Int.Linear.norm_dvd_gcd (ctx : Context) (k₁ : Int) (e₁ : Expr) (k₂ : Int) (p₂ : Poly) (g : Int) :
norm_dvd_gcd_cert k₁ e₁ k₂ p₂ g = true → (k₁ Expr.denote ctx e₁) = (k₂ Poly.denote' ctx p₂)
theorem Int.Linear.dvd_coeff (ctx : Context) (k₁ : Int) (p₁ : Poly) (k₂ : Int) (p₂ : Poly) (g : Int) :
dvd_coeff_cert k₁ p₁ k₂ p₂ g = truek₁ Poly.denote' ctx p₁k₂ Poly.denote' ctx p₂
def Int.Linear.dvd_elim_cert (k₁ : Int) (p₁ : Poly) (k₂ : Int) (p₂ : Poly) :
Equations
theorem Int.Linear.dvd_elim (ctx : Context) (k₁ : Int) (p₁ : Poly) (k₂ : Int) (p₂ : Poly) :
dvd_elim_cert k₁ p₁ k₂ p₂ = truek₁ Poly.denote' ctx p₁k₂ Poly.denote' ctx p₂
def Int.Linear.dvd_solve_combine_cert (d₁ : Int) (p₁ : Poly) (d₂ : Int) (p₂ : Poly) (d : Int) (p : Poly) (g α β : Int) :
Equations
theorem Int.Linear.dvd_solve_combine (ctx : Context) (d₁ : Int) (p₁ : Poly) (d₂ : Int) (p₂ : Poly) (d : Int) (p : Poly) (g α β : Int) :
dvd_solve_combine_cert d₁ p₁ d₂ p₂ d p g α β = trued₁ Poly.denote' ctx p₁d₂ Poly.denote' ctx p₂d Poly.denote' ctx p
def Int.Linear.dvd_solve_elim_cert (d₁ : Int) (p₁ : Poly) (d₂ : Int) (p₂ : Poly) (d : Int) (p : Poly) :
Equations
theorem Int.Linear.dvd_solve_elim (ctx : Context) (d₁ : Int) (p₁ : Poly) (d₂ : Int) (p₂ : Poly) (d : Int) (p : Poly) :
dvd_solve_elim_cert d₁ p₁ d₂ p₂ d p = trued₁ Poly.denote' ctx p₁d₂ Poly.denote' ctx p₂d Poly.denote' ctx p
theorem Int.Linear.dvd_norm (ctx : Context) (d : Int) (p₁ p₂ : Poly) :
(p₁.norm == p₂) = trued Poly.denote' ctx p₁d Poly.denote' ctx p₂
theorem Int.Linear.le_norm (ctx : Context) (p₁ p₂ : Poly) (h : (p₁.norm == p₂) = true) :
Poly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0
def Int.Linear.le_coeff_cert (p₁ p₂ : Poly) (k : Int) :
Equations
theorem Int.Linear.le_coeff (ctx : Context) (p₁ p₂ : Poly) (k : Int) :
le_coeff_cert p₁ p₂ k = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0
def Int.Linear.le_neg_cert (p₁ p₂ : Poly) :
Equations
theorem Int.Linear.le_neg (ctx : Context) (p₁ p₂ : Poly) :
le_neg_cert p₁ p₂ = true¬Poly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0
def Int.Linear.le_combine_cert (p₁ p₂ p₃ : Poly) :
Equations
theorem Int.Linear.le_combine (ctx : Context) (p₁ p₂ p₃ : Poly) :
le_combine_cert p₁ p₂ p₃ = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
def Int.Linear.le_combine_coeff_cert (p₁ p₂ p₃ : Poly) (k : Int) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.le_combine_coeff (ctx : Context) (p₁ p₂ p₃ : Poly) (k : Int) :
le_combine_coeff_cert p₁ p₂ p₃ k = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
theorem Int.Linear.le_unsat (ctx : Context) (p : Poly) :
p.isUnsatLe = truePoly.denote' ctx p 0False
theorem Int.Linear.eq_norm (ctx : Context) (p₁ p₂ : Poly) (h : (p₁.norm == p₂) = true) :
Poly.denote' ctx p₁ = 0Poly.denote' ctx p₂ = 0
Equations
theorem Int.Linear.eq_coeff (ctx : Context) (p p' : Poly) (k : Int) :
eq_coeff_cert p p' k = truePoly.denote' ctx p = 0Poly.denote' ctx p' = 0
theorem Int.Linear.eq_unsat (ctx : Context) (p : Poly) :
p.isUnsatEq = truePoly.denote' ctx p = 0False
theorem Int.Linear.eq_unsat_coeff (ctx : Context) (p : Poly) (k : Int) :
Equations
def Int.Linear.dvd_of_eq_cert (x : Var) (p₁ : Poly) (d₂ : Int) (p₂ : Poly) :
Equations
theorem Int.Linear.dvd_of_eq (ctx : Context) (x : Var) (p₁ : Poly) (d₂ : Int) (p₂ : Poly) :
dvd_of_eq_cert x p₁ d₂ p₂ = truePoly.denote' ctx p₁ = 0d₂ Poly.denote' ctx p₂
def Int.Linear.eq_dvd_subst_cert (x : Var) (p₁ : Poly) (d₂ : Int) (p₂ : Poly) (d₃ : Int) (p₃ : Poly) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.eq_dvd_subst (ctx : Context) (x : Var) (p₁ : Poly) (d₂ : Int) (p₂ : Poly) (d₃ : Int) (p₃ : Poly) :
eq_dvd_subst_cert x p₁ d₂ p₂ d₃ p₃ = truePoly.denote' ctx p₁ = 0d₂ Poly.denote' ctx p₂d₃ Poly.denote' ctx p₃
def Int.Linear.eq_eq_subst_cert (x : Var) (p₁ p₂ p₃ : Poly) :
Equations
theorem Int.Linear.eq_eq_subst (ctx : Context) (x : Var) (p₁ p₂ p₃ : Poly) :
eq_eq_subst_cert x p₁ p₂ p₃ = truePoly.denote' ctx p₁ = 0Poly.denote' ctx p₂ = 0Poly.denote' ctx p₃ = 0
def Int.Linear.eq_le_subst_nonneg_cert (x : Var) (p₁ p₂ p₃ : Poly) :
Equations
theorem Int.Linear.eq_le_subst_nonneg (ctx : Context) (x : Var) (p₁ p₂ p₃ : Poly) :
eq_le_subst_nonneg_cert x p₁ p₂ p₃ = truePoly.denote' ctx p₁ = 0Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
def Int.Linear.eq_le_subst_nonpos_cert (x : Var) (p₁ p₂ p₃ : Poly) :
Equations
theorem Int.Linear.eq_le_subst_nonpos (ctx : Context) (x : Var) (p₁ p₂ p₃ : Poly) :
eq_le_subst_nonpos_cert x p₁ p₂ p₃ = truePoly.denote' ctx p₁ = 0Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
def Int.Linear.eq_of_core_cert (p₁ p₂ p₃ : Poly) :
Equations
theorem Int.Linear.eq_of_core (ctx : Context) (p₁ p₂ p₃ : Poly) :
eq_of_core_cert p₁ p₂ p₃ = truePoly.denote' ctx p₁ = Poly.denote' ctx p₂Poly.denote' ctx p₃ = 0
theorem Int.Linear.diseq_norm (ctx : Context) (p₁ p₂ : Poly) (h : (p₁.norm == p₂) = true) :
Poly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0
theorem Int.Linear.diseq_coeff (ctx : Context) (p p' : Poly) (k : Int) :
eq_coeff_cert p p' k = truePoly.denote' ctx p 0Poly.denote' ctx p' 0
theorem Int.Linear.diseq_neg (ctx : Context) (p p' : Poly) :
(p' == p.mul (-1)) = truePoly.denote' ctx p 0Poly.denote' ctx p' 0
def Int.Linear.diseq_eq_subst_cert (x : Var) (p₁ p₂ p₃ : Poly) :
Equations
theorem Int.Linear.eq_diseq_subst (ctx : Context) (x : Var) (p₁ p₂ p₃ : Poly) :
diseq_eq_subst_cert x p₁ p₂ p₃ = truePoly.denote' ctx p₁ = 0Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
theorem Int.Linear.diseq_of_core (ctx : Context) (p₁ p₂ p₃ : Poly) :
eq_of_core_cert p₁ p₂ p₃ = truePoly.denote' ctx p₁ Poly.denote' ctx p₂Poly.denote' ctx p₃ 0
Equations
theorem Int.Linear.eq_of_le_ge (ctx : Context) (p₁ p₂ : Poly) :
eq_of_le_ge_cert p₁ p₂ = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0Poly.denote' ctx p₁ = 0
def Int.Linear.le_of_le_diseq_cert (p₁ p₂ p₃ : Poly) :
Equations
theorem Int.Linear.le_of_le_diseq (ctx : Context) (p₁ p₂ p₃ : Poly) :
le_of_le_diseq_cert p₁ p₂ p₃ = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
def Int.Linear.diseq_split_cert (p₁ p₂ p₃ : Poly) :
Equations
theorem Int.Linear.diseq_split (ctx : Context) (p₁ p₂ p₃ : Poly) :
diseq_split_cert p₁ p₂ p₃ = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0 Poly.denote' ctx p₃ 0
theorem Int.Linear.diseq_split_resolve (ctx : Context) (p₁ p₂ p₃ : Poly) :
diseq_split_cert p₁ p₂ p₃ = truePoly.denote' ctx p₁ 0¬Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
def Int.Linear.OrOver (n : Nat) (p : NatProp) :
Equations
theorem Int.Linear.orOver_one {p : NatProp} :
OrOver 1 pp 0
theorem Int.Linear.orOver_resolve {n : Nat} {p : NatProp} :
OrOver (n + 1) p¬p nOrOver n p
theorem Int.Linear.orOver_cases {n : Nat} {p : NatProp} :
OrOver (n + 1) pOrOver_cases_type n p
def Int.Linear.cooper_dvd_left_cert (p₁ p₂ p₃ : Poly) (d : Int) (n : Nat) :
Equations
  • One or more equations did not get rendered due to their size.
Equations
def Int.Linear.cooper_dvd_left_split (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (k : Nat) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.cooper_dvd_left (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (n : Nat) :
cooper_dvd_left_cert p₁ p₂ p₃ d n = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0d Poly.denote' ctx p₃OrOver n (cooper_dvd_left_split ctx p₁ p₂ p₃ d)
def Int.Linear.cooper_dvd_left_split_ineq_cert (p₁ p₂ : Poly) (k b : Int) (p' : Poly) :
Equations
theorem Int.Linear.cooper_dvd_left_split_ineq (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (k : Nat) (b : Int) (p' : Poly) :
cooper_dvd_left_split ctx p₁ p₂ p₃ d kcooper_dvd_left_split_ineq_cert p₁ p₂ (↑k) b p' = truePoly.denote' ctx p' 0
theorem Int.Linear.cooper_dvd_left_split_dvd1 (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (k : Nat) (a : Int) (p' : Poly) :
cooper_dvd_left_split ctx p₁ p₂ p₃ d kcooper_dvd_left_split_dvd1_cert p₁ p' a k = truea Poly.denote' ctx p'
def Int.Linear.cooper_dvd_left_split_dvd2_cert (p₁ p₃ : Poly) (d : Int) (k : Nat) (d' : Int) (p' : Poly) :
Equations
theorem Int.Linear.cooper_dvd_left_split_dvd2 (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (k : Nat) (d' : Int) (p' : Poly) :
cooper_dvd_left_split ctx p₁ p₂ p₃ d kcooper_dvd_left_split_dvd2_cert p₁ p₃ d k d' p' = trued' Poly.denote' ctx p'
def Int.Linear.cooper_left_cert (p₁ p₂ : Poly) (n : Nat) :
Equations
  • One or more equations did not get rendered due to their size.
def Int.Linear.cooper_left_split (ctx : Context) (p₁ p₂ : Poly) (k : Nat) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.cooper_left (ctx : Context) (p₁ p₂ : Poly) (n : Nat) :
cooper_left_cert p₁ p₂ n = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0OrOver n (cooper_left_split ctx p₁ p₂)
def Int.Linear.cooper_left_split_ineq_cert (p₁ p₂ : Poly) (k b : Int) (p' : Poly) :
Equations
theorem Int.Linear.cooper_left_split_ineq (ctx : Context) (p₁ p₂ : Poly) (k : Nat) (b : Int) (p' : Poly) :
cooper_left_split ctx p₁ p₂ kcooper_left_split_ineq_cert p₁ p₂ (↑k) b p' = truePoly.denote' ctx p' 0
theorem Int.Linear.cooper_left_split_dvd (ctx : Context) (p₁ p₂ : Poly) (k : Nat) (a : Int) (p' : Poly) :
cooper_left_split ctx p₁ p₂ kcooper_left_split_dvd_cert p₁ p' a k = truea Poly.denote' ctx p'
def Int.Linear.cooper_dvd_right_cert (p₁ p₂ p₃ : Poly) (d : Int) (n : Nat) :
Equations
  • One or more equations did not get rendered due to their size.
def Int.Linear.cooper_dvd_right_split (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (k : Nat) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.cooper_dvd_right (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (n : Nat) :
cooper_dvd_right_cert p₁ p₂ p₃ d n = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0d Poly.denote' ctx p₃OrOver n (cooper_dvd_right_split ctx p₁ p₂ p₃ d)
Equations
theorem Int.Linear.cooper_dvd_right_split_ineq (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (k : Nat) (a : Int) (p' : Poly) :
cooper_dvd_right_split ctx p₁ p₂ p₃ d kcooper_dvd_right_split_ineq_cert p₁ p₂ (↑k) a p' = truePoly.denote' ctx p' 0
theorem Int.Linear.cooper_dvd_right_split_dvd1 (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (k : Nat) (b : Int) (p' : Poly) :
cooper_dvd_right_split ctx p₁ p₂ p₃ d kcooper_dvd_right_split_dvd1_cert p₂ p' b k = trueb Poly.denote' ctx p'
def Int.Linear.cooper_dvd_right_split_dvd2_cert (p₂ p₃ : Poly) (d : Int) (k : Nat) (d' : Int) (p' : Poly) :
Equations
theorem Int.Linear.cooper_dvd_right_split_dvd2 (ctx : Context) (p₁ p₂ p₃ : Poly) (d : Int) (k : Nat) (d' : Int) (p' : Poly) :
cooper_dvd_right_split ctx p₁ p₂ p₃ d kcooper_dvd_right_split_dvd2_cert p₂ p₃ d k d' p' = trued' Poly.denote' ctx p'
def Int.Linear.cooper_right_cert (p₁ p₂ : Poly) (n : Nat) :
Equations
  • One or more equations did not get rendered due to their size.
def Int.Linear.cooper_right_split (ctx : Context) (p₁ p₂ : Poly) (k : Nat) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.cooper_right (ctx : Context) (p₁ p₂ : Poly) (n : Nat) :
cooper_right_cert p₁ p₂ n = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0OrOver n (cooper_right_split ctx p₁ p₂)
def Int.Linear.cooper_right_split_ineq_cert (p₁ p₂ : Poly) (k a : Int) (p' : Poly) :
Equations
theorem Int.Linear.cooper_right_split_ineq (ctx : Context) (p₁ p₂ : Poly) (k : Nat) (a : Int) (p' : Poly) :
cooper_right_split ctx p₁ p₂ kcooper_right_split_ineq_cert p₁ p₂ (↑k) a p' = truePoly.denote' ctx p' 0
theorem Int.Linear.cooper_right_split_dvd (ctx : Context) (p₁ p₂ : Poly) (k : Nat) (b : Int) (p' : Poly) :
cooper_right_split ctx p₁ p₂ kcooper_right_split_dvd_cert p₂ p' b k = trueb Poly.denote' ctx p'
@[reducible, inline]
abbrev Int.Linear.Poly.casesOnAdd (p : Poly) (k : IntVarPolyBool) :
Equations
@[reducible, inline]
abbrev Int.Linear.Poly.casesOnNum (p : Poly) (k : IntBool) :
Equations
def Int.Linear.cooper_unsat_cert (p₁ p₂ p₃ : Poly) (d α β : Int) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.cooper_unsat (ctx : Context) (p₁ p₂ p₃ : Poly) (d α β : Int) :
cooper_unsat_cert p₁ p₂ p₃ d α β = truePoly.denote' ctx p₁ 0Poly.denote' ctx p₂ 0d Poly.denote' ctx p₃False
theorem Int.Linear.ediv_emod (x y : Int) :
-1 * x + y * (x / y) + x % y = 0
theorem Int.Linear.emod_nonneg (x y : Int) :
(y != 0) = true-1 * (x % y) 0
Equations
theorem Int.Linear.emod_le (x y n : Int) :
emod_le_cert y n = truex % y + n 0
theorem Int.Linear.natCast_sub (x y : Nat) :
(x - y) = if y + -1 * x 0 then x + -1 * y else 0
def Int.Linear.dvd_le_tight_cert (d : Int) (p₁ p₂ p₃ : Poly) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.dvd_le_tight (ctx : Context) (d : Int) (p₁ p₂ p₃ : Poly) :
dvd_le_tight_cert d p₁ p₂ p₃ = trued Poly.denote' ctx p₁Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
def Int.Linear.dvd_neg_le_tight_cert (d : Int) (p₁ p₂ p₃ : Poly) :
Equations
  • One or more equations did not get rendered due to their size.
theorem Int.Linear.dvd_neg_le_tight (ctx : Context) (d : Int) (p₁ p₂ p₃ : Poly) :
dvd_neg_le_tight_cert d p₁ p₂ p₃ = trued Poly.denote' ctx p₁Poly.denote' ctx p₂ 0Poly.denote' ctx p₃ 0
theorem Int.Linear.le_norm_expr (ctx : Context) (lhs rhs : Expr) (p : Poly) :
norm_eq_cert lhs rhs p = trueExpr.denote ctx lhs Expr.denote ctx rhsPoly.denote' ctx p 0
Equations
theorem Int.Linear.not_le_norm_expr (ctx : Context) (lhs rhs : Expr) (p : Poly) :
not_le_norm_expr_cert lhs rhs p = true¬Expr.denote ctx lhs Expr.denote ctx rhsPoly.denote' ctx p 0
theorem Int.Linear.dvd_norm_expr (ctx : Context) (d : Int) (e : Expr) (p : Poly) :
(p == e.norm) = trued Expr.denote ctx ed Poly.denote' ctx p
theorem Int.Linear.eq_norm_expr (ctx : Context) (lhs rhs : Expr) (p : Poly) :
norm_eq_cert lhs rhs p = trueExpr.denote ctx lhs = Expr.denote ctx rhsPoly.denote' ctx p = 0
theorem Int.Linear.not_eq_norm_expr (ctx : Context) (lhs rhs : Expr) (p : Poly) :
norm_eq_cert lhs rhs p = true¬Expr.denote ctx lhs = Expr.denote ctx rhs¬Poly.denote' ctx p = 0
theorem Int.Linear.of_not_dvd (a b : Int) :
(a != 0) = true¬a bb % a > 0
Equations
theorem Int.Linear.le_of_le (ctx : Context) (p q : Poly) (k : Nat) :
le_of_le_cert p q k = truePoly.denote' ctx p 0Poly.denote' ctx q 0
Equations
theorem Int.Linear.not_le_of_le (ctx : Context) (p q : Poly) (k : Nat) :
not_le_of_le_cert p q k = truePoly.denote' ctx p 0¬Poly.denote' ctx q 0
theorem Int.not_le_eq (a b : Int) :
(¬a b) = (b + 1 a)
theorem Int.not_ge_eq (a b : Int) :
(¬a b) = (a + 1 b)
theorem Int.not_lt_eq (a b : Int) :
(¬a < b) = (b a)
theorem Int.not_gt_eq (a b : Int) :
(¬a > b) = (a b)