Lattice operations on finsets of products #
This file is concerned with folding binary lattice operations over finsets.
sup #
@[simp]
inf #
@[simp]
theorem
Finset.prodMk_sup'_sup'
{ι : Type u_7}
{κ : Type u_8}
{α : Type u_9}
{β : Type u_10}
[SemilatticeSup α]
[SemilatticeSup β]
{s : Finset ι}
{t : Finset κ}
(hs : s.Nonempty)
(ht : t.Nonempty)
(f : ι → α)
(g : κ → β)
:
See also Finset.sup'_prodMap
.
theorem
Finset.sup'_prodMap
{ι : Type u_7}
{κ : Type u_8}
{α : Type u_9}
{β : Type u_10}
[SemilatticeSup α]
[SemilatticeSup β]
{s : Finset ι}
{t : Finset κ}
(hst : (s ×ˢ t).Nonempty)
(f : ι → α)
(g : κ → β)
:
See also Finset.prodMk_sup'_sup'
.
theorem
Finset.prodMk_inf'_inf'
{ι : Type u_7}
{κ : Type u_8}
{α : Type u_9}
{β : Type u_10}
[SemilatticeInf α]
[SemilatticeInf β]
{s : Finset ι}
{t : Finset κ}
(hs : s.Nonempty)
(ht : t.Nonempty)
(f : ι → α)
(g : κ → β)
:
See also Finset.inf'_prodMap
.
theorem
Finset.inf'_prodMap
{ι : Type u_7}
{κ : Type u_8}
{α : Type u_9}
{β : Type u_10}
[SemilatticeInf α]
[SemilatticeInf β]
{s : Finset ι}
{t : Finset κ}
(hst : (s ×ˢ t).Nonempty)
(f : ι → α)
(g : κ → β)
:
See also Finset.prodMk_inf'_inf'
.