@@ -2463,15 +2463,14 @@ begin
2463
2463
exact is_unit.mul (u h (mem_cons_self h t)) (prod_is_unit (λ m mt, u m (mem_cons_of_mem h mt)))
2464
2464
end
2465
2465
2466
- -- `to_additive` chokes on the next few lemmas, so we do them by hand below
2467
- @[simp]
2466
+ @[simp, to_additive]
2468
2467
lemma prod_take_mul_prod_drop :
2469
2468
∀ (L : list α) (i : ℕ), (L.take i).prod * (L.drop i).prod = L.prod
2470
2469
| [] i := by simp
2471
2470
| L 0 := by simp
2472
2471
| (h :: t) (n+1 ) := by { dsimp, rw [prod_cons, prod_cons, mul_assoc, prod_take_mul_prod_drop], }
2473
2472
2474
- @[simp]
2473
+ @[simp, to_additive ]
2475
2474
lemma prod_take_succ :
2476
2475
∀ (L : list α) (i : ℕ) (p), (L.take (i + 1 )).prod = (L.take i).prod * L.nth_le i p
2477
2476
| [] i p := by cases p
@@ -2482,6 +2481,7 @@ lemma prod_take_succ :
2482
2481
lemma length_pos_of_prod_ne_one (L : list α) (h : L.prod ≠ 1 ) : 0 < L.length :=
2483
2482
by { cases L, { simp at h, cases h, }, { simp, }, }
2484
2483
2484
+ @[to_additive]
2485
2485
lemma prod_update_nth : ∀ (L : list α) (n : ℕ) (a : α),
2486
2486
(L.update_nth n a).prod =
2487
2487
(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⁻¹))
2505
2505
lemma prod_reverse_noncomm : ∀ (L : list α), L.reverse.prod = (L.map (λ x, x⁻¹)).prod⁻¹ :=
2506
2506
by simp [prod_inv_reverse]
2507
2507
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
+
2508
2516
end group
2509
2517
2510
2518
section comm_group
@@ -2516,21 +2524,21 @@ lemma prod_inv : ∀ (L : list α), L.prod⁻¹ = (L.map (λ x, x⁻¹)).prod
2516
2524
| [] := by simp
2517
2525
| (x :: xs) := by simp [mul_comm, prod_inv xs]
2518
2526
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
2527
2540
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
2534
2542
2535
2543
lemma eq_of_sum_take_eq [add_left_cancel_monoid α] {L L' : list α} (h : L.length = L'.length)
2536
2544
(h' : ∀ i ≤ L.length, (L.take i).sum = (L'.take i).sum) : L = L' :=
0 commit comments