diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 17a48ead2f2e6..0000e9c39a588 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -1045,7 +1045,6 @@ theorem blocksFun_sigmaCompositionAux (a : Composition n) (b : Composition a.len rw [get_of_eq (get_splitWrtComposition _ _ _), get_drop', get_take']; rfl #align composition.blocks_fun_sigma_composition_aux Composition.blocksFun_sigmaCompositionAux -set_option linter.deprecated false in /-- Auxiliary lemma to prove that the composition of formal multilinear series is associative. Consider a composition `a` of `n` and a composition `b` of `a.length`. Grouping together some @@ -1076,7 +1075,7 @@ theorem sizeUpTo_sizeUpTo_add (a : Composition n) (b : Composition a.length) {i take (sum (take i b.blocks)) (take (sum (take (i + 1) b.blocks)) a.blocks) := by rw [take_take, min_eq_left] apply monotone_sum_take _ (Nat.le_succ _) - rw [this, nthLe_map', nthLe, get_splitWrtComposition, ← + rw [this, get_map, get_splitWrtComposition, ← take_append_drop (sum (take i b.blocks)) (take (sum (take (Nat.succ i) b.blocks)) a.blocks), sum_append] congr diff --git a/Mathlib/Combinatorics/Composition.lean b/Mathlib/Combinatorics/Composition.lean index 1db57d92972f7..84d5d0f4979d9 100644 --- a/Mathlib/Combinatorics/Composition.lean +++ b/Mathlib/Combinatorics/Composition.lean @@ -225,8 +225,6 @@ theorem sizeUpTo_succ {i : ℕ} (h : i < c.length) : c.sizeUpTo (i + 1) = c.sizeUpTo i + c.blocks.get ⟨i, h⟩ := by simp only [sizeUpTo] rw [sum_take_succ _ _ h] - -- Porting note: didn't used to need `rfl` - rfl #align composition.size_up_to_succ Composition.sizeUpTo_succ theorem sizeUpTo_succ' (i : Fin c.length) : diff --git a/Mathlib/Data/List/BigOperators/Basic.lean b/Mathlib/Data/List/BigOperators/Basic.lean index 65bc366cd25c6..1c71ebd5a4233 100644 --- a/Mathlib/Data/List/BigOperators/Basic.lean +++ b/Mathlib/Data/List/BigOperators/Basic.lean @@ -170,14 +170,12 @@ theorem prod_take_mul_prod_drop : ∀ (L : List M) (i : ℕ), (L.take i).prod * @[to_additive (attr := simp)] theorem prod_take_succ : - ∀ (L : List M) (i : ℕ) (p), (L.take (i + 1)).prod = (L.take i).prod * L.nthLe i p + ∀ (L : List M) (i : ℕ) (p), (L.take (i + 1)).prod = (L.take i).prod * L.get ⟨i, p⟩ | [], i, p => by cases p | h :: t, 0, _ => rfl | h :: t, n + 1, p => by dsimp - rw [prod_cons, prod_cons, prod_take_succ t n (Nat.lt_of_succ_lt_succ p), mul_assoc, - nthLe_cons, dif_neg (Nat.add_one_ne_zero _)] - simp + rw [prod_cons, prod_cons, prod_take_succ t n (Nat.lt_of_succ_lt_succ p), mul_assoc] #align list.prod_take_succ List.prod_take_succ #align list.sum_take_succ List.sum_take_succ @@ -422,14 +420,13 @@ theorem prod_reverse_noncomm : ∀ L : List G, L.reverse.prod = (L.map fun x => #align list.prod_reverse_noncomm List.prod_reverse_noncomm #align list.sum_reverse_noncomm List.sum_reverse_noncomm -set_option linter.deprecated false in /-- Counterpart to `List.prod_take_succ` when we have an inverse operation -/ @[to_additive (attr := simp) "Counterpart to `List.sum_take_succ` when we have a negation operation"] theorem prod_drop_succ : - ∀ (L : List G) (i : ℕ) (p), (L.drop (i + 1)).prod = (L.nthLe i p)⁻¹ * (L.drop i).prod + ∀ (L : List G) (i : ℕ) (p), (L.drop (i + 1)).prod = (L.get ⟨i, p⟩)⁻¹ * (L.drop i).prod | [], i, p => False.elim (Nat.not_lt_zero _ p) - | x :: xs, 0, _ => by simp [nthLe] + | x :: xs, 0, _ => by simp | x :: xs, i + 1, p => prod_drop_succ xs i _ #align list.prod_drop_succ List.prod_drop_succ #align list.sum_drop_succ List.sum_drop_succ @@ -451,7 +448,7 @@ theorem prod_inv : ∀ L : List G, L.prod⁻¹ = (L.map fun x => x⁻¹).prod /-- Alternative version of `List.prod_set` when the list is over a group -/ @[to_additive "Alternative version of `List.sum_set` when the list is over a group"] theorem prod_set' (L : List G) (n : ℕ) (a : G) : - (L.set n a).prod = L.prod * if hn : n < L.length then (L.nthLe n hn)⁻¹ * a else 1 := by + (L.set n a).prod = L.prod * if hn : n < L.length then (L.get ⟨n, hn⟩)⁻¹ * a else 1 := by refine (prod_set L n a).trans ?_ split_ifs with hn · rw [mul_comm _ a, mul_assoc a, prod_drop_succ L n hn, mul_comm _ (drop n L).prod, ← diff --git a/Mathlib/Data/List/Join.lean b/Mathlib/Data/List/Join.lean index c00207fa48c82..49652b41e51db 100644 --- a/Mathlib/Data/List/Join.lean +++ b/Mathlib/Data/List/Join.lean @@ -153,11 +153,10 @@ theorem drop_take_succ_join_eq_nthLe (L : List (List α)) {i : ℕ} (hi : i < L. simp [take_sum_join, this, drop_sum_join, drop_take_succ_eq_cons_nthLe _ hi] #align list.drop_take_succ_join_eq_nth_le List.drop_take_succ_join_eq_nthLe -set_option linter.deprecated false in /-- Auxiliary lemma to control elements in a join. -/ @[deprecated] theorem sum_take_map_length_lt1 (L : List (List α)) {i j : ℕ} (hi : i < L.length) - (hj : j < (nthLe L i hi).length) : + (hj : j < (L.get ⟨i, hi⟩).length) : ((L.map length).take i).sum + j < ((L.map length).take (i + 1)).sum := by simp [hi, sum_take_succ, hj] #align list.sum_take_map_length_lt1 List.sum_take_map_length_lt1