@@ -7,7 +7,7 @@ import Mathlib.Data.Set.List
7
7
import Mathlib.Data.List.Perm
8
8
import Mathlib.Init.Quot -- Porting note: added import
9
9
10
- #align_import data.multiset.basic from "leanprover-community/mathlib" @"06a655b5fcfbda03502f9158bbf6c0f1400886f9 "
10
+ #align_import data.multiset.basic from "leanprover-community/mathlib" @"65a1391a0106c9204fe45bc73a039f056558cb83 "
11
11
12
12
/-!
13
13
# Multisets
@@ -2091,6 +2091,10 @@ theorem filter_filter (q) [DecidablePred q] (s : Multiset α) :
2091
2091
Quot.inductionOn s fun l => by simp
2092
2092
#align multiset.filter_filter Multiset.filter_filter
2093
2093
2094
+ lemma filter_comm (q) [DecidablePred q] (s : Multiset α) :
2095
+ filter p (filter q s) = filter q (filter p s) := by simp [and_comm]
2096
+ #align multiset.filter_comm Multiset.filter_comm
2097
+
2094
2098
theorem filter_add_filter (q) [DecidablePred q] (s : Multiset α) :
2095
2099
filter p s + filter q s = filter (fun a => p a ∨ q a) s + filter (fun a => p a ∧ q a) s :=
2096
2100
Multiset.induction_on s rfl fun a s IH => by by_cases p a <;> by_cases q a <;> simp [*]
@@ -2108,6 +2112,12 @@ theorem map_filter (f : β → α) (s : Multiset β) : filter p (map f s) = map
2108
2112
Quot.inductionOn s fun l => by simp [List.map_filter]; rfl
2109
2113
#align multiset.map_filter Multiset.map_filter
2110
2114
2115
+ lemma map_filter' {f : α → β} (hf : Injective f) (s : Multiset α)
2116
+ [DecidablePred fun b => ∃ a, p a ∧ f a = b] :
2117
+ (s.filter p).map f = (s.map f).filter fun b => ∃ a, p a ∧ f a = b := by
2118
+ simp [(· ∘ ·), map_filter, hf.eq_iff]
2119
+ #align multiset.map_filter' Multiset.map_filter'
2120
+
2111
2121
/-! ### Simultaneously filter and map elements of a multiset -/
2112
2122
2113
2123
@@ -2312,6 +2322,25 @@ theorem countP_map (f : α → β) (s : Multiset α) (p : β → Prop) [Decidabl
2312
2322
add_comm]
2313
2323
#align multiset.countp_map Multiset.countP_map
2314
2324
2325
+ -- porting note: `Lean.Internal.coeM` forces us to type-ascript `{a // a ∈ s}`
2326
+ lemma countP_attach (s : Multiset α) : s.attach.countP (fun a : {a // a ∈ s} ↦ p a) = s.countP p :=
2327
+ Quotient.inductionOn s fun l => by
2328
+ simp only [quot_mk_to_coe, coe_countP]
2329
+ -- porting note: was
2330
+ -- rw [quot_mk_to_coe, coe_attach, coe_countP]
2331
+ -- exact List.countP_attach _ _
2332
+ rw [coe_attach]
2333
+ refine (coe_countP _ _).trans ?_
2334
+ convert List.countP_attach _ _
2335
+ rfl
2336
+ #align multiset.countp_attach Multiset.countP_attach
2337
+
2338
+ lemma filter_attach (s : Multiset α) (p : α → Prop ) [DecidablePred p] :
2339
+ (s.attach.filter fun a : {a // a ∈ s} ↦ p ↑a) =
2340
+ (s.filter p).attach.map (Subtype.map id fun _ ↦ Multiset.mem_of_mem_filter) :=
2341
+ Quotient.inductionOn s fun l ↦ congr_arg _ (List.filter_attach l p)
2342
+ #align multiset.filter_attach Multiset.filter_attach
2343
+
2315
2344
variable {p}
2316
2345
2317
2346
theorem countP_pos {s} : 0 < countP p s ↔ ∃ a ∈ s, p a :=
@@ -2348,7 +2377,7 @@ end
2348
2377
2349
2378
section
2350
2379
2351
- variable [DecidableEq α]
2380
+ variable [DecidableEq α] {s : Multiset α}
2352
2381
2353
2382
/-- `count a s` is the multiplicity of `a` in `s`. -/
2354
2383
def count (a : α) : Multiset α → ℕ :=
@@ -2421,6 +2450,11 @@ theorem count_nsmul (a : α) (n s) : count a (n • s) = n * count a s := by
2421
2450
induction n <;> simp [*, succ_nsmul', succ_mul, zero_nsmul]
2422
2451
#align multiset.count_nsmul Multiset.count_nsmul
2423
2452
2453
+ @[simp]
2454
+ lemma count_attach (a : {x // x ∈ s}) : s.attach.count a = s.count ↑a :=
2455
+ Eq.trans (countP_congr rfl fun _ _ => by simp [Subtype.ext_iff]) <| countP_attach _ _
2456
+ #align multiset.count_attach Multiset.count_attach
2457
+
2424
2458
theorem count_pos {a : α} {s : Multiset α} : 0 < count a s ↔ a ∈ s := by simp [count, countP_pos]
2425
2459
#align multiset.count_pos Multiset.count_pos
2426
2460
@@ -2577,14 +2611,6 @@ theorem count_map_eq_count' [DecidableEq β] (f : α → β) (s : Multiset α) (
2577
2611
contradiction
2578
2612
#align multiset.count_map_eq_count' Multiset.count_map_eq_count'
2579
2613
2580
- @[simp]
2581
- theorem attach_count_eq_count_coe (m : Multiset α) (a) : m.attach.count a = m.count (a : α) :=
2582
- calc
2583
- m.attach.count a = (m.attach.map (Subtype.val : _ → α)).count (a : α) :=
2584
- (Multiset.count_map_eq_count' _ _ Subtype.coe_injective _).symm
2585
- _ = m.count (a : α) := congr_arg _ m.attach_map_val
2586
- #align multiset.attach_count_eq_count_coe Multiset.attach_count_eq_count_coe
2587
-
2588
2614
theorem filter_eq' (s : Multiset α) (b : α) : s.filter (· = b) = replicate (count b s) b :=
2589
2615
Quotient.inductionOn s <| fun l => by
2590
2616
simp only [quot_mk_to_coe, coe_filter, mem_coe, coe_count]
@@ -2874,6 +2900,17 @@ theorem map_injective {f : α → β} (hf : Function.Injective f) :
2874
2900
Function.Injective (Multiset.map f) := fun _x _y => (map_eq_map hf).1
2875
2901
#align multiset.map_injective Multiset.map_injective
2876
2902
2903
+ lemma filter_attach' (s : Multiset α) (p : {a // a ∈ s} → Prop ) [DecidableEq α]
2904
+ [DecidablePred p] :
2905
+ s.attach.filter p =
2906
+ (s.filter fun x ↦ ∃ h, p ⟨x, h⟩).attach.map (Subtype.map id fun x ↦ mem_of_mem_filter) := by
2907
+ classical
2908
+ refine' Multiset.map_injective Subtype.val_injective _
2909
+ rw [map_filter' _ Subtype.val_injective]
2910
+ simp only [Function.comp, Subtype.exists, coe_mk,
2911
+ exists_and_right, exists_eq_right, attach_map_val, map_map, map_coe, id.def]
2912
+ #align multiset.filter_attach' Multiset.filter_attach'
2913
+
2877
2914
end Map
2878
2915
2879
2916
section Quot
0 commit comments