Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 12515db

Browse files
committed
feat(data/list): product of list.update_nth in terms of inverses (#9800)
Expression for the product of `l.update_nth n x` in terms of inverses instead of `take` and `drop`, assuming a group instead of a monoid.
1 parent c20f08e commit 12515db

File tree

4 files changed

+77
-17
lines changed

4 files changed

+77
-17
lines changed

src/data/list/basic.lean

Lines changed: 25 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -2463,15 +2463,14 @@ begin
24632463
exact is_unit.mul (u h (mem_cons_self h t)) (prod_is_unit (λ m mt, u m (mem_cons_of_mem h mt)))
24642464
end
24652465

2466-
-- `to_additive` chokes on the next few lemmas, so we do them by hand below
2467-
@[simp]
2466+
@[simp, to_additive]
24682467
lemma prod_take_mul_prod_drop :
24692468
∀ (L : list α) (i : ℕ), (L.take i).prod * (L.drop i).prod = L.prod
24702469
| [] i := by simp
24712470
| L 0 := by simp
24722471
| (h :: t) (n+1) := by { dsimp, rw [prod_cons, prod_cons, mul_assoc, prod_take_mul_prod_drop], }
24732472

2474-
@[simp]
2473+
@[simp, to_additive]
24752474
lemma prod_take_succ :
24762475
∀ (L : list α) (i : ℕ) (p), (L.take (i + 1)).prod = (L.take i).prod * L.nth_le i p
24772476
| [] i p := by cases p
@@ -2482,6 +2481,7 @@ lemma prod_take_succ :
24822481
lemma length_pos_of_prod_ne_one (L : list α) (h : L.prod ≠ 1) : 0 < L.length :=
24832482
by { cases L, { simp at h, cases h, }, { simp, }, }
24842483

2484+
@[to_additive]
24852485
lemma prod_update_nth : ∀ (L : list α) (n : ℕ) (a : α),
24862486
(L.update_nth n a).prod =
24872487
(L.take n).prod * (if n < L.length then a else 1) * (L.drop (n + 1)).prod
@@ -2505,6 +2505,14 @@ lemma prod_inv_reverse : ∀ (L : list α), L.prod⁻¹ = (L.map (λ x, x⁻¹))
25052505
lemma prod_reverse_noncomm : ∀ (L : list α), L.reverse.prod = (L.map (λ x, x⁻¹)).prod⁻¹ :=
25062506
by simp [prod_inv_reverse]
25072507

2508+
/-- Counterpart to `list.prod_take_succ` when we have an inverse operation -/
2509+
@[simp, to_additive /-"Counterpart to `list.sum_take_succ` when we have an negation operation"-/]
2510+
lemma prod_drop_succ :
2511+
∀ (L : list α) (i : ℕ) (p), (L.drop (i + 1)).prod = (L.nth_le i p)⁻¹ * (L.drop i).prod
2512+
| [] i p := false.elim (nat.not_lt_zero _ p)
2513+
| (x :: xs) 0 p := by simp
2514+
| (x :: xs) (i + 1) p := prod_drop_succ xs i _
2515+
25082516
end group
25092517

25102518
section comm_group
@@ -2516,21 +2524,21 @@ lemma prod_inv : ∀ (L : list α), L.prod⁻¹ = (L.map (λ x, x⁻¹)).prod
25162524
| [] := by simp
25172525
| (x :: xs) := by simp [mul_comm, prod_inv xs]
25182526

2519-
end comm_group
2520-
2521-
@[simp]
2522-
lemma sum_take_add_sum_drop [add_monoid α] :
2523-
∀ (L : list α) (i : ℕ), (L.take i).sum + (L.drop i).sum = L.sum
2524-
| [] i := by simp
2525-
| L 0 := by simp
2526-
| (h :: t) (n+1) := by { dsimp, rw [sum_cons, sum_cons, add_assoc, sum_take_add_sum_drop], }
2527+
/-- Alternative version of `list.prod_update_nth` when the list is over a group -/
2528+
@[to_additive /-"Alternative version of `list.sum_update_nth` when the list is over a group"-/]
2529+
lemma prod_update_nth' (L : list α) (n : ℕ) (a : α) :
2530+
(L.update_nth n a).prod =
2531+
L.prod * (if hn : n < L.length then (L.nth_le n hn)⁻¹ * a else 1) :=
2532+
begin
2533+
refine (prod_update_nth L n a).trans _,
2534+
split_ifs with hn hn,
2535+
{ rw [mul_comm _ a, mul_assoc a, prod_drop_succ L n hn, mul_comm _ (drop n L).prod,
2536+
← mul_assoc (take n L).prod, prod_take_mul_prod_drop, mul_comm a, mul_assoc] },
2537+
{ simp only [take_all_of_le (le_of_not_lt hn), prod_nil, mul_one,
2538+
drop_eq_nil_of_le ((le_of_not_lt hn).trans n.le_succ)] }
2539+
end
25272540

2528-
@[simp]
2529-
lemma sum_take_succ [add_monoid α] :
2530-
∀ (L : list α) (i : ℕ) (p), (L.take (i + 1)).sum = (L.take i).sum + L.nth_le i p
2531-
| [] i p := by cases p
2532-
| (h :: t) 0 _ := by simp
2533-
| (h :: t) (n+1) _ := by { dsimp, rw [sum_cons, sum_cons, sum_take_succ, add_assoc], }
2541+
end comm_group
25342542

25352543
lemma eq_of_sum_take_eq [add_left_cancel_monoid α] {L L' : list α} (h : L.length = L'.length)
25362544
(h' : ∀ i ≤ L.length, (L.take i).sum = (L'.take i).sum) : L = L' :=

src/data/list/zip.lean

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -381,4 +381,26 @@ end
381381

382382
end distrib
383383

384+
section comm_monoid
385+
386+
variables [comm_monoid α]
387+
388+
@[to_additive]
389+
lemma prod_mul_prod_eq_prod_zip_with_mul_prod_drop : ∀ (L L' : list α), L.prod * L'.prod =
390+
(zip_with (*) L L').prod * (L.drop L'.length).prod * (L'.drop L.length).prod
391+
| [] ys := by simp
392+
| xs [] := by simp
393+
| (x :: xs) (y :: ys) := begin
394+
simp only [drop, length, zip_with_cons_cons, prod_cons],
395+
rw [mul_assoc x, mul_comm xs.prod, mul_assoc y, mul_comm ys.prod,
396+
prod_mul_prod_eq_prod_zip_with_mul_prod_drop xs ys, mul_assoc, mul_assoc, mul_assoc, mul_assoc]
397+
end
398+
399+
@[to_additive]
400+
lemma prod_mul_prod_eq_prod_zip_with_of_length_eq (L L' : list α) (h : L.length = L'.length) :
401+
L.prod * L'.prod = (zip_with (*) L L').prod :=
402+
(prod_mul_prod_eq_prod_zip_with_mul_prod_drop L L').trans (by simp [h])
403+
404+
end comm_monoid
405+
384406
end list

src/data/vector/basic.lean

Lines changed: 24 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -441,6 +441,10 @@ section update_nth
441441
def update_nth (v : vector α n) (i : fin n) (a : α) : vector α n :=
442442
⟨v.1.update_nth i.1 a, by rw [list.update_nth_length, v.2]⟩
443443

444+
@[simp] lemma to_list_update_nth (v : vector α n) (i : fin n) (a : α) :
445+
(v.update_nth i a).to_list = v.to_list.update_nth i a :=
446+
rfl
447+
444448
@[simp] lemma nth_update_nth_same (v : vector α n) (i : fin n) (a : α) :
445449
(v.update_nth i a).nth i = a :=
446450
by cases v; cases i; simp [vector.update_nth, vector.nth_eq_nth_le]
@@ -454,6 +458,26 @@ lemma nth_update_nth_eq_if {v : vector α n} {i j : fin n} (a : α) :
454458
(v.update_nth i a).nth j = if i = j then a else v.nth j :=
455459
by split_ifs; try {simp *}; try {rw nth_update_nth_of_ne}; assumption
456460

461+
@[to_additive]
462+
lemma prod_update_nth [monoid α] (v : vector α n) (i : fin n) (a : α) :
463+
(v.update_nth i a).to_list.prod =
464+
(v.take i).to_list.prod * a * (v.drop (i + 1)).to_list.prod :=
465+
begin
466+
refine (list.prod_update_nth v.to_list i a).trans _,
467+
have : ↑i < v.to_list.length := lt_of_lt_of_le i.2 (le_of_eq v.2.symm),
468+
simp [this],
469+
end
470+
471+
@[to_additive]
472+
lemma prod_update_nth' [comm_group α] (v : vector α n) (i : fin n) (a : α) :
473+
(v.update_nth i a).to_list.prod =
474+
v.to_list.prod * (v.nth i)⁻¹ * a :=
475+
begin
476+
refine (list.prod_update_nth' v.to_list i a).trans _,
477+
have : ↑i < v.to_list.length := lt_of_lt_of_le i.2 (le_of_eq v.2.symm),
478+
simp [this, nth_eq_nth_le, mul_assoc],
479+
end
480+
457481
end update_nth
458482

459483
end vector

src/data/vector/zip.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -40,6 +40,12 @@ lemma zip_with_tail (x : vector α n) (y : vector β n) :
4040
(vector.zip_with f x y).tail = vector.zip_with f x.tail y.tail :=
4141
by { ext, simp [nth_tail], }
4242

43+
@[to_additive]
44+
lemma prod_mul_prod_eq_prod_zip_with [comm_monoid α] (x y : vector α n) :
45+
x.to_list.prod * y.to_list.prod = (vector.zip_with (*) x y).to_list.prod :=
46+
list.prod_mul_prod_eq_prod_zip_with_of_length_eq x.to_list y.to_list
47+
((to_list_length x).trans (to_list_length y).symm)
48+
4349
end zip_with
4450

4551
end vector

0 commit comments

Comments
 (0)