diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index c6b71e7d575b2..3ed82f9bb48f7 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -259,6 +259,18 @@ by rw [finset.prod, ← multiset.coe_prod, ← multiset.coe_map, finset.coe_to_l end to_list +@[to_additive] +lemma _root_.equiv.perm.prod_comp (σ : equiv.perm α) (s : finset α) (f : α → β) + (hs : {a | σ a ≠ a} ⊆ s) : + (∏ x in s, f (σ x)) = ∏ x in s, f x := +by { convert (prod_map _ σ.to_embedding _).symm, exact (map_perm hs).symm } + +@[to_additive] +lemma _root_.equiv.perm.prod_comp' (σ : equiv.perm α) (s : finset α) (f : α → α → β) + (hs : {a | σ a ≠ a} ⊆ s) : + (∏ x in s, f (σ x) x) = ∏ x in s, f x (σ.symm x) := +by { convert σ.prod_comp s (λ x, f x (σ.symm x)) hs, ext, rw equiv.symm_apply_apply } + end comm_monoid end finset diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 8840fd57b313b..ba99c2cd7a052 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -1890,6 +1890,20 @@ mem_map.trans $ by simp only [exists_prop]; refl b ∈ s.map f.to_embedding ↔ f.symm b ∈ s := by { rw mem_map, exact ⟨by { rintro ⟨a, H, rfl⟩, simpa }, λ h, ⟨_, h, by simp⟩⟩ } +/-- If the only elements outside `s` are those left fixed by `σ`, then mapping by `σ` has no effect. +-/ +lemma map_perm {σ : equiv.perm α} (hs : {a | σ a ≠ a} ⊆ s) : s.map (σ : α ↪ α) = s := +begin + ext i, + rw mem_map, + obtain hi | hi := eq_or_ne (σ i) i, + { refine ⟨_, λ h, ⟨i, h, hi⟩⟩, + rintro ⟨j, hj, h⟩, + rwa σ.injective (hi.trans h.symm) }, + { refine iff_of_true ⟨σ.symm i, hs $ λ h, hi _, σ.apply_symm_apply _⟩ (hs hi), + convert congr_arg σ h; exact (σ.apply_symm_apply _).symm } +end + theorem mem_map' (f : α ↪ β) {a} {s : finset α} : f a ∈ s.map f ↔ a ∈ s := mem_map_of_injective f.2 diff --git a/src/data/finset/card.lean b/src/data/finset/card.lean index d083427cfbd3f..ffd13c0b7bb32 100644 --- a/src/data/finset/card.lean +++ b/src/data/finset/card.lean @@ -183,6 +183,9 @@ card_le_of_subset $ filter_subset _ _ lemma eq_of_subset_of_card_le {s t : finset α} (h : s ⊆ t) (h₂ : t.card ≤ s.card) : s = t := eq_of_veq $ multiset.eq_of_le_of_card_le (val_le_iff.mpr h) h₂ +lemma map_eq_of_subset {f : α ↪ α} (hs : s.map f ⊆ s) : s.map f = s := +eq_of_subset_of_card_le hs (card_map _).ge + lemma filter_card_eq {p : α → Prop} [decidable_pred p] (h : (s.filter p).card = s.card) (x : α) (hx : x ∈ s) : p x := diff --git a/src/linear_algebra/matrix/determinant.lean b/src/linear_algebra/matrix/determinant.lean index 620c4cfe8cdea..32c5f3f10d9e8 100644 --- a/src/linear_algebra/matrix/determinant.lean +++ b/src/linear_algebra/matrix/determinant.lean @@ -156,7 +156,7 @@ calc det (M ⬝ N) = ∑ p : n → n, ∑ σ : perm n, ε σ * ∏ i, (M (σ i) sum_congr rfl (λ σ _, fintype.sum_equiv (equiv.mul_right σ⁻¹) _ _ (λ τ, have ∏ j, M (τ j) (σ j) = ∏ j, M ((τ * σ⁻¹) j) j, - by { rw ← σ⁻¹.prod_comp, simp only [equiv.perm.coe_mul, apply_inv_self] }, + by { rw ← (σ⁻¹ : _ ≃ _).prod_comp, simp only [equiv.perm.coe_mul, apply_inv_self] }, have h : ε σ * ε (τ * σ⁻¹) = ε τ := calc ε σ * ε (τ * σ⁻¹) = ε ((τ * σ⁻¹) * σ) : by { rw [mul_comm, sign_mul (τ * σ⁻¹)], simp only [int.cast_mul, units.coe_mul] } diff --git a/src/number_theory/sum_four_squares.lean b/src/number_theory/sum_four_squares.lean index 27c0e130cb56f..3683cd8f7ad18 100644 --- a/src/number_theory/sum_four_squares.lean +++ b/src/number_theory/sum_four_squares.lean @@ -99,7 +99,7 @@ let ⟨x, hx⟩ := h01 in let ⟨y, hy⟩ := h23 in ← int.sq_add_sq_of_two_mul_sq_add_sq hy.symm, ← mul_right_inj' (show (2 : ℤ) ≠ 0, from dec_trivial), ← h, mul_add, ← hx, ← hy], have : ∑ x, f (σ x)^2 = ∑ x, f x^2, - { conv_rhs { rw ← σ.sum_comp } }, + { conv_rhs { rw ←equiv.sum_comp σ } }, have fin4univ : (univ : finset (fin 4)).1 = 0 ::ₘ 1 ::ₘ 2 ::ₘ 3 ::ₘ 0, from dec_trivial, simpa [finset.sum_eq_multiset_sum, fin4univ, multiset.sum_cons, f, add_assoc] end⟩