Documentation

Mathlib.Algebra.BigOperators.Intervals

Results about big operators over intervals #

We prove results about big operators over intervals.

theorem Finset.mul_prod_Ico_eq_prod_Icc {α : Type u_1} {M : Type u_2} [PartialOrder α] [CommMonoid M] {f : αM} {a b : α} [LocallyFiniteOrder α] (h : a b) :
f b * xIco a b, f x = xIcc a b, f x
theorem Finset.add_sum_Ico_eq_sum_Icc {α : Type u_1} {M : Type u_2} [PartialOrder α] [AddCommMonoid M] {f : αM} {a b : α} [LocallyFiniteOrder α] (h : a b) :
f b + xIco a b, f x = xIcc a b, f x
theorem Finset.prod_Ico_mul_eq_prod_Icc {α : Type u_1} {M : Type u_2} [PartialOrder α] [CommMonoid M] {f : αM} {a b : α} [LocallyFiniteOrder α] (h : a b) :
(∏ xIco a b, f x) * f b = xIcc a b, f x
theorem Finset.sum_Ico_add_eq_sum_Icc {α : Type u_1} {M : Type u_2} [PartialOrder α] [AddCommMonoid M] {f : αM} {a b : α} [LocallyFiniteOrder α] (h : a b) :
xIco a b, f x + f b = xIcc a b, f x
theorem Finset.mul_prod_Ioc_eq_prod_Icc {α : Type u_1} {M : Type u_2} [PartialOrder α] [CommMonoid M] {f : αM} {a b : α} [LocallyFiniteOrder α] (h : a b) :
f a * xIoc a b, f x = xIcc a b, f x
theorem Finset.add_sum_Ioc_eq_sum_Icc {α : Type u_1} {M : Type u_2} [PartialOrder α] [AddCommMonoid M] {f : αM} {a b : α} [LocallyFiniteOrder α] (h : a b) :
f a + xIoc a b, f x = xIcc a b, f x
theorem Finset.prod_Ioc_mul_eq_prod_Icc {α : Type u_1} {M : Type u_2} [PartialOrder α] [CommMonoid M] {f : αM} {a b : α} [LocallyFiniteOrder α] (h : a b) :
(∏ xIoc a b, f x) * f a = xIcc a b, f x
theorem Finset.sum_Ioc_add_eq_sum_Icc {α : Type u_1} {M : Type u_2} [PartialOrder α] [AddCommMonoid M] {f : αM} {a b : α} [LocallyFiniteOrder α] (h : a b) :
xIoc a b, f x + f a = xIcc a b, f x
theorem Finset.mul_prod_Ioi_eq_prod_Ici {α : Type u_1} {M : Type u_2} [PartialOrder α] [CommMonoid M] {f : αM} [LocallyFiniteOrderTop α] (a : α) :
f a * xIoi a, f x = xIci a, f x
theorem Finset.add_sum_Ioi_eq_sum_Ici {α : Type u_1} {M : Type u_2} [PartialOrder α] [AddCommMonoid M] {f : αM} [LocallyFiniteOrderTop α] (a : α) :
f a + xIoi a, f x = xIci a, f x
theorem Finset.prod_Ioi_mul_eq_prod_Ici {α : Type u_1} {M : Type u_2} [PartialOrder α] [CommMonoid M] {f : αM} [LocallyFiniteOrderTop α] (a : α) :
(∏ xIoi a, f x) * f a = xIci a, f x
theorem Finset.sum_Ioi_add_eq_sum_Ici {α : Type u_1} {M : Type u_2} [PartialOrder α] [AddCommMonoid M] {f : αM} [LocallyFiniteOrderTop α] (a : α) :
xIoi a, f x + f a = xIci a, f x
theorem Finset.mul_prod_Iio_eq_prod_Iic {α : Type u_1} {M : Type u_2} [PartialOrder α] [CommMonoid M] {f : αM} [LocallyFiniteOrderBot α] (a : α) :
f a * xIio a, f x = xIic a, f x
theorem Finset.add_sum_Iio_eq_sum_Iic {α : Type u_1} {M : Type u_2} [PartialOrder α] [AddCommMonoid M] {f : αM} [LocallyFiniteOrderBot α] (a : α) :
f a + xIio a, f x = xIic a, f x
theorem Finset.prod_Iio_mul_eq_prod_Iic {α : Type u_1} {M : Type u_2} [PartialOrder α] [CommMonoid M] {f : αM} [LocallyFiniteOrderBot α] (a : α) :
(∏ xIio a, f x) * f a = xIic a, f x
theorem Finset.sum_Iio_add_eq_sum_Iic {α : Type u_1} {M : Type u_2} [PartialOrder α] [AddCommMonoid M] {f : αM} [LocallyFiniteOrderBot α] (a : α) :
xIio a, f x + f a = xIic a, f x
theorem Finset.prod_prod_Ioi_mul_eq_prod_prod_off_diag {α : Type u_1} {M : Type u_2} [Fintype α] [LinearOrder α] [LocallyFiniteOrderTop α] [LocallyFiniteOrderBot α] [CommMonoid M] (f : ααM) :
i : α, jIoi i, f j i * f i j = i : α, j{i}, f j i
theorem Finset.sum_sum_Ioi_add_eq_sum_sum_off_diag {α : Type u_1} {M : Type u_2} [Fintype α] [LinearOrder α] [LocallyFiniteOrderTop α] [LocallyFiniteOrderBot α] [AddCommMonoid M] (f : ααM) :
i : α, jIoi i, (f j i + f i j) = i : α, j{i}, f j i
theorem Finset.prod_Ico_add' {α : Type u_1} {M : Type u_2} [CommMonoid M] [OrderedCancelAddCommMonoid α] [ExistsAddOfLE α] [LocallyFiniteOrder α] (f : αM) (a b c : α) :
xIco a b, f (x + c) = xIco (a + c) (b + c), f x
theorem Finset.sum_Ico_add' {α : Type u_1} {M : Type u_2} [AddCommMonoid M] [OrderedCancelAddCommMonoid α] [ExistsAddOfLE α] [LocallyFiniteOrder α] (f : αM) (a b c : α) :
xIco a b, f (x + c) = xIco (a + c) (b + c), f x
theorem Finset.prod_Ico_add {α : Type u_1} {M : Type u_2} [CommMonoid M] [OrderedCancelAddCommMonoid α] [ExistsAddOfLE α] [LocallyFiniteOrder α] (f : αM) (a b c : α) :
xIco a b, f (c + x) = xIco (a + c) (b + c), f x
theorem Finset.sum_Ico_add {α : Type u_1} {M : Type u_2} [AddCommMonoid M] [OrderedCancelAddCommMonoid α] [ExistsAddOfLE α] [LocallyFiniteOrder α] (f : αM) (a b c : α) :
xIco a b, f (c + x) = xIco (a + c) (b + c), f x
@[simp]
theorem Finset.prod_Ico_add_right_sub_eq {α : Type u_1} {M : Type u_2} [CommMonoid M] {f : αM} [OrderedCancelAddCommMonoid α] [ExistsAddOfLE α] [LocallyFiniteOrder α] [Sub α] [OrderedSub α] (a b c : α) :
xIco (a + c) (b + c), f (x - c) = xIco a b, f x
@[simp]
theorem Finset.sum_Ico_add_right_sub_eq {α : Type u_1} {M : Type u_2} [AddCommMonoid M] {f : αM} [OrderedCancelAddCommMonoid α] [ExistsAddOfLE α] [LocallyFiniteOrder α] [Sub α] [OrderedSub α] (a b c : α) :
xIco (a + c) (b + c), f (x - c) = xIco a b, f x
theorem Finset.prod_Ico_succ_top {M : Type u_2} [CommMonoid M] {a b : } (hab : a b) (f : M) :
kIco a (b + 1), f k = (∏ kIco a b, f k) * f b
theorem Finset.sum_Ico_succ_top {M : Type u_2} [AddCommMonoid M] {a b : } (hab : a b) (f : M) :
kIco a (b + 1), f k = kIco a b, f k + f b
theorem Finset.prod_eq_prod_Ico_succ_bot {M : Type u_2} [CommMonoid M] {a b : } (hab : a < b) (f : M) :
kIco a b, f k = f a * kIco (a + 1) b, f k
theorem Finset.sum_eq_sum_Ico_succ_bot {M : Type u_2} [AddCommMonoid M] {a b : } (hab : a < b) (f : M) :
kIco a b, f k = f a + kIco (a + 1) b, f k
theorem Finset.prod_Ico_consecutive {M : Type u_2} [CommMonoid M] (f : M) {m n k : } (hmn : m n) (hnk : n k) :
(∏ iIco m n, f i) * iIco n k, f i = iIco m k, f i
theorem Finset.sum_Ico_consecutive {M : Type u_2} [AddCommMonoid M] (f : M) {m n k : } (hmn : m n) (hnk : n k) :
iIco m n, f i + iIco n k, f i = iIco m k, f i
theorem Finset.prod_Ioc_consecutive {M : Type u_2} [CommMonoid M] (f : M) {m n k : } (hmn : m n) (hnk : n k) :
(∏ iIoc m n, f i) * iIoc n k, f i = iIoc m k, f i
theorem Finset.sum_Ioc_consecutive {M : Type u_2} [AddCommMonoid M] (f : M) {m n k : } (hmn : m n) (hnk : n k) :
iIoc m n, f i + iIoc n k, f i = iIoc m k, f i
theorem Finset.prod_Ioc_succ_top {M : Type u_2} [CommMonoid M] {a b : } (hab : a b) (f : M) :
kIoc a (b + 1), f k = (∏ kIoc a b, f k) * f (b + 1)
theorem Finset.sum_Ioc_succ_top {M : Type u_2} [AddCommMonoid M] {a b : } (hab : a b) (f : M) :
kIoc a (b + 1), f k = kIoc a b, f k + f (b + 1)
theorem Finset.prod_Icc_succ_top {M : Type u_2} [CommMonoid M] {a b : } (hab : a b + 1) (f : M) :
kIcc a (b + 1), f k = (∏ kIcc a b, f k) * f (b + 1)
theorem Finset.sum_Icc_succ_top {M : Type u_2} [AddCommMonoid M] {a b : } (hab : a b + 1) (f : M) :
kIcc a (b + 1), f k = kIcc a b, f k + f (b + 1)
theorem Finset.prod_range_mul_prod_Ico {M : Type u_2} [CommMonoid M] (f : M) {m n : } (h : m n) :
(∏ krange m, f k) * kIco m n, f k = krange n, f k
theorem Finset.sum_range_add_sum_Ico {M : Type u_2} [AddCommMonoid M] (f : M) {m n : } (h : m n) :
krange m, f k + kIco m n, f k = krange n, f k
theorem Finset.prod_range_eq_mul_Ico {M : Type u_2} [CommMonoid M] (f : M) {n : } (hn : 0 < n) :
xrange n, f x = f 0 * xIco 1 n, f x
theorem Finset.sum_range_eq_add_Ico {M : Type u_2} [AddCommMonoid M] (f : M) {n : } (hn : 0 < n) :
xrange n, f x = f 0 + xIco 1 n, f x
theorem Finset.prod_Ico_eq_mul_inv {δ : Type u_3} [CommGroup δ] (f : δ) {m n : } (h : m n) :
kIco m n, f k = (∏ krange n, f k) * (∏ krange m, f k)⁻¹
theorem Finset.sum_Ico_eq_add_neg {δ : Type u_3} [AddCommGroup δ] (f : δ) {m n : } (h : m n) :
kIco m n, f k = krange n, f k + -krange m, f k
theorem Finset.prod_Ico_eq_div {δ : Type u_3} [CommGroup δ] (f : δ) {m n : } (h : m n) :
kIco m n, f k = (∏ krange n, f k) / krange m, f k
theorem Finset.sum_Ico_eq_sub {δ : Type u_3} [AddCommGroup δ] (f : δ) {m n : } (h : m n) :
kIco m n, f k = krange n, f k - krange m, f k
theorem Finset.prod_range_div_prod_range {α : Type u_3} [CommGroup α] {f : α} {n m : } (hnm : n m) :
(∏ krange m, f k) / krange n, f k = kfilter (fun (k : ) => n k) (range m), f k
theorem Finset.sum_range_sub_sum_range {α : Type u_3} [AddCommGroup α] {f : α} {n m : } (hnm : n m) :
krange m, f k - krange n, f k = kfilter (fun (k : ) => n k) (range m), f k
theorem Finset.sum_Ico_Ico_comm {M : Type u_3} [AddCommMonoid M] (a b : ) (f : M) :
iIco a b, jIco i b, f i j = jIco a b, iIco a (j + 1), f i j

The two ways of summing over (i, j) in the range a ≤ i ≤ j < b are equal.

theorem Finset.sum_Ico_Ico_comm' {M : Type u_3} [AddCommMonoid M] (a b : ) (f : M) :
iIco a b, jIco (i + 1) b, f i j = jIco a b, iIco a j, f i j

The two ways of summing over (i, j) in the range a ≤ i < j < b are equal.

theorem Finset.prod_Ico_eq_prod_range {M : Type u_2} [CommMonoid M] (f : M) (m n : ) :
kIco m n, f k = krange (n - m), f (m + k)
theorem Finset.sum_Ico_eq_sum_range {M : Type u_2} [AddCommMonoid M] (f : M) (m n : ) :
kIco m n, f k = krange (n - m), f (m + k)
theorem Finset.prod_Ico_reflect {M : Type u_2} [CommMonoid M] (f : M) (k : ) {m n : } (h : m n + 1) :
jIco k m, f (n - j) = jIco (n + 1 - m) (n + 1 - k), f j
theorem Finset.sum_Ico_reflect {δ : Type u_3} [AddCommMonoid δ] (f : δ) (k : ) {m n : } (h : m n + 1) :
jIco k m, f (n - j) = jIco (n + 1 - m) (n + 1 - k), f j
theorem Finset.prod_range_reflect {M : Type u_2} [CommMonoid M] (f : M) (n : ) :
jrange n, f (n - 1 - j) = jrange n, f j
theorem Finset.sum_range_reflect {δ : Type u_3} [AddCommMonoid δ] (f : δ) (n : ) :
jrange n, f (n - 1 - j) = jrange n, f j
@[simp]
theorem Finset.prod_Ico_id_eq_factorial (n : ) :
xIco 1 (n + 1), x = n.factorial
@[simp]
theorem Finset.prod_range_add_one_eq_factorial (n : ) :
xrange n, (x + 1) = n.factorial
theorem Finset.sum_range_id_mul_two (n : ) :
(∑ irange n, i) * 2 = n * (n - 1)

Gauss' summation formula

theorem Finset.sum_range_id (n : ) :
irange n, i = n * (n - 1) / 2

Gauss' summation formula

theorem Finset.prod_range_diag_flip {M : Type u_2} [CommMonoid M] (n : ) (f : M) :
mrange n, krange (m + 1), f k (m - k) = mrange n, krange (n - m), f m k
theorem Finset.sum_range_diag_flip {M : Type u_2} [AddCommMonoid M] (n : ) (f : M) :
mrange n, krange (m + 1), f k (m - k) = mrange n, krange (n - m), f m k
theorem Finset.prod_range_succ_div_prod {M : Type u_3} (f : M) {n : } [CommGroup M] :
(∏ irange (n + 1), f i) / irange n, f i = f n
theorem Finset.sum_range_succ_sub_sum {M : Type u_3} (f : M) {n : } [AddCommGroup M] :
irange (n + 1), f i - irange n, f i = f n
theorem Finset.prod_range_succ_div_top {M : Type u_3} (f : M) {n : } [CommGroup M] :
(∏ irange (n + 1), f i) / f n = irange n, f i
theorem Finset.sum_range_succ_sub_top {M : Type u_3} (f : M) {n : } [AddCommGroup M] :
irange (n + 1), f i - f n = irange n, f i
theorem Finset.prod_Ico_div_bot {M : Type u_3} (f : M) {m n : } [CommGroup M] (hmn : m < n) :
(∏ iIco m n, f i) / f m = iIco (m + 1) n, f i
theorem Finset.sum_Ico_sub_bot {M : Type u_3} (f : M) {m n : } [AddCommGroup M] (hmn : m < n) :
iIco m n, f i - f m = iIco (m + 1) n, f i
theorem Finset.prod_Ico_succ_div_top {M : Type u_3} (f : M) {m n : } [CommGroup M] (hmn : m n) :
(∏ iIco m (n + 1), f i) / f n = iIco m n, f i
theorem Finset.sum_Ico_succ_sub_top {M : Type u_3} (f : M) {m n : } [AddCommGroup M] (hmn : m n) :
iIco m (n + 1), f i - f n = iIco m n, f i