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

Commit 6ae9f00

Browse files
committed
feat(ring_theory/polynomial/symmetric): degrees_esymm (#6718)
A lot of API also added for finset, finsupp, multiset, powerset_len Co-authored-by: Hanting Zhang <hantingzhang03@gmail.com>
1 parent 34a3317 commit 6ae9f00

File tree

11 files changed

+215
-6
lines changed

11 files changed

+215
-6
lines changed

src/algebra/big_operators/basic.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -91,6 +91,10 @@ theorem prod_eq_fold [comm_monoid β] (s : finset α) (f : α → β) :
9191
(∏ x in s, f x) = s.fold (*) 1 f :=
9292
rfl
9393

94+
@[simp] lemma sum_multiset_singleton (s : finset α) :
95+
s.sum (λ x, x ::ₘ 0) = s.val :=
96+
by simp [sum_eq_multiset_sum]
97+
9498
end finset
9599

96100

src/data/finset/lattice.lean

Lines changed: 10 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -670,9 +670,16 @@ lemma supr_finset_image {f : γ → α} {g : α → β} {s : finset γ} :
670670
(⨆ x ∈ s.image f, g x) = (⨆ y ∈ s, g (f y)) :=
671671
by rw [← supr_coe, coe_image, supr_image, supr_coe]
672672

673-
lemma sup_finset_image {f : γ → α} {g : α → β} {s : finset γ} :
674-
s.sup (g ∘ f) = (s.image f).sup g :=
675-
by { simp_rw [sup_eq_supr, comp_app], rw supr_finset_image, }
673+
lemma sup_finset_image {β γ : Type*} [semilattice_sup_bot β]
674+
(f : γ → α) (g : α → β) (s : finset γ) :
675+
(s.image f).sup g = s.sup (g ∘ f) :=
676+
begin
677+
classical,
678+
apply finset.induction_on s,
679+
{ simp },
680+
{ intros a s' ha ih,
681+
rw [sup_insert, image_insert, sup_insert, ih] }
682+
end
676683

677684
lemma infi_finset_image {f : γ → α} {g : α → β} {s : finset γ} :
678685
(⨅ x ∈ s.image f, g x) = (⨅ y ∈ s, g (f y)) :=

src/data/finset/powerset.lean

Lines changed: 57 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Mario Carneiro
55
-/
6-
import data.finset.basic
6+
import data.finset.lattice
77

88
/-!
99
# The powerset of a finset
@@ -106,6 +106,45 @@ theorem powerset_len_eq_filter {n} {s : finset α} :
106106
powerset_len n s = (powerset s).filter (λ x, x.card = n) :=
107107
by { ext, simp [mem_powerset_len] }
108108

109+
lemma powerset_len_succ_insert [decidable_eq α] {x : α} {s : finset α} (h : x ∉ s) (n : ℕ) :
110+
powerset_len n.succ (insert x s) = powerset_len n.succ s ∪ (powerset_len n s).image (insert x) :=
111+
begin
112+
rw [powerset_len_eq_filter, powerset_insert, filter_union, ←powerset_len_eq_filter],
113+
congr,
114+
rw [powerset_len_eq_filter, image_filter],
115+
congr' 1,
116+
ext t,
117+
simp only [mem_powerset, mem_filter, function.comp_app, and.congr_right_iff],
118+
intro ht,
119+
have : x ∉ t := λ H, h (ht H),
120+
simp [card_insert_of_not_mem this, nat.succ_inj']
121+
end
122+
123+
lemma powerset_len_nonempty {n : ℕ} {s : finset α} (h : n < s.card) :
124+
(powerset_len n s).nonempty :=
125+
begin
126+
classical,
127+
induction s using finset.induction_on with x s hx IH generalizing n,
128+
{ simpa using h },
129+
{ cases n,
130+
{ simp },
131+
{ rw [card_insert_of_not_mem hx, nat.succ_lt_succ_iff] at h,
132+
rw powerset_len_succ_insert hx,
133+
refine nonempty.mono _ ((IH h).image (insert x)),
134+
convert (subset_union_right _ _) } }
135+
end
136+
137+
@[simp] lemma powerset_len_self (s : finset α) :
138+
powerset_len s.card s = {s} :=
139+
begin
140+
ext,
141+
rw [mem_powerset_len, mem_singleton],
142+
split,
143+
{ exact λ ⟨hs, hc⟩, eq_of_subset_of_card_le hs hc.ge },
144+
{ rintro rfl,
145+
simp }
146+
end
147+
109148
lemma powerset_card_bUnion [decidable_eq (finset α)] (s : finset α) :
110149
finset.powerset s = (range (s.card + 1)).bUnion (λ i, powerset_len i s) :=
111150
begin
@@ -117,5 +156,22 @@ begin
117156
exact mem_powerset.mpr (mem_powerset_len.mp ha).1, }
118157
end
119158

159+
lemma powerset_len_sup [decidable_eq α] (u : finset α) (n : ℕ) (hn : n < u.card) :
160+
(powerset_len n.succ u).sup id = u :=
161+
begin
162+
apply le_antisymm,
163+
{ simp [mem_powerset_len, and.comm] },
164+
{ rw [sup_eq_bUnion, le_iff_subset, subset_iff],
165+
cases (nat.succ_le_of_lt hn).eq_or_lt with h' h',
166+
{ simp [h'] },
167+
{ intros x hx,
168+
simp only [mem_bUnion, exists_prop, id.def],
169+
obtain ⟨t, ht⟩ : ∃ t, t ∈ powerset_len n (u.erase x) := powerset_len_nonempty _,
170+
{ refine ⟨insert x t, _, mem_insert_self _ _⟩,
171+
rw [←insert_erase hx, powerset_len_succ_insert (not_mem_erase _ _)],
172+
exact mem_union_right _ (mem_image_of_mem _ ht) },
173+
{ rwa [card_erase_of_mem hx, nat.lt_pred_iff] } } }
174+
end
175+
120176
end powerset_len
121177
end finset

src/data/finsupp/basic.lean

Lines changed: 48 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -213,6 +213,9 @@ if_neg h
213213
lemma single_eq_update : ⇑(single a b) = function.update 0 a b :=
214214
by rw [single_eq_indicator, ← set.piecewise_eq_indicator, set.piecewise_singleton]
215215

216+
lemma single_eq_pi_single : ⇑(single a b) = pi.single a b :=
217+
single_eq_update
218+
216219
@[simp] lemma single_zero : (single a 0 : α →₀ M) = 0 :=
217220
coe_fn_injective $ by simpa only [single_eq_update, coe_zero]
218221
using function.update_eq_self a (0 : α → M)
@@ -284,6 +287,18 @@ lemma single_left_inj (h : b ≠ 0) :
284287
and_false, or_false, eq_self_iff_true, and_true] using H,
285288
λ H, by rw [H]⟩
286289

290+
lemma support_single_ne_bot (i : α) (h : b ≠ 0) :
291+
(single i b).support ≠ ⊥ :=
292+
begin
293+
have : i ∈ (single i b).support := by simpa using h,
294+
intro H,
295+
simpa [H]
296+
end
297+
298+
lemma support_single_disjoint {b' : M} (hb : b ≠ 0) (hb' : b' ≠ 0) {i j : α} :
299+
disjoint (single i b).support (single j b').support ↔ i ≠ j :=
300+
by simpa [support_single_ne_zero, hb, hb'] using ne_comm
301+
287302
@[simp] lemma single_eq_zero : single a b = 0 ↔ b = 0 :=
288303
by simp [ext_iff, single_eq_indicator]
289304

@@ -1090,6 +1105,25 @@ lemma multiset_sum_sum_index
10901105
multiset.induction_on f rfl $ assume a s ih,
10911106
by rw [multiset.sum_cons, multiset.map_cons, multiset.sum_cons, sum_add_index h₀ h₁, ih]
10921107

1108+
lemma support_sum_eq_bUnion {α : Type*} {ι : Type*} {M : Type*} [add_comm_monoid M]
1109+
{g : ι → α →₀ M} (s : finset ι) (h : ∀ i₁ i₂, i₁ ≠ i₂ → disjoint (g i₁).support (g i₂).support) :
1110+
(∑ i in s, g i).support = s.bUnion (λ i, (g i).support) :=
1111+
begin
1112+
apply finset.induction_on s,
1113+
{ simp },
1114+
{ intros i s hi,
1115+
simp only [hi, sum_insert, not_false_iff, bUnion_insert],
1116+
intro hs,
1117+
rw [finsupp.support_add_eq, hs],
1118+
rw [hs],
1119+
intros x hx,
1120+
simp only [mem_bUnion, exists_prop, inf_eq_inter, ne.def, mem_inter] at hx,
1121+
obtain ⟨hxi, j, hj, hxj⟩ := hx,
1122+
have hn : i ≠ j := λ H, hi (H.symm ▸ hj),
1123+
apply h _ _ hn,
1124+
simp [hxi, hxj] }
1125+
end
1126+
10931127
lemma multiset_map_sum [has_zero M] {f : α →₀ M} {m : β → γ} {h : α → M → multiset β} :
10941128
multiset.map m (f.sum h) = f.sum (λa b, (h a b).map m) :=
10951129
(f.support.sum_hom _).symm
@@ -1493,6 +1527,20 @@ lemma to_multiset_apply (f : α →₀ ℕ) : f.to_multiset = f.sum (λ a n, n
14931527
@[simp] lemma to_multiset_single (a : α) (n : ℕ) : to_multiset (single a n) = n •ℕ {a} :=
14941528
by rw [to_multiset_apply, sum_single_index]; apply zero_nsmul
14951529

1530+
lemma to_multiset_sum {ι : Type*} {f : ι → α →₀ ℕ} (s : finset ι) :
1531+
finsupp.to_multiset (∑ i in s, f i) = ∑ i in s, finsupp.to_multiset (f i) :=
1532+
begin
1533+
apply finset.induction_on s,
1534+
{ simp },
1535+
{ intros i s hi,
1536+
simp [hi] }
1537+
end
1538+
1539+
lemma to_multiset_sum_single {ι : Type*} (s : finset ι) (n : ℕ) :
1540+
finsupp.to_multiset (∑ i in s, single i n) = n •ℕ s.val :=
1541+
by simp_rw [to_multiset_sum, finsupp.to_multiset_single, multiset.singleton_eq_singleton,
1542+
sum_nsmul, sum_multiset_singleton]
1543+
14961544
lemma card_to_multiset (f : α →₀ ℕ) : f.to_multiset.card = f.sum (λa, id) :=
14971545
by simp [to_multiset_apply, add_monoid_hom.map_finsupp_sum, function.id_def]
14981546

src/data/fintype/basic.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1431,3 +1431,13 @@ def trunc_sigma_of_exists {α} [fintype α] {P : α → Prop} [decidable_pred P]
14311431
@trunc_of_nonempty_fintype (Σ' a, P a) (exists.elim h $ λ a ha, ⟨⟨a, ha⟩⟩) _
14321432

14331433
end trunc
1434+
1435+
namespace multiset
1436+
1437+
variables [fintype α] [decidable_eq α]
1438+
1439+
@[simp] lemma count_univ (a : α) :
1440+
count a finset.univ.val = 1 :=
1441+
count_eq_one_of_mem finset.univ.nodup (finset.mem_univ _)
1442+
1443+
end multiset

src/data/multiset/lattice.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -59,6 +59,15 @@ by rw [← sup_erase_dup, erase_dup_ext.2, sup_erase_dup, sup_add]; simp
5959
(ndinsert a s).sup = a ⊔ s.sup :=
6060
by rw [← sup_erase_dup, erase_dup_ext.2, sup_erase_dup, sup_cons]; simp
6161

62+
lemma nodup_sup_iff {α : Type*} [decidable_eq α] {m : multiset (multiset α) } :
63+
m.sup.nodup ↔ ∀ (a : multiset α), a ∈ m → a.nodup :=
64+
begin
65+
apply m.induction_on,
66+
{ simp },
67+
{ intros a s h,
68+
simp [h] }
69+
end
70+
6271
end sup
6372

6473
/-! ### inf -/

src/data/set/basic.lean

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1746,6 +1746,28 @@ begin
17461746
rw [hz x hx, hz y hy] }
17471747
end
17481748

1749+
@[simp] lemma pairwise_on_empty {α} (r : α → α → Prop) :
1750+
(∅ : set α).pairwise_on r :=
1751+
λ _, by simp
1752+
1753+
lemma pairwise_on_insert_of_symmetric {α} {s : set α} {a : α} {r : α → α → Prop}
1754+
(hr : symmetric r) :
1755+
(insert a s).pairwise_on r ↔ s.pairwise_on r ∧ ∀ b ∈ s, a ≠ b → r a b :=
1756+
begin
1757+
refine ⟨λ h, ⟨_, _⟩, λ h, _⟩,
1758+
{ exact h.mono (s.subset_insert a) },
1759+
{ intros b hb hn,
1760+
exact h a (s.mem_insert _) b (set.mem_insert_of_mem _ hb) hn },
1761+
{ intros b hb c hc hn,
1762+
rw [mem_insert_iff] at hb hc,
1763+
rcases hb with (rfl | hb);
1764+
rcases hc with (rfl | hc),
1765+
{ exact absurd rfl hn },
1766+
{ exact h.right _ hc hn },
1767+
{ exact hr (h.right _ hb hn.symm) },
1768+
{ exact h.left _ hb _ hc hn } }
1769+
end
1770+
17491771
end set
17501772

17511773
open set

src/data/support.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -257,4 +257,8 @@ begin
257257
simp
258258
end
259259

260+
lemma support_single_disjoint {b' : B} (hb : b ≠ 0) (hb' : b' ≠ 0) {i j : A} :
261+
disjoint (function.support (single i b)) (function.support (single j b')) ↔ i ≠ j :=
262+
by simpa [support_single, hb, hb'] using ne_comm
263+
260264
end pi

src/order/compactly_generated.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -144,7 +144,7 @@ begin
144144
classical,
145145
rw is_compact_element_iff_le_of_directed_Sup_le,
146146
intros d hemp hdir hsup,
147-
change f with id ∘ f, rw finset.sup_finset_image,
147+
change f with id ∘ f, rw finset.sup_finset_image,
148148
apply finset.sup_le_of_le_directed d hemp hdir,
149149
rintros x hx,
150150
obtain ⟨p, ⟨hps, rfl⟩⟩ := finset.mem_image.mp hx,

src/ring_theory/noetherian.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -289,7 +289,7 @@ begin
289289
{ suffices : u.sup id ≤ s, from le_antisymm husup this,
290290
rw [sSup, finset.sup_eq_Sup], exact Sup_le_Sup huspan, },
291291
obtain ⟨t, ⟨hts, rfl⟩⟩ := finset.subset_image_iff.mp huspan,
292-
rw [finset.sup_finset_image, function.comp.left_id, finset.sup_eq_supr, supr_rw,
292+
rw [finset.sup_finset_image, function.comp.left_id, finset.sup_eq_supr, supr_rw,
293293
←span_eq_supr_of_singleton_spans, eq_comm] at ssup,
294294
exact ⟨t, ssup⟩, },
295295
end

0 commit comments

Comments
 (0)