Get the i-th element (interpreted as 0 if the list is not long enough).
Instances For
@[simp]
theorem
Lean.Omega.IntList.get_cons_zero
{x : Int}
{xs : List Int}
 :
Lean.Omega.IntList.get (x :: xs) 0 = x
@[simp]
theorem
Lean.Omega.IntList.get_cons_succ
{x : Int}
{xs : List Int}
{i : Nat}
 :
Lean.Omega.IntList.get (x :: xs) (i + 1) = Lean.Omega.IntList.get xs i
theorem
Lean.Omega.IntList.get_map
{f : Int → Int}
{i : Nat}
{xs : Lean.Omega.IntList}
(h : f 0 = 0)
 :
Lean.Omega.IntList.get (List.map f xs) i = f (xs.get i)
theorem
Lean.Omega.IntList.get_of_length_le
{i : Nat}
{xs : Lean.Omega.IntList}
(h : List.length xs ≤ i)
 :
xs.get i = 0
Like List.set, but right-pad with zeroes as necessary first.
Equations
- Lean.Omega.IntList.set [] 0 y = [y]
- Lean.Omega.IntList.set [] i_2.succ y = 0 :: Lean.Omega.IntList.set [] i_2 y
- Lean.Omega.IntList.set (head :: xs_2) 0 y = y :: xs_2
- Lean.Omega.IntList.set (x :: xs_2) i_2.succ y = x :: Lean.Omega.IntList.set xs_2 i_2 y
Instances For
@[simp]
@[simp]
theorem
Lean.Omega.IntList.set_nil_succ
{i : Nat}
{y : Int}
 :
Lean.Omega.IntList.set [] (i + 1) y = 0 :: Lean.Omega.IntList.set [] i y
@[simp]
theorem
Lean.Omega.IntList.set_cons_zero
{x : Int}
{xs : List Int}
{y : Int}
 :
Lean.Omega.IntList.set (x :: xs) 0 y = y :: xs
@[simp]
theorem
Lean.Omega.IntList.set_cons_succ
{x : Int}
{xs : List Int}
{i : Nat}
{y : Int}
 :
Lean.Omega.IntList.set (x :: xs) (i + 1) y = x :: Lean.Omega.IntList.set xs i y
Returns the leading coefficient, i.e. the first non-zero entry.
Equations
- xs.leading = (List.find? (fun (x : Int) => !x == 0) xs).getD 0
Instances For
Equations
- Lean.Omega.IntList.instAdd = { add := Lean.Omega.IntList.add }
@[simp]
@[simp]
theorem
Lean.Omega.IntList.cons_add_cons
(x : Int)
(xs : Lean.Omega.IntList)
(y : Int)
(ys : Lean.Omega.IntList)
 :
Implementation of * on IntList.
Equations
- xs.mul ys = List.zipWith (fun (x x_1 : Int) => x * x_1) xs ys
Instances For
Equations
- Lean.Omega.IntList.instMul = { mul := Lean.Omega.IntList.mul }
theorem
Lean.Omega.IntList.mul_def
(xs : Lean.Omega.IntList)
(ys : Lean.Omega.IntList)
 :
xs * ys = List.zipWith (fun (x x_1 : Int) => x * x_1) xs ys
@[simp]
Implementation of negation on IntList.
Instances For
Equations
- Lean.Omega.IntList.instNeg = { neg := Lean.Omega.IntList.neg }
@[simp]
Equations
- Lean.Omega.IntList.instSub = { sub := Lean.Omega.IntList.sub }
Implementation of scalar multiplication by an integer on IntList.
Instances For
Equations
- Lean.Omega.IntList.instHMulInt = { hMul := fun (i : Int) (xs : Lean.Omega.IntList) => xs.smul i }
@[simp]
def
Lean.Omega.IntList.combo
(a : Int)
(xs : Lean.Omega.IntList)
(b : Int)
(ys : Lean.Omega.IntList)
 :
A linear combination of two IntLists.
Equations
- Lean.Omega.IntList.combo a xs b ys = List.zipWithAll (fun (x y : Option Int) => a * x.getD 0 + b * y.getD 0) xs ys
Instances For
theorem
Lean.Omega.IntList.combo_eq_smul_add_smul
(a : Int)
(xs : Lean.Omega.IntList)
(b : Int)
(ys : Lean.Omega.IntList)
 :
Lean.Omega.IntList.combo a xs b ys = a * xs + b * ys
theorem
Lean.Omega.IntList.mul_distrib_left
(xs : Lean.Omega.IntList)
(ys : Lean.Omega.IntList)
(zs : Lean.Omega.IntList)
 :
@[simp]
theorem
Lean.Omega.IntList.mul_smul_left
{i : Int}
{xs : Lean.Omega.IntList}
{ys : Lean.Omega.IntList}
 :
The sum of the entries of an IntList.
Equations
- xs.sum = List.foldr (fun (x x_1 : Int) => x + x_1) 0 xs
Instances For
@[simp]
theorem
Lean.Omega.IntList.sum_cons
{x : Int}
{xs : List Int}
 :
Lean.Omega.IntList.sum (x :: xs) = x + Lean.Omega.IntList.sum xs
@[simp]
The dot product of two IntLists.
Instances For
@[simp]
theorem
Lean.Omega.IntList.dot_cons₂
{x : Int}
{xs : List Int}
{y : Int}
{ys : List Int}
 :
Lean.Omega.IntList.dot (x :: xs) (y :: ys) = x * y + Lean.Omega.IntList.dot xs ys
@[simp]
theorem
Lean.Omega.IntList.dot_set_left
(xs : Lean.Omega.IntList)
(ys : Lean.Omega.IntList)
(i : Nat)
(z : Int)
 :
theorem
Lean.Omega.IntList.dot_distrib_left
(xs : Lean.Omega.IntList)
(ys : Lean.Omega.IntList)
(zs : Lean.Omega.IntList)
 :
@[simp]
@[simp]
theorem
Lean.Omega.IntList.dot_smul_left
(xs : Lean.Omega.IntList)
(ys : Lean.Omega.IntList)
(i : Int)
 :
theorem
Lean.Omega.IntList.dot_of_left_zero
{xs : Lean.Omega.IntList}
{ys : Lean.Omega.IntList}
(w : ∀ (x : Int), x ∈ xs → x = 0)
 :
xs.dot ys = 0
Division of an IntList by a integer.
Instances For
@[simp]
theorem
Lean.Omega.IntList.sdiv_cons
{x : Int}
{xs : List Int}
{g : Int}
 :
Lean.Omega.IntList.sdiv (x :: xs) g = x / g :: Lean.Omega.IntList.sdiv xs g
The gcd of the absolute values of the entries of an IntList.
Equations
- xs.gcd = List.foldr (fun (x : Int) (g : Nat) => x.natAbs.gcd g) 0 xs
Instances For
@[simp]
theorem
Lean.Omega.IntList.gcd_cons
{x : Int}
{xs : List Int}
 :
Lean.Omega.IntList.gcd (x :: xs) = x.natAbs.gcd (Lean.Omega.IntList.gcd xs)
theorem
Lean.Omega.IntList.gcd_cons_div_left
{x : Int}
{xs : List Int}
 :
↑(Lean.Omega.IntList.gcd (x :: xs)) ∣ x
theorem
Lean.Omega.IntList.gcd_cons_div_right
{x : Int}
{xs : List Int}
 :
Lean.Omega.IntList.gcd (x :: xs) ∣ Lean.Omega.IntList.gcd xs
theorem
Lean.Omega.IntList.gcd_cons_div_right'
{x : Int}
{xs : List Int}
 :
↑(Lean.Omega.IntList.gcd (x :: xs)) ∣ ↑(Lean.Omega.IntList.gcd xs)
theorem
Lean.Omega.IntList.dvd_gcd
(xs : Lean.Omega.IntList)
(c : Nat)
(w : ∀ {a : Int}, a ∈ xs → ↑c ∣ a)
 :
c ∣ xs.gcd
@[simp]
@[simp]
theorem
Lean.Omega.IntList.gcd_dvd_dot_left
(xs : Lean.Omega.IntList)
(ys : Lean.Omega.IntList)
 :
↑xs.gcd ∣ xs.dot ys
@[simp]
theorem
Lean.Omega.IntList.dot_eq_zero_of_left_eq_zero
{xs : Lean.Omega.IntList}
{ys : Lean.Omega.IntList}
(h : ∀ (x : Int), x ∈ xs → x = 0)
 :
xs.dot ys = 0
theorem
Lean.Omega.IntList.dot_sdiv_left
(xs : Lean.Omega.IntList)
(ys : Lean.Omega.IntList)
{d : Int}
(h : d ∣ ↑xs.gcd)
 :
theorem
Lean.Omega.IntList.bmod_length
(x : Lean.Omega.IntList)
(m : Nat)
 :
List.length (x.bmod m) ≤ List.length x
@[reducible, inline]
abbrev
Lean.Omega.IntList.bmod_dot_sub_dot_bmod
(m : Nat)
(a : Lean.Omega.IntList)
(b : Lean.Omega.IntList)
 :
The difference between the balanced mod of a dot product, and the dot product with balanced mod applied to each entry of the left factor.
Equations
- Lean.Omega.IntList.bmod_dot_sub_dot_bmod m a b = (a.dot b).bmod m - (a.bmod m).dot b
Instances For
theorem
Lean.Omega.IntList.dvd_bmod_dot_sub_dot_bmod
(m : Nat)
(xs : Lean.Omega.IntList)
(ys : Lean.Omega.IntList)
 :
↑m ∣ Lean.Omega.IntList.bmod_dot_sub_dot_bmod m xs ys