@@ -16,7 +16,7 @@ These are implemented as the quotient of a list by permutations.
16
16
We define the global infix notation `::ₘ` for `multiset.cons`.
17
17
-/
18
18
19
- open list subtype nat
19
+ open function list nat subtype
20
20
21
21
variables {α : Type *} {β : Type *} {γ : Type *}
22
22
@@ -1543,6 +1543,10 @@ le_antisymm (le_inter
1543
1543
filter p (filter q s) = filter (λ a, p a ∧ q a) s :=
1544
1544
quot.induction_on s $ λ l, congr_arg coe $ filter_filter p q l
1545
1545
1546
+ lemma filter_comm (q) [decidable_pred q] (s : multiset α) :
1547
+ filter p (filter q s) = filter q (filter p s) :=
1548
+ by simp [and_comm]
1549
+
1546
1550
theorem filter_add_filter (q) [decidable_pred q] (s : multiset α) :
1547
1551
filter p s + filter q s = filter (λ a, p a ∨ q a) s + filter (λ a, p a ∧ q a) s :=
1548
1552
multiset.induction_on s rfl $ λ a s IH,
@@ -1556,6 +1560,11 @@ theorem map_filter (f : β → α) (s : multiset β) :
1556
1560
filter p (map f s) = map f (filter (p ∘ f) s) :=
1557
1561
quot.induction_on s (λ l, by simp [map_filter])
1558
1562
1563
+ lemma map_filter' {f : α → β} (hf : injective f) (s : multiset α)
1564
+ [decidable_pred (λ b, ∃ a, p a ∧ f a = b)] :
1565
+ (s.filter p).map f = (s.map f).filter (λ b, ∃ a, p a ∧ f a = b) :=
1566
+ by simp [(∘), map_filter, hf.eq_iff]
1567
+
1559
1568
/-! ### Simultaneously filter and map elements of a multiset -/
1560
1569
1561
1570
/-- `filter_map f s` is a combination filter/map operation on `s`.
@@ -1704,6 +1713,18 @@ begin
1704
1713
card_singleton, add_comm] },
1705
1714
end
1706
1715
1716
+ @[simp] lemma countp_attach (s : multiset α) : s.attach.countp (λ a, p ↑a) = s.countp p :=
1717
+ quotient.induction_on s $ λ l, begin
1718
+ simp only [quot_mk_to_coe, coe_countp],
1719
+ rw [quot_mk_to_coe, coe_attach, coe_countp],
1720
+ exact list.countp_attach _ _,
1721
+ end
1722
+
1723
+ @[simp] lemma filter_attach (m : multiset α) (p : α → Prop ) [decidable_pred p] :
1724
+ (m.attach.filter (λ x, p ↑x)) =
1725
+ (m.filter p).attach.map (subtype.map id $ λ _, multiset.mem_of_mem_filter) :=
1726
+ quotient.induction_on m $ λ l, congr_arg coe (list.filter_attach l p)
1727
+
1707
1728
variable {p}
1708
1729
1709
1730
theorem countp_pos {s} : 0 < countp p s ↔ ∃ a ∈ s, p a :=
@@ -1720,7 +1741,7 @@ countp_pos.2 ⟨_, h, pa⟩
1720
1741
1721
1742
theorem countp_congr {s s' : multiset α} (hs : s = s')
1722
1743
{p p' : α → Prop } [decidable_pred p] [decidable_pred p']
1723
- (hp : ∀ x ∈ s, p x = p' x) : s.countp p = s'.countp p' :=
1744
+ (hp : ∀ x ∈ s, p x ↔ p' x) : s.countp p = s'.countp p' :=
1724
1745
quot.induction_on₂ s s' (λ l l' hs hp, begin
1725
1746
simp only [quot_mk_to_coe'', coe_eq_coe] at hs,
1726
1747
exact hs.countp_congr hp,
@@ -1731,7 +1752,7 @@ end
1731
1752
/-! ### Multiplicity of an element -/
1732
1753
1733
1754
section
1734
- variable [decidable_eq α]
1755
+ variables [decidable_eq α] {s : multiset α}
1735
1756
1736
1757
/-- `count a s` is the multiplicity of `a` in `s`. -/
1737
1758
def count (a : α) : multiset α → ℕ := countp (eq a)
@@ -1778,6 +1799,9 @@ def count_add_monoid_hom (a : α) : multiset α →+ ℕ := countp_add_monoid_ho
1778
1799
@[simp] theorem count_nsmul (a : α) (n s) : count a (n • s) = n * count a s :=
1779
1800
by induction n; simp [*, succ_nsmul', succ_mul, zero_nsmul]
1780
1801
1802
+ @[simp] lemma count_attach (a : {x // x ∈ s}) : s.attach.count a = s.count a :=
1803
+ eq.trans (countp_congr rfl $ λ _ _, subtype.ext_iff) $ countp_attach _ _
1804
+
1781
1805
theorem count_pos {a : α} {s : multiset α} : 0 < count a s ↔ a ∈ s :=
1782
1806
by simp [count, countp_pos]
1783
1807
@@ -1901,13 +1925,6 @@ begin
1901
1925
contradiction }
1902
1926
end
1903
1927
1904
- @[simp]
1905
- lemma attach_count_eq_count_coe (m : multiset α) (a) : m.attach.count a = m.count (a : α) :=
1906
- calc m.attach.count a
1907
- = (m.attach.map (coe : _ → α)).count (a : α) :
1908
- (multiset.count_map_eq_count' _ _ subtype.coe_injective _).symm
1909
- ... = m.count (a : α) : congr_arg _ m.attach_map_coe
1910
-
1911
1928
lemma filter_eq' (s : multiset α) (b : α) : s.filter (= b) = replicate (count b s) b :=
1912
1929
quotient.induction_on s $ λ l, congr_arg coe $ filter_eq' l b
1913
1930
@@ -2174,6 +2191,19 @@ theorem map_injective {f : α → β} (hf : function.injective f) :
2174
2191
function.injective (multiset.map f) :=
2175
2192
assume x y, (map_eq_map hf).1
2176
2193
2194
+ lemma filter_attach' (s : multiset α) (p : {a // a ∈ s} → Prop ) [decidable_eq α]
2195
+ [decidable_pred p] :
2196
+ s.attach.filter p =
2197
+ (s.filter $ λ x, ∃ h, p ⟨x, h⟩).attach.map (subtype.map id $ λ x hx,
2198
+ let ⟨h, _⟩ := of_mem_filter hx in h) :=
2199
+ begin
2200
+ classical,
2201
+ refine multiset.map_injective subtype.coe_injective _,
2202
+ simp only [function.comp, map_filter' _ subtype.coe_injective, subtype.exists, coe_mk,
2203
+ exists_and_distrib_right, exists_eq_right, attach_map_coe, map_map, map_coe, id.def],
2204
+ rw attach_map_coe,
2205
+ end
2206
+
2177
2207
end map
2178
2208
2179
2209
section quot
0 commit comments