From 205e212d2c42b0c781fc2d2929638282f7aa22fd Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Tue, 27 Feb 2024 22:30:14 +0100 Subject: [PATCH 001/106] [RFC] Complete homogeneous symmetric functions --- Mathlib/Data/Sym/Basic.lean | 39 +++++++++++++++ .../RingTheory/MvPolynomial/Symmetric.lean | 48 +++++++++++++++++-- 2 files changed, 83 insertions(+), 4 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 32ab2c914d5a2..4f43ba7ffd846 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -532,6 +532,45 @@ theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ Multiset.mem_add #align sym.mem_append_iff Sym.mem_append_iff +/-- An element of type α can be seen as an element of Sym α 1 --/ +def one {α : Type*} (i : α) : Sym α 1 := ⟨[i], by simp only [Multiset.coe_singleton, + Multiset.card_singleton]⟩ + +lemma one_injective : Function.Injective (Sym.one : α → Sym α 1) := by + intros i j h + have hi : (Sym.one i).1 = [i] := by + simp only [Sym.one, Multiset.coe_singleton] + have hj : (Sym.one j).1 = [j] := by + simp only [Sym.one, Multiset.coe_singleton] + rw [h] at hi + rw [hi] at hj + apply List.singleton_inj.mp + exact Multiset.coe_eq_singleton.mp hj + +lemma singleton_of_len_one (s : Sym α 1) : ∃! i, Sym.one i = s := by + have : ∃ i, s.toMultiset = {i} := by + simp_rw [Sym] at s + apply Multiset.card_eq_one.mp + exact s.2 + obtain ⟨i, hi⟩ := this + have : s = one i := by + simp only [Sym.one, Multiset.coe_singleton] + simp_rw [← hi] + rfl + use i + constructor + · exact this.symm + · intro j hj + apply Sym.one_injective + rw [hj] + exact this + +lemma one_bijective : Function.Bijective (Sym.one : α → Sym α 1) := by + rw [Function.bijective_iff_existsUnique] + exact Sym.singleton_of_len_one + +lemma one_equiv : α ≃ Sym α 1 := Equiv.ofBijective _ Sym.one_bijective + /-- Fill a term `m : Sym α (n - i)` with `i` copies of `a` to obtain a term of `Sym α n`. This is a convenience wrapper for `m.append (replicate i a)` that adjusts the term using `Sym.cast`. -/ diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 319ce08ad7a62..8bf355d05116b 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -5,6 +5,8 @@ Authors: Hanting Zhang, Johan Commelin -/ import Mathlib.Data.MvPolynomial.Rename import Mathlib.Data.MvPolynomial.CommRing +import Mathlib.Data.Set.Card +import Mathlib.Data.Sym.Basic import Mathlib.Algebra.Algebra.Subalgebra.Basic #align_import ring_theory.mv_polynomial.symmetric from "leanprover-community/mathlib"@"2f5b500a507264de86d666a5f87ddb976e2d8de4" @@ -29,6 +31,8 @@ We also prove some basic facts about them. + `esymm σ R n` is the `n`th elementary symmetric polynomial in `MvPolynomial σ R`. ++ `hsymm σ R n` is the `n`th complete homogeneous symmetric polynomial in `MvPolynomial σ R`. + + `psum σ R n` is the degree-`n` power sum in `MvPolynomial σ R`, i.e. the sum of monomials `(X i)^n` over `i ∈ σ`. @@ -157,12 +161,12 @@ end CommRing end IsSymmetric +variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] + section ElementarySymmetric open Finset -variable (σ R) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] - /-- The `n`th elementary symmetric `MvPolynomial σ R`. -/ def esymm (n : ℕ) : MvPolynomial σ R := ∑ t in powersetCard n univ, ∏ i in t, X i @@ -281,12 +285,48 @@ theorem degrees_esymm [Nontrivial R] (n : ℕ) (hpos : 0 < n) (hn : n ≤ Fintyp end ElementarySymmetric +section CompleteHomogeneousSymmetric + +open Classical Set + +/-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ +def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod + +lemma hsum_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl + +@[simp] +theorem hsymm_zero : hsymm σ R 0 = 1 := by + simp only [hsymm, Finset.univ_unique, Sym.eq_nil_of_card_zero, Sym.val_eq_coe, Sym.coe_nil, + Multiset.map_zero, Multiset.prod_zero, Finset.sum_const, Finset.card_singleton, one_smul] + +@[simp] +theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by + simp only [hsymm, Finset.univ_unique] + symm + apply Fintype.sum_equiv (Equiv.ofBijective _ Sym.one_bijective) + intro i + have : (Sym.one i).toMultiset = [i] := by rfl + simp only [Equiv.ofBijective_apply, Sym.val_eq_coe, this, Multiset.coe_singleton, + Multiset.map_singleton, Multiset.prod_singleton] + +theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by + simp_rw [hsymm, map_sum, ←Multiset.prod_hom'] + simp only [Sym.val_eq_coe, map_X] + +theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by + simp_rw [hsymm, map_sum, ← Multiset.prod_hom', rename_X] + apply Fintype.sum_equiv (Sym.equivCongr e) + simp only [Sym.val_eq_coe, Sym.equivCongr_apply, Sym.coe_map, Multiset.map_map, + Function.comp_apply, implies_true] + +theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n + +end CompleteHomogeneousSymmetric + section PowerSum open Finset -variable (σ R) [CommSemiring R] [Fintype σ] [Fintype τ] - /-- The degree-`n` power sum -/ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n From 357abb59309935da1c70ea5a142ad50030212768 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Tue, 27 Feb 2024 22:35:06 +0100 Subject: [PATCH 002/106] Linting --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 8bf355d05116b..65d45b04ce8fe 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -310,7 +310,7 @@ theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by Multiset.map_singleton, Multiset.prod_singleton] theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by - simp_rw [hsymm, map_sum, ←Multiset.prod_hom'] + simp_rw [hsymm, map_sum, ← Multiset.prod_hom'] simp only [Sym.val_eq_coe, map_X] theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by From ff506b76ba0be27974565aa82d5da0bf0d945e3b Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 28 Feb 2024 18:29:40 +0100 Subject: [PATCH 003/106] Implemented suggestions --- Mathlib/Data/Sym/Basic.lean | 54 +++++++------------ .../RingTheory/MvPolynomial/Symmetric.lean | 32 ++++++----- 2 files changed, 37 insertions(+), 49 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 4f43ba7ffd846..299284fcbe724 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -533,43 +533,27 @@ theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ #align sym.mem_append_iff Sym.mem_append_iff /-- An element of type α can be seen as an element of Sym α 1 --/ -def one {α : Type*} (i : α) : Sym α 1 := ⟨[i], by simp only [Multiset.coe_singleton, +def one (i : α) : Sym α 1 := ⟨[i], by simp only [Multiset.coe_singleton, Multiset.card_singleton]⟩ -lemma one_injective : Function.Injective (Sym.one : α → Sym α 1) := by - intros i j h - have hi : (Sym.one i).1 = [i] := by - simp only [Sym.one, Multiset.coe_singleton] - have hj : (Sym.one j).1 = [j] := by - simp only [Sym.one, Multiset.coe_singleton] - rw [h] at hi - rw [hi] at hj - apply List.singleton_inj.mp - exact Multiset.coe_eq_singleton.mp hj - -lemma singleton_of_len_one (s : Sym α 1) : ∃! i, Sym.one i = s := by - have : ∃ i, s.toMultiset = {i} := by - simp_rw [Sym] at s - apply Multiset.card_eq_one.mp - exact s.2 - obtain ⟨i, hi⟩ := this - have : s = one i := by - simp only [Sym.one, Multiset.coe_singleton] - simp_rw [← hi] - rfl - use i - constructor - · exact this.symm - · intro j hj - apply Sym.one_injective - rw [hj] - exact this - -lemma one_bijective : Function.Bijective (Sym.one : α → Sym α 1) := by - rw [Function.bijective_iff_existsUnique] - exact Sym.singleton_of_len_one - -lemma one_equiv : α ≃ Sym α 1 := Equiv.ofBijective _ Sym.one_bijective +lemma one_coe (i : α) : (one i : Multiset α) = {i} := rfl + +variable (α) in +lemma one_bijective : Function.Bijective (one : α → Sym α 1) := by + refine ⟨fun _ _ h ↦ Multiset.singleton_inj.1 <| by simp [← one_coe, h], fun x ↦ ?_⟩ + obtain ⟨i, hi⟩ := Multiset.card_eq_one.1 x.2 + exact ⟨i, ext <| by simp [one_coe, ← hi]⟩ + +open Multiset Classical in +@[simps apply] +noncomputable def one_equiv : α ≃ Sym α 1 where + toFun := fun i ↦ ⟨[i], by simp only [coe_singleton, card_singleton]⟩ + invFun := fun x ↦ choose <| card_eq_one.1 x.2 + left_inv := fun i ↦ by simpa using (choose_spec <| card_eq_one.1 + (⟨[i], by simp only [coe_singleton, card_singleton]⟩ : Sym α 1).2).symm + right_inv := fun x ↦ by + simp only [val_eq_coe, coe_singleton] + exact ext <| by simp_rw [← choose_spec <| card_eq_one.1 x.2]; rfl /-- Fill a term `m : Sym α (n - i)` with `i` copies of `a` to obtain a term of `Sym α n`. This is a convenience wrapper for `m.append (replicate i a)` that adjusts the term using diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 65d45b04ce8fe..37d7722f60201 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -5,9 +5,9 @@ Authors: Hanting Zhang, Johan Commelin -/ import Mathlib.Data.MvPolynomial.Rename import Mathlib.Data.MvPolynomial.CommRing -import Mathlib.Data.Set.Card import Mathlib.Data.Sym.Basic import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.Combinatorics.Partition #align_import ring_theory.mv_polynomial.symmetric from "leanprover-community/mathlib"@"2f5b500a507264de86d666a5f87ddb976e2d8de4" @@ -201,6 +201,10 @@ theorem esymm_zero : esymm σ R 0 = 1 := by simp only [esymm, powersetCard_zero, sum_singleton, prod_empty] #align mv_polynomial.esymm_zero MvPolynomial.esymm_zero +@[simp] +theorem esymm_one : esymm σ R 1 = ∑ i, X i := by + simp only [esymm, powersetCard_one, sum_map, Function.Embedding.coeFn_mk, prod_singleton] + theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by simp_rw [esymm, map_sum, map_prod, map_X] #align mv_polynomial.map_esymm MvPolynomial.map_esymm @@ -287,37 +291,37 @@ end ElementarySymmetric section CompleteHomogeneousSymmetric -open Classical Set +open Finset Multiset Sym + +variable [DecidableEq σ] [DecidableEq τ] /-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod -lemma hsum_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl +lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl @[simp] theorem hsymm_zero : hsymm σ R 0 = 1 := by - simp only [hsymm, Finset.univ_unique, Sym.eq_nil_of_card_zero, Sym.val_eq_coe, Sym.coe_nil, - Multiset.map_zero, Multiset.prod_zero, Finset.sum_const, Finset.card_singleton, one_smul] + simp only [hsymm, univ_unique, eq_nil_of_card_zero, val_eq_coe, Sym.coe_nil, Multiset.map_zero, + prod_zero, sum_const, Finset.card_singleton, one_smul] @[simp] theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by - simp only [hsymm, Finset.univ_unique] + simp only [hsymm, univ_unique] symm - apply Fintype.sum_equiv (Equiv.ofBijective _ Sym.one_bijective) + apply Fintype.sum_equiv (Equiv.ofBijective _ (one_bijective σ)) intro i - have : (Sym.one i).toMultiset = [i] := by rfl - simp only [Equiv.ofBijective_apply, Sym.val_eq_coe, this, Multiset.coe_singleton, + simp only [Equiv.ofBijective_apply, val_eq_coe, one_coe, Multiset.coe_singleton, Multiset.map_singleton, Multiset.prod_singleton] theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp_rw [hsymm, map_sum, ← Multiset.prod_hom'] - simp only [Sym.val_eq_coe, map_X] + simp only [val_eq_coe, map_X] theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by - simp_rw [hsymm, map_sum, ← Multiset.prod_hom', rename_X] - apply Fintype.sum_equiv (Sym.equivCongr e) - simp only [Sym.val_eq_coe, Sym.equivCongr_apply, Sym.coe_map, Multiset.map_map, - Function.comp_apply, implies_true] + simp_rw [hsymm, map_sum, ← prod_hom', rename_X] + apply Fintype.sum_equiv (equivCongr e) + simp only [val_eq_coe, equivCongr_apply, Sym.coe_map, Multiset.map_map, Function.comp_apply, implies_true] theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n From 710e4e077c72fc1072423f27de8fe0e576e5c51e Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 28 Feb 2024 18:37:21 +0100 Subject: [PATCH 004/106] Minor. --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 37d7722f60201..3ed4060e20c27 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -321,7 +321,8 @@ theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by simp_rw [hsymm, map_sum, ← prod_hom', rename_X] apply Fintype.sum_equiv (equivCongr e) - simp only [val_eq_coe, equivCongr_apply, Sym.coe_map, Multiset.map_map, Function.comp_apply, implies_true] + simp only [val_eq_coe, equivCongr_apply, Sym.coe_map, Multiset.map_map, Function.comp_apply, + implies_true] theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n From 9b15988a855dac84d3aadaee8293850cdb2a0670 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 28 Feb 2024 19:18:27 +0100 Subject: [PATCH 005/106] Simplifications --- Mathlib/Data/Sym/Basic.lean | 23 ++++--------------- .../RingTheory/MvPolynomial/Symmetric.lean | 5 ++-- 2 files changed, 6 insertions(+), 22 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 299284fcbe724..53926d7f57610 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -532,28 +532,13 @@ theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ Multiset.mem_add #align sym.mem_append_iff Sym.mem_append_iff -/-- An element of type α can be seen as an element of Sym α 1 --/ -def one (i : α) : Sym α 1 := ⟨[i], by simp only [Multiset.coe_singleton, - Multiset.card_singleton]⟩ - -lemma one_coe (i : α) : (one i : Multiset α) = {i} := rfl - -variable (α) in -lemma one_bijective : Function.Bijective (one : α → Sym α 1) := by - refine ⟨fun _ _ h ↦ Multiset.singleton_inj.1 <| by simp [← one_coe, h], fun x ↦ ?_⟩ - obtain ⟨i, hi⟩ := Multiset.card_eq_one.1 x.2 - exact ⟨i, ext <| by simp [one_coe, ← hi]⟩ - open Multiset Classical in @[simps apply] noncomputable def one_equiv : α ≃ Sym α 1 where - toFun := fun i ↦ ⟨[i], by simp only [coe_singleton, card_singleton]⟩ - invFun := fun x ↦ choose <| card_eq_one.1 x.2 - left_inv := fun i ↦ by simpa using (choose_spec <| card_eq_one.1 - (⟨[i], by simp only [coe_singleton, card_singleton]⟩ : Sym α 1).2).symm - right_inv := fun x ↦ by - simp only [val_eq_coe, coe_singleton] - exact ext <| by simp_rw [← choose_spec <| card_eq_one.1 x.2]; rfl + toFun := fun i ↦ ⟨{i}, by simp⟩ + invFun := fun x ↦ (card_eq_one.1 x.2).choose + left_inv := fun i ↦ by simpa using (card_eq_one.1 (⟨{i}, by simp⟩ : Sym α 1).2).choose_spec.symm + right_inv := fun x ↦ ext <| by simp_rw [← choose_spec <| card_eq_one.1 x.2]; rfl /-- Fill a term `m : Sym α (n - i)` with `i` copies of `a` to obtain a term of `Sym α n`. This is a convenience wrapper for `m.append (replicate i a)` that adjusts the term using diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 3ed4060e20c27..e54d419151d8d 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -309,10 +309,9 @@ theorem hsymm_zero : hsymm σ R 0 = 1 := by theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by simp only [hsymm, univ_unique] symm - apply Fintype.sum_equiv (Equiv.ofBijective _ (one_bijective σ)) + apply Fintype.sum_equiv one_equiv intro i - simp only [Equiv.ofBijective_apply, val_eq_coe, one_coe, Multiset.coe_singleton, - Multiset.map_singleton, Multiset.prod_singleton] + simp only [one_equiv_apply, Multiset.map_singleton, Multiset.prod_singleton] theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp_rw [hsymm, map_sum, ← Multiset.prod_hom'] From d4936f38e94ea73bee1109dd2f7b72c3a3564195 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 28 Feb 2024 21:06:19 +0100 Subject: [PATCH 006/106] Bases --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index e54d419151d8d..d82ccc9e60b06 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -172,6 +172,11 @@ def esymm (n : ℕ) : MvPolynomial σ R := ∑ t in powersetCard n univ, ∏ i in t, X i #align mv_polynomial.esymm MvPolynomial.esymm +lemma esymm_def (n : ℕ) : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl + +def esymm_mu {n : ℕ} (μ : Nat.Partition n) : MvPolynomial σ R := + Multiset.prod (μ.parts.map (esymm σ R)) + /-- The `n`th elementary symmetric `MvPolynomial σ R` is obtained by evaluating the `n`th elementary symmetric at the `Multiset` of the monomials -/ theorem esymm_eq_multiset_esymm : esymm σ R = (Finset.univ.val.map X).esymm := by @@ -300,6 +305,9 @@ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl +def hsymm_mu {n : ℕ} (μ : Nat.Partition n) : MvPolynomial σ R := + Multiset.prod (μ.parts.map (hsymm σ R)) + @[simp] theorem hsymm_zero : hsymm σ R 0 = 1 := by simp only [hsymm, univ_unique, eq_nil_of_card_zero, val_eq_coe, Sym.coe_nil, Multiset.map_zero, @@ -336,6 +344,9 @@ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl +def psum_mu {n : ℕ} (μ : Nat.Partition n) : MvPolynomial σ R := + Multiset.prod (μ.parts.map (psum σ R)) + @[simp] theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp only [psum, _root_.pow_zero, ← cast_card] From 6ddb1efa3fdc59c205310b8abfcc815aa0a56e1d Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 28 Feb 2024 23:04:01 +0100 Subject: [PATCH 007/106] Computability of oneEquiv --- Mathlib/Data/Sym/Basic.lean | 16 ++++++++++------ Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 4 ++-- 2 files changed, 12 insertions(+), 8 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 53926d7f57610..b7294efc2b529 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -532,13 +532,17 @@ theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ Multiset.mem_add #align sym.mem_append_iff Sym.mem_append_iff -open Multiset Classical in @[simps apply] -noncomputable def one_equiv : α ≃ Sym α 1 where - toFun := fun i ↦ ⟨{i}, by simp⟩ - invFun := fun x ↦ (card_eq_one.1 x.2).choose - left_inv := fun i ↦ by simpa using (card_eq_one.1 (⟨{i}, by simp⟩ : Sym α 1).2).choose_spec.symm - right_inv := fun x ↦ ext <| by simp_rw [← choose_spec <| card_eq_one.1 x.2]; rfl +def oneEquiv : α ≃ Sym α 1 where + toFun a := ⟨{a}, by simp⟩ + invFun s := Quotient.liftOn (Equiv.subtypeQuotientEquivQuotientSubtype + (·.length = 1) _ (fun l ↦ Iff.rfl) (fun l l' ↦ by rfl) s) + (fun l ↦ l.1.head <| List.length_pos.mp <| by simp) + fun ⟨l, h⟩ ⟨l', h'⟩ _perm ↦ by + obtain ⟨a, rfl⟩ := List.length_eq_one.mp h' + exact List.eq_of_mem_singleton (_perm.mem_iff.mp <| List.head_mem _) + left_inv a := by rfl + right_inv := by rintro ⟨⟨l⟩, h⟩; obtain ⟨a, rfl⟩ := List.length_eq_one.mp h; rfl /-- Fill a term `m : Sym α (n - i)` with `i` copies of `a` to obtain a term of `Sym α n`. This is a convenience wrapper for `m.append (replicate i a)` that adjusts the term using diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index d82ccc9e60b06..099ac96301f3a 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -317,9 +317,9 @@ theorem hsymm_zero : hsymm σ R 0 = 1 := by theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by simp only [hsymm, univ_unique] symm - apply Fintype.sum_equiv one_equiv + apply Fintype.sum_equiv oneEquiv intro i - simp only [one_equiv_apply, Multiset.map_singleton, Multiset.prod_singleton] + simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton] theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp_rw [hsymm, map_sum, ← Multiset.prod_hom'] From df119fa81a53df870314d929a21d9746588f00d5 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 28 Feb 2024 23:53:09 +0100 Subject: [PATCH 008/106] Minor --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 099ac96301f3a..d25d0d593fc40 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -25,6 +25,8 @@ We also prove some basic facts about them. * `MvPolynomial.esymm` +* `MvPolynomial.hsymm` + * `MvPolynomial.psum` ## Notation From 56c0091a6bf5d60ccf0c4f9d3450eacf90f0bafc Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Tue, 5 Mar 2024 18:11:40 +0100 Subject: [PATCH 009/106] Simp lemmas about products --- .../RingTheory/MvPolynomial/Symmetric.lean | 62 +++++++++++++++++-- 1 file changed, 56 insertions(+), 6 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index d25d0d593fc40..cc6dce0ad85e4 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -165,6 +165,22 @@ end IsSymmetric variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] +/- Multiplicativity on partitions -/ +def muProduct {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : MvPolynomial σ R := + Multiset.prod (μ.parts.map f) + +lemma muProduct_def {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : + muProduct σ R f μ = Multiset.prod (μ.parts.map f) := rfl + +@[simp] +theorem muProduct_zero (f : ℕ → MvPolynomial σ R) : muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by + simp only [muProduct, Nat.Partition.partition_zero_parts, Multiset.map_zero, Multiset.prod_zero] + +@[simp] +theorem muProduct_onePart (n : ℕ) (f : ℕ → MvPolynomial σ R) (npos : n > 0): muProduct σ R f (Nat.Partition.indiscrete n) = f n := by + rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] + linarith + section ElementarySymmetric open Finset @@ -176,8 +192,8 @@ def esymm (n : ℕ) : MvPolynomial σ R := lemma esymm_def (n : ℕ) : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl -def esymm_mu {n : ℕ} (μ : Nat.Partition n) : MvPolynomial σ R := - Multiset.prod (μ.parts.map (esymm σ R)) +def esymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := + muProduct σ R (esymm σ R) μ /-- The `n`th elementary symmetric `MvPolynomial σ R` is obtained by evaluating the `n`th elementary symmetric at the `Multiset` of the monomials -/ @@ -212,6 +228,16 @@ theorem esymm_zero : esymm σ R 0 = 1 := by theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp only [esymm, powersetCard_one, sum_map, Function.Embedding.coeFn_mk, prod_singleton] +@[simp] +theorem esymmMu_zero : esymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by + simp only [esymmMu, esymm_zero, muProduct_zero] + +@[simp] +theorem esymmMu_onePart (n : ℕ) : esymmMu σ R (Nat.Partition.indiscrete n) = esymm σ R n := by + cases n + · simp only [esymmMu, esymm_zero, muProduct_zero] + · simp only [esymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_onePart] + theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by simp_rw [esymm, map_sum, map_prod, map_X] #align mv_polynomial.map_esymm MvPolynomial.map_esymm @@ -307,8 +333,11 @@ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl -def hsymm_mu {n : ℕ} (μ : Nat.Partition n) : MvPolynomial σ R := - Multiset.prod (μ.parts.map (hsymm σ R)) +def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := + muProduct σ R (hsymm σ R) μ + +lemma hsymmMu_def {n : ℕ} (μ : n.Partition) : hsymmMu σ R μ = + Multiset.prod (μ.parts.map (hsymm σ R)) := rfl @[simp] theorem hsymm_zero : hsymm σ R 0 = 1 := by @@ -323,6 +352,16 @@ theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by intro i simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton] +@[simp] +theorem hsymmMu_zero : hsymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by + simp only [hsymmMu, hsymm_zero, muProduct_zero] + +@[simp] +theorem hsymmMu_onePart (n : ℕ) : hsymmMu σ R (Nat.Partition.indiscrete n) = hsymm σ R n := by + cases n + · simp only [hsymmMu, hsymm_zero, muProduct_zero] + · simp only [hsymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_onePart] + theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp_rw [hsymm, map_sum, ← Multiset.prod_hom'] simp only [val_eq_coe, map_X] @@ -346,8 +385,11 @@ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl -def psum_mu {n : ℕ} (μ : Nat.Partition n) : MvPolynomial σ R := - Multiset.prod (μ.parts.map (psum σ R)) +def psumMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := + muProduct σ R (psum σ R) μ + +lemma psumMu_def {n : ℕ} (μ : n.Partition) : psumMu σ R μ = + Multiset.prod (μ.parts.map (psum σ R)) := rfl @[simp] theorem psum_zero : psum σ R 0 = Fintype.card σ := by @@ -358,6 +400,14 @@ theorem psum_zero : psum σ R 0 = Fintype.card σ := by theorem psum_one : psum σ R 1 = ∑ i, X i := by simp only [psum, _root_.pow_one] +@[simp] +theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by + simp only [psumMu, muProduct_zero] + +@[simp] +theorem psumMu_onePart (n : ℕ) (npos : n > 0) : psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by + simp only [psumMu, npos, muProduct_onePart] + @[simp] theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by simp_rw [psum, map_sum, map_pow, rename_X, e.sum_comp (X · ^ n)] From b247fc261597f6228344ea4dc3330e298dd6e231 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Tue, 5 Mar 2024 18:27:47 +0100 Subject: [PATCH 010/106] Linting Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index cc6dce0ad85e4..98f1ada0115d8 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -170,7 +170,7 @@ def muProduct {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : MvP Multiset.prod (μ.parts.map f) lemma muProduct_def {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : - muProduct σ R f μ = Multiset.prod (μ.parts.map f) := rfl + muProduct σ R f μ = Multiset.prod (μ.parts.map f) := rfl @[simp] theorem muProduct_zero (f : ℕ → MvPolynomial σ R) : muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by From 48586d861cf03c226db486cf77d1a0aa4a9cd43a Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Tue, 5 Mar 2024 18:28:06 +0100 Subject: [PATCH 011/106] Linting Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 98f1ada0115d8..0f27376479c3a 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -337,7 +337,7 @@ def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ lemma hsymmMu_def {n : ℕ} (μ : n.Partition) : hsymmMu σ R μ = - Multiset.prod (μ.parts.map (hsymm σ R)) := rfl + Multiset.prod (μ.parts.map (hsymm σ R)) := rfl @[simp] theorem hsymm_zero : hsymm σ R 0 = 1 := by From fca66143ac34b81d656e1c9baf6778df9f902c4d Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Tue, 5 Mar 2024 18:28:22 +0100 Subject: [PATCH 012/106] Linting Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 0f27376479c3a..bb5e06353ec17 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -389,7 +389,7 @@ def psumMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (psum σ R) μ lemma psumMu_def {n : ℕ} (μ : n.Partition) : psumMu σ R μ = - Multiset.prod (μ.parts.map (psum σ R)) := rfl + Multiset.prod (μ.parts.map (psum σ R)) := rfl @[simp] theorem psum_zero : psum σ R 0 = Fintype.card σ := by From 465113141aee6e2205a9e066e429309efd97e7db Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 6 Mar 2024 13:03:44 +0100 Subject: [PATCH 013/106] Linting --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 9 ++++++--- 1 file changed, 6 insertions(+), 3 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index bb5e06353ec17..669aecb2b8c09 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -173,11 +173,13 @@ lemma muProduct_def {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) muProduct σ R f μ = Multiset.prod (μ.parts.map f) := rfl @[simp] -theorem muProduct_zero (f : ℕ → MvPolynomial σ R) : muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by +theorem muProduct_zero (f : ℕ → MvPolynomial σ R) : +muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by simp only [muProduct, Nat.Partition.partition_zero_parts, Multiset.map_zero, Multiset.prod_zero] @[simp] -theorem muProduct_onePart (n : ℕ) (f : ℕ → MvPolynomial σ R) (npos : n > 0): muProduct σ R f (Nat.Partition.indiscrete n) = f n := by +theorem muProduct_onePart (n : ℕ) (f : ℕ → MvPolynomial σ R) (npos : n > 0) : +muProduct σ R f (Nat.Partition.indiscrete n) = f n := by rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] linarith @@ -405,7 +407,8 @@ theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp only [psumMu, muProduct_zero] @[simp] -theorem psumMu_onePart (n : ℕ) (npos : n > 0) : psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by +theorem psumMu_onePart (n : ℕ) (npos : n > 0) : +psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by simp only [psumMu, npos, muProduct_onePart] @[simp] From ac0a48b734a16274a18dd957bbdb43af5f8eb90d Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 6 Mar 2024 13:09:00 +0100 Subject: [PATCH 014/106] Linting Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 669aecb2b8c09..1fba8176d6558 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -174,7 +174,7 @@ lemma muProduct_def {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) @[simp] theorem muProduct_zero (f : ℕ → MvPolynomial σ R) : -muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by + muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by simp only [muProduct, Nat.Partition.partition_zero_parts, Multiset.map_zero, Multiset.prod_zero] @[simp] From 62c37cce9a8901644aff9f2ea94a1f69fabc36a9 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 6 Mar 2024 13:09:14 +0100 Subject: [PATCH 015/106] Linting Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 1fba8176d6558..e036d6e78dfca 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -179,7 +179,7 @@ theorem muProduct_zero (f : ℕ → MvPolynomial σ R) : @[simp] theorem muProduct_onePart (n : ℕ) (f : ℕ → MvPolynomial σ R) (npos : n > 0) : -muProduct σ R f (Nat.Partition.indiscrete n) = f n := by + muProduct σ R f (Nat.Partition.indiscrete n) = f n := by rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] linarith From 58ba617435ca56b998c2f1d492cdec5f5d13c689 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Wed, 6 Mar 2024 13:09:27 +0100 Subject: [PATCH 016/106] Linting Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index e036d6e78dfca..a2494981e4eb6 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -408,7 +408,7 @@ theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by @[simp] theorem psumMu_onePart (n : ℕ) (npos : n > 0) : -psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by + psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by simp only [psumMu, npos, muProduct_onePart] @[simp] From 43f0f2d8e9fe4fa40f06a524dbc4d3b74ce3caa1 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 12:17:59 +0100 Subject: [PATCH 017/106] Monomial symmetric polynomials --- Mathlib/Combinatorics/Partition.lean | 27 +++++++ Mathlib/Data/Multiset/Dedup.lean | 19 +++++ .../RingTheory/MvPolynomial/Symmetric.lean | 80 +++++++++++++++++-- 3 files changed, 120 insertions(+), 6 deletions(-) diff --git a/Mathlib/Combinatorics/Partition.lean b/Mathlib/Combinatorics/Partition.lean index 3713601fa5f64..15a3333dba75c 100644 --- a/Mathlib/Combinatorics/Partition.lean +++ b/Mathlib/Combinatorics/Partition.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ import Mathlib.Combinatorics.Composition +import Mathlib.Data.Multiset.Dedup import Mathlib.Data.Nat.Parity import Mathlib.Tactic.ApplyFun @@ -102,6 +103,32 @@ def ofSums (n : ℕ) (l : Multiset ℕ) (hl : l.sum = n) : Partition n where def ofMultiset (l : Multiset ℕ) : Partition l.sum := ofSums _ l rfl #align nat.partition.of_multiset Nat.Partition.ofMultiset +/-- An element `s` of `Sym σ n` induces a partition given by its multiplicities. +-/ +def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition where + parts := (s.1.dedup).map s.1.count + parts_pos := by simp [Multiset.count_pos] + parts_sum := by + simp only [Sym.val_eq_coe, card_eq_sum_count] + exact s.2 + +lemma ofSymEquiv {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by + simp only [Nat.Partition.ofSym, Nat.Partition.ext_iff] + have dedup_eq : (s.map e).1.dedup = s.1.dedup.map e := by + simp only [Sym.val_eq_coe, Sym.coe_map] + rw [← Multiset.dedup_map_dedup_eq] + apply Multiset.Nodup.dedup + rw [Multiset.nodup_map_iff_of_injective e.injective] + exact Multiset.nodup_dedup s.1 + rw [dedup_eq] + have : (fun x ↦ Multiset.count x s.1) = fun x ↦ Multiset.count (e x) (s.map e).1 := by + funext x + rw [← Multiset.count_map_eq_count' e] + congr + exact e.injective + rw [this] + simp only [Sym.val_eq_coe, Sym.coe_map, Multiset.map_map, Function.comp_apply] + /-- The partition of exactly one part. -/ def indiscrete (n : ℕ) : Partition n := ofSums n {n} rfl #align nat.partition.indiscrete_partition Nat.Partition.indiscrete diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 16c42af709b70..8600f32221c07 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -142,6 +142,25 @@ theorem Nodup.le_dedup_iff_le {s t : Multiset α} (hno : s.Nodup) : s ≤ t.dedu simp [le_dedup, hno] #align multiset.nodup.le_dedup_iff_le Multiset.Nodup.le_dedup_iff_le +theorem card_eq_sum_count {α : Type*} {s : Multiset α} [DecidableEq α] : sum (Multiset.map (fun x ↦ count x ↑s) (s.dedup)) = card s := by + induction' s using Multiset.induction_on with a s ih + · simp + · simp_rw [count_cons, sum_map_add, card_cons] + by_cases h : a ∈ s + · rw [dedup_cons_of_mem h] + congr + rw [Multiset.sum_map_eq_nsmul_single a] + · simp only [Multiset.nodup_dedup, ↓reduceIte, smul_eq_mul, mul_one, Multiset.mem_dedup, h, Multiset.count_eq_one_of_mem] + · simp only [ne_eq, mem_dedup, ite_eq_right_iff, one_ne_zero, imp_false] + tauto + · congr + · rw [dedup_cons_of_not_mem h] + simp only [map_cons, sum_cons, ih, count_eq_zero_of_not_mem h, zero_add] + · rw [Multiset.sum_map_eq_nsmul_single a] + · simp [h] + · simp only [ne_eq, mem_dedup, ite_eq_right_iff, one_ne_zero, imp_false] + tauto + end Multiset theorem Multiset.Nodup.le_nsmul_iff_le {α : Type*} {s t : Multiset α} {n : ℕ} (h : s.Nodup) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index a2494981e4eb6..676ca85f72128 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -3,11 +3,12 @@ Copyright (c) 2020 Hanting Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Hanting Zhang, Johan Commelin -/ +import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.Combinatorics.Partition +import Mathlib.Data.Fintype.Perm import Mathlib.Data.MvPolynomial.Rename import Mathlib.Data.MvPolynomial.CommRing import Mathlib.Data.Sym.Basic -import Mathlib.Algebra.Algebra.Subalgebra.Basic -import Mathlib.Combinatorics.Partition #align_import ring_theory.mv_polynomial.symmetric from "leanprover-community/mathlib"@"2f5b500a507264de86d666a5f87ddb976e2d8de4" @@ -166,6 +167,11 @@ end IsSymmetric variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] /- Multiplicativity on partitions -/ + +/-- + Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, + `muProduct` computes the product of applying each function in `f` to the parts of `μ`. +-/ def muProduct {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : MvPolynomial σ R := Multiset.prod (μ.parts.map f) @@ -194,6 +200,9 @@ def esymm (n : ℕ) : MvPolynomial σ R := lemma esymm_def (n : ℕ) : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl +/-- +`esymmMu` is the product of the symmetric polynomials `esymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. +-/ def esymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (esymm σ R) μ @@ -230,7 +239,6 @@ theorem esymm_zero : esymm σ R 0 = 1 := by theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp only [esymm, powersetCard_one, sum_map, Function.Embedding.coeFn_mk, prod_singleton] -@[simp] theorem esymmMu_zero : esymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp only [esymmMu, esymm_zero, muProduct_zero] @@ -335,6 +343,9 @@ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl +/-- +`hsymmMu` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. +-/ def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ @@ -354,7 +365,6 @@ theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by intro i simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton] -@[simp] theorem hsymmMu_zero : hsymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp only [hsymmMu, hsymm_zero, muProduct_zero] @@ -387,6 +397,9 @@ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl +/-- +`psumMu` is the product of the symmetric polynomials `psum μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. +-/ def psumMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (psum σ R) μ @@ -395,8 +408,8 @@ lemma psumMu_def {n : ℕ} (μ : n.Partition) : psumMu σ R μ = @[simp] theorem psum_zero : psum σ R 0 = Fintype.card σ := by - simp only [psum, _root_.pow_zero, ← cast_card] - exact rfl + simp [psum, _root_.pow_zero, ← cast_card] + rfl @[simp] theorem psum_one : psum σ R 1 = ∑ i, X i := by @@ -419,4 +432,59 @@ theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum σ R n) := rename_psum _ end PowerSum +section MonomialSymmetric +-- open Multiset + +variable [DecidableEq σ] [DecidableEq τ] + +/-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. -/ +def msymm {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := + ∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod + +lemma msymm_def {n : ℕ} (μ : n.Partition) : msymm σ R μ = ∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod := rfl + +@[simp] +theorem msymm_zero : msymm σ R (Nat.Partition.indiscrete 0) = 1 := by + rw [msymm, Fintype.sum_subsingleton] + swap + exact ⟨(Sym.nil : Sym σ 0), by rfl⟩ + simp only [Sym.val_eq_coe, Sym.coe_nil, Multiset.map_zero, Multiset.prod_zero] + +@[simp] +theorem msymm_one : msymm σ R (Nat.Partition.indiscrete 1) = ∑ i, X i := by + symm + let f : σ ≃ { a // Nat.Partition.ofSym a = Nat.Partition.indiscrete 1 } := { + toFun := fun a => ⟨Sym.oneEquiv a, by ext; simp⟩ + invFun := fun a => Sym.oneEquiv.symm a.1 + left_inv := by intro a; simp only [Sym.oneEquiv_apply]; rfl + right_inv := by intro a; simp only [Equiv.apply_symm_apply, Subtype.coe_eta] + } + apply Fintype.sum_equiv f + intro i + have : (X i : MvPolynomial σ R) = Multiset.prod (Multiset.map X {i}) := by simp only [Multiset.map_singleton, Multiset.prod_singleton] + rw [this] + rfl + +@[simp] +theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by + simp only [msymm._eq_1, Sym.val_eq_coe, map_sum] + let f : {a : Sym σ n // Nat.Partition.ofSym a = μ} ≃ {a : Sym τ n // Nat.Partition.ofSym a = μ} := { + toFun := fun a => ⟨Sym.equivCongr e a, by + simp only [Sym.equivCongr_apply, Nat.Partition.ofSymEquiv, a.2]⟩ + invFun := fun a => ⟨Sym.equivCongr e.symm a, by + simp only [Sym.equivCongr_apply, Nat.Partition.ofSymEquiv, a.2]⟩ + left_inv := by intro a; simp + right_inv := by intro a; simp + } + apply Fintype.sum_equiv f + intro a + rw [← Multiset.prod_hom, Multiset.map_map] + congr + simp only [Sym.equivCongr_apply, Equiv.coe_fn_mk, Sym.coe_map, Multiset.map_map, + Function.comp_apply, rename_X] + +theorem msymm_isSymmetric {n : ℕ} (μ : n.Partition) : IsSymmetric (msymm σ R μ) := rename_msymm _ _ μ + +end MonomialSymmetric + end MvPolynomial From dd462208bcac4371c340725fabbd2e6bd9c0670a Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 15:38:43 +0100 Subject: [PATCH 018/106] Linting --- Mathlib/Combinatorics/Partition.lean | 3 ++- Mathlib/Data/Multiset/Dedup.lean | 6 +++-- .../RingTheory/MvPolynomial/Symmetric.lean | 24 ++++++++++++------- 3 files changed, 22 insertions(+), 11 deletions(-) diff --git a/Mathlib/Combinatorics/Partition.lean b/Mathlib/Combinatorics/Partition.lean index 15a3333dba75c..84f483acba6a5 100644 --- a/Mathlib/Combinatorics/Partition.lean +++ b/Mathlib/Combinatorics/Partition.lean @@ -112,7 +112,8 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w simp only [Sym.val_eq_coe, card_eq_sum_count] exact s.2 -lemma ofSymEquiv {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by +lemma ofSymEquiv {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : + Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by simp only [Nat.Partition.ofSym, Nat.Partition.ext_iff] have dedup_eq : (s.map e).1.dedup = s.1.dedup.map e := by simp only [Sym.val_eq_coe, Sym.coe_map] diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 8600f32221c07..76a3e7114e5f5 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -142,7 +142,8 @@ theorem Nodup.le_dedup_iff_le {s t : Multiset α} (hno : s.Nodup) : s ≤ t.dedu simp [le_dedup, hno] #align multiset.nodup.le_dedup_iff_le Multiset.Nodup.le_dedup_iff_le -theorem card_eq_sum_count {α : Type*} {s : Multiset α} [DecidableEq α] : sum (Multiset.map (fun x ↦ count x ↑s) (s.dedup)) = card s := by +theorem card_eq_sum_count {α : Type*} {s : Multiset α} [DecidableEq α] : + sum (Multiset.map (fun x ↦ count x ↑s) (s.dedup)) = card s := by induction' s using Multiset.induction_on with a s ih · simp · simp_rw [count_cons, sum_map_add, card_cons] @@ -150,7 +151,8 @@ theorem card_eq_sum_count {α : Type*} {s : Multiset α} [DecidableEq α] : sum · rw [dedup_cons_of_mem h] congr rw [Multiset.sum_map_eq_nsmul_single a] - · simp only [Multiset.nodup_dedup, ↓reduceIte, smul_eq_mul, mul_one, Multiset.mem_dedup, h, Multiset.count_eq_one_of_mem] + · simp only [Multiset.nodup_dedup, ↓reduceIte, smul_eq_mul, mul_one, Multiset.mem_dedup, h, + Multiset.count_eq_one_of_mem] · simp only [ne_eq, mem_dedup, ite_eq_right_iff, one_ne_zero, imp_false] tauto · congr diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 676ca85f72128..9ae0dbd42635f 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -201,7 +201,8 @@ def esymm (n : ℕ) : MvPolynomial σ R := lemma esymm_def (n : ℕ) : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl /-- -`esymmMu` is the product of the symmetric polynomials `esymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. +`esymmMu` is the product of the symmetric polynomials `esymm μᵢ`, +where `μ = (μ₁, μ₂, ...)` is a partition. -/ def esymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (esymm σ R) μ @@ -344,7 +345,8 @@ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl /-- -`hsymmMu` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. +`hsymmMu` is the product of the symmetric polynomials `hsymm μᵢ`, +where `μ = (μ₁, μ₂, ...)` is a partition. -/ def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ @@ -398,7 +400,8 @@ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl /-- -`psumMu` is the product of the symmetric polynomials `psum μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. +`psumMu` is the product of the symmetric polynomials `psum μᵢ`, +where `μ = (μ₁, μ₂, ...)` is a partition. -/ def psumMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (psum σ R) μ @@ -441,7 +444,8 @@ variable [DecidableEq σ] [DecidableEq τ] def msymm {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := ∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod -lemma msymm_def {n : ℕ} (μ : n.Partition) : msymm σ R μ = ∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod := rfl +lemma msymm_def {n : ℕ} (μ : n.Partition) : msymm σ R μ = +∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod := rfl @[simp] theorem msymm_zero : msymm σ R (Nat.Partition.indiscrete 0) = 1 := by @@ -461,14 +465,17 @@ theorem msymm_one : msymm σ R (Nat.Partition.indiscrete 1) = ∑ i, X i := by } apply Fintype.sum_equiv f intro i - have : (X i : MvPolynomial σ R) = Multiset.prod (Multiset.map X {i}) := by simp only [Multiset.map_singleton, Multiset.prod_singleton] + have : (X i : MvPolynomial σ R) = Multiset.prod (Multiset.map X {i}) := by + simp only [Multiset.map_singleton, Multiset.prod_singleton] rw [this] rfl @[simp] -theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by +theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : + rename e (msymm σ R μ) = msymm τ R μ := by simp only [msymm._eq_1, Sym.val_eq_coe, map_sum] - let f : {a : Sym σ n // Nat.Partition.ofSym a = μ} ≃ {a : Sym τ n // Nat.Partition.ofSym a = μ} := { + let f : {a : Sym σ n // Nat.Partition.ofSym a = μ} ≃ {a : Sym τ n // Nat.Partition.ofSym a = μ} + := { toFun := fun a => ⟨Sym.equivCongr e a, by simp only [Sym.equivCongr_apply, Nat.Partition.ofSymEquiv, a.2]⟩ invFun := fun a => ⟨Sym.equivCongr e.symm a, by @@ -483,7 +490,8 @@ theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rename e (ms simp only [Sym.equivCongr_apply, Equiv.coe_fn_mk, Sym.coe_map, Multiset.map_map, Function.comp_apply, rename_X] -theorem msymm_isSymmetric {n : ℕ} (μ : n.Partition) : IsSymmetric (msymm σ R μ) := rename_msymm _ _ μ +theorem msymm_isSymmetric {n : ℕ} (μ : n.Partition) : IsSymmetric (msymm σ R μ) := + rename_msymm _ _ μ end MonomialSymmetric From 1b6da2e034e71eef272add534193c75782bfa510 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 15:39:36 +0100 Subject: [PATCH 019/106] Update Mathlib/Combinatorics/Partition.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/Combinatorics/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Partition.lean b/Mathlib/Combinatorics/Partition.lean index 84f483acba6a5..b97ea33d831f1 100644 --- a/Mathlib/Combinatorics/Partition.lean +++ b/Mathlib/Combinatorics/Partition.lean @@ -113,7 +113,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w exact s.2 lemma ofSymEquiv {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : - Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by + Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by simp only [Nat.Partition.ofSym, Nat.Partition.ext_iff] have dedup_eq : (s.map e).1.dedup = s.1.dedup.map e := by simp only [Sym.val_eq_coe, Sym.coe_map] From 7b2fbad3358af507e17c0535371b2a70657f6a40 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 15:39:57 +0100 Subject: [PATCH 020/106] Update Mathlib/Data/Multiset/Dedup.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/Data/Multiset/Dedup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 76a3e7114e5f5..a86223a8d6d65 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -143,7 +143,7 @@ theorem Nodup.le_dedup_iff_le {s t : Multiset α} (hno : s.Nodup) : s ≤ t.dedu #align multiset.nodup.le_dedup_iff_le Multiset.Nodup.le_dedup_iff_le theorem card_eq_sum_count {α : Type*} {s : Multiset α} [DecidableEq α] : - sum (Multiset.map (fun x ↦ count x ↑s) (s.dedup)) = card s := by + sum (Multiset.map (fun x ↦ count x ↑s) (s.dedup)) = card s := by induction' s using Multiset.induction_on with a s ih · simp · simp_rw [count_cons, sum_map_add, card_cons] From 42e5ee49bc7354919a0baf6a19a12100737a065c Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 15:40:08 +0100 Subject: [PATCH 021/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 9ae0dbd42635f..dd21794b2a9a8 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -445,7 +445,7 @@ def msymm {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := ∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod lemma msymm_def {n : ℕ} (μ : n.Partition) : msymm σ R μ = -∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod := rfl + ∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod := rfl @[simp] theorem msymm_zero : msymm σ R (Nat.Partition.indiscrete 0) = 1 := by From 5121d1c9954aa6fdc0da54bf282fa19f7ec87630 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 15:40:17 +0100 Subject: [PATCH 022/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index dd21794b2a9a8..3834f4e3b4181 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -472,7 +472,7 @@ theorem msymm_one : msymm σ R (Nat.Partition.indiscrete 1) = ∑ i, X i := by @[simp] theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : - rename e (msymm σ R μ) = msymm τ R μ := by + rename e (msymm σ R μ) = msymm τ R μ := by simp only [msymm._eq_1, Sym.val_eq_coe, map_sum] let f : {a : Sym σ n // Nat.Partition.ofSym a = μ} ≃ {a : Sym τ n // Nat.Partition.ofSym a = μ} := { From dea7a3de6801826ee7c8c0011cc38f35dea69d97 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 15:42:42 +0100 Subject: [PATCH 023/106] Linting --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 3834f4e3b4181..116c1c4b1dc66 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -474,8 +474,8 @@ theorem msymm_one : msymm σ R (Nat.Partition.indiscrete 1) = ∑ i, X i := by theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by simp only [msymm._eq_1, Sym.val_eq_coe, map_sum] - let f : {a : Sym σ n // Nat.Partition.ofSym a = μ} ≃ {a : Sym τ n // Nat.Partition.ofSym a = μ} - := { + let f : {a : Sym σ n // Nat.Partition.ofSym a = μ} ≃ {a : Sym τ n // Nat.Partition.ofSym a = μ} := + { toFun := fun a => ⟨Sym.equivCongr e a, by simp only [Sym.equivCongr_apply, Nat.Partition.ofSymEquiv, a.2]⟩ invFun := fun a => ⟨Sym.equivCongr e.symm a, by From e232e2b1f322f2168085dd31422a435fcafb7436 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 16:30:54 +0100 Subject: [PATCH 024/106] Linting --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 116c1c4b1dc66..98d8f4462ec26 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -5,7 +5,6 @@ Authors: Hanting Zhang, Johan Commelin -/ import Mathlib.Algebra.Algebra.Subalgebra.Basic import Mathlib.Combinatorics.Partition -import Mathlib.Data.Fintype.Perm import Mathlib.Data.MvPolynomial.Rename import Mathlib.Data.MvPolynomial.CommRing import Mathlib.Data.Sym.Basic From d6a34bcc32db34ee80ddf9dbff05716271762fdb Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 19:07:07 +0100 Subject: [PATCH 025/106] Docstring --- Mathlib/Data/Sym/Basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index b7294efc2b529..7df979ad13698 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -532,6 +532,9 @@ theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ Multiset.mem_add #align sym.mem_append_iff Sym.mem_append_iff +/-- +Defines an equivalence between `α` and `Sym α 1`, where `α` is a type and `Sym α 1` is the type of symmetric lists of length 1. +-/ @[simps apply] def oneEquiv : α ≃ Sym α 1 where toFun a := ⟨{a}, by simp⟩ From 31ad0f74d7ff1bb9b1b1f717cba2bf1bacda7dea Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 10 Mar 2024 19:58:54 +0100 Subject: [PATCH 026/106] More linting. Will it ever be over? --- Mathlib/Data/Sym/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 7df979ad13698..6d2920c784bb8 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -533,7 +533,7 @@ theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ #align sym.mem_append_iff Sym.mem_append_iff /-- -Defines an equivalence between `α` and `Sym α 1`, where `α` is a type and `Sym α 1` is the type of symmetric lists of length 1. +Defines an equivalence between `α` and `Sym α 1`. -/ @[simps apply] def oneEquiv : α ≃ Sym α 1 where From e1b1ae6ff199d8ada00c75e0771c1a3d5f196f82 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 13 Mar 2024 12:32:11 +0100 Subject: [PATCH 027/106] fix build --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index b083cca4e2ebd..f6225baaf6d3f 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -487,7 +487,7 @@ theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rw [← Multiset.prod_hom, Multiset.map_map] congr simp only [Sym.equivCongr_apply, Equiv.coe_fn_mk, Sym.coe_map, Multiset.map_map, - Function.comp_apply, rename_X] + Function.comp_apply, rename_X, f] theorem msymm_isSymmetric {n : ℕ} (μ : n.Partition) : IsSymmetric (msymm σ R μ) := rename_msymm _ _ μ From 52f2b8bf7d41f3d50a591e0f212460c223204728 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 16 Mar 2024 16:32:53 +0100 Subject: [PATCH 028/106] Cleanup --- Mathlib/Combinatorics/Partition.lean | 26 ++++++----------- Mathlib/Data/Multiset/Dedup.lean | 28 +++++-------------- .../RingTheory/MvPolynomial/Symmetric.lean | 4 +++ 3 files changed, 20 insertions(+), 38 deletions(-) diff --git a/Mathlib/Combinatorics/Partition.lean b/Mathlib/Combinatorics/Partition.lean index 7d59d3e7efdf8..71f49d6cc677b 100644 --- a/Mathlib/Combinatorics/Partition.lean +++ b/Mathlib/Combinatorics/Partition.lean @@ -106,29 +106,21 @@ def ofMultiset (l : Multiset ℕ) : Partition l.sum := ofSums _ l rfl /-- An element `s` of `Sym σ n` induces a partition given by its multiplicities. -/ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition where - parts := (s.1.dedup).map s.1.count + parts := (s.1.toFinset).val.map s.1.count parts_pos := by simp [Multiset.count_pos] parts_sum := by - simp only [Sym.val_eq_coe, card_eq_sum_count] + have : sum (map (fun a ↦ count a s.1) (toFinset s.1).val) = Finset.sum (toFinset s.1) fun a ↦ count a s.1 := rfl + rw [this, Multiset.toFinset_sum_count_eq] exact s.2 lemma ofSymEquiv {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : - Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by - simp only [Nat.Partition.ofSym, Nat.Partition.ext_iff] - have dedup_eq : (s.map e).1.dedup = s.1.dedup.map e := by - simp only [Sym.val_eq_coe, Sym.coe_map] - rw [← Multiset.dedup_map_dedup_eq] - apply Multiset.Nodup.dedup - rw [Multiset.nodup_map_iff_of_injective e.injective] - exact Multiset.nodup_dedup s.1 - rw [dedup_eq] - have : (fun x ↦ Multiset.count x s.1) = fun x ↦ Multiset.count (e x) (s.map e).1 := by - funext x + Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by + simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] + rw [Multiset.dedup_map_dedup_eq_injective] + simp only [map_map, Function.comp_apply] + congr; funext i rw [← Multiset.count_map_eq_count' e] - congr - exact e.injective - rw [this] - simp only [Sym.val_eq_coe, Sym.coe_map, Multiset.map_map, Function.comp_apply] + all_goals exact e.injective /-- The partition of exactly one part. -/ def indiscrete (n : ℕ) : Partition n := ofSums n {n} rfl diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index a86223a8d6d65..e18bd546129f4 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -132,6 +132,13 @@ theorem dedup_map_dedup_eq [DecidableEq β] (f : α → β) (s : Multiset α) : simp [dedup_ext] #align multiset.dedup_map_dedup_eq Multiset.dedup_map_dedup_eq +theorem dedup_map_dedup_eq_injective [DecidableEq β] (f : α → β) (hinj : Function.Injective f) (s : Multiset α) : + dedup (map f s) = map f (dedup s) := by + rw [← dedup_map_dedup_eq] + apply Multiset.dedup_eq_self.mpr + rw [Multiset.nodup_map_iff_of_injective hinj] + exact Multiset.nodup_dedup s + @[simp] theorem dedup_nsmul {s : Multiset α} {n : ℕ} (h0 : n ≠ 0) : (n • s).dedup = s.dedup := by ext a @@ -142,27 +149,6 @@ theorem Nodup.le_dedup_iff_le {s t : Multiset α} (hno : s.Nodup) : s ≤ t.dedu simp [le_dedup, hno] #align multiset.nodup.le_dedup_iff_le Multiset.Nodup.le_dedup_iff_le -theorem card_eq_sum_count {α : Type*} {s : Multiset α} [DecidableEq α] : - sum (Multiset.map (fun x ↦ count x ↑s) (s.dedup)) = card s := by - induction' s using Multiset.induction_on with a s ih - · simp - · simp_rw [count_cons, sum_map_add, card_cons] - by_cases h : a ∈ s - · rw [dedup_cons_of_mem h] - congr - rw [Multiset.sum_map_eq_nsmul_single a] - · simp only [Multiset.nodup_dedup, ↓reduceIte, smul_eq_mul, mul_one, Multiset.mem_dedup, h, - Multiset.count_eq_one_of_mem] - · simp only [ne_eq, mem_dedup, ite_eq_right_iff, one_ne_zero, imp_false] - tauto - · congr - · rw [dedup_cons_of_not_mem h] - simp only [map_cons, sum_cons, ih, count_eq_zero_of_not_mem h, zero_add] - · rw [Multiset.sum_map_eq_nsmul_single a] - · simp [h] - · simp only [ne_eq, mem_dedup, ite_eq_right_iff, one_ne_zero, imp_false] - tauto - end Multiset theorem Multiset.Nodup.le_nsmul_iff_le {α : Type*} {s t : Multiset α} {n : ℕ} (h : s.Nodup) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index f6225baaf6d3f..42a3b557f86ef 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -29,6 +29,8 @@ We also prove some basic facts about them. * `MvPolynomial.psum` +* `MvPolynomial.msymm` + ## Notation + `esymm σ R n` is the `n`th elementary symmetric polynomial in `MvPolynomial σ R`. @@ -38,6 +40,8 @@ We also prove some basic facts about them. + `psum σ R n` is the degree-`n` power sum in `MvPolynomial σ R`, i.e. the sum of monomials `(X i)^n` over `i ∈ σ`. ++ `msymm σ R μ` is the monomial symmetric polynomial whose exponents set are the parts of `μ ⊢ n` in `MvPolynomial σ R`. + As in other polynomial files, we typically use the notation: + `σ τ : Type*` (indexing the variables) From a3499f30d70dabcd06b517c8a38742a7c27a32ec Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 16 Mar 2024 16:33:56 +0100 Subject: [PATCH 029/106] Update Mathlib/Combinatorics/Partition.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/Combinatorics/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Partition.lean b/Mathlib/Combinatorics/Partition.lean index 71f49d6cc677b..5c2b3f583ddd1 100644 --- a/Mathlib/Combinatorics/Partition.lean +++ b/Mathlib/Combinatorics/Partition.lean @@ -114,7 +114,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w exact s.2 lemma ofSymEquiv {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : - Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by + Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] rw [Multiset.dedup_map_dedup_eq_injective] simp only [map_map, Function.comp_apply] From 8a97d351ebe494df4cc3a0ee57bbdddad1feb7d7 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 16 Mar 2024 16:37:01 +0100 Subject: [PATCH 030/106] Linting --- Mathlib/Combinatorics/Partition.lean | 17 +++++++++-------- Mathlib/Data/Multiset/Dedup.lean | 4 ++-- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 3 ++- 3 files changed, 13 insertions(+), 11 deletions(-) diff --git a/Mathlib/Combinatorics/Partition.lean b/Mathlib/Combinatorics/Partition.lean index 71f49d6cc677b..a29eea0b0d4ed 100644 --- a/Mathlib/Combinatorics/Partition.lean +++ b/Mathlib/Combinatorics/Partition.lean @@ -109,18 +109,19 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w parts := (s.1.toFinset).val.map s.1.count parts_pos := by simp [Multiset.count_pos] parts_sum := by - have : sum (map (fun a ↦ count a s.1) (toFinset s.1).val) = Finset.sum (toFinset s.1) fun a ↦ count a s.1 := rfl + have : sum (map (fun a ↦ count a s.1) (toFinset s.1).val) = + Finset.sum (toFinset s.1) fun a ↦ count a s.1 := rfl rw [this, Multiset.toFinset_sum_count_eq] exact s.2 lemma ofSymEquiv {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : - Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by - simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] - rw [Multiset.dedup_map_dedup_eq_injective] - simp only [map_map, Function.comp_apply] - congr; funext i - rw [← Multiset.count_map_eq_count' e] - all_goals exact e.injective + Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by + simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] + rw [Multiset.dedup_map_dedup_eq_injective] + simp only [map_map, Function.comp_apply] + congr; funext i + rw [← Multiset.count_map_eq_count' e] + all_goals exact e.injective /-- The partition of exactly one part. -/ def indiscrete (n : ℕ) : Partition n := ofSums n {n} rfl diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index e18bd546129f4..f8c91ad1709f7 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -132,8 +132,8 @@ theorem dedup_map_dedup_eq [DecidableEq β] (f : α → β) (s : Multiset α) : simp [dedup_ext] #align multiset.dedup_map_dedup_eq Multiset.dedup_map_dedup_eq -theorem dedup_map_dedup_eq_injective [DecidableEq β] (f : α → β) (hinj : Function.Injective f) (s : Multiset α) : - dedup (map f s) = map f (dedup s) := by +theorem dedup_map_dedup_eq_injective [DecidableEq β] (f : α → β) (hinj : Function.Injective f) + (s : Multiset α) : dedup (map f s) = map f (dedup s) := by rw [← dedup_map_dedup_eq] apply Multiset.dedup_eq_self.mpr rw [Multiset.nodup_map_iff_of_injective hinj] diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 42a3b557f86ef..658b402a346cf 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -40,7 +40,8 @@ We also prove some basic facts about them. + `psum σ R n` is the degree-`n` power sum in `MvPolynomial σ R`, i.e. the sum of monomials `(X i)^n` over `i ∈ σ`. -+ `msymm σ R μ` is the monomial symmetric polynomial whose exponents set are the parts of `μ ⊢ n` in `MvPolynomial σ R`. ++ `msymm σ R μ` is the monomial symmetric polynomial whose exponents set are the parts + of `μ ⊢ n` in `MvPolynomial σ R`. As in other polynomial files, we typically use the notation: From 937f636268701f4f800aa354661e1e27d8758ce0 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 28 Mar 2024 14:36:34 +0100 Subject: [PATCH 031/106] Fix --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 6520b93ad900c..cdaf0a221616f 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Hanting Zhang, Johan Commelin -/ import Mathlib.Algebra.Algebra.Subalgebra.Basic -import Mathlib.Combinatorics.Partition +import Mathlib.Combinatorics.Enumerative.Partition import Mathlib.Data.MvPolynomial.Rename import Mathlib.Data.MvPolynomial.CommRing import Mathlib.Data.Sym.Basic From ec9d09350ebc311bb190ccc79ca10a1bb1f5ffaf Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:30:04 +0100 Subject: [PATCH 032/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Combinatorics/Enumerative/Partition.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index a87f86ae417aa..17e194e7c1d15 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -103,8 +103,7 @@ def ofSums (n : ℕ) (l : Multiset ℕ) (hl : l.sum = n) : Partition n where def ofMultiset (l : Multiset ℕ) : Partition l.sum := ofSums _ l rfl #align nat.partition.of_multiset Nat.Partition.ofMultiset -/-- An element `s` of `Sym σ n` induces a partition given by its multiplicities. --/ +/-- An element `s` of `Sym σ n` induces a partition given by its multiplicities. -/ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition where parts := (s.1.toFinset).val.map s.1.count parts_pos := by simp [Multiset.count_pos] From f098543a06e5e5198a96cafa585c0cca7e22827e Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:30:26 +0100 Subject: [PATCH 033/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 17e194e7c1d15..580e87b4f1976 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -105,7 +105,7 @@ def ofMultiset (l : Multiset ℕ) : Partition l.sum := ofSums _ l rfl /-- An element `s` of `Sym σ n` induces a partition given by its multiplicities. -/ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition where - parts := (s.1.toFinset).val.map s.1.count + parts := s.1.dedup.map s.1.count parts_pos := by simp [Multiset.count_pos] parts_sum := by have : sum (map (fun a ↦ count a s.1) (toFinset s.1).val) = From 6b5bbba23a7858414923cbc38bdcc11bbd065d99 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:30:47 +0100 Subject: [PATCH 034/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 580e87b4f1976..20d798a39f9cf 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -108,7 +108,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w parts := s.1.dedup.map s.1.count parts_pos := by simp [Multiset.count_pos] parts_sum := by - have : sum (map (fun a ↦ count a s.1) (toFinset s.1).val) = + have : sum (map (fun a ↦ count a s.1) s.1.dedup) = Finset.sum (toFinset s.1) fun a ↦ count a s.1 := rfl rw [this, Multiset.toFinset_sum_count_eq] exact s.2 From 85e93b7063af95f9d920bd77b6db82860c6998ba Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:31:19 +0100 Subject: [PATCH 035/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 20d798a39f9cf..0eefff85d78bd 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -113,7 +113,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w rw [this, Multiset.toFinset_sum_count_eq] exact s.2 -lemma ofSymEquiv {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : +lemma ofSym_map {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] rw [Multiset.dedup_map_dedup_eq_injective] From 7b7a95ede12439c7e6cd2c25d443920e7a2a3ef6 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:31:48 +0100 Subject: [PATCH 036/106] Update Mathlib/Data/Multiset/Dedup.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Data/Multiset/Dedup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 55028f3eabe6c..10bb96faee172 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -122,7 +122,7 @@ theorem dedup_map_dedup_eq [DecidableEq β] (f : α → β) (s : Multiset α) : simp [dedup_ext] #align multiset.dedup_map_dedup_eq Multiset.dedup_map_dedup_eq -theorem dedup_map_dedup_eq_injective [DecidableEq β] (f : α → β) (hinj : Function.Injective f) +theorem dedup_map_of_injective [DecidableEq β] {f : α → β} (hinj : Function.Injective f) (s : Multiset α) : dedup (map f s) = map f (dedup s) := by rw [← dedup_map_dedup_eq] apply Multiset.dedup_eq_self.mpr From 3a4e072a360b624dc55abcc123987e9583577184 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:32:03 +0100 Subject: [PATCH 037/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index cdaf0a221616f..b63142c725bc0 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -190,7 +190,7 @@ def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] -/- Multiplicativity on partitions -/ +/-! ### Multiplicativity on partitions -/ /-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, From 32465669a3b3d14a43771a7551f35901b0175f99 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:32:15 +0100 Subject: [PATCH 038/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index b63142c725bc0..f15b6d4cc3abc 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -203,7 +203,7 @@ lemma muProduct_def {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) muProduct σ R f μ = Multiset.prod (μ.parts.map f) := rfl @[simp] -theorem muProduct_zero (f : ℕ → MvPolynomial σ R) : +theorem muProduct_indiscrete_zero (f : ℕ → MvPolynomial σ R) : muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by simp only [muProduct, Nat.Partition.partition_zero_parts, Multiset.map_zero, Multiset.prod_zero] From c1f714983ba065db68f04b9a7a00885489bbff1f Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:32:35 +0100 Subject: [PATCH 039/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index f15b6d4cc3abc..2cb7e3f2715db 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -192,10 +192,8 @@ variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype /-! ### Multiplicativity on partitions -/ -/-- - Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, - `muProduct` computes the product of applying each function in `f` to the parts of `μ`. --/ +/-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, +`muProduct` computes the product of applying each function in `f` to the parts of `μ`. -/ def muProduct {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : MvPolynomial σ R := Multiset.prod (μ.parts.map f) From 61fe1c72f3d2d955fea4406754b8798856c9dd2e Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:32:46 +0100 Subject: [PATCH 040/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 2cb7e3f2715db..5c79d0d05a7a2 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -513,7 +513,7 @@ theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : Function.comp_apply, rename_X, f] theorem msymm_isSymmetric {n : ℕ} (μ : n.Partition) : IsSymmetric (msymm σ R μ) := - rename_msymm _ _ μ + rename_msymm _ _ μ end MonomialSymmetric From 35453565370f86d4ac1727fdaadfde5b43eda5d6 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 10:33:01 +0100 Subject: [PATCH 041/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 5c79d0d05a7a2..eebe02b8ae799 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -206,7 +206,7 @@ theorem muProduct_indiscrete_zero (f : ℕ → MvPolynomial σ R) : simp only [muProduct, Nat.Partition.partition_zero_parts, Multiset.map_zero, Multiset.prod_zero] @[simp] -theorem muProduct_onePart (n : ℕ) (f : ℕ → MvPolynomial σ R) (npos : n > 0) : +theorem muProduct_indiscrete_of_pos (n : ℕ) (f : ℕ → MvPolynomial σ R) (npos : n > 0) : muProduct σ R f (Nat.Partition.indiscrete n) = f n := by rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] linarith From 5632270d9503f972c4afbcddd4a202672810ae0e Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 15:29:35 +0100 Subject: [PATCH 042/106] Fix --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 0eefff85d78bd..c579b9fa9e653 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -116,7 +116,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w lemma ofSym_map {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] - rw [Multiset.dedup_map_dedup_eq_injective] + rw [Multiset.dedup_map_of_injective] simp only [map_map, Function.comp_apply] congr; funext i rw [← Multiset.count_map_eq_count' e] From 877bcef922255942b394e539412d7634cda93531 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sat, 30 Mar 2024 17:14:29 +0100 Subject: [PATCH 043/106] Rename --- .../RingTheory/MvPolynomial/Symmetric.lean | 20 +++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index eebe02b8ae799..d0c8f772ce0c8 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -263,13 +263,13 @@ theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp only [esymm, powersetCard_one, sum_map, Function.Embedding.coeFn_mk, prod_singleton] theorem esymmMu_zero : esymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by - simp only [esymmMu, esymm_zero, muProduct_zero] + simp only [esymmMu, esymm_zero, muProduct_indiscrete_zero] @[simp] theorem esymmMu_onePart (n : ℕ) : esymmMu σ R (Nat.Partition.indiscrete n) = esymm σ R n := by cases n - · simp only [esymmMu, esymm_zero, muProduct_zero] - · simp only [esymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_onePart] + · simp only [esymmMu, esymm_zero, muProduct_indiscrete_zero] + · simp only [esymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_indiscrete_of_pos] theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by simp_rw [esymm, map_sum, map_prod, map_X] @@ -390,13 +390,13 @@ theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton] theorem hsymmMu_zero : hsymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by - simp only [hsymmMu, hsymm_zero, muProduct_zero] + simp only [hsymmMu, hsymm_zero, muProduct_indiscrete_zero] @[simp] theorem hsymmMu_onePart (n : ℕ) : hsymmMu σ R (Nat.Partition.indiscrete n) = hsymm σ R n := by cases n - · simp only [hsymmMu, hsymm_zero, muProduct_zero] - · simp only [hsymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_onePart] + · simp only [hsymmMu, hsymm_zero, muProduct_indiscrete_zero] + · simp only [hsymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_indiscrete_of_pos] theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp_rw [hsymm, map_sum, ← Multiset.prod_hom'] @@ -442,12 +442,12 @@ theorem psum_one : psum σ R 1 = ∑ i, X i := by @[simp] theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by - simp only [psumMu, muProduct_zero] + simp only [psumMu, muProduct_indiscrete_zero] @[simp] theorem psumMu_onePart (n : ℕ) (npos : n > 0) : psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by - simp only [psumMu, npos, muProduct_onePart] + simp only [psumMu, npos, muProduct_indiscrete_of_pos] @[simp] theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by @@ -499,9 +499,9 @@ theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : let f : {a : Sym σ n // Nat.Partition.ofSym a = μ} ≃ {a : Sym τ n // Nat.Partition.ofSym a = μ} := { toFun := fun a => ⟨Sym.equivCongr e a, by - simp only [Sym.equivCongr_apply, Nat.Partition.ofSymEquiv, a.2]⟩ + simp only [Sym.equivCongr_apply, Nat.Partition.ofSym_map, a.2]⟩ invFun := fun a => ⟨Sym.equivCongr e.symm a, by - simp only [Sym.equivCongr_apply, Nat.Partition.ofSymEquiv, a.2]⟩ + simp only [Sym.equivCongr_apply, Nat.Partition.ofSym_map, a.2]⟩ left_inv := by intro a; simp right_inv := by intro a; simp } From beaaf8bbe955a067dbbb1d5acaa5d046a6070bc7 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 12 Apr 2024 07:58:04 +0200 Subject: [PATCH 044/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 650af20ac1f23..c6586348935d9 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -178,7 +178,6 @@ end CommRing end IsSymmetric - /-- `MvPolynomial.rename` induces an isomorphism between the symmetric subalgebras. -/ @[simps!] def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : From ace3facdfb5cc2ca05f141007222825c9bdf3249 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 12:41:10 +0200 Subject: [PATCH 045/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean Co-authored-by: Riccardo Brasca --- Mathlib/Combinatorics/Enumerative/Partition.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index c579b9fa9e653..b00700772f6f5 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -116,11 +116,10 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w lemma ofSym_map {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] - rw [Multiset.dedup_map_of_injective] + rw [Multiset.dedup_map_of_injective e.injective] simp only [map_map, Function.comp_apply] congr; funext i - rw [← Multiset.count_map_eq_count' e] - all_goals exact e.injective + rw [← Multiset.count_map_eq_count' e _ e.injective] /-- The partition of exactly one part. -/ def indiscrete (n : ℕ) : Partition n := ofSums n {n} rfl From 6c7ba56fbafe5251d9c5b1a14f123d740592518d Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 12:41:46 +0200 Subject: [PATCH 046/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Riccardo Brasca --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index c6586348935d9..e06429c596dd8 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -267,9 +267,7 @@ theorem esymmMu_zero : esymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by @[simp] theorem esymmMu_onePart (n : ℕ) : esymmMu σ R (Nat.Partition.indiscrete n) = esymm σ R n := by - cases n - · simp only [esymmMu, esymm_zero, muProduct_indiscrete_zero] - · simp only [esymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_indiscrete_of_pos] + cases n <;> simp [esymmMu] theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by simp_rw [esymm, map_sum, map_prod, map_X] From b487fffe4cae85fe2f21d9a195cd9fbc2b792e39 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 12:42:24 +0200 Subject: [PATCH 047/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Riccardo Brasca --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index e06429c596dd8..208fd71c577d8 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -397,8 +397,7 @@ theorem hsymmMu_onePart (n : ℕ) : hsymmMu σ R (Nat.Partition.indiscrete n) = · simp only [hsymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_indiscrete_of_pos] theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by - simp_rw [hsymm, map_sum, ← Multiset.prod_hom'] - simp only [val_eq_coe, map_X] + simp [hsymm, ← Multiset.prod_hom'] theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by simp_rw [hsymm, map_sum, ← prod_hom', rename_X] From a86cefecfd721bf4bced4e79d31e15383db59937 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 16:15:42 +0200 Subject: [PATCH 048/106] Suggestions I'm not sure that msymm_one is better now --- .../Combinatorics/Enumerative/Partition.lean | 20 +++- .../RingTheory/MvPolynomial/Symmetric.lean | 97 +++++++------------ 2 files changed, 52 insertions(+), 65 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index b00700772f6f5..e06bc756171de 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -113,7 +113,9 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w rw [this, Multiset.toFinset_sum_count_eq] exact s.2 -lemma ofSym_map {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : σ ≃ τ) (s : Sym σ n) : +variable {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] + +lemma ofSym_map (e : σ ≃ τ) (s : Sym σ n) : Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] rw [Multiset.dedup_map_of_injective e.injective] @@ -121,6 +123,13 @@ lemma ofSym_map {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] (e : congr; funext i rw [← Multiset.count_map_eq_count' e _ e.injective] +def ofSym_shape_equiv (μ : Partition n) (e : σ ≃ τ) : + {x : Sym σ n // Nat.Partition.ofSym x = μ} ≃ {x : Sym τ n // Nat.Partition.ofSym x = μ} where + toFun := fun x => ⟨Sym.equivCongr e x, by simp [ofSym_map, x.2]⟩ + invFun := fun x => ⟨Sym.equivCongr e.symm x, by simp [ofSym_map, x.2]⟩ + left_inv := by intro x; simp + right_inv := by intro x; simp + /-- The partition of exactly one part. -/ def indiscrete (n : ℕ) : Partition n := ofSums n {n} rfl #align nat.partition.indiscrete_partition Nat.Partition.indiscrete @@ -145,6 +154,15 @@ instance UniquePartitionZero : Unique (Partition 0) where instance UniquePartitionOne : Unique (Partition 1) where uniq _ := Partition.ext _ _ <| by simp +lemma ofSym_one (s : Sym σ 1) : Nat.Partition.ofSym s = Nat.Partition.indiscrete 1 := by + ext; simp + +def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ { a : Sym σ 1 // Nat.Partition.ofSym a = Nat.Partition.indiscrete 1 } where + toFun := fun a => ⟨Sym.oneEquiv a, by ext; simp⟩ + invFun := fun a => Sym.oneEquiv.symm a.1 + left_inv := by intro a; simp only [Sym.oneEquiv_apply]; rfl + right_inv := by intro a; simp only [Equiv.apply_symm_apply, Subtype.coe_eta] + /-- The number of times a positive integer `i` appears in the partition `ofSums n l hl` is the same as the number of times it appears in the multiset `l`. (For `i = 0`, `Partition.non_zero` combined with `Multiset.count_eq_zero_of_not_mem` gives that diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 208fd71c577d8..04b1c054cfd24 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -15,8 +15,8 @@ import Mathlib.Data.Sym.Basic /-! # Symmetric Polynomials and Elementary Symmetric Polynomials -This file defines symmetric `MvPolynomial`s and elementary symmetric `MvPolynomial`s. -We also prove some basic facts about them. +This file defines symmetric `MvPolynomial`s and the bases of elementary, complete homogeneous, +power sum, and monomial symmetric `MvPolynomial`s. We also prove some basic facts about them. ## Main declarations @@ -84,8 +84,7 @@ end Multiset namespace MvPolynomial -variable {σ : Type*} {R : Type*} -variable {τ : Type*} {S : Type*} +variable {σ τ : Type*} {R S : Type*} /-- A `MvPolynomial φ` is symmetric if it is invariant under permutations of its variables by the `rename` operation -/ @@ -190,27 +189,33 @@ def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] + +section Partitions + /-! ### Multiplicativity on partitions -/ +variable {n : ℕ} (f : ℕ → MvPolynomial σ R) + /-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, `muProduct` computes the product of applying each function in `f` to the parts of `μ`. -/ def muProduct {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : MvPolynomial σ R := Multiset.prod (μ.parts.map f) -lemma muProduct_def {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : +lemma muProduct_def (μ : n.Partition) : muProduct σ R f μ = Multiset.prod (μ.parts.map f) := rfl @[simp] -theorem muProduct_indiscrete_zero (f : ℕ → MvPolynomial σ R) : - muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by - simp only [muProduct, Nat.Partition.partition_zero_parts, Multiset.map_zero, Multiset.prod_zero] +theorem muProduct_indiscrete_zero : + muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by simp [muProduct] @[simp] -theorem muProduct_indiscrete_of_pos (n : ℕ) (f : ℕ → MvPolynomial σ R) (npos : n > 0) : +theorem muProduct_indiscrete_of_pos (npos : n > 0) : muProduct σ R f (Nat.Partition.indiscrete n) = f n := by rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] linarith +end Partitions + section ElementarySymmetric open Finset @@ -254,16 +259,13 @@ theorem esymm_eq_sum_monomial (n : ℕ) : #align mv_polynomial.esymm_eq_sum_monomial MvPolynomial.esymm_eq_sum_monomial @[simp] -theorem esymm_zero : esymm σ R 0 = 1 := by - simp only [esymm, powersetCard_zero, sum_singleton, prod_empty] +theorem esymm_zero : esymm σ R 0 = 1 := by simp [esymm] #align mv_polynomial.esymm_zero MvPolynomial.esymm_zero @[simp] -theorem esymm_one : esymm σ R 1 = ∑ i, X i := by - simp only [esymm, powersetCard_one, sum_map, Function.Embedding.coeFn_mk, prod_singleton] +theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp [esymm, powersetCard_one] -theorem esymmMu_zero : esymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by - simp only [esymmMu, esymm_zero, muProduct_indiscrete_zero] +theorem esymmMu_zero : esymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp [esymmMu] @[simp] theorem esymmMu_onePart (n : ℕ) : esymmMu σ R (Nat.Partition.indiscrete n) = esymm σ R n := by @@ -368,33 +370,26 @@ lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := `hsymmMu` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := - muProduct σ R (hsymm σ R) μ +def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ lemma hsymmMu_def {n : ℕ} (μ : n.Partition) : hsymmMu σ R μ = Multiset.prod (μ.parts.map (hsymm σ R)) := rfl @[simp] -theorem hsymm_zero : hsymm σ R 0 = 1 := by - simp only [hsymm, univ_unique, eq_nil_of_card_zero, val_eq_coe, Sym.coe_nil, Multiset.map_zero, - prod_zero, sum_const, Finset.card_singleton, one_smul] +theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] @[simp] theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by simp only [hsymm, univ_unique] symm apply Fintype.sum_equiv oneEquiv - intro i - simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton] + simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton, implies_true] -theorem hsymmMu_zero : hsymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by - simp only [hsymmMu, hsymm_zero, muProduct_indiscrete_zero] +theorem hsymmMu_zero : hsymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp [hsymmMu] @[simp] theorem hsymmMu_onePart (n : ℕ) : hsymmMu σ R (Nat.Partition.indiscrete n) = hsymm σ R n := by - cases n - · simp only [hsymmMu, hsymm_zero, muProduct_indiscrete_zero] - · simp only [hsymmMu, gt_iff_lt, Nat.zero_lt_succ, muProduct_indiscrete_of_pos] + cases n <;> simp [hsymmMu] theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp [hsymm, ← Multiset.prod_hom'] @@ -402,8 +397,7 @@ theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by simp_rw [hsymm, map_sum, ← prod_hom', rename_X] apply Fintype.sum_equiv (equivCongr e) - simp only [val_eq_coe, equivCongr_apply, Sym.coe_map, Multiset.map_map, Function.comp_apply, - implies_true] + simp theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n @@ -429,20 +423,17 @@ lemma psumMu_def {n : ℕ} (μ : n.Partition) : psumMu σ R μ = Multiset.prod (μ.parts.map (psum σ R)) := rfl @[simp] -theorem psum_zero : psum σ R 0 = Fintype.card σ := by - simp [psum, _root_.pow_zero, ← cast_card] - rfl +theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum]; rfl @[simp] -theorem psum_one : psum σ R 1 = ∑ i, X i := by - simp only [psum, _root_.pow_one] +theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] @[simp] theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp only [psumMu, muProduct_indiscrete_zero] @[simp] -theorem psumMu_onePart (n : ℕ) (npos : n > 0) : +theorem psumMu_onePart {n : ℕ} (npos : n > 0) : psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by simp only [psumMu, npos, muProduct_indiscrete_of_pos] @@ -455,7 +446,6 @@ theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum σ R n) := rename_psum _ end PowerSum section MonomialSymmetric --- open Multiset variable [DecidableEq σ] [DecidableEq τ] @@ -468,46 +458,25 @@ lemma msymm_def {n : ℕ} (μ : n.Partition) : msymm σ R μ = @[simp] theorem msymm_zero : msymm σ R (Nat.Partition.indiscrete 0) = 1 := by - rw [msymm, Fintype.sum_subsingleton] - swap - exact ⟨(Sym.nil : Sym σ 0), by rfl⟩ - simp only [Sym.val_eq_coe, Sym.coe_nil, Multiset.map_zero, Multiset.prod_zero] + rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), by rfl⟩] + simp @[simp] theorem msymm_one : msymm σ R (Nat.Partition.indiscrete 1) = ∑ i, X i := by symm - let f : σ ≃ { a // Nat.Partition.ofSym a = Nat.Partition.indiscrete 1 } := { - toFun := fun a => ⟨Sym.oneEquiv a, by ext; simp⟩ - invFun := fun a => Sym.oneEquiv.symm a.1 - left_inv := by intro a; simp only [Sym.oneEquiv_apply]; rfl - right_inv := by intro a; simp only [Equiv.apply_symm_apply, Subtype.coe_eta] - } - apply Fintype.sum_equiv f + apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_onePart σ) intro i - have : (X i : MvPolynomial σ R) = Multiset.prod (Multiset.map X {i}) := by - simp only [Multiset.map_singleton, Multiset.prod_singleton] - rw [this] - rfl + show X i = Multiset.prod (Multiset.map X {i}) + simp only [Multiset.map_singleton, Multiset.prod_singleton] @[simp] theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by simp only [msymm._eq_1, Sym.val_eq_coe, map_sum] - let f : {a : Sym σ n // Nat.Partition.ofSym a = μ} ≃ {a : Sym τ n // Nat.Partition.ofSym a = μ} := - { - toFun := fun a => ⟨Sym.equivCongr e a, by - simp only [Sym.equivCongr_apply, Nat.Partition.ofSym_map, a.2]⟩ - invFun := fun a => ⟨Sym.equivCongr e.symm a, by - simp only [Sym.equivCongr_apply, Nat.Partition.ofSym_map, a.2]⟩ - left_inv := by intro a; simp - right_inv := by intro a; simp - } - apply Fintype.sum_equiv f + apply Fintype.sum_equiv (Nat.Partition.ofSym_shape_equiv μ e) intro a rw [← Multiset.prod_hom, Multiset.map_map] - congr - simp only [Sym.equivCongr_apply, Equiv.coe_fn_mk, Sym.coe_map, Multiset.map_map, - Function.comp_apply, rename_X, f] + simp [Nat.Partition.ofSym_shape_equiv] theorem msymm_isSymmetric {n : ℕ} (μ : n.Partition) : IsSymmetric (msymm σ R μ) := rename_msymm _ _ μ From 2443cdd7424dad3efb59ef85951f23e70a616274 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 16:20:04 +0200 Subject: [PATCH 049/106] Linting --- Mathlib/Combinatorics/Enumerative/Partition.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index e06bc756171de..bb837f37ee5ef 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -157,7 +157,8 @@ instance UniquePartitionOne : Unique (Partition 1) where lemma ofSym_one (s : Sym σ 1) : Nat.Partition.ofSym s = Nat.Partition.indiscrete 1 := by ext; simp -def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ { a : Sym σ 1 // Nat.Partition.ofSym a = Nat.Partition.indiscrete 1 } where +def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ + { a : Sym σ 1 // Nat.Partition.ofSym a = Nat.Partition.indiscrete 1 } where toFun := fun a => ⟨Sym.oneEquiv a, by ext; simp⟩ invFun := fun a => Sym.oneEquiv.symm a.1 left_inv := by intro a; simp only [Sym.oneEquiv_apply]; rfl From 5378e04e2d0a73d667945e04ec12eb80dde72367 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 16:37:40 +0200 Subject: [PATCH 050/106] Remove show --- Mathlib/Combinatorics/Enumerative/Partition.lean | 3 +++ Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 4 +--- 2 files changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index bb837f37ee5ef..7f3ef18b39ea3 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -164,6 +164,9 @@ def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ left_inv := by intro a; simp only [Sym.oneEquiv_apply]; rfl right_inv := by intro a; simp only [Equiv.apply_symm_apply, Subtype.coe_eta] +@[simp] +lemma ofSym_equiv_onePart_apply (i : σ) : ((Nat.Partition.ofSym_equiv_onePart σ) i : Multiset σ) = {i} := rfl + /-- The number of times a positive integer `i` appears in the partition `ofSums n l hl` is the same as the number of times it appears in the multiset `l`. (For `i = 0`, `Partition.non_zero` combined with `Multiset.count_eq_zero_of_not_mem` gives that diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 04b1c054cfd24..6e6778fb9511e 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -465,9 +465,7 @@ theorem msymm_zero : msymm σ R (Nat.Partition.indiscrete 0) = 1 := by theorem msymm_one : msymm σ R (Nat.Partition.indiscrete 1) = ∑ i, X i := by symm apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_onePart σ) - intro i - show X i = Multiset.prod (Multiset.map X {i}) - simp only [Multiset.map_singleton, Multiset.prod_singleton] + simp @[simp] theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : From e18d906698f0f5717067a8ea4ec2af4e54c11e9a Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 16:42:24 +0200 Subject: [PATCH 051/106] Linting --- Mathlib/Combinatorics/Enumerative/Partition.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 7f3ef18b39ea3..c705ba50a4c47 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -165,7 +165,8 @@ def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ right_inv := by intro a; simp only [Equiv.apply_symm_apply, Subtype.coe_eta] @[simp] -lemma ofSym_equiv_onePart_apply (i : σ) : ((Nat.Partition.ofSym_equiv_onePart σ) i : Multiset σ) = {i} := rfl +lemma ofSym_equiv_onePart_apply (i : σ) : + ((Nat.Partition.ofSym_equiv_onePart σ) i : Multiset σ) = {i} := rfl /-- The number of times a positive integer `i` appears in the partition `ofSums n l hl` is the same as the number of times it appears in the multiset `l`. From 4517acb4357b5cc97dfba76dada44af3d4d33be2 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 16:44:37 +0200 Subject: [PATCH 052/106] Linting --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index c705ba50a4c47..8b450c0a892d6 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -166,7 +166,7 @@ def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ @[simp] lemma ofSym_equiv_onePart_apply (i : σ) : - ((Nat.Partition.ofSym_equiv_onePart σ) i : Multiset σ) = {i} := rfl + ((Nat.Partition.ofSym_equiv_onePart σ) i : Multiset σ) = {i} := rfl /-- The number of times a positive integer `i` appears in the partition `ofSums n l hl` is the same as the number of times it appears in the multiset `l`. From 977788b2025e38e6da712944396f03f96966f482 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 10 May 2024 17:54:11 +0200 Subject: [PATCH 053/106] Linting --- Mathlib/Combinatorics/Enumerative/Partition.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 8b450c0a892d6..c4f8e191a83c3 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -123,6 +123,8 @@ lemma ofSym_map (e : σ ≃ τ) (s : Sym σ n) : congr; funext i rw [← Multiset.count_map_eq_count' e _ e.injective] +/-- An equivalence between `σ` and `τ` induces an equivalence between the subtypes of `Sym σ n` and +`Sym τ n` corresponding to a given partition. -/ def ofSym_shape_equiv (μ : Partition n) (e : σ ≃ τ) : {x : Sym σ n // Nat.Partition.ofSym x = μ} ≃ {x : Sym τ n // Nat.Partition.ofSym x = μ} where toFun := fun x => ⟨Sym.equivCongr e x, by simp [ofSym_map, x.2]⟩ @@ -157,6 +159,7 @@ instance UniquePartitionOne : Unique (Partition 1) where lemma ofSym_one (s : Sym σ 1) : Nat.Partition.ofSym s = Nat.Partition.indiscrete 1 := by ext; simp +/-- The equivalence between `σ` and `1`-tuples of elements of σ -/ def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ { a : Sym σ 1 // Nat.Partition.ofSym a = Nat.Partition.indiscrete 1 } where toFun := fun a => ⟨Sym.oneEquiv a, by ext; simp⟩ From 4c78fb9a398ea93b3bda61085c9f1c98910c99c4 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:04:10 +0200 Subject: [PATCH 054/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index c4f8e191a83c3..a31c761626a3f 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -109,7 +109,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w parts_pos := by simp [Multiset.count_pos] parts_sum := by have : sum (map (fun a ↦ count a s.1) s.1.dedup) = - Finset.sum (toFinset s.1) fun a ↦ count a s.1 := rfl + ∑ a ∈ toFinset s.1, count a s.1 := rfl rw [this, Multiset.toFinset_sum_count_eq] exact s.2 From c236dc6b66f0511502e39ff6a034b32427c155d1 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:04:24 +0200 Subject: [PATCH 055/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 6e6778fb9511e..edf07fdd11a10 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -472,7 +472,7 @@ theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by simp only [msymm._eq_1, Sym.val_eq_coe, map_sum] apply Fintype.sum_equiv (Nat.Partition.ofSym_shape_equiv μ e) - intro a + intro rw [← Multiset.prod_hom, Multiset.map_map] simp [Nat.Partition.ofSym_shape_equiv] From 98dce0c21089828b5aba409adf25fe20ddf7bc2d Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:04:38 +0200 Subject: [PATCH 056/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index a31c761626a3f..e85ffbab412ac 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -126,7 +126,7 @@ lemma ofSym_map (e : σ ≃ τ) (s : Sym σ n) : /-- An equivalence between `σ` and `τ` induces an equivalence between the subtypes of `Sym σ n` and `Sym τ n` corresponding to a given partition. -/ def ofSym_shape_equiv (μ : Partition n) (e : σ ≃ τ) : - {x : Sym σ n // Nat.Partition.ofSym x = μ} ≃ {x : Sym τ n // Nat.Partition.ofSym x = μ} where + {x : Sym σ n // ofSym x = μ} ≃ {x : Sym τ n // ofSym x = μ} where toFun := fun x => ⟨Sym.equivCongr e x, by simp [ofSym_map, x.2]⟩ invFun := fun x => ⟨Sym.equivCongr e.symm x, by simp [ofSym_map, x.2]⟩ left_inv := by intro x; simp From 08711988068cc1f1d7e9f767360906302bb1401b Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:04:46 +0200 Subject: [PATCH 057/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index e85ffbab412ac..3a7761be1bcb1 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -116,7 +116,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w variable {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] lemma ofSym_map (e : σ ≃ τ) (s : Sym σ n) : - Nat.Partition.ofSym (s.map e) = Nat.Partition.ofSym s := by + ofSym (s.map e) = ofSym s := by simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] rw [Multiset.dedup_map_of_injective e.injective] simp only [map_map, Function.comp_apply] From ef52886e3e0c15a2a2471620170a3f3821404a2a Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:05:19 +0200 Subject: [PATCH 058/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 3a7761be1bcb1..4275854e02af3 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -156,7 +156,7 @@ instance UniquePartitionZero : Unique (Partition 0) where instance UniquePartitionOne : Unique (Partition 1) where uniq _ := Partition.ext _ _ <| by simp -lemma ofSym_one (s : Sym σ 1) : Nat.Partition.ofSym s = Nat.Partition.indiscrete 1 := by +lemma ofSym_one (s : Sym σ 1) : ofSym s = indiscrete 1 := by ext; simp /-- The equivalence between `σ` and `1`-tuples of elements of σ -/ From d83c1216e16f9e3ad408f7f399d43980e1a436cf Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:05:40 +0200 Subject: [PATCH 059/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index edf07fdd11a10..7e1dc0f8a94d7 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -420,7 +420,7 @@ def psumMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (psum σ R) μ lemma psumMu_def {n : ℕ} (μ : n.Partition) : psumMu σ R μ = - Multiset.prod (μ.parts.map (psum σ R)) := rfl + (μ.parts.map (psum σ R)).prod := rfl @[simp] theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum]; rfl From 69d7db952fb989cd64927e211aa4dde5be3cf7be Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:05:54 +0200 Subject: [PATCH 060/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 7e1dc0f8a94d7..7d435040b0c95 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -435,7 +435,7 @@ theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by @[simp] theorem psumMu_onePart {n : ℕ} (npos : n > 0) : psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by - simp only [psumMu, npos, muProduct_indiscrete_of_pos] + simp [psumMu, npos] @[simp] theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by From 21d0e7432a57be752d51f85ef8d98b9cc72adb8d Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:06:13 +0200 Subject: [PATCH 061/106] Update Mathlib/Data/Multiset/Dedup.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Data/Multiset/Dedup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 10bb96faee172..57d938aea57b0 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -122,7 +122,7 @@ theorem dedup_map_dedup_eq [DecidableEq β] (f : α → β) (s : Multiset α) : simp [dedup_ext] #align multiset.dedup_map_dedup_eq Multiset.dedup_map_dedup_eq -theorem dedup_map_of_injective [DecidableEq β] {f : α → β} (hinj : Function.Injective f) +theorem dedup_map_of_injective [DecidableEq β] {f : α → β} (hinj : f.Injective) (s : Multiset α) : dedup (map f s) = map f (dedup s) := by rw [← dedup_map_dedup_eq] apply Multiset.dedup_eq_self.mpr From d248eeaa6d2ad4e9e6ac1b69275a0c6f7a59b246 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:06:34 +0200 Subject: [PATCH 062/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 4275854e02af3..d5d938af89db4 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -161,7 +161,7 @@ lemma ofSym_one (s : Sym σ 1) : ofSym s = indiscrete 1 := by /-- The equivalence between `σ` and `1`-tuples of elements of σ -/ def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ - { a : Sym σ 1 // Nat.Partition.ofSym a = Nat.Partition.indiscrete 1 } where + { a : Sym σ 1 // ofSym a = indiscrete 1 } where toFun := fun a => ⟨Sym.oneEquiv a, by ext; simp⟩ invFun := fun a => Sym.oneEquiv.symm a.1 left_inv := by intro a; simp only [Sym.oneEquiv_apply]; rfl From be785a6c553fa039cb35497cb110def1414efebc Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:06:58 +0200 Subject: [PATCH 063/106] Update Mathlib/Data/Multiset/Dedup.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Data/Multiset/Dedup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 57d938aea57b0..55306b7012ce4 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -125,7 +125,7 @@ theorem dedup_map_dedup_eq [DecidableEq β] (f : α → β) (s : Multiset α) : theorem dedup_map_of_injective [DecidableEq β] {f : α → β} (hinj : f.Injective) (s : Multiset α) : dedup (map f s) = map f (dedup s) := by rw [← dedup_map_dedup_eq] - apply Multiset.dedup_eq_self.mpr + apply dedup_eq_self.mpr rw [Multiset.nodup_map_iff_of_injective hinj] exact Multiset.nodup_dedup s From 1dd283f00486762ecc1cbc64b2e1686270c1d9f5 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:07:23 +0200 Subject: [PATCH 064/106] Update Mathlib/Data/Multiset/Dedup.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Data/Multiset/Dedup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 55306b7012ce4..d5a15f1b1d60e 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -126,7 +126,7 @@ theorem dedup_map_of_injective [DecidableEq β] {f : α → β} (hinj : f.Inject (s : Multiset α) : dedup (map f s) = map f (dedup s) := by rw [← dedup_map_dedup_eq] apply dedup_eq_self.mpr - rw [Multiset.nodup_map_iff_of_injective hinj] + rw [nodup_map_iff_of_injective hinj] exact Multiset.nodup_dedup s @[simp] From f8d7f16ccc6a6cb69b0a993aba8fbab46f7d6f87 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:07:52 +0200 Subject: [PATCH 065/106] Update Mathlib/Data/Sym/Basic.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Data/Sym/Basic.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 2d76dc664caa2..8c0a0a4fb467f 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -535,9 +535,7 @@ theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ Multiset.mem_add #align sym.mem_append_iff Sym.mem_append_iff -/-- -Defines an equivalence between `α` and `Sym α 1`. --/ +/-- Defines an equivalence between `α` and `Sym α 1`. -/ @[simps apply] def oneEquiv : α ≃ Sym α 1 where toFun a := ⟨{a}, by simp⟩ From e3fa73753399fb69f2acbf5b1f4e966e07527200 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:08:06 +0200 Subject: [PATCH 066/106] Update Mathlib/Data/Multiset/Dedup.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Data/Multiset/Dedup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index d5a15f1b1d60e..86ec35b420ae4 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -127,7 +127,7 @@ theorem dedup_map_of_injective [DecidableEq β] {f : α → β} (hinj : f.Inject rw [← dedup_map_dedup_eq] apply dedup_eq_self.mpr rw [nodup_map_iff_of_injective hinj] - exact Multiset.nodup_dedup s + exact nodup_dedup s @[simp] theorem dedup_nsmul {s : Multiset α} {n : ℕ} (h0 : n ≠ 0) : (n • s).dedup = s.dedup := by From d0cad776c470b27c45ba15ddcd09ab24ec5b2259 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:08:34 +0200 Subject: [PATCH 067/106] Update Mathlib/Data/Sym/Basic.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Data/Sym/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 8c0a0a4fb467f..b939ea9f4f1bf 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -539,8 +539,8 @@ theorem mem_append_iff {s' : Sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ @[simps apply] def oneEquiv : α ≃ Sym α 1 where toFun a := ⟨{a}, by simp⟩ - invFun s := Quotient.liftOn (Equiv.subtypeQuotientEquivQuotientSubtype - (·.length = 1) _ (fun l ↦ Iff.rfl) (fun l l' ↦ by rfl) s) + invFun s := (Equiv.subtypeQuotientEquivQuotientSubtype + (·.length = 1) _ (fun l ↦ Iff.rfl) (fun l l' ↦ by rfl) s).liftOn (fun l ↦ l.1.head <| List.length_pos.mp <| by simp) fun ⟨l, h⟩ ⟨l', h'⟩ _perm ↦ by obtain ⟨a, rfl⟩ := List.length_eq_one.mp h' From 0b47a3abe08567d21464d95247fe26fdd541bc48 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:08:58 +0200 Subject: [PATCH 068/106] Update Mathlib/Data/Sym/Basic.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Data/Sym/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index b939ea9f4f1bf..c2bd7249e402e 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -542,8 +542,8 @@ def oneEquiv : α ≃ Sym α 1 where invFun s := (Equiv.subtypeQuotientEquivQuotientSubtype (·.length = 1) _ (fun l ↦ Iff.rfl) (fun l l' ↦ by rfl) s).liftOn (fun l ↦ l.1.head <| List.length_pos.mp <| by simp) - fun ⟨l, h⟩ ⟨l', h'⟩ _perm ↦ by - obtain ⟨a, rfl⟩ := List.length_eq_one.mp h' + fun ⟨_, _⟩ ⟨_, h⟩ _perm ↦ by + obtain ⟨a, rfl⟩ := List.length_eq_one.mp h exact List.eq_of_mem_singleton (_perm.mem_iff.mp <| List.head_mem _) left_inv a := by rfl right_inv := by rintro ⟨⟨l⟩, h⟩; obtain ⟨a, rfl⟩ := List.length_eq_one.mp h; rfl From 686339e3340f5d06ad6b9ecbd7347ebb4e81fa83 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:10:01 +0200 Subject: [PATCH 069/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 7d435040b0c95..15bfd53f7d96c 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -6,9 +6,7 @@ Authors: Hanting Zhang, Johan Commelin import Mathlib.Algebra.Algebra.Subalgebra.Basic import Mathlib.Algebra.MvPolynomial.CommRing -import Mathlib.Algebra.MvPolynomial.Rename import Mathlib.Combinatorics.Enumerative.Partition -import Mathlib.Data.Sym.Basic #align_import ring_theory.mv_polynomial.symmetric from "leanprover-community/mathlib"@"2f5b500a507264de86d666a5f87ddb976e2d8de4" From a5303ceb688170a612f16cc2265f22ac4d7bd820 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:10:18 +0200 Subject: [PATCH 070/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 15bfd53f7d96c..0943fed520821 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -196,7 +196,7 @@ variable {n : ℕ} (f : ℕ → MvPolynomial σ R) /-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, `muProduct` computes the product of applying each function in `f` to the parts of `μ`. -/ -def muProduct {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : MvPolynomial σ R := +def muProduct (μ : n.Partition) : MvPolynomial σ R := Multiset.prod (μ.parts.map f) lemma muProduct_def (μ : n.Partition) : From 01bd97b89e236a5585f5d4c96a80982b48693909 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:10:32 +0200 Subject: [PATCH 071/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 0943fed520821..57a30f0528a75 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -197,7 +197,7 @@ variable {n : ℕ} (f : ℕ → MvPolynomial σ R) /-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, `muProduct` computes the product of applying each function in `f` to the parts of `μ`. -/ def muProduct (μ : n.Partition) : MvPolynomial σ R := - Multiset.prod (μ.parts.map f) + (μ.parts.map f).prod lemma muProduct_def (μ : n.Partition) : muProduct σ R f μ = Multiset.prod (μ.parts.map f) := rfl From 914f69e4875a54eaf0cf961e3c20c91771ca78cf Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:10:43 +0200 Subject: [PATCH 072/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 57a30f0528a75..a79929816d320 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -200,7 +200,7 @@ def muProduct (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map f).prod lemma muProduct_def (μ : n.Partition) : - muProduct σ R f μ = Multiset.prod (μ.parts.map f) := rfl + muProduct σ R f μ = (μ.parts.map f).prod := rfl @[simp] theorem muProduct_indiscrete_zero : From 29de762fc8ac5b26319dfdbfc59169d7ba97ece5 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:10:57 +0200 Subject: [PATCH 073/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index a79929816d320..03dcb8535393a 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -371,7 +371,7 @@ where `μ = (μ₁, μ₂, ...)` is a partition. def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ lemma hsymmMu_def {n : ℕ} (μ : n.Partition) : hsymmMu σ R μ = - Multiset.prod (μ.parts.map (hsymm σ R)) := rfl + (μ.parts.map (hsymm σ R)).prod := rfl @[simp] theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] From a78b7ca58effdc09644c2a0b633def8f2c445592 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:11:14 +0200 Subject: [PATCH 074/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 03dcb8535393a..58662ef166729 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -364,10 +364,8 @@ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl -/-- -`hsymmMu` is the product of the symmetric polynomials `hsymm μᵢ`, -where `μ = (μ₁, μ₂, ...)` is a partition. --/ +/-- `hsymmMu` is the product of the symmetric polynomials `hsymm μᵢ`, +where `μ = (μ₁, μ₂, ...)` is a partition. -/ def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ lemma hsymmMu_def {n : ℕ} (μ : n.Partition) : hsymmMu σ R μ = From c62773b2ba8bed9a2bc468c5c4210a6cadcab44f Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:11:31 +0200 Subject: [PATCH 075/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 58662ef166729..eda002009502e 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -408,10 +408,8 @@ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl -/-- -`psumMu` is the product of the symmetric polynomials `psum μᵢ`, -where `μ = (μ₁, μ₂, ...)` is a partition. --/ +/-- `psumMu` is the product of the symmetric polynomials `psum μᵢ`, +where `μ = (μ₁, μ₂, ...)` is a partition. -/ def psumMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (psum σ R) μ From 9d4664a4600904fb02860b18872c3e703206b1bc Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:11:43 +0200 Subject: [PATCH 076/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Combinatorics/Enumerative/Partition.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index d5d938af89db4..515073effa30e 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ import Mathlib.Combinatorics.Enumerative.Composition -import Mathlib.Data.Multiset.Dedup import Mathlib.Data.Nat.Parity import Mathlib.Tactic.ApplyFun From e1d91a921787f11317aa2e96a449eec10a2c960d Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:11:56 +0200 Subject: [PATCH 077/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index eda002009502e..82937dc1dec3e 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -424,7 +424,7 @@ theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] @[simp] theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by - simp only [psumMu, muProduct_indiscrete_zero] + rw [psumMu, muProduct_indiscrete_zero] @[simp] theorem psumMu_onePart {n : ℕ} (npos : n > 0) : From 160272b33f95331318049ffebe66638a692178d6 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 11:38:23 +0200 Subject: [PATCH 078/106] Stylistic changes --- .../Combinatorics/Enumerative/Partition.lean | 2 +- .../RingTheory/MvPolynomial/Symmetric.lean | 22 ++++++++++--------- 2 files changed, 13 insertions(+), 11 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 515073effa30e..e1c9321631330 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -108,7 +108,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w parts_pos := by simp [Multiset.count_pos] parts_sum := by have : sum (map (fun a ↦ count a s.1) s.1.dedup) = - ∑ a ∈ toFinset s.1, count a s.1 := rfl + (toFinset s.1).sum fun a ↦ count a s.1 := rfl rw [this, Multiset.toFinset_sum_count_eq] exact s.2 diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 82937dc1dec3e..1c085b86ae1b0 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -376,8 +376,8 @@ theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] @[simp] theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by - simp only [hsymm, univ_unique] symm + rw [hsymm] apply Fintype.sum_equiv oneEquiv simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton, implies_true] @@ -441,34 +441,36 @@ end PowerSum section MonomialSymmetric +open Nat.Partition + variable [DecidableEq σ] [DecidableEq τ] /-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. -/ def msymm {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := - ∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod + ∑ s : {a : Sym σ n // ofSym a = μ}, (s.1.1.map X).prod lemma msymm_def {n : ℕ} (μ : n.Partition) : msymm σ R μ = - ∑ s : {a : Sym σ n // Nat.Partition.ofSym a = μ}, (s.1.1.map X).prod := rfl + ∑ s : {a : Sym σ n // ofSym a = μ}, (s.1.1.map X).prod := rfl @[simp] -theorem msymm_zero : msymm σ R (Nat.Partition.indiscrete 0) = 1 := by +theorem msymm_zero : msymm σ R (indiscrete 0) = 1 := by rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), by rfl⟩] simp @[simp] -theorem msymm_one : msymm σ R (Nat.Partition.indiscrete 1) = ∑ i, X i := by +theorem msymm_one : msymm σ R (indiscrete 1) = ∑ i, X i := by symm - apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_onePart σ) + apply Fintype.sum_equiv (ofSym_equiv_onePart σ) simp @[simp] theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by - simp only [msymm._eq_1, Sym.val_eq_coe, map_sum] - apply Fintype.sum_equiv (Nat.Partition.ofSym_shape_equiv μ e) + rw [msymm._eq_1, map_sum] + apply Fintype.sum_equiv (ofSym_shape_equiv μ e) intro - rw [← Multiset.prod_hom, Multiset.map_map] - simp [Nat.Partition.ofSym_shape_equiv] + rw [← Multiset.prod_hom, Multiset.map_map, ofSym_shape_equiv] + simp theorem msymm_isSymmetric {n : ℕ} (μ : n.Partition) : IsSymmetric (msymm σ R μ) := rename_msymm _ _ μ From cd3d2f730a6943270da5ecd6b5f3a5eab0556fff Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 13 May 2024 10:22:12 +0000 Subject: [PATCH 079/106] use notation --- Mathlib/Combinatorics/Enumerative/Partition.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 4550c7898bb98..90cd5b4fcdde2 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -107,8 +107,7 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w parts := s.1.dedup.map s.1.count parts_pos := by simp [Multiset.count_pos] parts_sum := by - have : sum (map (fun a ↦ count a s.1) s.1.dedup) = - (toFinset s.1).sum fun a ↦ count a s.1 := rfl + have : sum (map (fun a ↦ count a s.1) s.1.dedup) = ∑ a ∈ s.1.toFinset, count a s.1 := rfl rw [this, Multiset.toFinset_sum_count_eq] exact s.2 From 9869d7927cbacc3b7424a94593e18e8bc8ce680e Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 14:13:41 +0200 Subject: [PATCH 080/106] Fixes --- Mathlib/Combinatorics/Enumerative/Partition.lean | 4 ++-- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 8 +++----- 2 files changed, 5 insertions(+), 7 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 90cd5b4fcdde2..3fc0595aea4ee 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -107,8 +107,8 @@ def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition w parts := s.1.dedup.map s.1.count parts_pos := by simp [Multiset.count_pos] parts_sum := by - have : sum (map (fun a ↦ count a s.1) s.1.dedup) = ∑ a ∈ s.1.toFinset, count a s.1 := rfl - rw [this, Multiset.toFinset_sum_count_eq] + show ∑ a ∈ s.1.toFinset, count a s.1 = n + rw [toFinset_sum_count_eq] exact s.2 variable {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 9471007a8f04c..62226712a77a9 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -377,7 +377,6 @@ theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] @[simp] theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by symm - rw [hsymm] apply Fintype.sum_equiv oneEquiv simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton, implies_true] @@ -417,7 +416,7 @@ lemma psumMu_def {n : ℕ} (μ : n.Partition) : psumMu σ R μ = (μ.parts.map (psum σ R)).prod := rfl @[simp] -theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum]; rfl +theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum] @[simp] theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] @@ -428,8 +427,7 @@ theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by @[simp] theorem psumMu_onePart {n : ℕ} (npos : n > 0) : - psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by - simp [psumMu, npos] + psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by simp [psumMu, npos] @[simp] theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by @@ -466,7 +464,7 @@ theorem msymm_one : msymm σ R (indiscrete 1) = ∑ i, X i := by @[simp] theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by - rw [msymm._eq_1, map_sum] + rw [msymm, map_sum] apply Fintype.sum_equiv (ofSym_shape_equiv μ e) intro rw [← Multiset.prod_hom, Multiset.map_map, ofSym_shape_equiv] From 73dae1d0b5f2a25188b99b04b7416089f3d0229b Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 14:43:03 +0200 Subject: [PATCH 081/106] Add some variables --- .../RingTheory/MvPolynomial/Symmetric.lean | 73 +++++++++---------- 1 file changed, 36 insertions(+), 37 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 62226712a77a9..cc61233955ceb 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -82,7 +82,7 @@ end Multiset namespace MvPolynomial -variable {σ τ : Type*} {R S : Type*} +variable (n : ℕ) {σ τ : Type*} {R S : Type*} /-- A `MvPolynomial φ` is symmetric if it is invariant under permutations of its variables by the `rename` operation -/ @@ -186,17 +186,17 @@ def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : (AlgHom.ext <| fun p => Subtype.ext <| by simp) variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] - +variable (n : ℕ) section Partitions /-! ### Multiplicativity on partitions -/ -variable {n : ℕ} (f : ℕ → MvPolynomial σ R) +variable (f : ℕ → MvPolynomial σ R) /-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, `muProduct` computes the product of applying each function in `f` to the parts of `μ`. -/ -def muProduct (μ : n.Partition) : MvPolynomial σ R := +def muProduct {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map f).prod lemma muProduct_def (μ : n.Partition) : @@ -219,11 +219,11 @@ section ElementarySymmetric open Finset /-- The `n`th elementary symmetric `MvPolynomial σ R`. -/ -def esymm (n : ℕ) : MvPolynomial σ R := +def esymm: MvPolynomial σ R := ∑ t in powersetCard n univ, ∏ i in t, X i #align mv_polynomial.esymm MvPolynomial.esymm -lemma esymm_def (n : ℕ) : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl +lemma esymm_def : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl /-- `esymmMu` is the product of the symmetric polynomials `esymm μᵢ`, @@ -238,19 +238,19 @@ theorem esymm_eq_multiset_esymm : esymm σ R = (univ.val.map X).esymm := by exact funext fun n => (esymm_map_val X _ n).symm #align mv_polynomial.esymm_eq_multiset_esymm MvPolynomial.esymm_eq_multiset_esymm -theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (f : σ → S) (n : ℕ) : +theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (f : σ → S) : aeval f (esymm σ R n) = (univ.val.map f).esymm n := by simp_rw [esymm, aeval_sum, aeval_prod, aeval_X, esymm_map_val] #align mv_polynomial.aeval_esymm_eq_multiset_esymm MvPolynomial.aeval_esymm_eq_multiset_esymm /-- We can define `esymm σ R n` by summing over a subtype instead of over `powerset_len`. -/ -theorem esymm_eq_sum_subtype (n : ℕ) : +theorem esymm_eq_sum_subtype : esymm σ R n = ∑ t : { s : Finset σ // s.card = n }, ∏ i in (t : Finset σ), X i := sum_subtype _ (fun _ => mem_powersetCard_univ) _ #align mv_polynomial.esymm_eq_sum_subtype MvPolynomial.esymm_eq_sum_subtype /-- We can define `esymm σ R n` as a sum over explicit monomials -/ -theorem esymm_eq_sum_monomial (n : ℕ) : +theorem esymm_eq_sum_monomial : esymm σ R n = ∑ t in powersetCard n univ, monomial (∑ i in t, Finsupp.single i 1) 1 := by simp_rw [monomial_sum_one] rfl @@ -266,14 +266,14 @@ theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp [esymm, powersetCard_on theorem esymmMu_zero : esymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp [esymmMu] @[simp] -theorem esymmMu_onePart (n : ℕ) : esymmMu σ R (Nat.Partition.indiscrete n) = esymm σ R n := by +theorem esymmMu_onePart : esymmMu σ R (Nat.Partition.indiscrete n) = esymm σ R n := by cases n <;> simp [esymmMu] -theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by +theorem map_esymm (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by simp_rw [esymm, map_sum, map_prod, map_X] #align mv_polynomial.map_esymm MvPolynomial.map_esymm -theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm σ R n) = esymm τ R n := +theorem rename_esymm (e : σ ≃ τ) : rename e (esymm σ R n) = esymm τ R n := calc rename e (esymm σ R n) = ∑ x in powersetCard n univ, ∏ i in x, X (e i) := by simp_rw [esymm, map_sum, map_prod, rename_X] @@ -285,12 +285,10 @@ theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm σ R n) = esymm _ = ∑ t in powersetCard n univ, ∏ i in t, X i := by rw [map_univ_equiv] #align mv_polynomial.rename_esymm MvPolynomial.rename_esymm -theorem esymm_isSymmetric (n : ℕ) : IsSymmetric (esymm σ R n) := by - intro - rw [rename_esymm] +theorem esymm_isSymmetric : IsSymmetric (esymm σ R n) := rename_esymm _ _ n #align mv_polynomial.esymm_is_symmetric MvPolynomial.esymm_isSymmetric -theorem support_esymm'' (n : ℕ) [DecidableEq σ] [Nontrivial R] : +theorem support_esymm'' [DecidableEq σ] [Nontrivial R] : (esymm σ R n).support = (powersetCard n (univ : Finset σ)).biUnion fun t => (Finsupp.single (∑ i : σ in t, Finsupp.single i 1) (1 : R)).support := by @@ -315,7 +313,7 @@ theorem support_esymm'' (n : ℕ) [DecidableEq σ] [Nontrivial R] : all_goals intro x y; simp [Finsupp.support_single_disjoint] #align mv_polynomial.support_esymm'' MvPolynomial.support_esymm'' -theorem support_esymm' (n : ℕ) [DecidableEq σ] [Nontrivial R] : +theorem support_esymm' [DecidableEq σ] [Nontrivial R] : (esymm σ R n).support = (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i : σ in t, Finsupp.single i 1} := by rw [support_esymm''] @@ -324,14 +322,14 @@ theorem support_esymm' (n : ℕ) [DecidableEq σ] [Nontrivial R] : exact Finsupp.support_single_ne_zero _ one_ne_zero #align mv_polynomial.support_esymm' MvPolynomial.support_esymm' -theorem support_esymm (n : ℕ) [DecidableEq σ] [Nontrivial R] : +theorem support_esymm [DecidableEq σ] [Nontrivial R] : (esymm σ R n).support = (powersetCard n (univ : Finset σ)).image fun t => ∑ i : σ in t, Finsupp.single i 1 := by rw [support_esymm'] exact biUnion_singleton #align mv_polynomial.support_esymm MvPolynomial.support_esymm -theorem degrees_esymm [Nontrivial R] (n : ℕ) (hpos : 0 < n) (hn : n ≤ Fintype.card σ) : +theorem degrees_esymm [Nontrivial R] (hpos : 0 < n) (hn : n ≤ Fintype.card σ) : (esymm σ R n).degrees = (univ : Finset σ).val := by classical have : @@ -357,19 +355,18 @@ section CompleteHomogeneousSymmetric open Finset Multiset Sym -variable [DecidableEq σ] [DecidableEq τ] +variable [DecidableEq σ] [DecidableEq τ] (n : ℕ) /-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ -def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod +def hsymm : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod -lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl +lemma hsymm_def : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl /-- `hsymmMu` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ -lemma hsymmMu_def {n : ℕ} (μ : n.Partition) : hsymmMu σ R μ = - (μ.parts.map (hsymm σ R)).prod := rfl +lemma hsymmMu_def {n : ℕ} (μ : n.Partition) : hsymmMu σ R μ = (μ.parts.map (hsymm σ R)).prod := rfl @[simp] theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] @@ -383,18 +380,18 @@ theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by theorem hsymmMu_zero : hsymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp [hsymmMu] @[simp] -theorem hsymmMu_onePart (n : ℕ) : hsymmMu σ R (Nat.Partition.indiscrete n) = hsymm σ R n := by +theorem hsymmMu_onePart : hsymmMu σ R (Nat.Partition.indiscrete n) = hsymm σ R n := by cases n <;> simp [hsymmMu] -theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by +theorem map_hsymm (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp [hsymm, ← Multiset.prod_hom'] -theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by +theorem rename_hsymm (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by simp_rw [hsymm, map_sum, ← prod_hom', rename_X] apply Fintype.sum_equiv (equivCongr e) simp -theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n +theorem hsymm_isSymmetric : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n end CompleteHomogeneousSymmetric @@ -402,10 +399,12 @@ section PowerSum open Finset +variable (n : ℕ) + /-- The degree-`n` power sum -/ -def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n +def psum : MvPolynomial σ R := ∑ i, X i ^ n -lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl +lemma psum_def : psum σ R n = ∑ i, X i ^ n := rfl /-- `psumMu` is the product of the symmetric polynomials `psum μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ @@ -430,10 +429,10 @@ theorem psumMu_onePart {n : ℕ} (npos : n > 0) : psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by simp [psumMu, npos] @[simp] -theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by +theorem rename_psum (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by simp_rw [psum, map_sum, map_pow, rename_X, e.sum_comp (X · ^ n)] -theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum σ R n) := rename_psum _ _ n +theorem psum_isSymmetric : IsSymmetric (psum σ R n) := rename_psum _ _ n end PowerSum @@ -441,13 +440,13 @@ section MonomialSymmetric open Nat.Partition -variable [DecidableEq σ] [DecidableEq τ] +variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} (μ : n.Partition) /-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. -/ -def msymm {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := +def msymm : MvPolynomial σ R := ∑ s : {a : Sym σ n // ofSym a = μ}, (s.1.1.map X).prod -lemma msymm_def {n : ℕ} (μ : n.Partition) : msymm σ R μ = +lemma msymm_def : msymm σ R μ = ∑ s : {a : Sym σ n // ofSym a = μ}, (s.1.1.map X).prod := rfl @[simp] @@ -462,7 +461,7 @@ theorem msymm_one : msymm σ R (indiscrete 1) = ∑ i, X i := by simp @[simp] -theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : +theorem rename_msymm (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by rw [msymm, map_sum] apply Fintype.sum_equiv (ofSym_shape_equiv μ e) @@ -470,7 +469,7 @@ theorem rename_msymm {n : ℕ} (μ : n.Partition) (e : σ ≃ τ) : rw [← Multiset.prod_hom, Multiset.map_map, ofSym_shape_equiv] simp -theorem msymm_isSymmetric {n : ℕ} (μ : n.Partition) : IsSymmetric (msymm σ R μ) := +theorem msymm_isSymmetric : IsSymmetric (msymm σ R μ) := rename_msymm _ _ μ end MonomialSymmetric From 40a90fa69a2472032b5a589d725ba2b9a6df685b Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 14:45:00 +0200 Subject: [PATCH 082/106] Minor --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index cc61233955ceb..9440f59f3e94c 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -186,7 +186,6 @@ def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : (AlgHom.ext <| fun p => Subtype.ext <| by simp) variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] -variable (n : ℕ) section Partitions @@ -218,6 +217,8 @@ section ElementarySymmetric open Finset +variable (n : ℕ) + /-- The `n`th elementary symmetric `MvPolynomial σ R`. -/ def esymm: MvPolynomial σ R := ∑ t in powersetCard n univ, ∏ i in t, X i From 3d101ea87cf5a59e7063e3379b9125aa36642483 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 13 May 2024 21:00:17 +0200 Subject: [PATCH 083/106] Update Mathlib/Combinatorics/Enumerative/Partition.lean Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 3fc0595aea4ee..7df1f69e33413 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -167,7 +167,7 @@ def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ @[simp] lemma ofSym_equiv_onePart_apply (i : σ) : - ((Nat.Partition.ofSym_equiv_onePart σ) i : Multiset σ) = {i} := rfl + ((ofSym_equiv_onePart σ) i : Multiset σ) = {i} := rfl /-- The number of times a positive integer `i` appears in the partition `ofSums n l hl` is the same as the number of times it appears in the multiset `l`. From eac3cf9ad92733e702b26a8338e0f24d8313c7a7 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 16 Jun 2024 13:44:40 +0200 Subject: [PATCH 084/106] Update Mathlib/Data/Sym/Basic.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Data/Sym/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 111861df9f6f2..b94a068cbc28f 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -546,9 +546,9 @@ def oneEquiv : α ≃ Sym α 1 where invFun s := (Equiv.subtypeQuotientEquivQuotientSubtype (·.length = 1) _ (fun l ↦ Iff.rfl) (fun l l' ↦ by rfl) s).liftOn (fun l ↦ l.1.head <| List.length_pos.mp <| by simp) - fun ⟨_, _⟩ ⟨_, h⟩ _perm ↦ by + fun ⟨_, _⟩ ⟨_, h⟩ ↦ fun perm ↦ by obtain ⟨a, rfl⟩ := List.length_eq_one.mp h - exact List.eq_of_mem_singleton (_perm.mem_iff.mp <| List.head_mem _) + exact List.eq_of_mem_singleton (perm.mem_iff.mp <| List.head_mem _) left_inv a := by rfl right_inv := by rintro ⟨⟨l⟩, h⟩; obtain ⟨a, rfl⟩ := List.length_eq_one.mp h; rfl From 67b1910e0ff35e622614114324f9e57a1699bf6e Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 17 Jun 2024 11:41:39 +0200 Subject: [PATCH 085/106] Update Mathlib/Data/Multiset/Dedup.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/Data/Multiset/Dedup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 86ec35b420ae4..7e40dcbd9451c 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -122,7 +122,7 @@ theorem dedup_map_dedup_eq [DecidableEq β] (f : α → β) (s : Multiset α) : simp [dedup_ext] #align multiset.dedup_map_dedup_eq Multiset.dedup_map_dedup_eq -theorem dedup_map_of_injective [DecidableEq β] {f : α → β} (hinj : f.Injective) +theorem dedup_map [DecidableEq β] {f : α → β} (hinj : f.Injective) (s : Multiset α) : dedup (map f s) = map f (dedup s) := by rw [← dedup_map_dedup_eq] apply dedup_eq_self.mpr From cf0f59febd58380ad1827f597e0be04eea80c6a9 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 17 Jun 2024 11:41:56 +0200 Subject: [PATCH 086/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 63551f3b0b07c..59b7cbc58d2e6 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -201,7 +201,7 @@ lemma muProduct_def (μ : n.Partition) : @[simp] theorem muProduct_indiscrete_zero : - muProduct σ R f (Nat.Partition.indiscrete 0) = 1 := by simp [muProduct] + muProduct σ R f (.indiscrete 0) = 1 := by simp [muProduct] @[simp] theorem muProduct_indiscrete_of_pos (npos : n > 0) : From df57d4a843fba97dd51f820b122aeb3fcf68346a Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 17 Jun 2024 11:44:33 +0200 Subject: [PATCH 087/106] Update Mathlib/RingTheory/MvPolynomial/Symmetric.lean MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 59b7cbc58d2e6..b0f6e7695dd6d 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -424,7 +424,7 @@ theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by rw [psumMu, muProduct_indiscrete_zero] @[simp] -theorem psumMu_onePart {n : ℕ} (npos : n > 0) : +theorem psumMu_indiscrete {n : ℕ} (npos : n > 0) : psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by simp [psumMu, npos] @[simp] From 81bd5600203178c2c739aafdf99371c2acf2f33a Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 10:32:28 +0200 Subject: [PATCH 088/106] Fix errors --- .../Combinatorics/Enumerative/Partition.lean | 2 +- .../RingTheory/MvPolynomial/Symmetric.lean | 34 ++++++++----------- 2 files changed, 16 insertions(+), 20 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 8f2309a1f4bc7..ba549b1f3201d 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -113,7 +113,7 @@ variable {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] lemma ofSym_map (e : σ ≃ τ) (s : Sym σ n) : ofSym (s.map e) = ofSym s := by simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] - rw [Multiset.dedup_map_of_injective e.injective] + rw [Multiset.dedup_map e.injective] simp only [map_map, Function.comp_apply] congr; funext i rw [← Multiset.count_map_eq_count' e _ e.injective] diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 56c6223dda11e..a6a4029872165 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Hanting Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Hanting Zhang, Johan Commelin -/ - import Mathlib.Algebra.Algebra.Subalgebra.Basic import Mathlib.Algebra.MvPolynomial.CommRing import Mathlib.Combinatorics.Enumerative.Partition @@ -193,8 +192,7 @@ variable (f : ℕ → MvPolynomial σ R) /-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, `muProduct` computes the product of applying each function in `f` to the parts of `μ`. -/ -def muProduct {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := - (μ.parts.map f).prod +def muProduct {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map f).prod lemma muProduct_def (μ : n.Partition) : muProduct σ R f μ = (μ.parts.map f).prod := rfl @@ -205,7 +203,7 @@ theorem muProduct_indiscrete_zero : @[simp] theorem muProduct_indiscrete_of_pos (npos : n > 0) : - muProduct σ R f (Nat.Partition.indiscrete n) = f n := by + muProduct σ R f (.indiscrete n) = f n := by rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] linarith @@ -262,10 +260,10 @@ theorem esymm_zero : esymm σ R 0 = 1 := by simp [esymm] @[simp] theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp [esymm, powersetCard_one] -theorem esymmMu_zero : esymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp [esymmMu] +theorem esymmMu_zero : esymmMu σ R (.indiscrete 0) = 1 := by simp [esymmMu] @[simp] -theorem esymmMu_onePart : esymmMu σ R (Nat.Partition.indiscrete n) = esymm σ R n := by +theorem esymmMu_onePart : esymmMu σ R (.indiscrete n) = esymm σ R n := by cases n <;> simp [esymmMu] theorem map_esymm (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by @@ -376,10 +374,10 @@ theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by apply Fintype.sum_equiv oneEquiv simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton, implies_true] -theorem hsymmMu_zero : hsymmMu σ R (Nat.Partition.indiscrete 0) = 1 := by simp [hsymmMu] +theorem hsymmMu_zero : hsymmMu σ R (.indiscrete 0) = 1 := by simp [hsymmMu] @[simp] -theorem hsymmMu_onePart : hsymmMu σ R (Nat.Partition.indiscrete n) = hsymm σ R n := by +theorem hsymmMu_onePart : hsymmMu σ R (.indiscrete n) = hsymm σ R n := by cases n <;> simp [hsymmMu] theorem map_hsymm (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by @@ -420,12 +418,12 @@ theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum] theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] @[simp] -theorem psumMu_zero : psumMu σ R (Nat.Partition.indiscrete 0) = 1 := by +theorem psumMu_zero : psumMu σ R (.indiscrete 0) = 1 := by rw [psumMu, muProduct_indiscrete_zero] @[simp] theorem psumMu_indiscrete {n : ℕ} (npos : n > 0) : - psumMu σ R (Nat.Partition.indiscrete n) = psum σ R n := by simp [psumMu, npos] + psumMu σ R (.indiscrete n) = psum σ R n := by simp [psumMu, npos] @[simp] theorem rename_psum (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by @@ -437,35 +435,33 @@ end PowerSum section MonomialSymmetric -open Nat.Partition - variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} (μ : n.Partition) /-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. -/ def msymm : MvPolynomial σ R := - ∑ s : {a : Sym σ n // ofSym a = μ}, (s.1.1.map X).prod + ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod lemma msymm_def : msymm σ R μ = - ∑ s : {a : Sym σ n // ofSym a = μ}, (s.1.1.map X).prod := rfl + ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod := rfl @[simp] -theorem msymm_zero : msymm σ R (indiscrete 0) = 1 := by +theorem msymm_zero : msymm σ R (.indiscrete 0) = 1 := by rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), by rfl⟩] simp @[simp] -theorem msymm_one : msymm σ R (indiscrete 1) = ∑ i, X i := by +theorem msymm_one : msymm σ R (.indiscrete 1) = ∑ i, X i := by symm - apply Fintype.sum_equiv (ofSym_equiv_onePart σ) + apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_onePart σ) simp @[simp] theorem rename_msymm (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by rw [msymm, map_sum] - apply Fintype.sum_equiv (ofSym_shape_equiv μ e) + apply Fintype.sum_equiv (Nat.Partition.ofSym_shape_equiv μ e) intro - rw [← Multiset.prod_hom, Multiset.map_map, ofSym_shape_equiv] + rw [← Multiset.prod_hom, Multiset.map_map, Nat.Partition.ofSym_shape_equiv] simp theorem msymm_isSymmetric : IsSymmetric (msymm σ R μ) := From 262fa8e5fe9afffa9966232943ef95546515d221 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 14:49:02 +0200 Subject: [PATCH 089/106] Merge --- Mathlib/Combinatorics/Enumerative/Partition.lean | 2 +- Mathlib/Data/Multiset/Dedup.lean | 7 ------- 2 files changed, 1 insertion(+), 8 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index ba549b1f3201d..8f2309a1f4bc7 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -113,7 +113,7 @@ variable {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] lemma ofSym_map (e : σ ≃ τ) (s : Sym σ n) : ofSym (s.map e) = ofSym s := by simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] - rw [Multiset.dedup_map e.injective] + rw [Multiset.dedup_map_of_injective e.injective] simp only [map_map, Function.comp_apply] congr; funext i rw [← Multiset.count_map_eq_count' e _ e.injective] diff --git a/Mathlib/Data/Multiset/Dedup.lean b/Mathlib/Data/Multiset/Dedup.lean index 8a6dfdafa77c7..fdcbb371bc59c 100644 --- a/Mathlib/Data/Multiset/Dedup.lean +++ b/Mathlib/Data/Multiset/Dedup.lean @@ -127,13 +127,6 @@ theorem dedup_map_dedup_eq [DecidableEq β] (f : α → β) (s : Multiset α) : simp [dedup_ext] #align multiset.dedup_map_dedup_eq Multiset.dedup_map_dedup_eq -theorem dedup_map [DecidableEq β] {f : α → β} (hinj : f.Injective) - (s : Multiset α) : dedup (map f s) = map f (dedup s) := by - rw [← dedup_map_dedup_eq] - apply dedup_eq_self.mpr - rw [nodup_map_iff_of_injective hinj] - exact nodup_dedup s - @[simp] theorem dedup_nsmul {s : Multiset α} {n : ℕ} (h0 : n ≠ 0) : (n • s).dedup = s.dedup := by ext a From f4c4f307a5368d420854d19bf821069d11f02969 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 17:59:05 +0200 Subject: [PATCH 090/106] Updates --- .../RingTheory/MvPolynomial/Symmetric.lean | 113 ++++++++---------- 1 file changed, 51 insertions(+), 62 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index a6a4029872165..bb7fe0728b1ad 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -79,7 +79,7 @@ end Multiset namespace MvPolynomial -variable (n : ℕ) {σ τ : Type*} {R S : Type*} +variable {σ τ : Type*} {R S : Type*} /-- A `MvPolynomial φ` is symmetric if it is invariant under permutations of its variables by the `rename` operation -/ @@ -87,18 +87,14 @@ def IsSymmetric [CommSemiring R] (φ : MvPolynomial σ R) : Prop := ∀ e : Perm σ, rename e φ = φ #align mv_polynomial.is_symmetric MvPolynomial.IsSymmetric -variable (σ R) - /-- The subalgebra of symmetric `MvPolynomial`s. -/ -def symmetricSubalgebra [CommSemiring R] : Subalgebra R (MvPolynomial σ R) where +def symmetricSubalgebra (σ R : Type*) [CommSemiring R] : Subalgebra R (MvPolynomial σ R) where carrier := setOf IsSymmetric algebraMap_mem' r e := rename_C e r mul_mem' ha hb e := by rw [map_mul, ha, hb] add_mem' ha hb e := by rw [map_add, ha, hb] #align mv_polynomial.symmetric_subalgebra MvPolynomial.symmetricSubalgebra -variable {σ R} - @[simp] theorem mem_symmetricSubalgebra [CommSemiring R] (p : MvPolynomial σ R) : p ∈ symmetricSubalgebra σ R ↔ p.IsSymmetric := @@ -188,21 +184,20 @@ section Partitions /-! ### Multiplicativity on partitions -/ -variable (f : ℕ → MvPolynomial σ R) - /-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, `muProduct` computes the product of applying each function in `f` to the parts of `μ`. -/ -def muProduct {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map f).prod +def muProduct {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : MvPolynomial σ R := + (μ.parts.map f).prod -lemma muProduct_def (μ : n.Partition) : +lemma muProduct_def {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : muProduct σ R f μ = (μ.parts.map f).prod := rfl @[simp] -theorem muProduct_indiscrete_zero : +theorem muProduct_indiscrete_zero (f : ℕ → MvPolynomial σ R) : muProduct σ R f (.indiscrete 0) = 1 := by simp [muProduct] @[simp] -theorem muProduct_indiscrete_of_pos (npos : n > 0) : +theorem muProduct_indiscrete_of_pos {n : ℕ} (npos : n > 0) (f : ℕ → MvPolynomial σ R) : muProduct σ R f (.indiscrete n) = f n := by rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] linarith @@ -213,20 +208,18 @@ section ElementarySymmetric open Finset -variable (n : ℕ) - /-- The `n`th elementary symmetric `MvPolynomial σ R`. -/ def esymm (n : ℕ) : MvPolynomial σ R := ∑ t ∈ powersetCard n univ, ∏ i ∈ t, X i #align mv_polynomial.esymm MvPolynomial.esymm -lemma esymm_def : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl +lemma esymm_def (n : ℕ) : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl /-- -`esymmMu` is the product of the symmetric polynomials `esymm μᵢ`, +`esymmPart` is the product of the symmetric polynomials `esymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def esymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := +def esymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (esymm σ R) μ /-- The `n`th elementary symmetric `MvPolynomial σ R` is obtained by evaluating the @@ -235,7 +228,7 @@ theorem esymm_eq_multiset_esymm : esymm σ R = (univ.val.map X).esymm := by exact funext fun n => (esymm_map_val X _ n).symm #align mv_polynomial.esymm_eq_multiset_esymm MvPolynomial.esymm_eq_multiset_esymm -theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (f : σ → S) : +theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (n : ℕ) (f : σ → S) : aeval f (esymm σ R n) = (univ.val.map f).esymm n := by simp_rw [esymm, aeval_sum, aeval_prod, aeval_X, esymm_map_val] #align mv_polynomial.aeval_esymm_eq_multiset_esymm MvPolynomial.aeval_esymm_eq_multiset_esymm @@ -260,17 +253,17 @@ theorem esymm_zero : esymm σ R 0 = 1 := by simp [esymm] @[simp] theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp [esymm, powersetCard_one] -theorem esymmMu_zero : esymmMu σ R (.indiscrete 0) = 1 := by simp [esymmMu] +theorem esymmPart_zero : esymmPart σ R (.indiscrete 0) = 1 := by simp [esymmPart] @[simp] -theorem esymmMu_onePart : esymmMu σ R (.indiscrete n) = esymm σ R n := by - cases n <;> simp [esymmMu] +theorem esymmPart_onePart (n : ℕ) : esymmPart σ R (.indiscrete n) = esymm σ R n := by + cases n <;> simp [esymmPart] -theorem map_esymm (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by +theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by simp_rw [esymm, map_sum, map_prod, map_X] #align mv_polynomial.map_esymm MvPolynomial.map_esymm -theorem rename_esymm (e : σ ≃ τ) : rename e (esymm σ R n) = esymm τ R n := +theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm σ R n) = esymm τ R n := calc rename e (esymm σ R n) = ∑ x ∈ powersetCard n univ, ∏ i ∈ x, X (e i) := by simp_rw [esymm, map_sum, map_prod, rename_X] @@ -282,10 +275,10 @@ theorem rename_esymm (e : σ ≃ τ) : rename e (esymm σ R n) = esymm τ R n := _ = ∑ t ∈ powersetCard n univ, ∏ i ∈ t, X i := by rw [map_univ_equiv] #align mv_polynomial.rename_esymm MvPolynomial.rename_esymm -theorem esymm_isSymmetric : IsSymmetric (esymm σ R n) := rename_esymm _ _ n +theorem esymm_isSymmetric (n : ℕ) : IsSymmetric (esymm σ R n) := rename_esymm _ _ n #align mv_polynomial.esymm_is_symmetric MvPolynomial.esymm_isSymmetric -theorem support_esymm'' [DecidableEq σ] [Nontrivial R] : +theorem support_esymm'' [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = (powersetCard n (univ : Finset σ)).biUnion fun t => (Finsupp.single (∑ i ∈ t, Finsupp.single i 1) (1 : R)).support := by @@ -310,23 +303,21 @@ theorem support_esymm'' [DecidableEq σ] [Nontrivial R] : all_goals intro x y; simp [Finsupp.support_single_disjoint] #align mv_polynomial.support_esymm'' MvPolynomial.support_esymm'' -theorem support_esymm' [DecidableEq σ] [Nontrivial R] : - (esymm σ R n).support = - (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i ∈ t, Finsupp.single i 1} := by +theorem support_esymm' [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = + (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i ∈ t, Finsupp.single i 1} := by rw [support_esymm''] congr funext exact Finsupp.support_single_ne_zero _ one_ne_zero #align mv_polynomial.support_esymm' MvPolynomial.support_esymm' -theorem support_esymm [DecidableEq σ] [Nontrivial R] : - (esymm σ R n).support = - (powersetCard n (univ : Finset σ)).image fun t => ∑ i ∈ t, Finsupp.single i 1 := by +theorem support_esymm [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = + (powersetCard n (univ : Finset σ)).image fun t => ∑ i ∈ t, Finsupp.single i 1 := by rw [support_esymm'] exact biUnion_singleton #align mv_polynomial.support_esymm MvPolynomial.support_esymm -theorem degrees_esymm [Nontrivial R] (hpos : 0 < n) (hn : n ≤ Fintype.card σ) : +theorem degrees_esymm [Nontrivial R] {n : ℕ} (hpos : 0 < n) (hn : n ≤ Fintype.card σ) : (esymm σ R n).degrees = (univ : Finset σ).val := by classical have : @@ -352,18 +343,18 @@ section CompleteHomogeneousSymmetric open Finset Multiset Sym -variable [DecidableEq σ] [DecidableEq τ] (n : ℕ) +variable [DecidableEq σ] [DecidableEq τ] /-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ -def hsymm : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod +def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod -lemma hsymm_def : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl +lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl -/-- `hsymmMu` is the product of the symmetric polynomials `hsymm μᵢ`, +/-- `hsymmPart` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def hsymmMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ +def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ -lemma hsymmMu_def {n : ℕ} (μ : n.Partition) : hsymmMu σ R μ = (μ.parts.map (hsymm σ R)).prod := rfl +lemma hsymmPart_def {n : ℕ} (μ : n.Partition) : hsymmPart σ R μ = (μ.parts.map (hsymm σ R)).prod := rfl @[simp] theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] @@ -374,21 +365,21 @@ theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by apply Fintype.sum_equiv oneEquiv simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton, implies_true] -theorem hsymmMu_zero : hsymmMu σ R (.indiscrete 0) = 1 := by simp [hsymmMu] +theorem hsymmPart_zero : hsymmPart σ R (.indiscrete 0) = 1 := by simp [hsymmPart] @[simp] -theorem hsymmMu_onePart : hsymmMu σ R (.indiscrete n) = hsymm σ R n := by - cases n <;> simp [hsymmMu] +theorem hsymmPart_onePart (n : ℕ) : hsymmPart σ R (.indiscrete n) = hsymm σ R n := by + cases n <;> simp [hsymmPart] -theorem map_hsymm (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by +theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp [hsymm, ← Multiset.prod_hom'] -theorem rename_hsymm (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by +theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by simp_rw [hsymm, map_sum, ← prod_hom', rename_X] apply Fintype.sum_equiv (equivCongr e) simp -theorem hsymm_isSymmetric : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n +theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n end CompleteHomogeneousSymmetric @@ -396,19 +387,17 @@ section PowerSum open Finset -variable (n : ℕ) - /-- The degree-`n` power sum -/ -def psum : MvPolynomial σ R := ∑ i, X i ^ n +def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n -lemma psum_def : psum σ R n = ∑ i, X i ^ n := rfl +lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl -/-- `psumMu` is the product of the symmetric polynomials `psum μᵢ`, +/-- `psumPart` is the product of the symmetric polynomials `psum μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def psumMu {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := +def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (psum σ R) μ -lemma psumMu_def {n : ℕ} (μ : n.Partition) : psumMu σ R μ = +lemma psumPart_def {n : ℕ} (μ : n.Partition) : psumPart σ R μ = (μ.parts.map (psum σ R)).prod := rfl @[simp] @@ -418,30 +407,30 @@ theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum] theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] @[simp] -theorem psumMu_zero : psumMu σ R (.indiscrete 0) = 1 := by - rw [psumMu, muProduct_indiscrete_zero] +theorem psumPart_zero : psumPart σ R (.indiscrete 0) = 1 := by + rw [psumPart, muProduct_indiscrete_zero] @[simp] -theorem psumMu_indiscrete {n : ℕ} (npos : n > 0) : - psumMu σ R (.indiscrete n) = psum σ R n := by simp [psumMu, npos] +theorem psumPart_indiscrete {n : ℕ} (npos : n > 0) : + psumPart σ R (.indiscrete n) = psum σ R n := by simp [psumPart, npos] @[simp] -theorem rename_psum (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by +theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by simp_rw [psum, map_sum, map_pow, rename_X, e.sum_comp (X · ^ n)] -theorem psum_isSymmetric : IsSymmetric (psum σ R n) := rename_psum _ _ n +theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum σ R n) := rename_psum _ _ n end PowerSum section MonomialSymmetric -variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} (μ : n.Partition) +variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} /-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. -/ -def msymm : MvPolynomial σ R := +def msymm (μ : n.Partition) : MvPolynomial σ R := ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod -lemma msymm_def : msymm σ R μ = +lemma msymm_def (μ : n.Partition) : msymm σ R μ = ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod := rfl @[simp] @@ -456,7 +445,7 @@ theorem msymm_one : msymm σ R (.indiscrete 1) = ∑ i, X i := by simp @[simp] -theorem rename_msymm (e : σ ≃ τ) : +theorem rename_msymm (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by rw [msymm, map_sum] apply Fintype.sum_equiv (Nat.Partition.ofSym_shape_equiv μ e) @@ -464,7 +453,7 @@ theorem rename_msymm (e : σ ≃ τ) : rw [← Multiset.prod_hom, Multiset.map_map, Nat.Partition.ofSym_shape_equiv] simp -theorem msymm_isSymmetric : IsSymmetric (msymm σ R μ) := +theorem msymm_isSymmetric (μ : n.Partition) : IsSymmetric (msymm σ R μ) := rename_msymm _ _ μ end MonomialSymmetric From bcf74cbf06108309cedf0c842845fddc44e767ec Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 18:03:53 +0200 Subject: [PATCH 091/106] npos --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index bb7fe0728b1ad..73772a873f9d4 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -197,10 +197,10 @@ theorem muProduct_indiscrete_zero (f : ℕ → MvPolynomial σ R) : muProduct σ R f (.indiscrete 0) = 1 := by simp [muProduct] @[simp] -theorem muProduct_indiscrete_of_pos {n : ℕ} (npos : n > 0) (f : ℕ → MvPolynomial σ R) : +theorem muProduct_indiscrete_of_pos {n : ℕ} (npos : n ≠ 0) (f : ℕ → MvPolynomial σ R) : muProduct σ R f (.indiscrete n) = f n := by rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] - linarith + exact npos end Partitions @@ -354,7 +354,8 @@ lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := where `μ = (μ₁, μ₂, ...)` is a partition. -/ def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ -lemma hsymmPart_def {n : ℕ} (μ : n.Partition) : hsymmPart σ R μ = (μ.parts.map (hsymm σ R)).prod := rfl +lemma hsymmPart_def {n : ℕ} (μ : n.Partition) : hsymmPart σ R μ = (μ.parts.map (hsymm σ R)).prod := + rfl @[simp] theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] @@ -411,7 +412,7 @@ theorem psumPart_zero : psumPart σ R (.indiscrete 0) = 1 := by rw [psumPart, muProduct_indiscrete_zero] @[simp] -theorem psumPart_indiscrete {n : ℕ} (npos : n > 0) : +theorem psumPart_indiscrete {n : ℕ} (npos : n ≠ 0) : psumPart σ R (.indiscrete n) = psum σ R n := by simp [psumPart, npos] @[simp] From 6a04bc8f9677193d75b15a2cdec261c2cf260e97 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 18:12:12 +0200 Subject: [PATCH 092/106] Got rid of products --- .../RingTheory/MvPolynomial/Symmetric.lean | 40 +++---------------- 1 file changed, 6 insertions(+), 34 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 73772a873f9d4..0e2cad0c2e51f 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -180,30 +180,6 @@ def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] -section Partitions - -/-! ### Multiplicativity on partitions -/ - -/-- Given a sequence of `MvPolynomial` functions `f` and a partition `μ` of size `n`, -`muProduct` computes the product of applying each function in `f` to the parts of `μ`. -/ -def muProduct {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : MvPolynomial σ R := - (μ.parts.map f).prod - -lemma muProduct_def {n : ℕ} (f : ℕ → MvPolynomial σ R) (μ : n.Partition) : - muProduct σ R f μ = (μ.parts.map f).prod := rfl - -@[simp] -theorem muProduct_indiscrete_zero (f : ℕ → MvPolynomial σ R) : - muProduct σ R f (.indiscrete 0) = 1 := by simp [muProduct] - -@[simp] -theorem muProduct_indiscrete_of_pos {n : ℕ} (npos : n ≠ 0) (f : ℕ → MvPolynomial σ R) : - muProduct σ R f (.indiscrete n) = f n := by - rw [muProduct, Nat.Partition.indiscrete_parts, Multiset.map_singleton, Multiset.prod_singleton] - exact npos - -end Partitions - section ElementarySymmetric open Finset @@ -219,8 +195,7 @@ lemma esymm_def (n : ℕ) : esymm σ R n = ∑ t in powersetCard n univ, ∏ i i `esymmPart` is the product of the symmetric polynomials `esymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def esymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := - muProduct σ R (esymm σ R) μ +def esymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (esymm σ R)).prod /-- The `n`th elementary symmetric `MvPolynomial σ R` is obtained by evaluating the `n`th elementary symmetric at the `Multiset` of the monomials -/ @@ -348,11 +323,11 @@ variable [DecidableEq σ] [DecidableEq τ] /-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod -lemma hsymm_def (n : ℕ) : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl +lemma hsymm_def {n : ℕ} : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl /-- `hsymmPart` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := muProduct σ R (hsymm σ R) μ +def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (hsymm σ R)).prod lemma hsymmPart_def {n : ℕ} (μ : n.Partition) : hsymmPart σ R μ = (μ.parts.map (hsymm σ R)).prod := rfl @@ -395,11 +370,9 @@ lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl /-- `psumPart` is the product of the symmetric polynomials `psum μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := - muProduct σ R (psum σ R) μ +def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (psum σ R)).prod -lemma psumPart_def {n : ℕ} (μ : n.Partition) : psumPart σ R μ = - (μ.parts.map (psum σ R)).prod := rfl +lemma psumPart_def {n : ℕ} (μ : n.Partition) : psumPart σ R μ = (μ.parts.map (psum σ R)).prod := rfl @[simp] theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum] @@ -408,8 +381,7 @@ theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum] theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] @[simp] -theorem psumPart_zero : psumPart σ R (.indiscrete 0) = 1 := by - rw [psumPart, muProduct_indiscrete_zero] +theorem psumPart_zero : psumPart σ R (.indiscrete 0) = 1 := by simp [psumPart] @[simp] theorem psumPart_indiscrete {n : ℕ} (npos : n ≠ 0) : From d6a9e77eb52e908c5d9864452ed4f36752bb3d15 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 18:28:39 +0200 Subject: [PATCH 093/106] Remove explicit variables --- .../RingTheory/MvPolynomial/Symmetric.lean | 104 +++++++++--------- 1 file changed, 51 insertions(+), 53 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 0e2cad0c2e51f..08b352213c30f 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -178,7 +178,7 @@ def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : (AlgHom.ext <| fun p => Subtype.ext <| by simp) (AlgHom.ext <| fun p => Subtype.ext <| by simp) -variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] +variable [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] section ElementarySymmetric @@ -189,58 +189,59 @@ def esymm (n : ℕ) : MvPolynomial σ R := ∑ t ∈ powersetCard n univ, ∏ i ∈ t, X i #align mv_polynomial.esymm MvPolynomial.esymm -lemma esymm_def (n : ℕ) : esymm σ R n = ∑ t in powersetCard n univ, ∏ i in t, X i := rfl +lemma esymm_def (n : ℕ) : + esymm n = ∑ t in powersetCard n univ, ∏ i in t, (X i : MvPolynomial σ R) := rfl /-- `esymmPart` is the product of the symmetric polynomials `esymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def esymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (esymm σ R)).prod +def esymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (esymm)).prod /-- The `n`th elementary symmetric `MvPolynomial σ R` is obtained by evaluating the `n`th elementary symmetric at the `Multiset` of the monomials -/ -theorem esymm_eq_multiset_esymm : esymm σ R = (univ.val.map X).esymm := by +theorem esymm_eq_multiset_esymm : esymm = (univ.val.map (X : σ → MvPolynomial σ R)).esymm := by exact funext fun n => (esymm_map_val X _ n).symm #align mv_polynomial.esymm_eq_multiset_esymm MvPolynomial.esymm_eq_multiset_esymm theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (n : ℕ) (f : σ → S) : - aeval f (esymm σ R n) = (univ.val.map f).esymm n := by + aeval f (esymm n : MvPolynomial σ R) = (univ.val.map f).esymm n := by simp_rw [esymm, aeval_sum, aeval_prod, aeval_X, esymm_map_val] #align mv_polynomial.aeval_esymm_eq_multiset_esymm MvPolynomial.aeval_esymm_eq_multiset_esymm /-- We can define `esymm σ R n` by summing over a subtype instead of over `powerset_len`. -/ -theorem esymm_eq_sum_subtype (n : ℕ) : - esymm σ R n = ∑ t : { s : Finset σ // s.card = n }, ∏ i ∈ (t : Finset σ), X i := +theorem esymm_eq_sum_subtype (n : ℕ) : esymm n = ∑ t : { s : Finset σ // s.card = n }, + ∏ i ∈ (t : Finset σ), (X i : MvPolynomial σ R) := sum_subtype _ (fun _ => mem_powersetCard_univ) _ #align mv_polynomial.esymm_eq_sum_subtype MvPolynomial.esymm_eq_sum_subtype /-- We can define `esymm σ R n` as a sum over explicit monomials -/ -theorem esymm_eq_sum_monomial (n : ℕ) : - esymm σ R n = ∑ t ∈ powersetCard n univ, monomial (∑ i ∈ t, Finsupp.single i 1) 1 := by +theorem esymm_eq_sum_monomial (n : ℕ) : (esymm n : MvPolynomial σ R) = ∑ t ∈ powersetCard n univ, + monomial (∑ i ∈ t, Finsupp.single i 1) 1 := by simp_rw [monomial_sum_one] rfl #align mv_polynomial.esymm_eq_sum_monomial MvPolynomial.esymm_eq_sum_monomial @[simp] -theorem esymm_zero : esymm σ R 0 = 1 := by simp [esymm] +theorem esymm_zero : esymm 0 = (1 : MvPolynomial σ R) := by simp [esymm] #align mv_polynomial.esymm_zero MvPolynomial.esymm_zero @[simp] -theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp [esymm, powersetCard_one] +theorem esymm_one : esymm 1 = ∑ i, (X i : MvPolynomial σ R) := by simp [esymm, powersetCard_one] -theorem esymmPart_zero : esymmPart σ R (.indiscrete 0) = 1 := by simp [esymmPart] +theorem esymmPart_zero : esymmPart (.indiscrete 0) = (1 : MvPolynomial σ R) := by simp [esymmPart] @[simp] -theorem esymmPart_onePart (n : ℕ) : esymmPart σ R (.indiscrete n) = esymm σ R n := by +theorem esymmPart_onePart (n : ℕ) : esymmPart (.indiscrete n) = (esymm n : MvPolynomial σ R) := by cases n <;> simp [esymmPart] -theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by +theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm n) = (esymm n : MvPolynomial σ S) := by simp_rw [esymm, map_sum, map_prod, map_X] #align mv_polynomial.map_esymm MvPolynomial.map_esymm -theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm σ R n) = esymm τ R n := +theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm n) = (esymm n : MvPolynomial τ R) := calc - rename e (esymm σ R n) = ∑ x ∈ powersetCard n univ, ∏ i ∈ x, X (e i) := by + rename e (esymm n) = ∑ x ∈ powersetCard n univ, ∏ i ∈ x, X (e i) := by simp_rw [esymm, map_sum, map_prod, rename_X] _ = ∑ t ∈ powersetCard n (univ.map e.toEmbedding), ∏ i ∈ t, X i := by simp [powersetCard_map, -map_univ_equiv] @@ -250,12 +251,11 @@ theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm σ R n) = esymm _ = ∑ t ∈ powersetCard n univ, ∏ i ∈ t, X i := by rw [map_univ_equiv] #align mv_polynomial.rename_esymm MvPolynomial.rename_esymm -theorem esymm_isSymmetric (n : ℕ) : IsSymmetric (esymm σ R n) := rename_esymm _ _ n +theorem esymm_isSymmetric (n : ℕ) : IsSymmetric (esymm n : MvPolynomial σ R) := rename_esymm n #align mv_polynomial.esymm_is_symmetric MvPolynomial.esymm_isSymmetric theorem support_esymm'' [DecidableEq σ] [Nontrivial R] (n : ℕ) : - (esymm σ R n).support = - (powersetCard n (univ : Finset σ)).biUnion fun t => + (esymm n : MvPolynomial σ R).support = (powersetCard n (univ : Finset σ)).biUnion fun t => (Finsupp.single (∑ i ∈ t, Finsupp.single i 1) (1 : R)).support := by rw [esymm_eq_sum_monomial] simp only [← single_eq_monomial] @@ -278,22 +278,23 @@ theorem support_esymm'' [DecidableEq σ] [Nontrivial R] (n : ℕ) : all_goals intro x y; simp [Finsupp.support_single_disjoint] #align mv_polynomial.support_esymm'' MvPolynomial.support_esymm'' -theorem support_esymm' [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = - (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i ∈ t, Finsupp.single i 1} := by +theorem support_esymm' [DecidableEq σ] [Nontrivial R] (n : ℕ) : + (esymm n : MvPolynomial σ R).support = + (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i ∈ t, Finsupp.single i 1} := by rw [support_esymm''] congr funext exact Finsupp.support_single_ne_zero _ one_ne_zero #align mv_polynomial.support_esymm' MvPolynomial.support_esymm' -theorem support_esymm [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = +theorem support_esymm [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm n : MvPolynomial σ R).support = (powersetCard n (univ : Finset σ)).image fun t => ∑ i ∈ t, Finsupp.single i 1 := by rw [support_esymm'] exact biUnion_singleton #align mv_polynomial.support_esymm MvPolynomial.support_esymm -theorem degrees_esymm [Nontrivial R] {n : ℕ} (hpos : 0 < n) (hn : n ≤ Fintype.card σ) : - (esymm σ R n).degrees = (univ : Finset σ).val := by +theorem degrees_esymm [Nontrivial R] {n : ℕ} (hpos : n ≠ 0) (hn : n ≤ Fintype.card σ) : + (esymm n : MvPolynomial σ R).degrees = (univ : Finset σ).val := by classical have : (Finsupp.toMultiset ∘ fun t : Finset σ => ∑ i ∈ t, Finsupp.single i 1) = val := by @@ -308,7 +309,7 @@ theorem degrees_esymm [Nontrivial R] {n : ℕ} (hpos : 0 < n) (hn : n ≤ Fintyp congr · rfl rw [← this] - obtain ⟨k, rfl⟩ := Nat.exists_eq_succ_of_ne_zero hpos.ne' + obtain ⟨k, rfl⟩ := Nat.exists_eq_succ_of_ne_zero hpos simpa using powersetCard_sup _ _ (Nat.lt_of_succ_le hn) #align mv_polynomial.degrees_esymm MvPolynomial.degrees_esymm @@ -323,39 +324,37 @@ variable [DecidableEq σ] [DecidableEq τ] /-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod -lemma hsymm_def {n : ℕ} : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl - /-- `hsymmPart` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (hsymm σ R)).prod +def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map hsymm).prod -lemma hsymmPart_def {n : ℕ} (μ : n.Partition) : hsymmPart σ R μ = (μ.parts.map (hsymm σ R)).prod := - rfl +lemma hsymmPart_def {n : ℕ} (μ : n.Partition) : + (hsymmPart μ : MvPolynomial σ R) = (μ.parts.map hsymm).prod := rfl @[simp] -theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] +theorem hsymm_zero : hsymm 0 = (1 : MvPolynomial σ R) := by simp [hsymm, eq_nil_of_card_zero] @[simp] -theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by +theorem hsymm_one : hsymm 1 = ∑ i, (X i : MvPolynomial σ R) := by symm apply Fintype.sum_equiv oneEquiv simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton, implies_true] -theorem hsymmPart_zero : hsymmPart σ R (.indiscrete 0) = 1 := by simp [hsymmPart] +theorem hsymmPart_zero : hsymmPart (.indiscrete 0) = (1 : MvPolynomial σ R) := by simp [hsymmPart] @[simp] -theorem hsymmPart_onePart (n : ℕ) : hsymmPart σ R (.indiscrete n) = hsymm σ R n := by +theorem hsymmPart_onePart (n : ℕ) : hsymmPart (.indiscrete n) = (hsymm n : MvPolynomial σ R) := by cases n <;> simp [hsymmPart] -theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by +theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm n) = (hsymm n : MvPolynomial σ S) := by simp [hsymm, ← Multiset.prod_hom'] -theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by +theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm n) = (hsymm n : MvPolynomial τ R) := by simp_rw [hsymm, map_sum, ← prod_hom', rename_X] apply Fintype.sum_equiv (equivCongr e) simp -theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n +theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm n : MvPolynomial σ R) := rename_hsymm n end CompleteHomogeneousSymmetric @@ -366,32 +365,31 @@ open Finset /-- The degree-`n` power sum -/ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n -lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl - /-- `psumPart` is the product of the symmetric polynomials `psum μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (psum σ R)).prod +def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map psum).prod -lemma psumPart_def {n : ℕ} (μ : n.Partition) : psumPart σ R μ = (μ.parts.map (psum σ R)).prod := rfl +lemma psumPart_def {n : ℕ} (μ : n.Partition) : + (psumPart μ : MvPolynomial σ R) = (μ.parts.map psum).prod := rfl @[simp] -theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum] +theorem psum_zero : psum 0 = (Fintype.card σ : MvPolynomial σ R) := by simp [psum] @[simp] -theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] +theorem psum_one : psum 1 = ∑ i, (X i : MvPolynomial σ R) := by simp [psum] @[simp] -theorem psumPart_zero : psumPart σ R (.indiscrete 0) = 1 := by simp [psumPart] +theorem psumPart_zero : psumPart (.indiscrete 0) = (1 : MvPolynomial σ R) := by simp [psumPart] @[simp] theorem psumPart_indiscrete {n : ℕ} (npos : n ≠ 0) : - psumPart σ R (.indiscrete n) = psum σ R n := by simp [psumPart, npos] + psumPart (.indiscrete n) = (psum n : MvPolynomial σ R) := by simp [psumPart, npos] @[simp] -theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by +theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum n) = (psum n : MvPolynomial τ R) := by simp_rw [psum, map_sum, map_pow, rename_X, e.sum_comp (X · ^ n)] -theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum σ R n) := rename_psum _ _ n +theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum n : MvPolynomial σ R) := rename_psum n end PowerSum @@ -403,31 +401,31 @@ variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} def msymm (μ : n.Partition) : MvPolynomial σ R := ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod -lemma msymm_def (μ : n.Partition) : msymm σ R μ = +lemma msymm_def (μ : n.Partition) : (msymm μ : MvPolynomial σ R) = ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod := rfl @[simp] -theorem msymm_zero : msymm σ R (.indiscrete 0) = 1 := by +theorem msymm_zero : msymm (.indiscrete 0) = (1 : MvPolynomial σ R) := by rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), by rfl⟩] simp @[simp] -theorem msymm_one : msymm σ R (.indiscrete 1) = ∑ i, X i := by +theorem msymm_one : msymm (.indiscrete 1) = (∑ i, X i : MvPolynomial σ R) := by symm apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_onePart σ) simp @[simp] theorem rename_msymm (μ : n.Partition) (e : σ ≃ τ) : - rename e (msymm σ R μ) = msymm τ R μ := by + rename e (msymm μ) = (msymm μ : MvPolynomial τ R) := by rw [msymm, map_sum] apply Fintype.sum_equiv (Nat.Partition.ofSym_shape_equiv μ e) intro rw [← Multiset.prod_hom, Multiset.map_map, Nat.Partition.ofSym_shape_equiv] simp -theorem msymm_isSymmetric (μ : n.Partition) : IsSymmetric (msymm σ R μ) := - rename_msymm _ _ μ +theorem msymm_isSymmetric (μ : n.Partition) : IsSymmetric (msymm μ : MvPolynomial σ R) := + rename_msymm μ end MonomialSymmetric From fda6c6dbb13202cfda756af8577b2ea53c5f93ea Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 18:53:19 +0200 Subject: [PATCH 094/106] Fix build fail --- Mathlib/RingTheory/Polynomial/Vieta.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/Polynomial/Vieta.lean b/Mathlib/RingTheory/Polynomial/Vieta.lean index e978315c3175b..b8777f1c54e66 100644 --- a/Mathlib/RingTheory/Polynomial/Vieta.lean +++ b/Mathlib/RingTheory/Polynomial/Vieta.lean @@ -168,11 +168,11 @@ the symmetric polynomials `esymm σ R j`. -/ theorem MvPolynomial.prod_C_add_X_eq_sum_esymm : (∏ i : σ, (Polynomial.X + Polynomial.C (MvPolynomial.X i))) = ∑ j ∈ range (card σ + 1), Polynomial.C - (MvPolynomial.esymm σ R j) * Polynomial.X ^ (card σ - j) := by + (esymm j : MvPolynomial σ R) * Polynomial.X ^ (card σ - j) := by let s := Finset.univ.val.map fun i : σ => (MvPolynomial.X i : MvPolynomial σ R) have : Fintype.card σ = Multiset.card s := by rw [Multiset.card_map, ← Finset.card_univ, Finset.card_def] - simp_rw [this, MvPolynomial.esymm_eq_multiset_esymm σ R, Finset.prod_eq_multiset_prod] + simp_rw [this, MvPolynomial.esymm_eq_multiset_esymm, Finset.prod_eq_multiset_prod] convert Multiset.prod_X_add_C_eq_sum_esymm s simp_rw [s, Multiset.map_map, Function.comp_apply] set_option linter.uppercaseLean3 false in @@ -180,12 +180,12 @@ set_option linter.uppercaseLean3 false in theorem MvPolynomial.prod_X_add_C_coeff (k : ℕ) (h : k ≤ card σ) : (∏ i : σ, (Polynomial.X + Polynomial.C (MvPolynomial.X i)) : Polynomial _).coeff k = - MvPolynomial.esymm σ R (card σ - k) := by + (esymm (card σ - k) : MvPolynomial σ R) := by let s := Finset.univ.val.map fun i => (MvPolynomial.X i : MvPolynomial σ R) have : Fintype.card σ = Multiset.card s := by rw [Multiset.card_map, ← Finset.card_univ, Finset.card_def] rw [this] at h ⊢ - rw [MvPolynomial.esymm_eq_multiset_esymm σ R, Finset.prod_eq_multiset_prod] + rw [MvPolynomial.esymm_eq_multiset_esymm, Finset.prod_eq_multiset_prod] convert Multiset.prod_X_add_C_coeff s h dsimp simp_rw [s, Multiset.map_map, Function.comp_apply] From 1371a3fedb13b7dc1672954513661ff367e91ebf Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 19:03:56 +0200 Subject: [PATCH 095/106] More fixes --- .../MvPolynomial/NewtonIdentities.lean | 29 ++++++++++--------- 1 file changed, 15 insertions(+), 14 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index 3f216f57d4abe..168b6ea4cec1d 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -189,7 +189,7 @@ private theorem esymm_summand_to_weight (k : ℕ) (A : Finset σ) (h : A ∈ pow ∑ j ∈ A, weight σ R k (A, j) = k * (-1) ^ k * (∏ i ∈ A, X i : MvPolynomial σ R) := by simp [weight, mem_powersetCard_univ.mp h, mul_assoc] -private theorem esymm_to_weight (k : ℕ) : k * esymm σ R k = +private theorem esymm_to_weight (k : ℕ) : k * (esymm k : MvPolynomial σ R) = (-1) ^ k * ∑ t ∈ filter (fun t ↦ card t.fst = k) (pairs σ k), weight σ R k t := by rw [esymm, sum_filter_pairs_eq_sum_powersetCard_sum σ R k (fun t ↦ weight σ R k t), sum_congr rfl (esymm_summand_to_weight σ R k), mul_comm (k : MvPolynomial σ R) ((-1) ^ k), @@ -197,8 +197,8 @@ private theorem esymm_to_weight (k : ℕ) : k * esymm σ R k = private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha : a ∈ antidiagonal k) : ∑ A ∈ powersetCard a.fst univ, ∑ j, weight σ R k (A, j) = - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by - simp only [esymm, psum_def, weight, ← mul_assoc, mul_sum] + (-1) ^ a.fst * esymm a.fst * psum a.snd := by + simp only [esymm, psum, weight, ← mul_assoc, mul_sum] rw [sum_comm] refine sum_congr rfl fun x _ ↦ ?_ rw [sum_mul] @@ -207,7 +207,7 @@ private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha private theorem esymm_mul_psum_to_weight (k : ℕ) : ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = + (-1) ^ a.fst * esymm a.fst * psum a.snd = ∑ t ∈ filter (fun t ↦ card t.fst < k) (pairs σ k), weight σ R k t := by rw [← sum_congr rfl (fun a ha ↦ esymm_mul_psum_summand_to_weight σ R k a (mem_filter.mp ha).left), sum_filter_pairs_eq_sum_filter_antidiagonal_powersetCard_sum σ R k] @@ -218,9 +218,9 @@ variable (σ : Type*) [Fintype σ] (R : Type*) [CommRing R] /-- **Newton's identities** give a recurrence relation for the kth elementary symmetric polynomial in terms of lower degree elementary symmetric polynomials and power sums. -/ -theorem mul_esymm_eq_sum (k : ℕ) : k * esymm σ R k = +theorem mul_esymm_eq_sum (k : ℕ) : k * (esymm k : MvPolynomial σ R) = (-1) ^ (k + 1) * ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by + (-1) ^ a.fst * esymm a.fst * psum a.snd := by classical rw [NewtonIdentities.esymm_to_weight σ R k, NewtonIdentities.esymm_mul_psum_to_weight σ R k, eq_comm, ← sub_eq_zero, sub_eq_add_neg, neg_mul_eq_neg_mul, @@ -232,27 +232,28 @@ theorem mul_esymm_eq_sum (k : ℕ) : k * esymm σ R k = neg_one_pow_mul_eq_zero_iff.mpr rfl] theorem sum_antidiagonal_card_esymm_psum_eq_zero : - ∑ a ∈ antidiagonal (Fintype.card σ), (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = 0 := by + ∑ a ∈ antidiagonal (Fintype.card σ), ((-1) ^ a.fst : MvPolynomial σ R) + * esymm a.fst * psum a.snd = 0 := by let k := Fintype.card σ suffices (-1 : MvPolynomial σ R) ^ (k + 1) * - ∑ a ∈ antidiagonal k, (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = 0 by + ∑ a ∈ antidiagonal k, (-1) ^ a.fst * esymm a.fst * psum a.snd = 0 by simpa using this simp [← sum_filter_add_sum_filter_not (antidiagonal k) (fun a ↦ a.fst < k), ← mul_esymm_eq_sum, - mul_add, ← mul_assoc, ← pow_add, mul_comm ↑k (esymm σ R k)] + mul_add, ← mul_assoc, ← pow_add, mul_comm ↑k (esymm k)] /-- A version of Newton's identities which may be more useful in the case that we know the values of the elementary symmetric polynomials and would like to calculate the values of the power sums. -/ -theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum σ R k = - (-1) ^ (k + 1) * k * esymm σ R k - +theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum k = + ((-1) ^ (k + 1) * k : MvPolynomial σ R) * esymm k - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst ∈ Set.Ioo 0 k), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by + (-1) ^ a.fst * esymm a.fst * psum a.snd := by simp only [Set.Ioo, Set.mem_setOf_eq, and_comm] have hesymm := mul_esymm_eq_sum σ R k rw [← (sum_filter_add_sum_filter_not ((antidiagonal k).filter (fun a ↦ a.fst < k)) - (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd))] at hesymm + (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm a.fst * psum a.snd))] at hesymm have sub_both_sides := congrArg (· - (-1 : MvPolynomial σ R) ^ (k + 1) * ∑ a ∈ ((antidiagonal k).filter (fun a ↦ a.fst < k)).filter (fun a ↦ 0 < a.fst), - (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd) hesymm + (-1) ^ a.fst * esymm a.fst * psum a.snd) hesymm simp only [left_distrib, add_sub_cancel_left] at sub_both_sides have sub_both_sides := congrArg ((-1 : MvPolynomial σ R) ^ (k + 1) * ·) sub_both_sides simp only [mul_sub_left_distrib, ← mul_assoc, ← pow_add, Even.neg_one_pow ⟨k + 1, rfl⟩, one_mul, From dcbf7a89eda0c42f97a1c7d21bd93d0b2b83abfb Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 19:07:05 +0200 Subject: [PATCH 096/106] Remove redundant lemmas --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 12 ------------ 1 file changed, 12 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 08b352213c30f..c7dc2b11d38a2 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -189,9 +189,6 @@ def esymm (n : ℕ) : MvPolynomial σ R := ∑ t ∈ powersetCard n univ, ∏ i ∈ t, X i #align mv_polynomial.esymm MvPolynomial.esymm -lemma esymm_def (n : ℕ) : - esymm n = ∑ t in powersetCard n univ, ∏ i in t, (X i : MvPolynomial σ R) := rfl - /-- `esymmPart` is the product of the symmetric polynomials `esymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. @@ -328,9 +325,6 @@ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod where `μ = (μ₁, μ₂, ...)` is a partition. -/ def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map hsymm).prod -lemma hsymmPart_def {n : ℕ} (μ : n.Partition) : - (hsymmPart μ : MvPolynomial σ R) = (μ.parts.map hsymm).prod := rfl - @[simp] theorem hsymm_zero : hsymm 0 = (1 : MvPolynomial σ R) := by simp [hsymm, eq_nil_of_card_zero] @@ -369,9 +363,6 @@ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n where `μ = (μ₁, μ₂, ...)` is a partition. -/ def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map psum).prod -lemma psumPart_def {n : ℕ} (μ : n.Partition) : - (psumPart μ : MvPolynomial σ R) = (μ.parts.map psum).prod := rfl - @[simp] theorem psum_zero : psum 0 = (Fintype.card σ : MvPolynomial σ R) := by simp [psum] @@ -401,9 +392,6 @@ variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} def msymm (μ : n.Partition) : MvPolynomial σ R := ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod -lemma msymm_def (μ : n.Partition) : (msymm μ : MvPolynomial σ R) = - ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod := rfl - @[simp] theorem msymm_zero : msymm (.indiscrete 0) = (1 : MvPolynomial σ R) := by rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), by rfl⟩] From 7977955ac327b871cf9a46eab7fe760cb6b6be5c Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 19:21:52 +0200 Subject: [PATCH 097/106] Revert "More fixes" This reverts commit 1371a3fedb13b7dc1672954513661ff367e91ebf. --- .../MvPolynomial/NewtonIdentities.lean | 29 +++++++++---------- 1 file changed, 14 insertions(+), 15 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index 168b6ea4cec1d..3f216f57d4abe 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -189,7 +189,7 @@ private theorem esymm_summand_to_weight (k : ℕ) (A : Finset σ) (h : A ∈ pow ∑ j ∈ A, weight σ R k (A, j) = k * (-1) ^ k * (∏ i ∈ A, X i : MvPolynomial σ R) := by simp [weight, mem_powersetCard_univ.mp h, mul_assoc] -private theorem esymm_to_weight (k : ℕ) : k * (esymm k : MvPolynomial σ R) = +private theorem esymm_to_weight (k : ℕ) : k * esymm σ R k = (-1) ^ k * ∑ t ∈ filter (fun t ↦ card t.fst = k) (pairs σ k), weight σ R k t := by rw [esymm, sum_filter_pairs_eq_sum_powersetCard_sum σ R k (fun t ↦ weight σ R k t), sum_congr rfl (esymm_summand_to_weight σ R k), mul_comm (k : MvPolynomial σ R) ((-1) ^ k), @@ -197,8 +197,8 @@ private theorem esymm_to_weight (k : ℕ) : k * (esymm k : MvPolynomial σ R) = private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha : a ∈ antidiagonal k) : ∑ A ∈ powersetCard a.fst univ, ∑ j, weight σ R k (A, j) = - (-1) ^ a.fst * esymm a.fst * psum a.snd := by - simp only [esymm, psum, weight, ← mul_assoc, mul_sum] + (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by + simp only [esymm, psum_def, weight, ← mul_assoc, mul_sum] rw [sum_comm] refine sum_congr rfl fun x _ ↦ ?_ rw [sum_mul] @@ -207,7 +207,7 @@ private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha private theorem esymm_mul_psum_to_weight (k : ℕ) : ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), - (-1) ^ a.fst * esymm a.fst * psum a.snd = + (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = ∑ t ∈ filter (fun t ↦ card t.fst < k) (pairs σ k), weight σ R k t := by rw [← sum_congr rfl (fun a ha ↦ esymm_mul_psum_summand_to_weight σ R k a (mem_filter.mp ha).left), sum_filter_pairs_eq_sum_filter_antidiagonal_powersetCard_sum σ R k] @@ -218,9 +218,9 @@ variable (σ : Type*) [Fintype σ] (R : Type*) [CommRing R] /-- **Newton's identities** give a recurrence relation for the kth elementary symmetric polynomial in terms of lower degree elementary symmetric polynomials and power sums. -/ -theorem mul_esymm_eq_sum (k : ℕ) : k * (esymm k : MvPolynomial σ R) = +theorem mul_esymm_eq_sum (k : ℕ) : k * esymm σ R k = (-1) ^ (k + 1) * ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), - (-1) ^ a.fst * esymm a.fst * psum a.snd := by + (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by classical rw [NewtonIdentities.esymm_to_weight σ R k, NewtonIdentities.esymm_mul_psum_to_weight σ R k, eq_comm, ← sub_eq_zero, sub_eq_add_neg, neg_mul_eq_neg_mul, @@ -232,28 +232,27 @@ theorem mul_esymm_eq_sum (k : ℕ) : k * (esymm k : MvPolynomial σ R) = neg_one_pow_mul_eq_zero_iff.mpr rfl] theorem sum_antidiagonal_card_esymm_psum_eq_zero : - ∑ a ∈ antidiagonal (Fintype.card σ), ((-1) ^ a.fst : MvPolynomial σ R) - * esymm a.fst * psum a.snd = 0 := by + ∑ a ∈ antidiagonal (Fintype.card σ), (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = 0 := by let k := Fintype.card σ suffices (-1 : MvPolynomial σ R) ^ (k + 1) * - ∑ a ∈ antidiagonal k, (-1) ^ a.fst * esymm a.fst * psum a.snd = 0 by + ∑ a ∈ antidiagonal k, (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = 0 by simpa using this simp [← sum_filter_add_sum_filter_not (antidiagonal k) (fun a ↦ a.fst < k), ← mul_esymm_eq_sum, - mul_add, ← mul_assoc, ← pow_add, mul_comm ↑k (esymm k)] + mul_add, ← mul_assoc, ← pow_add, mul_comm ↑k (esymm σ R k)] /-- A version of Newton's identities which may be more useful in the case that we know the values of the elementary symmetric polynomials and would like to calculate the values of the power sums. -/ -theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum k = - ((-1) ^ (k + 1) * k : MvPolynomial σ R) * esymm k - +theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum σ R k = + (-1) ^ (k + 1) * k * esymm σ R k - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst ∈ Set.Ioo 0 k), - (-1) ^ a.fst * esymm a.fst * psum a.snd := by + (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by simp only [Set.Ioo, Set.mem_setOf_eq, and_comm] have hesymm := mul_esymm_eq_sum σ R k rw [← (sum_filter_add_sum_filter_not ((antidiagonal k).filter (fun a ↦ a.fst < k)) - (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm a.fst * psum a.snd))] at hesymm + (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd))] at hesymm have sub_both_sides := congrArg (· - (-1 : MvPolynomial σ R) ^ (k + 1) * ∑ a ∈ ((antidiagonal k).filter (fun a ↦ a.fst < k)).filter (fun a ↦ 0 < a.fst), - (-1) ^ a.fst * esymm a.fst * psum a.snd) hesymm + (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd) hesymm simp only [left_distrib, add_sub_cancel_left] at sub_both_sides have sub_both_sides := congrArg ((-1 : MvPolynomial σ R) ^ (k + 1) * ·) sub_both_sides simp only [mul_sub_left_distrib, ← mul_assoc, ← pow_add, Even.neg_one_pow ⟨k + 1, rfl⟩, one_mul, From 389f4724981d9f3965d0460d374446fce0c15f77 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 19:22:00 +0200 Subject: [PATCH 098/106] Revert "Fix build fail" This reverts commit fda6c6dbb13202cfda756af8577b2ea53c5f93ea. --- Mathlib/RingTheory/Polynomial/Vieta.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/RingTheory/Polynomial/Vieta.lean b/Mathlib/RingTheory/Polynomial/Vieta.lean index b8777f1c54e66..e978315c3175b 100644 --- a/Mathlib/RingTheory/Polynomial/Vieta.lean +++ b/Mathlib/RingTheory/Polynomial/Vieta.lean @@ -168,11 +168,11 @@ the symmetric polynomials `esymm σ R j`. -/ theorem MvPolynomial.prod_C_add_X_eq_sum_esymm : (∏ i : σ, (Polynomial.X + Polynomial.C (MvPolynomial.X i))) = ∑ j ∈ range (card σ + 1), Polynomial.C - (esymm j : MvPolynomial σ R) * Polynomial.X ^ (card σ - j) := by + (MvPolynomial.esymm σ R j) * Polynomial.X ^ (card σ - j) := by let s := Finset.univ.val.map fun i : σ => (MvPolynomial.X i : MvPolynomial σ R) have : Fintype.card σ = Multiset.card s := by rw [Multiset.card_map, ← Finset.card_univ, Finset.card_def] - simp_rw [this, MvPolynomial.esymm_eq_multiset_esymm, Finset.prod_eq_multiset_prod] + simp_rw [this, MvPolynomial.esymm_eq_multiset_esymm σ R, Finset.prod_eq_multiset_prod] convert Multiset.prod_X_add_C_eq_sum_esymm s simp_rw [s, Multiset.map_map, Function.comp_apply] set_option linter.uppercaseLean3 false in @@ -180,12 +180,12 @@ set_option linter.uppercaseLean3 false in theorem MvPolynomial.prod_X_add_C_coeff (k : ℕ) (h : k ≤ card σ) : (∏ i : σ, (Polynomial.X + Polynomial.C (MvPolynomial.X i)) : Polynomial _).coeff k = - (esymm (card σ - k) : MvPolynomial σ R) := by + MvPolynomial.esymm σ R (card σ - k) := by let s := Finset.univ.val.map fun i => (MvPolynomial.X i : MvPolynomial σ R) have : Fintype.card σ = Multiset.card s := by rw [Multiset.card_map, ← Finset.card_univ, Finset.card_def] rw [this] at h ⊢ - rw [MvPolynomial.esymm_eq_multiset_esymm, Finset.prod_eq_multiset_prod] + rw [MvPolynomial.esymm_eq_multiset_esymm σ R, Finset.prod_eq_multiset_prod] convert Multiset.prod_X_add_C_coeff s h dsimp simp_rw [s, Multiset.map_map, Function.comp_apply] From 881be751395cabe1f19ff56efd7ae93a542a774b Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Thu, 11 Jul 2024 19:23:56 +0200 Subject: [PATCH 099/106] Revert "Remove explicit variables" --- .../RingTheory/MvPolynomial/Symmetric.lean | 92 ++++++++++--------- 1 file changed, 48 insertions(+), 44 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index c7dc2b11d38a2..83b5198e547f4 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -178,7 +178,7 @@ def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : (AlgHom.ext <| fun p => Subtype.ext <| by simp) (AlgHom.ext <| fun p => Subtype.ext <| by simp) -variable [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] +variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] section ElementarySymmetric @@ -193,52 +193,52 @@ def esymm (n : ℕ) : MvPolynomial σ R := `esymmPart` is the product of the symmetric polynomials `esymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def esymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (esymm)).prod +def esymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (esymm σ R)).prod /-- The `n`th elementary symmetric `MvPolynomial σ R` is obtained by evaluating the `n`th elementary symmetric at the `Multiset` of the monomials -/ -theorem esymm_eq_multiset_esymm : esymm = (univ.val.map (X : σ → MvPolynomial σ R)).esymm := by +theorem esymm_eq_multiset_esymm : esymm σ R = (univ.val.map X).esymm := by exact funext fun n => (esymm_map_val X _ n).symm #align mv_polynomial.esymm_eq_multiset_esymm MvPolynomial.esymm_eq_multiset_esymm theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (n : ℕ) (f : σ → S) : - aeval f (esymm n : MvPolynomial σ R) = (univ.val.map f).esymm n := by + aeval f (esymm σ R n) = (univ.val.map f).esymm n := by simp_rw [esymm, aeval_sum, aeval_prod, aeval_X, esymm_map_val] #align mv_polynomial.aeval_esymm_eq_multiset_esymm MvPolynomial.aeval_esymm_eq_multiset_esymm /-- We can define `esymm σ R n` by summing over a subtype instead of over `powerset_len`. -/ -theorem esymm_eq_sum_subtype (n : ℕ) : esymm n = ∑ t : { s : Finset σ // s.card = n }, - ∏ i ∈ (t : Finset σ), (X i : MvPolynomial σ R) := +theorem esymm_eq_sum_subtype (n : ℕ) : + esymm σ R n = ∑ t : { s : Finset σ // s.card = n }, ∏ i ∈ (t : Finset σ), X i := sum_subtype _ (fun _ => mem_powersetCard_univ) _ #align mv_polynomial.esymm_eq_sum_subtype MvPolynomial.esymm_eq_sum_subtype /-- We can define `esymm σ R n` as a sum over explicit monomials -/ -theorem esymm_eq_sum_monomial (n : ℕ) : (esymm n : MvPolynomial σ R) = ∑ t ∈ powersetCard n univ, - monomial (∑ i ∈ t, Finsupp.single i 1) 1 := by +theorem esymm_eq_sum_monomial (n : ℕ) : + esymm σ R n = ∑ t ∈ powersetCard n univ, monomial (∑ i ∈ t, Finsupp.single i 1) 1 := by simp_rw [monomial_sum_one] rfl #align mv_polynomial.esymm_eq_sum_monomial MvPolynomial.esymm_eq_sum_monomial @[simp] -theorem esymm_zero : esymm 0 = (1 : MvPolynomial σ R) := by simp [esymm] +theorem esymm_zero : esymm σ R 0 = 1 := by simp [esymm] #align mv_polynomial.esymm_zero MvPolynomial.esymm_zero @[simp] -theorem esymm_one : esymm 1 = ∑ i, (X i : MvPolynomial σ R) := by simp [esymm, powersetCard_one] +theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp [esymm, powersetCard_one] -theorem esymmPart_zero : esymmPart (.indiscrete 0) = (1 : MvPolynomial σ R) := by simp [esymmPart] +theorem esymmPart_zero : esymmPart σ R (.indiscrete 0) = 1 := by simp [esymmPart] @[simp] -theorem esymmPart_onePart (n : ℕ) : esymmPart (.indiscrete n) = (esymm n : MvPolynomial σ R) := by +theorem esymmPart_onePart (n : ℕ) : esymmPart σ R (.indiscrete n) = esymm σ R n := by cases n <;> simp [esymmPart] -theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm n) = (esymm n : MvPolynomial σ S) := by +theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by simp_rw [esymm, map_sum, map_prod, map_X] #align mv_polynomial.map_esymm MvPolynomial.map_esymm -theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm n) = (esymm n : MvPolynomial τ R) := +theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm σ R n) = esymm τ R n := calc - rename e (esymm n) = ∑ x ∈ powersetCard n univ, ∏ i ∈ x, X (e i) := by + rename e (esymm σ R n) = ∑ x ∈ powersetCard n univ, ∏ i ∈ x, X (e i) := by simp_rw [esymm, map_sum, map_prod, rename_X] _ = ∑ t ∈ powersetCard n (univ.map e.toEmbedding), ∏ i ∈ t, X i := by simp [powersetCard_map, -map_univ_equiv] @@ -248,11 +248,12 @@ theorem rename_esymm (n : ℕ) (e : σ ≃ τ) : rename e (esymm n) = (esymm n : _ = ∑ t ∈ powersetCard n univ, ∏ i ∈ t, X i := by rw [map_univ_equiv] #align mv_polynomial.rename_esymm MvPolynomial.rename_esymm -theorem esymm_isSymmetric (n : ℕ) : IsSymmetric (esymm n : MvPolynomial σ R) := rename_esymm n +theorem esymm_isSymmetric (n : ℕ) : IsSymmetric (esymm σ R n) := rename_esymm _ _ n #align mv_polynomial.esymm_is_symmetric MvPolynomial.esymm_isSymmetric theorem support_esymm'' [DecidableEq σ] [Nontrivial R] (n : ℕ) : - (esymm n : MvPolynomial σ R).support = (powersetCard n (univ : Finset σ)).biUnion fun t => + (esymm σ R n).support = + (powersetCard n (univ : Finset σ)).biUnion fun t => (Finsupp.single (∑ i ∈ t, Finsupp.single i 1) (1 : R)).support := by rw [esymm_eq_sum_monomial] simp only [← single_eq_monomial] @@ -275,23 +276,22 @@ theorem support_esymm'' [DecidableEq σ] [Nontrivial R] (n : ℕ) : all_goals intro x y; simp [Finsupp.support_single_disjoint] #align mv_polynomial.support_esymm'' MvPolynomial.support_esymm'' -theorem support_esymm' [DecidableEq σ] [Nontrivial R] (n : ℕ) : - (esymm n : MvPolynomial σ R).support = - (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i ∈ t, Finsupp.single i 1} := by +theorem support_esymm' [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = + (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i ∈ t, Finsupp.single i 1} := by rw [support_esymm''] congr funext exact Finsupp.support_single_ne_zero _ one_ne_zero #align mv_polynomial.support_esymm' MvPolynomial.support_esymm' -theorem support_esymm [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm n : MvPolynomial σ R).support = +theorem support_esymm [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = (powersetCard n (univ : Finset σ)).image fun t => ∑ i ∈ t, Finsupp.single i 1 := by rw [support_esymm'] exact biUnion_singleton #align mv_polynomial.support_esymm MvPolynomial.support_esymm -theorem degrees_esymm [Nontrivial R] {n : ℕ} (hpos : n ≠ 0) (hn : n ≤ Fintype.card σ) : - (esymm n : MvPolynomial σ R).degrees = (univ : Finset σ).val := by +theorem degrees_esymm [Nontrivial R] {n : ℕ} (hpos : 0 < n) (hn : n ≤ Fintype.card σ) : + (esymm σ R n).degrees = (univ : Finset σ).val := by classical have : (Finsupp.toMultiset ∘ fun t : Finset σ => ∑ i ∈ t, Finsupp.single i 1) = val := by @@ -306,7 +306,7 @@ theorem degrees_esymm [Nontrivial R] {n : ℕ} (hpos : n ≠ 0) (hn : n ≤ Fint congr · rfl rw [← this] - obtain ⟨k, rfl⟩ := Nat.exists_eq_succ_of_ne_zero hpos + obtain ⟨k, rfl⟩ := Nat.exists_eq_succ_of_ne_zero hpos.ne' simpa using powersetCard_sup _ _ (Nat.lt_of_succ_le hn) #align mv_polynomial.degrees_esymm MvPolynomial.degrees_esymm @@ -321,34 +321,36 @@ variable [DecidableEq σ] [DecidableEq τ] /-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod +lemma hsymm_def {n : ℕ} : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl + /-- `hsymmPart` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map hsymm).prod +def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (hsymm σ R)).prod @[simp] -theorem hsymm_zero : hsymm 0 = (1 : MvPolynomial σ R) := by simp [hsymm, eq_nil_of_card_zero] +theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] @[simp] -theorem hsymm_one : hsymm 1 = ∑ i, (X i : MvPolynomial σ R) := by +theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by symm apply Fintype.sum_equiv oneEquiv simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton, implies_true] -theorem hsymmPart_zero : hsymmPart (.indiscrete 0) = (1 : MvPolynomial σ R) := by simp [hsymmPart] +theorem hsymmPart_zero : hsymmPart σ R (.indiscrete 0) = 1 := by simp [hsymmPart] @[simp] -theorem hsymmPart_onePart (n : ℕ) : hsymmPart (.indiscrete n) = (hsymm n : MvPolynomial σ R) := by +theorem hsymmPart_onePart (n : ℕ) : hsymmPart σ R (.indiscrete n) = hsymm σ R n := by cases n <;> simp [hsymmPart] -theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm n) = (hsymm n : MvPolynomial σ S) := by +theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by simp [hsymm, ← Multiset.prod_hom'] -theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm n) = (hsymm n : MvPolynomial τ R) := by +theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by simp_rw [hsymm, map_sum, ← prod_hom', rename_X] apply Fintype.sum_equiv (equivCongr e) simp -theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm n : MvPolynomial σ R) := rename_hsymm n +theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n end CompleteHomogeneousSymmetric @@ -359,28 +361,30 @@ open Finset /-- The degree-`n` power sum -/ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n +lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl + /-- `psumPart` is the product of the symmetric polynomials `psum μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ -def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map psum).prod +def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (psum σ R)).prod @[simp] -theorem psum_zero : psum 0 = (Fintype.card σ : MvPolynomial σ R) := by simp [psum] +theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum] @[simp] -theorem psum_one : psum 1 = ∑ i, (X i : MvPolynomial σ R) := by simp [psum] +theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] @[simp] -theorem psumPart_zero : psumPart (.indiscrete 0) = (1 : MvPolynomial σ R) := by simp [psumPart] +theorem psumPart_zero : psumPart σ R (.indiscrete 0) = 1 := by simp [psumPart] @[simp] theorem psumPart_indiscrete {n : ℕ} (npos : n ≠ 0) : - psumPart (.indiscrete n) = (psum n : MvPolynomial σ R) := by simp [psumPart, npos] + psumPart σ R (.indiscrete n) = psum σ R n := by simp [psumPart, npos] @[simp] -theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum n) = (psum n : MvPolynomial τ R) := by +theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by simp_rw [psum, map_sum, map_pow, rename_X, e.sum_comp (X · ^ n)] -theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum n : MvPolynomial σ R) := rename_psum n +theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum σ R n) := rename_psum _ _ n end PowerSum @@ -393,27 +397,27 @@ def msymm (μ : n.Partition) : MvPolynomial σ R := ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod @[simp] -theorem msymm_zero : msymm (.indiscrete 0) = (1 : MvPolynomial σ R) := by +theorem msymm_zero : msymm σ R (.indiscrete 0) = 1 := by rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), by rfl⟩] simp @[simp] -theorem msymm_one : msymm (.indiscrete 1) = (∑ i, X i : MvPolynomial σ R) := by +theorem msymm_one : msymm σ R (.indiscrete 1) = ∑ i, X i := by symm apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_onePart σ) simp @[simp] theorem rename_msymm (μ : n.Partition) (e : σ ≃ τ) : - rename e (msymm μ) = (msymm μ : MvPolynomial τ R) := by + rename e (msymm σ R μ) = msymm τ R μ := by rw [msymm, map_sum] apply Fintype.sum_equiv (Nat.Partition.ofSym_shape_equiv μ e) intro rw [← Multiset.prod_hom, Multiset.map_map, Nat.Partition.ofSym_shape_equiv] simp -theorem msymm_isSymmetric (μ : n.Partition) : IsSymmetric (msymm μ : MvPolynomial σ R) := - rename_msymm μ +theorem msymm_isSymmetric (μ : n.Partition) : IsSymmetric (msymm σ R μ) := + rename_msymm _ _ μ end MonomialSymmetric From 644640524eaf75eda19fb168656e874dffcaab1a Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 12 Jul 2024 00:16:49 +0200 Subject: [PATCH 100/106] Remove redundancies --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 83b5198e547f4..d21283d9f8f95 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -321,8 +321,6 @@ variable [DecidableEq σ] [DecidableEq τ] /-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod -lemma hsymm_def {n : ℕ} : hsymm σ R n = ∑ s : Sym σ n, (s.1.map X).prod := rfl - /-- `hsymmPart` is the product of the symmetric polynomials `hsymm μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (hsymm σ R)).prod @@ -361,8 +359,6 @@ open Finset /-- The degree-`n` power sum -/ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n -lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl - /-- `psumPart` is the product of the symmetric polynomials `psum μᵢ`, where `μ = (μ₁, μ₂, ...)` is a partition. -/ def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (psum σ R)).prod From b71624d717fae2ab38a01b3541c79dccf63eb58e Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 12 Jul 2024 00:38:34 +0200 Subject: [PATCH 101/106] Yeah right forgot this one --- Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index 3f216f57d4abe..d9f1a7bdaee22 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -198,7 +198,7 @@ private theorem esymm_to_weight (k : ℕ) : k * esymm σ R k = private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha : a ∈ antidiagonal k) : ∑ A ∈ powersetCard a.fst univ, ∑ j, weight σ R k (A, j) = (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by - simp only [esymm, psum_def, weight, ← mul_assoc, mul_sum] + simp only [esymm, psum, weight, ← mul_assoc, mul_sum] rw [sum_comm] refine sum_congr rfl fun x _ ↦ ?_ rw [sum_mul] From 82fa131fae751e2fcc84775cce4ff43239e79e59 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Fri, 12 Jul 2024 14:11:23 +0200 Subject: [PATCH 102/106] `onePart` to `indiscrete`, docstrings --- .../Combinatorics/Enumerative/Partition.lean | 6 +++--- .../RingTheory/MvPolynomial/Symmetric.lean | 19 ++++++++++++------- 2 files changed, 15 insertions(+), 10 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 8f2309a1f4bc7..73ae13a00953a 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -155,7 +155,7 @@ lemma ofSym_one (s : Sym σ 1) : ofSym s = indiscrete 1 := by ext; simp /-- The equivalence between `σ` and `1`-tuples of elements of σ -/ -def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ +def ofSym_equiv_indiscrete (σ : Type*) [DecidableEq σ] : σ ≃ { a : Sym σ 1 // ofSym a = indiscrete 1 } where toFun := fun a => ⟨Sym.oneEquiv a, by ext; simp⟩ invFun := fun a => Sym.oneEquiv.symm a.1 @@ -163,8 +163,8 @@ def ofSym_equiv_onePart (σ : Type*) [DecidableEq σ] : σ ≃ right_inv := by intro a; simp only [Equiv.apply_symm_apply, Subtype.coe_eta] @[simp] -lemma ofSym_equiv_onePart_apply (i : σ) : - ((ofSym_equiv_onePart σ) i : Multiset σ) = {i} := rfl +lemma ofSym_equiv_indiscrete_apply (i : σ) : + ((ofSym_equiv_indiscrete σ) i : Multiset σ) = {i} := rfl /-- The number of times a positive integer `i` appears in the partition `ofSums n l hl` is the same as the number of times it appears in the multiset `l`. diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index d21283d9f8f95..bfa5ff2761950 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -184,7 +184,8 @@ section ElementarySymmetric open Finset -/-- The `n`th elementary symmetric `MvPolynomial σ R`. -/ +/-- The `n`th elementary symmetric `MvPolynomial σ R`. +It is the sum over all the degree n squarefree monomials in `MvPolynomial σ R`. -/ def esymm (n : ℕ) : MvPolynomial σ R := ∑ t ∈ powersetCard n univ, ∏ i ∈ t, X i #align mv_polynomial.esymm MvPolynomial.esymm @@ -229,7 +230,7 @@ theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp [esymm, powersetCard_on theorem esymmPart_zero : esymmPart σ R (.indiscrete 0) = 1 := by simp [esymmPart] @[simp] -theorem esymmPart_onePart (n : ℕ) : esymmPart σ R (.indiscrete n) = esymm σ R n := by +theorem esymmPart_indiscrete (n : ℕ) : esymmPart σ R (.indiscrete n) = esymm σ R n := by cases n <;> simp [esymmPart] theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by @@ -318,7 +319,8 @@ open Finset Multiset Sym variable [DecidableEq σ] [DecidableEq τ] -/-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. -/ +/-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. +It is the sum over all the degree n monomials in `MvPolynomial σ R`. -/ def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod /-- `hsymmPart` is the product of the symmetric polynomials `hsymm μᵢ`, @@ -337,7 +339,7 @@ theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by theorem hsymmPart_zero : hsymmPart σ R (.indiscrete 0) = 1 := by simp [hsymmPart] @[simp] -theorem hsymmPart_onePart (n : ℕ) : hsymmPart σ R (.indiscrete n) = hsymm σ R n := by +theorem hsymmPart_indiscrete (n : ℕ) : hsymmPart σ R (.indiscrete n) = hsymm σ R n := by cases n <;> simp [hsymmPart] theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by @@ -356,7 +358,8 @@ section PowerSum open Finset -/-- The degree-`n` power sum -/ +/-- The degree-`n` power sum symmetric `MvPolynomial σ R`. +It is the sum over all the `n`-th powers of the variables. -/ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n /-- `psumPart` is the product of the symmetric polynomials `psum μᵢ`, @@ -388,7 +391,9 @@ section MonomialSymmetric variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} -/-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. -/ +/-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. +It is the sum over all the monomials in `MvPolynomial σ R` such that +the multiset of exponents is equal to the multiset of parts of μ. -/ def msymm (μ : n.Partition) : MvPolynomial σ R := ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod @@ -400,7 +405,7 @@ theorem msymm_zero : msymm σ R (.indiscrete 0) = 1 := by @[simp] theorem msymm_one : msymm σ R (.indiscrete 1) = ∑ i, X i := by symm - apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_onePart σ) + apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_indiscrete σ) simp @[simp] From acd703b4aac4a195d622e99b03d167aab58f75ee Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Mon, 15 Jul 2024 15:54:42 +0200 Subject: [PATCH 103/106] Remove extra lemma --- Mathlib/Combinatorics/Enumerative/Partition.lean | 12 ------------ Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 13 +++++++++++-- 2 files changed, 11 insertions(+), 14 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 73ae13a00953a..e04d2b53b75c5 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -154,18 +154,6 @@ instance UniquePartitionOne : Unique (Partition 1) where lemma ofSym_one (s : Sym σ 1) : ofSym s = indiscrete 1 := by ext; simp -/-- The equivalence between `σ` and `1`-tuples of elements of σ -/ -def ofSym_equiv_indiscrete (σ : Type*) [DecidableEq σ] : σ ≃ - { a : Sym σ 1 // ofSym a = indiscrete 1 } where - toFun := fun a => ⟨Sym.oneEquiv a, by ext; simp⟩ - invFun := fun a => Sym.oneEquiv.symm a.1 - left_inv := by intro a; simp only [Sym.oneEquiv_apply]; rfl - right_inv := by intro a; simp only [Equiv.apply_symm_apply, Subtype.coe_eta] - -@[simp] -lemma ofSym_equiv_indiscrete_apply (i : σ) : - ((ofSym_equiv_indiscrete σ) i : Multiset σ) = {i} := rfl - /-- The number of times a positive integer `i` appears in the partition `ofSums n l hl` is the same as the number of times it appears in the multiset `l`. (For `i = 0`, `Partition.non_zero` combined with `Multiset.count_eq_zero_of_not_mem` gives that diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index bfa5ff2761950..03cb4c5138ac9 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -404,9 +404,18 @@ theorem msymm_zero : msymm σ R (.indiscrete 0) = 1 := by @[simp] theorem msymm_one : msymm σ R (.indiscrete 1) = ∑ i, X i := by + have : (fun (x : Sym σ 1) ↦ x ∈ Set.univ) = + (fun x ↦ Nat.Partition.ofSym x = Nat.Partition.indiscrete 1) := by + simp_rw [Set.mem_univ, Nat.Partition.ofSym_one] symm - apply Fintype.sum_equiv (Nat.Partition.ofSym_equiv_indiscrete σ) - simp + rw [Fintype.sum_equiv (Equiv.trans Sym.oneEquiv (Equiv.Set.univ (Sym σ 1)).symm) + _ (fun s ↦ (s.1.1.map X).prod)] + · apply Fintype.sum_equiv (Equiv.subtypeEquivProp this) + intro x + congr + · intro x + rw [← Multiset.prod_singleton (X x), ← Multiset.map_singleton] + congr @[simp] theorem rename_msymm (μ : n.Partition) (e : σ ≃ τ) : From dd82e438fbfca1e7d6b62e7eaf4d13c0cbd934ae Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 4 Aug 2024 09:44:08 +0200 Subject: [PATCH 104/106] Remove duplicate --- Mathlib/Data/Sym/Basic.lean | 13 ------------- 1 file changed, 13 deletions(-) diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 5a130b73bad18..4612fa17121c9 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -476,19 +476,6 @@ def oneEquiv : α ≃ Sym α 1 where left_inv a := by rfl right_inv := by rintro ⟨⟨l⟩, h⟩; obtain ⟨a, rfl⟩ := List.length_eq_one.mp h; rfl -/-- Defines an equivalence between `α` and `Sym α 1`. -/ -@[simps apply] -def oneEquiv : α ≃ Sym α 1 where - toFun a := ⟨{a}, by simp⟩ - invFun s := (Equiv.subtypeQuotientEquivQuotientSubtype - (·.length = 1) _ (fun l ↦ Iff.rfl) (fun l l' ↦ by rfl) s).liftOn - (fun l ↦ l.1.head <| List.length_pos.mp <| by simp) - fun ⟨_, _⟩ ⟨_, h⟩ ↦ fun perm ↦ by - obtain ⟨a, rfl⟩ := List.length_eq_one.mp h - exact List.eq_of_mem_singleton (perm.mem_iff.mp <| List.head_mem _) - left_inv a := by rfl - right_inv := by rintro ⟨⟨l⟩, h⟩; obtain ⟨a, rfl⟩ := List.length_eq_one.mp h; rfl - /-- Fill a term `m : Sym α (n - i)` with `i` copies of `a` to obtain a term of `Sym α n`. This is a convenience wrapper for `m.append (replicate i a)` that adjusts the term using `Sym.cast`. -/ From 937f17101e4df1cce86b963e021d32d4f540e5b2 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 4 Aug 2024 09:49:17 +0200 Subject: [PATCH 105/106] Idk why there were duplicates --- .../Combinatorics/Enumerative/Partition.lean | 28 ------------------- 1 file changed, 28 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index 88a36d5e9004c..eaba453a5c731 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -120,34 +120,6 @@ def ofSymShapeEquiv (μ : Partition n) (e : σ ≃ τ) : left_inv := by intro x; simp right_inv := by intro x; simp -/-- An element `s` of `Sym σ n` induces a partition given by its multiplicities. -/ -def ofSym {n : ℕ} {σ : Type*} (s : Sym σ n) [DecidableEq σ] : n.Partition where - parts := s.1.dedup.map s.1.count - parts_pos := by simp [Multiset.count_pos] - parts_sum := by - show ∑ a ∈ s.1.toFinset, count a s.1 = n - rw [toFinset_sum_count_eq] - exact s.2 - -variable {n : ℕ} {σ τ : Type*} [DecidableEq σ] [DecidableEq τ] - -lemma ofSym_map (e : σ ≃ τ) (s : Sym σ n) : - ofSym (s.map e) = ofSym s := by - simp only [ofSym, Sym.val_eq_coe, Sym.coe_map, toFinset_val, mk.injEq] - rw [Multiset.dedup_map_of_injective e.injective] - simp only [map_map, Function.comp_apply] - congr; funext i - rw [← Multiset.count_map_eq_count' e _ e.injective] - -/-- An equivalence between `σ` and `τ` induces an equivalence between the subtypes of `Sym σ n` and -`Sym τ n` corresponding to a given partition. -/ -def ofSym_shape_equiv (μ : Partition n) (e : σ ≃ τ) : - {x : Sym σ n // ofSym x = μ} ≃ {x : Sym τ n // ofSym x = μ} where - toFun := fun x => ⟨Sym.equivCongr e x, by simp [ofSym_map, x.2]⟩ - invFun := fun x => ⟨Sym.equivCongr e.symm x, by simp [ofSym_map, x.2]⟩ - left_inv := by intro x; simp - right_inv := by intro x; simp - /-- The partition of exactly one part. -/ def indiscrete (n : ℕ) : Partition n := ofSums n {n} rfl From 3070c490a87b1d4bc98f36811ea844d8132ab2ff Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 4 Aug 2024 09:53:04 +0200 Subject: [PATCH 106/106] I wish I got all the corrections before splitting the PR --- Mathlib/RingTheory/MvPolynomial/Symmetric.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index f1112b2b30c0d..49097e138f0bb 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -394,9 +394,9 @@ theorem msymm_one : msymm σ R (.indiscrete 1) = ∑ i, X i := by theorem rename_msymm (μ : n.Partition) (e : σ ≃ τ) : rename e (msymm σ R μ) = msymm τ R μ := by rw [msymm, map_sum] - apply Fintype.sum_equiv (Nat.Partition.ofSym_shape_equiv μ e) + apply Fintype.sum_equiv (Nat.Partition.ofSymShapeEquiv μ e) intro - rw [← Multiset.prod_hom, Multiset.map_map, Nat.Partition.ofSym_shape_equiv] + rw [← Multiset.prod_hom, Multiset.map_map, Nat.Partition.ofSymShapeEquiv] simp theorem msymm_isSymmetric (μ : n.Partition) : IsSymmetric (msymm σ R μ) :=