Documentation

Mathlib.Order.CompleteLattice.Finset

Lattice operations on finsets #

This file is concerned with how big lattice or set operations behave when indexed by a finset.

See also Lattice.lean, which is concerned with folding binary lattice operations over a finset.

theorem iSup_eq_iSup_finset {α : Type u_2} {ι : Type u_5} [CompleteLattice α] (s : ια) :
⨆ (i : ι), s i = ⨆ (t : Finset ι), it, s i

Supremum of s i, i : ι, is equal to the supremum over t : Finset ι of suprema ⨆ i ∈ t, s i. This version assumes ι is a Type*. See iSup_eq_iSup_finset' for a version that works for ι : Sort*.

theorem iSup_eq_iSup_finset' {α : Type u_2} {ι' : Sort u_7} [CompleteLattice α] (s : ι'α) :
⨆ (i : ι'), s i = ⨆ (t : Finset (PLift ι')), it, s i.down

Supremum of s i, i : ι, is equal to the supremum over t : Finset ι of suprema ⨆ i ∈ t, s i. This version works for ι : Sort*. See iSup_eq_iSup_finset for a version that assumes ι : Type* but has no PLifts.

theorem iInf_eq_iInf_finset {α : Type u_2} {ι : Type u_5} [CompleteLattice α] (s : ια) :
⨅ (i : ι), s i = ⨅ (t : Finset ι), it, s i

Infimum of s i, i : ι, is equal to the infimum over t : Finset ι of infima ⨅ i ∈ t, s i. This version assumes ι is a Type*. See iInf_eq_iInf_finset' for a version that works for ι : Sort*.

theorem iInf_eq_iInf_finset' {α : Type u_2} {ι' : Sort u_7} [CompleteLattice α] (s : ι'α) :
⨅ (i : ι'), s i = ⨅ (t : Finset (PLift ι')), it, s i.down

Infimum of s i, i : ι, is equal to the infimum over t : Finset ι of infima ⨅ i ∈ t, s i. This version works for ι : Sort*. See iInf_eq_iInf_finset for a version that assumes ι : Type* but has no PLifts.

theorem Set.iUnion_eq_iUnion_finset {α : Type u_2} {ι : Type u_5} (s : ιSet α) :
⋃ (i : ι), s i = ⋃ (t : Finset ι), it, s i

Union of an indexed family of sets s : ι → Set α is equal to the union of the unions of finite subfamilies. This version assumes ι : Type*. See also iUnion_eq_iUnion_finset' for a version that works for ι : Sort*.

theorem Set.iUnion_eq_iUnion_finset' {α : Type u_2} {ι' : Sort u_7} (s : ι'Set α) :
⋃ (i : ι'), s i = ⋃ (t : Finset (PLift ι')), it, s i.down

Union of an indexed family of sets s : ι → Set α is equal to the union of the unions of finite subfamilies. This version works for ι : Sort*. See also iUnion_eq_iUnion_finset for a version that assumes ι : Type* but avoids PLifts in the right hand side.

theorem Set.iInter_eq_iInter_finset {α : Type u_2} {ι : Type u_5} (s : ιSet α) :
⋂ (i : ι), s i = ⋂ (t : Finset ι), it, s i

Intersection of an indexed family of sets s : ι → Set α is equal to the intersection of the intersections of finite subfamilies. This version assumes ι : Type*. See also iInter_eq_iInter_finset' for a version that works for ι : Sort*.

theorem Set.iInter_eq_iInter_finset' {α : Type u_2} {ι' : Sort u_7} (s : ι'Set α) :
⋂ (i : ι'), s i = ⋂ (t : Finset (PLift ι')), it, s i.down

Intersection of an indexed family of sets s : ι → Set α is equal to the intersection of the intersections of finite subfamilies. This version works for ι : Sort*. See also iInter_eq_iInter_finset for a version that assumes ι : Type* but avoids PLifts in the right hand side.

theorem Finset.maximal_iff_forall_insert {α : Type u_2} [DecidableEq α] {P : Finset αProp} {s : Finset α} (hP : ∀ ⦃s t : Finset α⦄, P ts tP s) :
Maximal P s P s xs, ¬P (insert x s)
theorem Finset.minimal_iff_forall_diff_singleton {α : Type u_2} [DecidableEq α] {P : Finset αProp} {s : Finset α} (hP : ∀ ⦃s t : Finset α⦄, P tt sP s) :
Minimal P s P s xs, ¬P (s.erase x)

Interaction with big lattice/set operations #

theorem Finset.iSup_coe {α : Type u_2} {β : Type u_3} [SupSet β] (f : αβ) (s : Finset α) :
xs, f x = xs, f x
theorem Finset.iInf_coe {α : Type u_2} {β : Type u_3} [InfSet β] (f : αβ) (s : Finset α) :
xs, f x = xs, f x
theorem Finset.iSup_singleton {α : Type u_2} {β : Type u_3} [CompleteLattice β] (a : α) (s : αβ) :
x{a}, s x = s a
theorem Finset.iInf_singleton {α : Type u_2} {β : Type u_3} [CompleteLattice β] (a : α) (s : αβ) :
x{a}, s x = s a
theorem Finset.iSup_option_toFinset {α : Type u_2} {β : Type u_3} [CompleteLattice β] (o : Option α) (f : αβ) :
xo.toFinset, f x = xo, f x
theorem Finset.iInf_option_toFinset {α : Type u_2} {β : Type u_3} [CompleteLattice β] (o : Option α) (f : αβ) :
xo.toFinset, f x = xo, f x
theorem Finset.iSup_union {α : Type u_2} {β : Type u_3} [CompleteLattice β] [DecidableEq α] {f : αβ} {s : Finset α} {t : Finset α} :
xs t, f x = (⨆ xs, f x) xt, f x
theorem Finset.iInf_union {α : Type u_2} {β : Type u_3} [CompleteLattice β] [DecidableEq α] {f : αβ} {s : Finset α} {t : Finset α} :
xs t, f x = (⨅ xs, f x) xt, f x
theorem Finset.iSup_insert {α : Type u_2} {β : Type u_3} [CompleteLattice β] [DecidableEq α] (a : α) (s : Finset α) (t : αβ) :
xinsert a s, t x = t a xs, t x
theorem Finset.iInf_insert {α : Type u_2} {β : Type u_3} [CompleteLattice β] [DecidableEq α] (a : α) (s : Finset α) (t : αβ) :
xinsert a s, t x = t a xs, t x
theorem Finset.iSup_finset_image {α : Type u_2} {β : Type u_3} {γ : Type u_4} [CompleteLattice β] [DecidableEq α] {f : γα} {g : αβ} {s : Finset γ} :
xFinset.image f s, g x = ys, g (f y)
theorem Finset.iInf_finset_image {α : Type u_2} {β : Type u_3} {γ : Type u_4} [CompleteLattice β] [DecidableEq α] {f : γα} {g : αβ} {s : Finset γ} :
xFinset.image f s, g x = ys, g (f y)
theorem Finset.iSup_insert_update {α : Type u_2} {β : Type u_3} [CompleteLattice β] [DecidableEq α] {x : α} {t : Finset α} (f : αβ) {s : β} (hx : xt) :
iinsert x t, Function.update f x s i = s it, f i
theorem Finset.iInf_insert_update {α : Type u_2} {β : Type u_3} [CompleteLattice β] [DecidableEq α] {x : α} {t : Finset α} (f : αβ) {s : β} (hx : xt) :
iinsert x t, Function.update f x s i = s it, f i
theorem Finset.iSup_biUnion {α : Type u_2} {β : Type u_3} {γ : Type u_4} [CompleteLattice β] [DecidableEq α] (s : Finset γ) (t : γFinset α) (f : αβ) :
ys.biUnion t, f y = xs, yt x, f y
theorem Finset.iInf_biUnion {α : Type u_2} {β : Type u_3} {γ : Type u_4} [CompleteLattice β] [DecidableEq α] (s : Finset γ) (t : γFinset α) (f : αβ) :
ys.biUnion t, f y = xs, yt x, f y
theorem Finset.set_biUnion_coe {α : Type u_2} {β : Type u_3} (s : Finset α) (t : αSet β) :
xs, t x = xs, t x
theorem Finset.set_biInter_coe {α : Type u_2} {β : Type u_3} (s : Finset α) (t : αSet β) :
xs, t x = xs, t x
theorem Finset.set_biUnion_singleton {α : Type u_2} {β : Type u_3} (a : α) (s : αSet β) :
x{a}, s x = s a
theorem Finset.set_biInter_singleton {α : Type u_2} {β : Type u_3} (a : α) (s : αSet β) :
x{a}, s x = s a
@[simp]
theorem Finset.set_biUnion_preimage_singleton {α : Type u_2} {β : Type u_3} (f : αβ) (s : Finset β) :
ys, f ⁻¹' {y} = f ⁻¹' s
theorem Finset.set_biUnion_option_toFinset {α : Type u_2} {β : Type u_3} (o : Option α) (f : αSet β) :
xo.toFinset, f x = xo, f x
theorem Finset.set_biInter_option_toFinset {α : Type u_2} {β : Type u_3} (o : Option α) (f : αSet β) :
xo.toFinset, f x = xo, f x
theorem Finset.subset_set_biUnion_of_mem {α : Type u_2} {β : Type u_3} {s : Finset α} {f : αSet β} {x : α} (h : x s) :
f x ys, f y
theorem Finset.set_biUnion_union {α : Type u_2} {β : Type u_3} [DecidableEq α] (s : Finset α) (t : Finset α) (u : αSet β) :
xs t, u x = (⋃ xs, u x) xt, u x
theorem Finset.set_biInter_inter {α : Type u_2} {β : Type u_3} [DecidableEq α] (s : Finset α) (t : Finset α) (u : αSet β) :
xs t, u x = (⋂ xs, u x) xt, u x
theorem Finset.set_biUnion_insert {α : Type u_2} {β : Type u_3} [DecidableEq α] (a : α) (s : Finset α) (t : αSet β) :
xinsert a s, t x = t a xs, t x
theorem Finset.set_biInter_insert {α : Type u_2} {β : Type u_3} [DecidableEq α] (a : α) (s : Finset α) (t : αSet β) :
xinsert a s, t x = t a xs, t x
theorem Finset.set_biUnion_finset_image {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq α] {f : γα} {g : αSet β} {s : Finset γ} :
xFinset.image f s, g x = ys, g (f y)
theorem Finset.set_biInter_finset_image {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq α] {f : γα} {g : αSet β} {s : Finset γ} :
xFinset.image f s, g x = ys, g (f y)
theorem Finset.set_biUnion_insert_update {α : Type u_2} {β : Type u_3} [DecidableEq α] {x : α} {t : Finset α} (f : αSet β) {s : Set β} (hx : xt) :
iinsert x t, Function.update f x s i = s it, f i
theorem Finset.set_biInter_insert_update {α : Type u_2} {β : Type u_3} [DecidableEq α] {x : α} {t : Finset α} (f : αSet β) {s : Set β} (hx : xt) :
iinsert x t, Function.update f x s i = s it, f i
theorem Finset.set_biUnion_biUnion {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq α] (s : Finset γ) (t : γFinset α) (f : αSet β) :
ys.biUnion t, f y = xs, yt x, f y
theorem Finset.set_biInter_biUnion {α : Type u_2} {β : Type u_3} {γ : Type u_4} [DecidableEq α] (s : Finset γ) (t : γFinset α) (f : αSet β) :
ys.biUnion t, f y = xs, yt x, f y