theorem
Lean.Omega.IntList.get_of_length_le
{i : Nat}
{xs : 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
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 }
Implementation of *
on IntList
.
Equations
- xs.mul ys = List.zipWith (fun (x1 x2 : Int) => x1 * x2) xs ys
Instances For
Equations
- Lean.Omega.IntList.instMul = { mul := Lean.Omega.IntList.mul }
theorem
Lean.Omega.IntList.mul_def
(xs ys : IntList)
:
xs * ys = List.zipWith (fun (x1 x2 : Int) => x1 * x2) xs ys
Equations
- Lean.Omega.IntList.instNeg = { neg := Lean.Omega.IntList.neg }
Equations
- Lean.Omega.IntList.instSub = { sub := Lean.Omega.IntList.sub }
Equations
- Lean.Omega.IntList.instHMulInt = { hMul := fun (i : Int) (xs : Lean.Omega.IntList) => xs.smul i }
The sum of the entries of an IntList
.
Equations
- xs.sum = List.foldr (fun (x1 x2 : Int) => x1 + x2) 0 xs
Instances For
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.bmod_length
(x : IntList)
(m : Nat)
:
List.length (x.bmod m) ≤ List.length x
@[reducible, inline]
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 ys : IntList)
:
↑m ∣ bmod_dot_sub_dot_bmod m xs ys