Documentation

Mathlib.Data.Finset.Card

Cardinality of a finite set #

This defines the cardinality of a Finset and provides induction principles for finsets.

Main declarations #

Induction principles #

def Finset.card {α : Type u_1} (s : Finset α) :

s.card is the number of elements of s, aka its cardinality.

Equations
  • s.card = Multiset.card s.val
Instances For
    theorem Finset.card_def {α : Type u_1} (s : Finset α) :
    s.card = Multiset.card s.val
    @[simp]
    theorem Finset.card_val {α : Type u_1} (s : Finset α) :
    Multiset.card s.val = s.card
    @[simp]
    theorem Finset.card_mk {α : Type u_1} {m : Multiset α} {nodup : m.Nodup} :
    { val := m, nodup := nodup }.card = Multiset.card m
    @[simp]
    theorem Finset.card_empty {α : Type u_1} :
    .card = 0
    theorem Finset.card_le_card {α : Type u_1} {s : Finset α} {t : Finset α} :
    s ts.card t.card
    theorem Finset.card_mono {α : Type u_1} :
    Monotone Finset.card
    @[simp]
    theorem Finset.card_eq_zero {α : Type u_1} {s : Finset α} :
    s.card = 0 s =
    theorem Finset.card_ne_zero {α : Type u_1} {s : Finset α} :
    s.card 0 s.Nonempty
    @[simp]
    theorem Finset.card_pos {α : Type u_1} {s : Finset α} :
    0 < s.card s.Nonempty
    @[simp]
    theorem Finset.one_le_card {α : Type u_1} {s : Finset α} :
    1 s.card s.Nonempty
    theorem Finset.Nonempty.card_pos {α : Type u_1} {s : Finset α} :
    s.Nonempty0 < s.card

    Alias of the reverse direction of Finset.card_pos.

    theorem Finset.Nonempty.card_ne_zero {α : Type u_1} {s : Finset α} :
    s.Nonemptys.card 0

    Alias of the reverse direction of Finset.card_ne_zero.

    theorem Finset.card_ne_zero_of_mem {α : Type u_1} {s : Finset α} {a : α} (h : a s) :
    s.card 0
    @[simp]
    theorem Finset.card_singleton {α : Type u_1} (a : α) :
    {a}.card = 1
    theorem Finset.card_singleton_inter {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] :
    ({a} s).card 1
    @[simp]
    theorem Finset.card_cons {α : Type u_1} {s : Finset α} {a : α} (h : as) :
    (Finset.cons a s h).card = s.card + 1
    @[simp]
    theorem Finset.card_insert_of_not_mem {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] (h : as) :
    (insert a s).card = s.card + 1
    theorem Finset.card_insert_of_mem {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] (h : a s) :
    (insert a s).card = s.card
    theorem Finset.card_insert_le {α : Type u_1} [DecidableEq α] (a : α) (s : Finset α) :
    (insert a s).card s.card + 1
    theorem Finset.card_le_two {α : Type u_1} [DecidableEq α] {a : α} {b : α} :
    {a, b}.card 2
    theorem Finset.card_le_three {α : Type u_1} [DecidableEq α] {a : α} {b : α} {c : α} :
    {a, b, c}.card 3
    theorem Finset.card_le_four {α : Type u_1} [DecidableEq α] {a : α} {b : α} {c : α} {d : α} :
    {a, b, c, d}.card 4
    theorem Finset.card_le_five {α : Type u_1} [DecidableEq α] {a : α} {b : α} {c : α} {d : α} {e : α} :
    {a, b, c, d, e}.card 5
    theorem Finset.card_le_six {α : Type u_1} [DecidableEq α] {a : α} {b : α} {c : α} {d : α} {e : α} {f : α} :
    {a, b, c, d, e, f}.card 6
    theorem Finset.card_insert_eq_ite {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] :
    (insert a s).card = if a s then s.card else s.card + 1

    If a ∈ s is known, see also Finset.card_insert_of_mem and Finset.card_insert_of_not_mem.

    @[simp]
    theorem Finset.card_pair_eq_one_or_two {α : Type u_1} {a : α} {b : α} [DecidableEq α] :
    {a, b}.card = 1 {a, b}.card = 2
    @[simp]
    theorem Finset.card_pair {α : Type u_1} {a : α} {b : α} [DecidableEq α] (h : a b) :
    {a, b}.card = 2
    @[deprecated Finset.card_pair]
    theorem Finset.card_doubleton {α : Type u_1} {a : α} {b : α} [DecidableEq α] (h : a b) :
    {a, b}.card = 2

    Alias of Finset.card_pair.

    @[simp]
    theorem Finset.card_erase_of_mem {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] :
    a s(s.erase a).card = s.card - 1

    $\#(s \setminus \{a\}) = \#s - 1$ if $a \in s$.

    @[simp]
    theorem Finset.cast_card_erase_of_mem {α : Type u_1} {a : α} [DecidableEq α] {R : Type u_4} [AddGroupWithOne R] {s : Finset α} (hs : a s) :
    (s.erase a).card = s.card - 1

    $\#(s \setminus \{a\}) = \#s - 1$ if $a \in s$. This result is casted to any additive group with 1, so that we don't have to work with -subtraction.

    @[simp]
    theorem Finset.card_erase_add_one {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] :
    a s(s.erase a).card + 1 = s.card
    theorem Finset.card_erase_lt_of_mem {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] :
    a s(s.erase a).card < s.card
    theorem Finset.card_erase_le {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] :
    (s.erase a).card s.card
    theorem Finset.pred_card_le_card_erase {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] :
    s.card - 1 (s.erase a).card
    theorem Finset.card_erase_eq_ite {α : Type u_1} {s : Finset α} {a : α} [DecidableEq α] :
    (s.erase a).card = if a s then s.card - 1 else s.card

    If a ∈ s is known, see also Finset.card_erase_of_mem and Finset.erase_eq_of_not_mem.

    @[simp]
    theorem Finset.card_range (n : ) :
    (Finset.range n).card = n
    @[simp]
    theorem Finset.card_attach {α : Type u_1} {s : Finset α} :
    s.attach.card = s.card
    theorem Multiset.card_toFinset {α : Type u_1} [DecidableEq α] (m : Multiset α) :
    m.toFinset.card = Multiset.card m.dedup
    theorem Multiset.toFinset_card_le {α : Type u_1} [DecidableEq α] (m : Multiset α) :
    m.toFinset.card Multiset.card m
    theorem Multiset.toFinset_card_of_nodup {α : Type u_1} [DecidableEq α] {m : Multiset α} (h : m.Nodup) :
    m.toFinset.card = Multiset.card m
    theorem Multiset.dedup_card_eq_card_iff_nodup {α : Type u_1} [DecidableEq α] {m : Multiset α} :
    Multiset.card m.dedup = Multiset.card m m.Nodup
    theorem Multiset.toFinset_card_eq_card_iff_nodup {α : Type u_1} [DecidableEq α] {m : Multiset α} :
    m.toFinset.card = Multiset.card m m.Nodup
    theorem List.card_toFinset {α : Type u_1} [DecidableEq α] (l : List α) :
    l.toFinset.card = l.dedup.length
    theorem List.toFinset_card_le {α : Type u_1} [DecidableEq α] (l : List α) :
    l.toFinset.card l.length
    theorem List.toFinset_card_of_nodup {α : Type u_1} [DecidableEq α] {l : List α} (h : l.Nodup) :
    l.toFinset.card = l.length
    @[simp]
    theorem Finset.length_toList {α : Type u_1} (s : Finset α) :
    s.toList.length = s.card
    theorem Finset.card_image_le {α : Type u_1} {β : Type u_2} {s : Finset α} {f : αβ} [DecidableEq β] :
    (Finset.image f s).card s.card
    theorem Finset.card_image_of_injOn {α : Type u_1} {β : Type u_2} {s : Finset α} {f : αβ} [DecidableEq β] (H : Set.InjOn f s) :
    (Finset.image f s).card = s.card
    theorem Finset.injOn_of_card_image_eq {α : Type u_1} {β : Type u_2} {s : Finset α} {f : αβ} [DecidableEq β] (H : (Finset.image f s).card = s.card) :
    Set.InjOn f s
    theorem Finset.card_image_iff {α : Type u_1} {β : Type u_2} {s : Finset α} {f : αβ} [DecidableEq β] :
    (Finset.image f s).card = s.card Set.InjOn f s
    theorem Finset.card_image_of_injective {α : Type u_1} {β : Type u_2} {f : αβ} [DecidableEq β] (s : Finset α) (H : Function.Injective f) :
    (Finset.image f s).card = s.card
    theorem Finset.fiber_card_ne_zero_iff_mem_image {α : Type u_1} {β : Type u_2} (s : Finset α) (f : αβ) [DecidableEq β] (y : β) :
    (Finset.filter (fun (x : α) => f x = y) s).card 0 y Finset.image f s
    theorem Finset.card_filter_le_iff {α : Type u_1} (s : Finset α) (P : αProp) [DecidablePred P] (n : ) :
    (Finset.filter P s).card n s's, n < s'.cardas', ¬P a
    @[simp]
    theorem Finset.card_map {α : Type u_1} {β : Type u_2} {s : Finset α} (f : α β) :
    (Finset.map f s).card = s.card
    @[simp]
    theorem Finset.card_subtype {α : Type u_1} (p : αProp) [DecidablePred p] (s : Finset α) :
    (Finset.subtype p s).card = (Finset.filter p s).card
    theorem Finset.card_filter_le {α : Type u_1} (s : Finset α) (p : αProp) [DecidablePred p] :
    (Finset.filter p s).card s.card
    theorem Finset.eq_of_subset_of_card_le {α : Type u_1} {s : Finset α} {t : Finset α} (h : s t) (h₂ : t.card s.card) :
    s = t
    theorem Finset.eq_of_superset_of_card_ge {α : Type u_1} {s : Finset α} {t : Finset α} (hst : s t) (hts : t.card s.card) :
    t = s
    theorem Finset.subset_iff_eq_of_card_le {α : Type u_1} {s : Finset α} {t : Finset α} (h : t.card s.card) :
    s t s = t
    theorem Finset.map_eq_of_subset {α : Type u_1} {s : Finset α} {f : α α} (hs : Finset.map f s s) :
    theorem Finset.filter_card_eq {α : Type u_1} {s : Finset α} {p : αProp} [DecidablePred p] (h : (Finset.filter p s).card = s.card) (x : α) (hx : x s) :
    p x
    theorem Finset.card_lt_card {α : Type u_1} {s : Finset α} {t : Finset α} (h : s t) :
    s.card < t.card
    theorem Finset.card_strictMono {α : Type u_1} :
    StrictMono Finset.card
    theorem Finset.card_eq_of_bijective {α : Type u_1} {s : Finset α} {n : } (f : (i : ) → i < nα) (hf : as, ∃ (i : ) (h : i < n), f i h = a) (hf' : ∀ (i : ) (h : i < n), f i h s) (f_inj : ∀ (i j : ) (hi : i < n) (hj : j < n), f i hi = f j hji = j) :
    s.card = n
    theorem Finset.card_bij {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (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 : bt, ∃ (a : α) (ha : a s), i a ha = b) :
    s.card = t.card

    Reorder a finset.

    The difference with Finset.card_bij' is that the bijection is specified as a surjective injection, rather than by an inverse function.

    The difference with Finset.card_nbij is that the bijection is allowed to use membership of the domain, rather than being a non-dependent function.

    @[deprecated Finset.card_bij]
    theorem Finset.card_congr {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (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 : bt, ∃ (a : α) (ha : a s), i a ha = b) :
    s.card = t.card

    Alias of Finset.card_bij.


    Reorder a finset.

    The difference with Finset.card_bij' is that the bijection is specified as a surjective injection, rather than by an inverse function.

    The difference with Finset.card_nbij is that the bijection is allowed to use membership of the domain, rather than being a non-dependent function.

    theorem Finset.card_bij' {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (i : (a : α) → a sβ) (j : (a : β) → a tα) (hi : ∀ (a : α) (ha : a s), i a ha t) (hj : ∀ (a : β) (ha : a t), j a ha s) (left_inv : ∀ (a : α) (ha : a s), j (i a ha) = a) (right_inv : ∀ (a : β) (ha : a t), i (j a ha) = a) :
    s.card = t.card

    Reorder a finset.

    The difference with Finset.card_bij is that the bijection is specified with an inverse, rather than as a surjective injection.

    The difference with Finset.card_nbij' is that the bijection and its inverse are allowed to use membership of the domains, rather than being non-dependent functions.

    theorem Finset.card_nbij {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (i : αβ) (hi : as, i a t) (i_inj : Set.InjOn i s) (i_surj : Set.SurjOn i s t) :
    s.card = t.card

    Reorder a finset.

    The difference with Finset.card_nbij' is that the bijection is specified as a surjective injection, rather than by an inverse function.

    The difference with Finset.card_bij is that the bijection is a non-dependent function, rather than being allowed to use membership of the domain.

    theorem Finset.card_nbij' {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (i : αβ) (j : βα) (hi : as, i a t) (hj : at, j a s) (left_inv : as, j (i a) = a) (right_inv : at, i (j a) = a) :
    s.card = t.card

    Reorder a finset.

    The difference with Finset.card_nbij is that the bijection is specified with an inverse, rather than as a surjective injection.

    The difference with Finset.card_bij' is that the bijection and its inverse are non-dependent functions, rather than being allowed to use membership of the domains.

    The difference with Finset.card_equiv is that bijectivity is only required to hold on the domains, rather than on the entire types.

    theorem Finset.card_equiv {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (e : α β) (hst : ∀ (i : α), i s e i t) :
    s.card = t.card

    Specialization of Finset.card_nbij' that automatically fills in most arguments.

    See Fintype.card_equiv for the version where s and t are univ.

    theorem Finset.card_bijective {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (e : αβ) (he : Function.Bijective e) (hst : ∀ (i : α), i s e i t) :
    s.card = t.card

    Specialization of Finset.card_nbij that automatically fills in most arguments.

    See Fintype.card_bijective for the version where s and t are univ.

    theorem Finset.card_le_card_of_injOn {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (f : αβ) (hf : as, f a t) (f_inj : Set.InjOn f s) :
    s.card t.card
    @[deprecated Finset.card_le_card_of_injOn]
    theorem Finset.card_le_card_of_inj_on {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (f : αβ) (hf : as, f a t) (f_inj : Set.InjOn f s) :
    s.card t.card

    Alias of Finset.card_le_card_of_injOn.

    theorem Finset.card_le_card_of_surjOn {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (f : αβ) (hf : Set.SurjOn f s t) :
    t.card s.card
    theorem Finset.exists_ne_map_eq_of_card_lt_of_maps_to {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (hc : t.card < s.card) {f : αβ} (hf : as, f a t) :
    xs, ys, x y f x = f y

    If there are more pigeons than pigeonholes, then there are two pigeons in the same pigeonhole.

    theorem Finset.le_card_of_inj_on_range {α : Type u_1} {s : Finset α} {n : } (f : α) (hf : i < n, f i s) (f_inj : i < n, j < n, f i = f ji = j) :
    n s.card
    theorem Finset.surj_on_of_inj_on_of_card_le {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (f : (a : α) → a sβ) (hf : ∀ (a : α) (ha : a s), f a ha t) (hinj : ∀ (a₁ a₂ : α) (ha₁ : a₁ s) (ha₂ : a₂ s), f a₁ ha₁ = f a₂ ha₂a₁ = a₂) (hst : t.card s.card) (b : β) :
    b t∃ (a : α) (ha : a s), b = f a ha
    theorem Finset.inj_on_of_surj_on_of_card_le {α : Type u_1} {β : Type u_2} {s : Finset α} {t : Finset β} (f : (a : α) → a sβ) (hf : ∀ (a : α) (ha : a s), f a ha t) (hsurj : bt, ∃ (a : α) (ha : a s), f a ha = b) (hst : s.card t.card) ⦃a₁ : α (ha₁ : a₁ s) ⦃a₂ : α (ha₂ : a₂ s) (ha₁a₂ : f a₁ ha₁ = f a₂ ha₂) :
    a₁ = a₂
    @[simp]
    theorem Finset.card_disjUnion {α : Type u_1} (s : Finset α) (t : Finset α) (h : Disjoint s t) :
    (s.disjUnion t h).card = s.card + t.card

    Lattice structure #

    theorem Finset.card_union_add_card_inter {α : Type u_1} [DecidableEq α] (s : Finset α) (t : Finset α) :
    (s t).card + (s t).card = s.card + t.card
    theorem Finset.card_inter_add_card_union {α : Type u_1} [DecidableEq α] (s : Finset α) (t : Finset α) :
    (s t).card + (s t).card = s.card + t.card
    theorem Finset.card_union {α : Type u_1} [DecidableEq α] (s : Finset α) (t : Finset α) :
    (s t).card = s.card + t.card - (s t).card
    theorem Finset.card_inter {α : Type u_1} [DecidableEq α] (s : Finset α) (t : Finset α) :
    (s t).card = s.card + t.card - (s t).card
    theorem Finset.card_union_le {α : Type u_1} [DecidableEq α] (s : Finset α) (t : Finset α) :
    (s t).card s.card + t.card
    theorem Finset.card_union_eq_card_add_card {α : Type u_1} {s : Finset α} {t : Finset α} [DecidableEq α] :
    (s t).card = s.card + t.card Disjoint s t
    @[simp]
    theorem Finset.card_union_of_disjoint {α : Type u_1} {s : Finset α} {t : Finset α} [DecidableEq α] :
    Disjoint s t(s t).card = s.card + t.card

    Alias of the reverse direction of Finset.card_union_eq_card_add_card.

    @[deprecated Finset.card_union_of_disjoint]
    theorem Finset.card_union_eq {α : Type u_1} {s : Finset α} {t : Finset α} [DecidableEq α] :
    Disjoint s t(s t).card = s.card + t.card

    Alias of the reverse direction of Finset.card_union_eq_card_add_card.


    Alias of the reverse direction of Finset.card_union_eq_card_add_card.

    @[deprecated Finset.card_union_of_disjoint]
    theorem Finset.card_disjoint_union {α : Type u_1} {s : Finset α} {t : Finset α} [DecidableEq α] :
    Disjoint s t(s t).card = s.card + t.card

    Alias of the reverse direction of Finset.card_union_eq_card_add_card.


    Alias of the reverse direction of Finset.card_union_eq_card_add_card.

    theorem Finset.cast_card_inter {α : Type u_1} {R : Type u_3} {s : Finset α} {t : Finset α} [DecidableEq α] [AddGroupWithOne R] :
    (s t).card = s.card + t.card - (s t).card
    theorem Finset.cast_card_union {α : Type u_1} {R : Type u_3} {s : Finset α} {t : Finset α} [DecidableEq α] [AddGroupWithOne R] :
    (s t).card = s.card + t.card - (s t).card
    theorem Finset.card_sdiff {α : Type u_1} {s : Finset α} {t : Finset α} [DecidableEq α] (h : s t) :
    (t \ s).card = t.card - s.card
    theorem Finset.cast_card_sdiff {α : Type u_1} {R : Type u_3} {s : Finset α} {t : Finset α} [DecidableEq α] [AddGroupWithOne R] (h : s t) :
    (t \ s).card = t.card - s.card
    theorem Finset.card_sdiff_add_card_eq_card {α : Type u_1} [DecidableEq α] {s : Finset α} {t : Finset α} (h : s t) :
    (t \ s).card + s.card = t.card
    theorem Finset.le_card_sdiff {α : Type u_1} [DecidableEq α] (s : Finset α) (t : Finset α) :
    t.card - s.card (t \ s).card
    theorem Finset.card_le_card_sdiff_add_card {α : Type u_1} {s : Finset α} {t : Finset α} [DecidableEq α] :
    s.card (s \ t).card + t.card
    theorem Finset.card_sdiff_add_card {α : Type u_1} {s : Finset α} {t : Finset α} [DecidableEq α] :
    (s \ t).card + t.card = (s t).card
    theorem Finset.card_sdiff_comm {α : Type u_1} {s : Finset α} {t : Finset α} [DecidableEq α] (h : s.card = t.card) :
    (s \ t).card = (t \ s).card
    @[simp]
    theorem Finset.card_sdiff_add_card_inter {α : Type u_1} [DecidableEq α] (s : Finset α) (t : Finset α) :
    (s \ t).card + (s t).card = s.card
    @[simp]
    theorem Finset.card_inter_add_card_sdiff {α : Type u_1} [DecidableEq α] (s : Finset α) (t : Finset α) :
    (s t).card + (s \ t).card = s.card
    theorem Finset.inter_nonempty_of_card_lt_card_add_card {α : Type u_1} {s : Finset α} {t : Finset α} {u : Finset α} [DecidableEq α] (hts : t s) (hus : u s) (hstu : s.card < t.card + u.card) :
    (t u).Nonempty

    Pigeonhole principle for two finsets inside an ambient finset.

    theorem Finset.filter_card_add_filter_neg_card_eq_card {α : Type u_1} {s : Finset α} (p : αProp) [DecidablePred p] [(x : α) → Decidable ¬p x] :
    (Finset.filter p s).card + (Finset.filter (fun (a : α) => ¬p a) s).card = s.card
    theorem Finset.exists_subsuperset_card_eq {α : Type u_1} {s : Finset α} {t : Finset α} {n : } (hst : s t) (hsn : s.card n) (hnt : n t.card) :
    ∃ (u : Finset α), s u u t u.card = n

    Given a subset s of a set t, of sizes at most and at least n respectively, there exists a set u of size n which is both a superset of s and a subset of t.

    theorem Finset.exists_subset_card_eq {α : Type u_1} {s : Finset α} {n : } (hns : n s.card) :
    ts, t.card = n

    We can shrink a set to any smaller size.

    @[deprecated Finset.exists_subsuperset_card_eq]
    theorem Finset.exists_intermediate_set {α : Type u_1} {A : Finset α} {B : Finset α} (i : ) (h₁ : i + B.card A.card) (h₂ : B A) :
    ∃ (C : Finset α), B C C A C.card = i + B.card

    Given a set A and a set B inside it, we can shrink A to any appropriate size, and keep B inside it.

    @[deprecated Finset.exists_subset_card_eq]
    theorem Finset.exists_smaller_set {α : Type u_1} (A : Finset α) (i : ) (h₁ : i A.card) :
    BA, B.card = i

    We can shrink A to any smaller size.

    theorem Finset.le_card_iff_exists_subset_card {α : Type u_1} {s : Finset α} {n : } :
    n s.card ts, t.card = n
    theorem Finset.exists_subset_or_subset_of_two_mul_lt_card {α : Type u_1} [DecidableEq α] {X : Finset α} {Y : Finset α} {n : } (hXY : 2 * n < (X Y).card) :
    ∃ (C : Finset α), n < C.card (C X C Y)

    Explicit description of a finset from its card #

    theorem Finset.card_eq_one {α : Type u_1} {s : Finset α} :
    s.card = 1 ∃ (a : α), s = {a}
    theorem Multiset.toFinset_card_eq_one_iff {α : Type u_1} [DecidableEq α] (s : Multiset α) :
    s.toFinset.card = 1 Multiset.card s 0 ∃ (a : α), s = Multiset.card s {a}
    theorem Finset.exists_eq_insert_iff {α : Type u_1} [DecidableEq α] {s : Finset α} {t : Finset α} :
    (as, insert a s = t) s t s.card + 1 = t.card
    theorem Finset.card_le_one {α : Type u_1} {s : Finset α} :
    s.card 1 as, bs, a = b
    theorem Finset.card_le_one_iff {α : Type u_1} {s : Finset α} :
    s.card 1 ∀ {a b : α}, a sb sa = b
    theorem Finset.card_le_one_iff_subsingleton_coe {α : Type u_1} {s : Finset α} :
    s.card 1 Subsingleton { x : α // x s }
    theorem Finset.card_le_one_iff_subset_singleton {α : Type u_1} {s : Finset α} [Nonempty α] :
    s.card 1 ∃ (x : α), s {x}
    theorem Finset.exists_mem_ne {α : Type u_1} {s : Finset α} (hs : 1 < s.card) (a : α) :
    bs, b a
    theorem Finset.card_le_one_of_subsingleton {α : Type u_1} [Subsingleton α] (s : Finset α) :
    s.card 1

    A Finset of a subsingleton type has cardinality at most one.

    theorem Finset.one_lt_card {α : Type u_1} {s : Finset α} :
    1 < s.card as, bs, a b
    theorem Finset.one_lt_card_iff {α : Type u_1} {s : Finset α} :
    1 < s.card ∃ (a : α) (b : α), a s b s a b
    theorem Finset.one_lt_card_iff_nontrivial {α : Type u_1} {s : Finset α} :
    1 < s.card s.Nontrivial
    @[deprecated Finset.one_lt_card_iff_nontrivial]
    theorem Finset.one_lt_card_iff_nontrivial_coe {α : Type u_1} {s : Finset α} :
    1 < s.card s.Nontrivial

    Alias of Finset.one_lt_card_iff_nontrivial.

    theorem Finset.exists_ne_of_one_lt_card {α : Type u_1} {s : Finset α} (hs : 1 < s.card) (a : α) :
    bs, b a
    theorem Finset.exists_of_one_lt_card_pi {ι : Type u_4} {α : ιType u_5} [(i : ι) → DecidableEq (α i)] {s : Finset ((i : ι) → α i)} (h : 1 < s.card) :
    ∃ (i : ι), 1 < (Finset.image (fun (x : (i : ι) → α i) => x i) s).card ∀ (ai : α i), Finset.filter (fun (x : (i : ι) → α i) => x i = ai) s s

    If a Finset in a Pi type is nontrivial (has at least two elements), then its projection to some factor is nontrivial, and the fibers of the projection are proper subsets.

    theorem Finset.card_eq_succ {α : Type u_1} {s : Finset α} {n : } [DecidableEq α] :
    s.card = n + 1 ∃ (a : α) (t : Finset α), at insert a t = s t.card = n
    theorem Finset.card_eq_two {α : Type u_1} {s : Finset α} [DecidableEq α] :
    s.card = 2 ∃ (x : α) (y : α), x y s = {x, y}
    theorem Finset.card_eq_three {α : Type u_1} {s : Finset α} [DecidableEq α] :
    s.card = 3 ∃ (x : α) (y : α) (z : α), x y x z y z s = {x, y, z}
    theorem Finset.two_lt_card_iff {α : Type u_1} {s : Finset α} :
    2 < s.card ∃ (a : α) (b : α) (c : α), a s b s c s a b a c b c
    theorem Finset.two_lt_card {α : Type u_1} {s : Finset α} :
    2 < s.card as, bs, cs, a b a c b c

    Inductions #

    @[irreducible]
    def Finset.strongInduction {α : Type u_1} {p : Finset αSort u_4} (H : (s : Finset α) → ((t : Finset α) → t sp t)p s) (s : Finset α) :
    p s

    Suppose that, given objects defined on all strict subsets of any finset s, one knows how to define an object on s. Then one can inductively define an object on all finsets, starting from the empty set and iterating. This can be used either to define data, or to prove properties.

    Equations
    Instances For
      theorem Finset.strongInduction_eq {α : Type u_1} {p : Finset αSort u_4} (H : (s : Finset α) → ((t : Finset α) → t sp t)p s) (s : Finset α) :
      Finset.strongInduction H s = H s fun (t : Finset α) (x : t s) => Finset.strongInduction H t
      def Finset.strongInductionOn {α : Type u_1} {p : Finset αSort u_4} (s : Finset α) :
      ((s : Finset α) → ((t : Finset α) → t sp t)p s)p s

      Analogue of strongInduction with order of arguments swapped.

      Equations
      Instances For
        theorem Finset.strongInductionOn_eq {α : Type u_1} {p : Finset αSort u_4} (s : Finset α) (H : (s : Finset α) → ((t : Finset α) → t sp t)p s) :
        s.strongInductionOn H = H s fun (t : Finset α) (x : t s) => t.strongInductionOn H
        theorem Finset.case_strong_induction_on {α : Type u_1} [DecidableEq α] {p : Finset αProp} (s : Finset α) (h₀ : p ) (h₁ : ∀ (a : α) (s : Finset α), as(ts, p t)p (insert a s)) :
        p s
        @[irreducible]
        theorem Finset.Nonempty.strong_induction {α : Type u_1} {p : (s : Finset α) → s.NonemptyProp} (h₀ : ∀ (a : α), p {a} ) (h₁ : ∀ ⦃s : Finset α⦄ (hs : s.Nontrivial), (∀ (t : Finset α) (ht : t.Nonempty), t sp t ht)p s ) ⦃s : Finset α (hs : s.Nonempty) :
        p s hs

        Suppose that, given objects defined on all nonempty strict subsets of any nontrivial finset s, one knows how to define an object on s. Then one can inductively define an object on all finsets, starting from singletons and iterating.

        TODO: Currently this can only be used to prove properties. Replace Finset.Nonempty.exists_eq_singleton_or_nontrivial with computational content in order to let p be Sort-valued.

        @[irreducible]
        def Finset.strongDownwardInduction {α : Type u_1} {p : Finset αSort u_4} {n : } (H : (t₁ : Finset α) → ({t₂ : Finset α} → t₂.card nt₁ t₂p t₂)t₁.card np t₁) (s : Finset α) :
        s.card np s

        Suppose that, given that p t can be defined on all supersets of s of cardinality less than n, one knows how to define p s. Then one can inductively define p s for all finsets s of cardinality less than n, starting from finsets of card n and iterating. This can be used either to define data, or to prove properties.

        Equations
        Instances For
          theorem Finset.strongDownwardInduction_eq {α : Type u_1} {n : } {p : Finset αSort u_4} (H : (t₁ : Finset α) → ({t₂ : Finset α} → t₂.card nt₁ t₂p t₂)t₁.card np t₁) (s : Finset α) :
          Finset.strongDownwardInduction H s = H s fun {t : Finset α} (ht : t.card n) (x : s t) => Finset.strongDownwardInduction H t ht
          def Finset.strongDownwardInductionOn {α : Type u_1} {n : } {p : Finset αSort u_4} (s : Finset α) (H : (t₁ : Finset α) → ({t₂ : Finset α} → t₂.card nt₁ t₂p t₂)t₁.card np t₁) :
          s.card np s

          Analogue of strongDownwardInduction with order of arguments swapped.

          Equations
          Instances For
            theorem Finset.strongDownwardInductionOn_eq {α : Type u_1} {n : } {p : Finset αSort u_4} (s : Finset α) (H : (t₁ : Finset α) → ({t₂ : Finset α} → t₂.card nt₁ t₂p t₂)t₁.card np t₁) :
            (fun (a : s.card n) => s.strongDownwardInductionOn H a) = H s fun {t : Finset α} (ht : t.card n) (x : s t) => t.strongDownwardInductionOn H ht
            theorem Finset.lt_wf {α : Type u_4} :
            @[deprecated Finset.card_le_card]
            theorem Finset.card_le_of_subset {α : Type u_1} {s : Finset α} {t : Finset α} :
            s ts.card t.card

            Alias of Finset.card_le_card.