Nodup s means that s has no duplicates, i.e. the multiplicity of
any element is at most 1.
Equations
- s.Nodup = Quot.liftOn s List.Nodup ⋯
Instances For
theorem
Multiset.Nodup.cons
{α : Type u_1}
{s : Multiset α}
{a : α}
(m : a ∉ s)
(n : s.Nodup)
 :
(a ::ₘ s).Nodup
theorem
Multiset.Nodup.of_cons
{α : Type u_1}
{s : Multiset α}
{a : α}
(h : (a ::ₘ s).Nodup)
 :
s.Nodup
theorem
Multiset.nodup_iff_count_le_one
{α : Type u_1}
[DecidableEq α]
{s : Multiset α}
 :
s.Nodup ↔ ∀ (a : α), Multiset.count a s ≤ 1
theorem
Multiset.nodup_iff_count_eq_one
{α : Type u_1}
{s : Multiset α}
[DecidableEq α]
 :
s.Nodup ↔ ∀ a ∈ s, Multiset.count a s = 1
@[simp]
theorem
Multiset.count_eq_one_of_mem
{α : Type u_1}
[DecidableEq α]
{a : α}
{s : Multiset α}
(d : s.Nodup)
(h : a ∈ s)
 :
Multiset.count a s = 1
theorem
Multiset.count_eq_of_nodup
{α : Type u_1}
[DecidableEq α]
{a : α}
{s : Multiset α}
(d : s.Nodup)
 :
Multiset.count a s = if a ∈ s then 1 else 0
theorem
Multiset.nodup_iff_pairwise
{α : Type u_4}
{s : Multiset α}
 :
s.Nodup ↔ Multiset.Pairwise (fun (x x_1 : α) => x ≠ x_1) s
theorem
Multiset.Nodup.pairwise
{α : Type u_1}
{r : α → α → Prop}
{s : Multiset α}
 :
(∀ a ∈ s, ∀ b ∈ s, a ≠ b → r a b) → s.Nodup → Multiset.Pairwise r s
theorem
Multiset.Pairwise.forall
{α : Type u_1}
{r : α → α → Prop}
{s : Multiset α}
(H : Symmetric r)
(hs : Multiset.Pairwise r s)
⦃a : α⦄
 :
theorem
Multiset.Nodup.of_map
{α : Type u_1}
{β : Type u_2}
{s : Multiset α}
(f : α → β)
 :
(Multiset.map f s).Nodup → s.Nodup
theorem
Multiset.Nodup.map_on
{α : Type u_1}
{β : Type u_2}
{s : Multiset α}
{f : α → β}
 :
(∀ x ∈ s, ∀ y ∈ s, f x = f y → x = y) → s.Nodup → (Multiset.map f s).Nodup
theorem
Multiset.Nodup.map
{α : Type u_1}
{β : Type u_2}
{f : α → β}
{s : Multiset α}
(hf : Function.Injective f)
 :
s.Nodup → (Multiset.map f s).Nodup
theorem
Multiset.nodup_map_iff_of_inj_on
{α : Type u_1}
{β : Type u_2}
{s : Multiset α}
{f : α → β}
(d : ∀ x ∈ s, ∀ y ∈ s, f x = f y → x = y)
 :
(Multiset.map f s).Nodup ↔ s.Nodup
theorem
Multiset.nodup_map_iff_of_injective
{α : Type u_1}
{β : Type u_2}
{s : Multiset α}
{f : α → β}
(d : Function.Injective f)
 :
(Multiset.map f s).Nodup ↔ s.Nodup
theorem
Multiset.inj_on_of_nodup_map
{α : Type u_1}
{β : Type u_2}
{f : α → β}
{s : Multiset α}
 :
(Multiset.map f s).Nodup → ∀ x ∈ s, ∀ y ∈ s, f x = f y → x = y
theorem
Multiset.nodup_map_iff_inj_on
{α : Type u_1}
{β : Type u_2}
{f : α → β}
{s : Multiset α}
(d : s.Nodup)
 :
(Multiset.map f s).Nodup ↔ ∀ x ∈ s, ∀ y ∈ s, f x = f y → x = y
theorem
Multiset.Nodup.filter
{α : Type u_1}
(p : α → Prop)
[DecidablePred p]
{s : Multiset α}
 :
s.Nodup → (Multiset.filter p s).Nodup
@[simp]
Alias of the reverse direction of Multiset.nodup_attach.
theorem
Multiset.Nodup.pmap
{α : Type u_1}
{β : Type u_2}
{p : α → Prop}
{f : (a : α) → p a → β}
{s : Multiset α}
{H : ∀ a ∈ s, p a}
(hf : ∀ (a : α) (ha : p a) (b : α) (hb : p b), f a ha = f b hb → a = b)
 :
s.Nodup → (Multiset.pmap f s H).Nodup
Equations
- s.nodupDecidable = Quotient.recOnSubsingleton s fun (l : List α) => l.nodupDecidable
theorem
Multiset.Nodup.erase_eq_filter
{α : Type u_1}
[DecidableEq α]
(a : α)
{s : Multiset α}
 :
s.Nodup → s.erase a = Multiset.filter (fun (x : α) => x ≠ a) s
theorem
Multiset.Nodup.erase
{α : Type u_1}
[DecidableEq α]
(a : α)
{l : Multiset α}
 :
l.Nodup → (l.erase a).Nodup
theorem
Multiset.Nodup.mem_erase_iff
{α : Type u_1}
[DecidableEq α]
{a : α}
{b : α}
{l : Multiset α}
(d : l.Nodup)
 :
theorem
Multiset.Nodup.not_mem_erase
{α : Type u_1}
[DecidableEq α]
{a : α}
{s : Multiset α}
(h : s.Nodup)
 :
a ∉ s.erase a
theorem
Multiset.Nodup.filterMap
{α : Type u_1}
{β : Type u_2}
{s : Multiset α}
(f : α → Option β)
(H : ∀ (a a' : α), ∀ b ∈ f a, b ∈ f a' → a = a')
 :
s.Nodup → (Multiset.filterMap f s).Nodup
theorem
Multiset.Nodup.inter_left
{α : Type u_1}
{s : Multiset α}
[DecidableEq α]
(t : Multiset α)
 :
s.Nodup → (s ∩ t).Nodup
theorem
Multiset.Nodup.inter_right
{α : Type u_1}
{t : Multiset α}
[DecidableEq α]
(s : Multiset α)
 :
t.Nodup → (s ∩ t).Nodup
@[simp]
theorem
Multiset.map_eq_map_of_bij_of_nodup
{α : Type u_1}
{β : Type u_2}
{γ : Type u_3}
(f : α → γ)
(g : β → γ)
{s : Multiset α}
{t : Multiset β}
(hs : s.Nodup)
(ht : t.Nodup)
(i : (a : α) → a ∈ s → β)
(hi : ∀ (a : α) (ha : a ∈ s), i a ha ∈ t)
(i_inj : ∀ (a₁ : α) (ha₁ : a₁ ∈ s) (a₂ : α) (ha₂ : a₂ ∈ s), i a₁ ha₁ = i a₂ ha₂ → a₁ = a₂)
(i_surj : ∀ b ∈ t, ∃ (a : α) (ha : a ∈ s), i a ha = b)
(h : ∀ (a : α) (ha : a ∈ s), f a = g (i a ha))
 :
Multiset.map f s = Multiset.map g t