Skip to content

Commit

Permalink
feat(data/finset/basic): disj_Union (#16421)
Browse files Browse the repository at this point in the history
`finset.disj_Union` is to `finset.bUnion` as:
* `finset.disj_union` is to `finset.union`
* `finset.cons` is to `insert`
* `finset.map` is to `finset.image`



Co-authored-by: Yaël Dillies <yael.dillies@gmail.com>
  • Loading branch information
eric-wieser and YaelDillies committed Oct 11, 2022
1 parent 0a0be05 commit ccf8aa0
Show file tree
Hide file tree
Showing 8 changed files with 137 additions and 13 deletions.
14 changes: 14 additions & 0 deletions src/algebra/big_operators/basic.lean
Expand Up @@ -261,6 +261,16 @@ attribute [congr] finset.sum_congr
lemma prod_disj_union (h) : ∏ x in s₁.disj_union s₂ h, f x = (∏ x in s₁, f x) * ∏ x in s₂, f x :=
by { refine eq.trans _ (fold_disj_union h), rw one_mul, refl }

@[to_additive]
lemma prod_disj_Union (s : finset ι) (t : ι → finset α) (h) :
∏ x in s.disj_Union t h, f x = ∏ i in s, ∏ x in t i, f x :=
begin
refine eq.trans _ (fold_disj_Union h),
dsimp [finset.prod, multiset.prod, multiset.fold, finset.disj_Union, finset.fold],
congr',
exact prod_const_one.symm,
end

@[to_additive]
lemma prod_union_inter [decidable_eq α] :
(∏ x in (s₁ ∪ s₂), f x) * (∏ x in (s₁ ∩ s₂), f x) = (∏ x in s₁, f x) * (∏ x in s₂, f x) :=
Expand Down Expand Up @@ -1469,6 +1479,10 @@ end comm_group
card (s.sigma t) = ∑ a in s, card (t a) :=
multiset.card_sigma _ _

@[simp] lemma card_disj_Union (s : finset α) (t : α → finset β) (h) :
(s.disj_Union t h).card = s.sum (λ i, (t i).card) :=
multiset.card_bind _ _

lemma card_bUnion [decidable_eq β] {s : finset α} {t : α → finset β}
(h : ∀ x ∈ s, ∀ y ∈ s, x ≠ y → disjoint (t x) (t y)) :
(s.bUnion t).card = ∑ u in s, card (t u) :=
Expand Down
80 changes: 80 additions & 0 deletions src/data/finset/basic.lean
Expand Up @@ -2623,6 +2623,78 @@ cons_eq_insert _ _ h ▸ to_list_cons _

end to_list

/-!
### disj_Union
This section is about the bounded union of a disjoint indexed family `t : α → finset β` of finite
sets over a finite set `s : finset α`. In most cases `finset.bUnion` should be preferred.
-/
section disj_Union

variables {s s₁ s₂ : finset α} {t t₁ t₂ : α → finset β}

/-- `disj_Union s f h` is the set such that `a ∈ disj_Union s f` iff `a ∈ f i` for some `i ∈ s`.
It is the same as `s.bUnion f`, but it does not require decidable equality on the type. The
hypothesis ensures that the sets are disjoint. -/
def disj_Union (s : finset α) (t : α → finset β)
(hf : (s : set α).pairwise $ λ a b, ∀ x, x ∈ t a → x ∉ t b) : finset β :=
⟨(s.val.bind (finset.val ∘ t)), multiset.nodup_bind.mpr
⟨λ a ha, (t a).nodup, s.nodup.pairwise hf⟩⟩

@[simp] theorem disj_Union_val (s : finset α) (t : α → finset β) (h) :
(s.disj_Union t h).1 = (s.1.bind (λ a, (t a).1)) := rfl

@[simp] theorem disj_Union_empty (t : α → finset β) : disj_Union ∅ t (by simp) = ∅ := rfl

@[simp] lemma mem_disj_Union {b : β} {h} :
b ∈ s.disj_Union t h ↔ ∃ a ∈ s, b ∈ t a :=
by simp only [mem_def, disj_Union_val, mem_bind, exists_prop]

@[simp, norm_cast] lemma coe_disj_Union {h} : (s.disj_Union t h : set β) = ⋃ x ∈ (s : set α), t x :=
by simp only [set.ext_iff, mem_disj_Union, set.mem_Union, iff_self, mem_coe, implies_true_iff]

@[simp] theorem disj_Union_cons (a : α) (s : finset α) (ha : a ∉ s) (f : α → finset β) (H) :
disj_Union (cons a s ha) f H = (f a).disj_union
(s.disj_Union f $
λ b hb c hc, H (mem_cons_of_mem hb) (mem_cons_of_mem hc))
(λ b hb h, let ⟨c, hc, h⟩ := mem_disj_Union.mp h in
H (mem_cons_self a s) (mem_cons_of_mem hc) (ne_of_mem_of_not_mem hc ha).symm b hb h) :=
eq_of_veq $ multiset.cons_bind _ _ _

@[simp] lemma singleton_disj_Union (a : α) {h} : finset.disj_Union {a} t h = t a :=
eq_of_veq $ multiset.singleton_bind _ _

theorem map_disj_Union {f : α ↪ β} {s : finset α} {t : β → finset γ} {h} :
(s.map f).disj_Union t h = s.disj_Union (λa, t (f a))
(λ a ha b hb hab, h (mem_map_of_mem _ ha) (mem_map_of_mem _ hb) (f.injective.ne hab)) :=
eq_of_veq $ multiset.bind_map _ _ _

theorem disj_Union_map {s : finset α} {t : α → finset β} {f : β ↪ γ} {h} :
(s.disj_Union t h).map f = s.disj_Union (λa, (t a).map f)
(λ a ha b hb hab x hxa hxb, begin
obtain ⟨xa, hfa, rfl⟩ := mem_map.mp hxa,
obtain ⟨xb, hfb, hfab⟩ := mem_map.mp hxb,
obtain rfl := f.injective hfab,
exact h ha hb hab _ hfa hfb,
end) :=
eq_of_veq $ multiset.map_bind _ _ _

lemma disj_Union_disj_Union (s : finset α) (f : α → finset β) (g : β → finset γ) (h1 h2) :
(s.disj_Union f h1).disj_Union g h2 =
s.attach.disj_Union (λ a, (f a).disj_Union g $
λ b hb c hc, h2 (mem_disj_Union.mpr ⟨_, a.prop, hb⟩) (mem_disj_Union.mpr ⟨_, a.prop, hc⟩))
(λ a ha b hb hab x hxa hxb, begin
obtain ⟨xa, hfa, hga⟩ := mem_disj_Union.mp hxa,
obtain ⟨xb, hfb, hgb⟩ := mem_disj_Union.mp hxb,
refine h2
(mem_disj_Union.mpr ⟨_, a.prop, hfa⟩) (mem_disj_Union.mpr ⟨_, b.prop, hfb⟩) _ _ hga hgb,
rintro rfl,
exact h1 a.prop b.prop (subtype.coe_injective.ne hab) _ hfa hfb,
end) :=
eq_of_veq $ multiset.bind_assoc.trans (multiset.attach_bind_coe _ _).symm

end disj_Union

section bUnion
/-!
### bUnion
Expand Down Expand Up @@ -2657,6 +2729,14 @@ ext $ λ x, by simp only [mem_bUnion, exists_prop, mem_union, mem_insert,
lemma bUnion_congr (hs : s₁ = s₂) (ht : ∀ a ∈ s₁, t₁ a = t₂ a) : s₁.bUnion t₁ = s₂.bUnion t₂ :=
ext $ λ x, by simp [hs, ht] { contextual := tt }

@[simp] lemma disj_Union_eq_bUnion (s : finset α) (f : α → finset β) (hf) :
s.disj_Union f hf = s.bUnion f :=
begin
dsimp [disj_Union, finset.bUnion, function.comp],
generalize_proofs h,
exact eq_of_veq h.dedup.symm,
end

theorem bUnion_subset {s' : finset β} : s.bUnion t ⊆ s' ↔ ∀ x ∈ s, t x ⊆ s' :=
by simp only [subset_iff, mem_bUnion]; exact
⟨λ H a ha b hb, H ⟨a, ha, hb⟩, λ H b ⟨a, ha, hb⟩, H a ha hb⟩
Expand Down
4 changes: 4 additions & 0 deletions src/data/finset/fold.lean
Expand Up @@ -74,6 +74,10 @@ theorem fold_disj_union {s₁ s₂ : finset α} {b₁ b₂ : β} (h) :
(s₁.disj_union s₂ h).fold op (b₁ * b₂) f = s₁.fold op b₁ f * s₂.fold op b₂ f :=
(congr_arg _ $ multiset.map_add _ _ _).trans (multiset.fold_add _ _ _ _ _)

theorem fold_disj_Union {ι : Type*} {s : finset ι} {t : ι → finset α} {b : ι → β} {b₀ : β} (h) :
(s.disj_Union t h).fold op (s.fold op b₀ b) f = s.fold op b₀ (λ i, (t i).fold op (b i) f) :=
(congr_arg _ $ multiset.map_bind _ _ _).trans (multiset.fold_bind _ _ _ _ _)

theorem fold_union_inter [decidable_eq α] {s₁ s₂ : finset α} {b₁ b₂ : β} :
(s₁ ∪ s₂).fold op b₁ f * (s₁ ∩ s₂).fold op b₂ f = s₁.fold op b₂ f * s₂.fold op b₁ f :=
by unfold fold; rw [← fold_add op, ← multiset.map_add, union_val,
Expand Down
12 changes: 10 additions & 2 deletions src/data/list/basic.lean
Expand Up @@ -2710,8 +2710,16 @@ theorem pmap_eq_map_attach {p : α → Prop} (f : Π a, p a → β)
(l H) : pmap f l H = l.attach.map (λ x, f x.1 (H _ x.2)) :=
by rw [attach, map_pmap]; exact pmap_congr l (λ _ _ _ _, rfl)

theorem attach_map_val (l : list α) : l.attach.map subtype.val = l :=
by rw [attach, map_pmap]; exact (pmap_eq_map _ _ _ _).trans (map_id l)
@[simp] lemma attach_map_coe' (l : list α) (f : α → β) : l.attach.map (λ i, f i) = l.map f :=
by rw [attach, map_pmap]; exact (pmap_eq_map _ _ _ _)

lemma attach_map_val' (l : list α) (f : α → β) : l.attach.map (λ i, f i.val) = l.map f :=
attach_map_coe' _ _

@[simp] lemma attach_map_coe (l : list α) : l.attach.map (coe : _ → α) = l :=
(attach_map_coe' _ _).trans l.map_id

lemma attach_map_val (l : list α) : l.attach.map subtype.val = l := attach_map_coe _

@[simp] theorem mem_attach (l : list α) : ∀ x, x ∈ l.attach | ⟨a, h⟩ :=
by have := mem_map.1 (by rw [attach_map_val]; exact h);
Expand Down
17 changes: 11 additions & 6 deletions src/data/multiset/basic.lean
Expand Up @@ -1077,11 +1077,19 @@ theorem map_pmap {p : α → Prop} (g : β → γ) (f : Π a, p a → β)
quot.induction_on s $ λ l H, congr_arg coe $ map_pmap g f l H

theorem pmap_eq_map_attach {p : α → Prop} (f : Π a, p a → β)
(s) : ∀ H, pmap f s H = s.attach.map (λ x, f x.1 (H _ x.2)) :=
(s) : ∀ H, pmap f s H = s.attach.map (λ x, f x (H _ x.prop)) :=
quot.induction_on s $ λ l H, congr_arg coe $ pmap_eq_map_attach f l H

theorem attach_map_val (s : multiset α) : s.attach.map subtype.val = s :=
quot.induction_on s $ λ l, congr_arg coe $ attach_map_val l
@[simp] lemma attach_map_coe' (s : multiset α) (f : α → β) : s.attach.map (λ i, f i) = s.map f :=
quot.induction_on s $ λ l, congr_arg coe $ attach_map_coe' l f

lemma attach_map_val' (s : multiset α) (f : α → β) : s.attach.map (λ i, f i.val) = s.map f :=
attach_map_coe' _ _

@[simp] lemma attach_map_coe (s : multiset α) : s.attach.map (coe : _ → α) = s :=
(attach_map_coe' _ _).trans s.map_id

lemma attach_map_val (s : multiset α) : s.attach.map subtype.val = s := attach_map_coe _

@[simp] theorem mem_attach (s : multiset α) : ∀ x, x ∈ s.attach :=
quot.induction_on s $ λ l, mem_attach _
Expand All @@ -1103,9 +1111,6 @@ lemma attach_cons (a : α) (m : multiset α) :
quotient.induction_on m $ assume l, congr_arg coe $ congr_arg (list.cons _) $
by rw [list.map_pmap]; exact list.pmap_congr _ (λ _ _ _ _, subtype.eq rfl)

@[simp]
lemma attach_map_coe (m : multiset α) : multiset.map (coe : _ → α) m.attach = m := m.attach_map_val

section decidable_pi_exists
variables {m : multiset α}

Expand Down
5 changes: 5 additions & 0 deletions src/data/multiset/bind.lean
Expand Up @@ -139,6 +139,11 @@ begin
rw count_bind, apply le_sum_of_mem,
rw mem_map, exact ⟨x, hx, rfl⟩
end

@[simp] theorem attach_bind_coe (s : multiset α) (f : α → multiset β) :
s.attach.bind (λ i, f i) = s.bind f :=
congr_arg join $ attach_map_coe' _ _

end bind

/-! ### Product of two multisets -/
Expand Down
8 changes: 8 additions & 0 deletions src/data/multiset/fold.lean
Expand Up @@ -55,6 +55,14 @@ multiset.induction_on s₂
(by rw [add_zero, fold_zero, ← fold_cons'_right, ← fold_cons_right op])
(by simp {contextual := tt}; cc)

theorem fold_bind {ι : Type*} (s : multiset ι) (t : ι → multiset α) (b : ι → α) (b₀ : α) :
(s.bind t).fold op ((s.map b).fold op b₀) = (s.map (λ i, (t i).fold op (b i))).fold op b₀ :=
begin
induction s using multiset.induction_on with a ha ih,
{ rw [zero_bind, map_zero, map_zero, fold_zero] },
{ rw [cons_bind, map_cons, map_cons, fold_cons_left, fold_cons_left, fold_add, ih] },
end

theorem fold_singleton (b a : α) : ({a} : multiset α).fold op b = a * b := foldr_singleton _ _ _ _

theorem fold_distrib {f g : β → α} (u₁ u₂ : α) (s : multiset β) :
Expand Down
10 changes: 5 additions & 5 deletions src/data/multiset/nodup.lean
Expand Up @@ -205,14 +205,14 @@ lemma map_eq_map_of_bij_of_nodup (f : α → γ) (g : β → γ) {s : multiset
(i_inj : ∀a₁ a₂ ha₁ ha₂, i a₁ ha₁ = i a₂ ha₂ → a₁ = a₂)
(i_surj : ∀b∈t, ∃a ha, b = i a ha) :
s.map f = t.map g :=
have t = s.attach.map (λ x, i x.1 x.2),
have t = s.attach.map (λ x, i x x.2),
from (ht.ext $ (nodup_attach.2 hs).map $
show injective (λ x : {x // x ∈ s}, i x.1 x.2), from λ x y hxy,
subtype.eq $ i_inj x.1 y.1 x.2 y.2 hxy).2
show injective (λ x : {x // x ∈ s}, i x x.2), from λ x y hxy,
subtype.ext $ i_inj x y x.2 y.2 hxy).2
(λ x, by simp only [mem_map, true_and, subtype.exists, eq_comm, mem_attach];
exact ⟨i_surj _, λ ⟨y, hy⟩, hy.snd.symm ▸ hi _ _⟩),
calc s.map f = s.pmap (λ x _, f x) (λ _, id) : by rw [pmap_eq_map]
... = s.attach.map (λ x, f x.1) : by rw [pmap_eq_map_attach]
calc s.map f = s.pmap (λ x _, f x) (λ _, id) : by rw [pmap_eq_map]
... = s.attach.map (λ x, f x) : by rw [pmap_eq_map_attach]
... = t.map g : by rw [this, multiset.map_map]; exact map_congr rfl (λ x _, h _ _)

end multiset

0 comments on commit ccf8aa0

Please sign in to comment.