(Semi-)lattices #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
Semilattices are partially ordered sets with join (greatest lower bound, or sup) or
meet (least upper bound, or inf) operations. Lattices are posets that are both
join-semilattices and meet-semilattices.
Distributive lattices are lattices which satisfy any of four equivalent distributivity properties,
of sup over inf, on the left or on the right.
Main declarations #
-
semilattice_sup: a type class for join semilattices -
semilattice_sup.mk': an alternative constructor forsemilattice_supvia proofs that⊔is commutative, associative and idempotent. -
semilattice_inf: a type class for meet semilattices -
semilattice_sup.mk': an alternative constructor forsemilattice_infvia proofs that⊓is commutative, associative and idempotent. -
lattice: a type class for lattices -
lattice.mk': an alternative constructor forlatticevia profs that⊔and⊓are commutative, associative and satisfy a pair of "absorption laws". -
distrib_lattice: a type class for distributive lattices.
Notations #
a ⊔ b: the supremum or join ofaandba ⊓ b: the infimum or meet ofaandb
TODO #
- (Semi-)lattice homomorphisms
- Alternative constructors for distributive lattices from the other distributive properties
Tags #
semilattice, lattice
Join-semilattices #
- sup : α → α → α
- le : α → α → Prop
- lt : α → α → Prop
- le_refl : ∀ (a : α), a ≤ a
- le_trans : ∀ (a b c : α), a ≤ b → b ≤ c → a ≤ c
- lt_iff_le_not_le : (∀ (a b : α), a < b ↔ a ≤ b ∧ ¬b ≤ a) . "order_laws_tac"
- le_antisymm : ∀ (a b : α), a ≤ b → b ≤ a → a = b
- le_sup_left : ∀ (a b : α), a ≤ a ⊔ b
- le_sup_right : ∀ (a b : α), b ≤ a ⊔ b
- sup_le : ∀ (a b c : α), a ≤ c → b ≤ c → a ⊔ b ≤ c
A semilattice_sup is a join-semilattice, that is, a partial order
with a join (a.k.a. lub / least upper bound, sup / supremum) operation
⊔ which is the least element larger than both factors.
Instances of this typeclass
- lattice.to_semilattice_sup
- canonically_linear_ordered_monoid.semilattice_sup
- canonically_linear_ordered_add_monoid.semilattice_sup
- idem_semiring.to_semilattice_sup
- order_dual.semilattice_sup
- pi.semilattice_sup
- prod.semilattice_sup
- with_bot.semilattice_sup
- with_top.semilattice_sup
- nat.subtype.semilattice_sup
- rat.semilattice_sup
- top_hom.semilattice_sup
- bot_hom.semilattice_sup
- sup_hom.semilattice_sup
- sup_bot_hom.semilattice_sup
- set.Ioc.semilattice_sup
- set.Ioi.semilattice_sup
- set.Iic.semilattice_sup
- set.Ici.semilattice_sup
- set.Icc.semilattice_sup
- order_hom.semilattice_sup
- part_enat.semilattice_sup
- fixed_points.function.fixed_points.semilattice_sup
- finsupp.semilattice_sup
- nonneg.semilattice_sup
- real.semilattice_sup
- nnreal.semilattice_sup
- ennreal.semilattice_sup
- group_seminorm.semilattice_sup
- add_group_seminorm.semilattice_sup
- nonarch_add_group_seminorm.semilattice_sup
- group_norm.semilattice_sup
- add_group_norm.semilattice_sup
- nonarch_add_group_norm.semilattice_sup
- seminorm.semilattice_sup
- topological_space.compacts.semilattice_sup
- topological_space.nonempty_compacts.semilattice_sup
- topological_space.positive_compacts.semilattice_sup
- topological_space.compact_opens.semilattice_sup
Instances of other typeclasses for semilattice_sup
- semilattice_sup.has_sizeof_inst
A type with a commutative, associative and idempotent binary sup operation has the structure of a
join-semilattice.
The partial order is defined so that a ≤ b unfolds to a ⊔ b = b; cf. sup_eq_right.
Equations
- semilattice_sup.mk' sup_comm sup_assoc sup_idem = {sup := has_sup.sup _inst_1, le := λ (a b : α), a ⊔ b = b, lt := partial_order.lt._default (λ (a b : α), a ⊔ b = b), le_refl := sup_idem, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _}
Equations
- order_dual.has_sup α = {sup := has_inf.inf _inst_1}
Equations
- order_dual.has_inf α = {inf := has_sup.sup _inst_1}
Alias of the reverse direction of sup_eq_left.
Alias of the reverse direction of sup_eq_right.
Alias of the forward direction of sup_eq_right.
Meet-semilattices #
- inf : α → α → α
- le : α → α → Prop
- lt : α → α → Prop
- le_refl : ∀ (a : α), a ≤ a
- le_trans : ∀ (a b c : α), a ≤ b → b ≤ c → a ≤ c
- lt_iff_le_not_le : (∀ (a b : α), a < b ↔ a ≤ b ∧ ¬b ≤ a) . "order_laws_tac"
- le_antisymm : ∀ (a b : α), a ≤ b → b ≤ a → a = b
- inf_le_left : ∀ (a b : α), a ⊓ b ≤ a
- inf_le_right : ∀ (a b : α), a ⊓ b ≤ b
- le_inf : ∀ (a b c : α), a ≤ b → a ≤ c → a ≤ b ⊓ c
A semilattice_inf is a meet-semilattice, that is, a partial order
with a meet (a.k.a. glb / greatest lower bound, inf / infimum) operation
⊓ which is the greatest element smaller than both factors.
Instances of this typeclass
- lattice.to_semilattice_inf
- order_dual.semilattice_inf
- pi.semilattice_inf
- prod.semilattice_inf
- with_bot.semilattice_inf
- with_top.semilattice_inf
- rat.semilattice_inf
- top_hom.semilattice_inf
- bot_hom.semilattice_inf
- inf_hom.semilattice_inf
- inf_top_hom.semilattice_inf
- set.Ico.semilattice_inf
- set.Iio.semilattice_inf
- set.Iic.semilattice_inf
- set.Ici.semilattice_inf
- set.Icc.semilattice_inf
- order_hom.semilattice_inf
- fixed_points.function.fixed_points.semilattice_inf
- linear_pmap.semilattice_inf
- finsupp.semilattice_inf
- group_topology.semilattice_inf
- add_group_topology.semilattice_inf
- nonneg.semilattice_inf
- real.semilattice_inf
- nnreal.semilattice_inf
- topological_space.compact_opens.semilattice_inf
- pequiv.semilattice_inf
Instances of other typeclasses for semilattice_inf
- semilattice_inf.has_sizeof_inst
Equations
- order_dual.semilattice_sup α = {sup := has_sup.sup (order_dual.has_sup α), le := partial_order.le (order_dual.partial_order α), lt := partial_order.lt (order_dual.partial_order α), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _}
Equations
- order_dual.semilattice_inf α = {inf := has_inf.inf (order_dual.has_inf α), le := partial_order.le (order_dual.partial_order α), lt := partial_order.lt (order_dual.partial_order α), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, inf_le_left := _, inf_le_right := _, le_inf := _}
Alias of the forward direction of inf_eq_left.
Alias of the reverse direction of inf_eq_left.
Alias of the reverse direction of inf_eq_right.
A type with a commutative, associative and idempotent binary inf operation has the structure of a
meet-semilattice.
The partial order is defined so that a ≤ b unfolds to b ⊓ a = a; cf. inf_eq_right.
Equations
- semilattice_inf.mk' inf_comm inf_assoc inf_idem = order_dual.semilattice_inf αᵒᵈ
Lattices #
- sup : α → α → α
- le : α → α → Prop
- lt : α → α → Prop
- le_refl : ∀ (a : α), a ≤ a
- le_trans : ∀ (a b c : α), a ≤ b → b ≤ c → a ≤ c
- lt_iff_le_not_le : (∀ (a b : α), a < b ↔ a ≤ b ∧ ¬b ≤ a) . "order_laws_tac"
- le_antisymm : ∀ (a b : α), a ≤ b → b ≤ a → a = b
- le_sup_left : ∀ (a b : α), a ≤ a ⊔ b
- le_sup_right : ∀ (a b : α), b ≤ a ⊔ b
- sup_le : ∀ (a b c : α), a ≤ c → b ≤ c → a ⊔ b ≤ c
- inf : α → α → α
- inf_le_left : ∀ (a b : α), a ⊓ b ≤ a
- inf_le_right : ∀ (a b : α), a ⊓ b ≤ b
- le_inf : ∀ (a b c : α), a ≤ b → a ≤ c → a ≤ b ⊓ c
A lattice is a join-semilattice which is also a meet-semilattice.
Instances of this typeclass
- distrib_lattice.to_lattice
- linear_order.to_lattice
- generalized_heyting_algebra.to_lattice
- generalized_coheyting_algebra.to_lattice
- complete_lattice.to_lattice
- conditionally_complete_lattice.to_lattice
- order_dual.lattice
- pi.lattice
- prod.lattice
- with_bot.lattice
- with_top.lattice
- with_zero.lattice
- rat.lattice
- fin.lattice
- multiset.lattice
- finset.lattice
- top_hom.lattice
- bot_hom.lattice
- set.Iic.lattice
- set.Ici.lattice
- set.Icc.lattice
- nat.lattice
- order_hom.lattice
- part_enat.lattice
- finsupp.lattice
- associates.lattice
- real.lattice
- group_seminorm.lattice
- add_group_seminorm.lattice
- seminorm.lattice
Instances of other typeclasses for lattice
- lattice.has_sizeof_inst
Equations
- order_dual.lattice α = {sup := semilattice_sup.sup (order_dual.semilattice_sup α), le := semilattice_sup.le (order_dual.semilattice_sup α), lt := semilattice_sup.lt (order_dual.semilattice_sup α), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := semilattice_inf.inf (order_dual.semilattice_inf α), inf_le_left := _, inf_le_right := _, le_inf := _}
The partial orders from semilattice_sup_mk' and semilattice_inf_mk' agree
if sup and inf satisfy the lattice absorption laws sup_inf_self (a ⊔ a ⊓ b = a)
and inf_sup_self (a ⊓ (a ⊔ b) = a).
A type with a pair of commutative and associative binary operations which satisfy two absorption laws relating the two operations has the structure of a lattice.
The partial order is defined so that a ≤ b unfolds to a ⊔ b = b; cf. sup_eq_right.
Equations
- lattice.mk' sup_comm sup_assoc inf_comm inf_assoc sup_inf_self inf_sup_self = have sup_idem : ∀ (b : α), b ⊔ b = b, from _, have inf_idem : ∀ (b : α), b ⊓ b = b, from _, let semilatt_inf_inst : semilattice_inf α := semilattice_inf.mk' inf_comm inf_assoc inf_idem, semilatt_sup_inst : semilattice_sup α := semilattice_sup.mk' sup_comm sup_assoc sup_idem, partial_order_inst : partial_order α := semilattice_sup.to_partial_order α in have partial_order_eq : partial_order_inst = semilattice_inf.to_partial_order α, from _, {sup := semilattice_sup.sup semilatt_sup_inst, le := partial_order.le partial_order_inst, lt := partial_order.lt partial_order_inst, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := semilattice_inf.inf semilatt_inf_inst, inf_le_left := _, inf_le_right := _, le_inf := _}
Distributivity laws #
Distributive lattices #
- sup : α → α → α
- le : α → α → Prop
- lt : α → α → Prop
- le_refl : ∀ (a : α), a ≤ a
- le_trans : ∀ (a b c : α), a ≤ b → b ≤ c → a ≤ c
- lt_iff_le_not_le : (∀ (a b : α), a < b ↔ a ≤ b ∧ ¬b ≤ a) . "order_laws_tac"
- le_antisymm : ∀ (a b : α), a ≤ b → b ≤ a → a = b
- le_sup_left : ∀ (a b : α), a ≤ a ⊔ b
- le_sup_right : ∀ (a b : α), b ≤ a ⊔ b
- sup_le : ∀ (a b c : α), a ≤ c → b ≤ c → a ⊔ b ≤ c
- inf : α → α → α
- inf_le_left : ∀ (a b : α), a ⊓ b ≤ a
- inf_le_right : ∀ (a b : α), a ⊓ b ≤ b
- le_inf : ∀ (a b c : α), a ≤ b → a ≤ c → a ≤ b ⊓ c
- le_sup_inf : ∀ (x y z : α), (x ⊔ y) ⊓ (x ⊔ z) ≤ x ⊔ y ⊓ z
A distributive lattice is a lattice that satisfies any of four
equivalent distributive properties (of sup over inf or inf over sup,
on the left or right).
The definition here chooses le_sup_inf: (x ⊔ y) ⊓ (x ⊔ z) ≤ x ⊔ (y ⊓ z). To prove distributivity
from the dual law, use distrib_lattice.of_inf_sup_le.
A classic example of a distributive lattice is the lattice of subsets of a set, and in fact this example is generic in the sense that every distributive lattice is realizable as a sublattice of a powerset lattice.
Instances of this typeclass
- linear_order.to_distrib_lattice
- generalized_heyting_algebra.to_distrib_lattice
- generalized_coheyting_algebra.to_distrib_lattice
- coheyting_algebra.to_distrib_lattice
- generalized_boolean_algebra.to_distrib_lattice
- boolean_algebra.to_distrib_lattice
- frame.to_distrib_lattice
- coframe.to_distrib_lattice
- order_dual.distrib_lattice
- nat.distrib_lattice
- pi.distrib_lattice
- prod.distrib_lattice
- bool.distrib_lattice
- complementeds.distrib_lattice
- with_bot.distrib_lattice
- with_top.distrib_lattice
- Prop.distrib_lattice
- rat.distrib_lattice
- multiset.distrib_lattice
- finset.distrib_lattice
- top_hom.distrib_lattice
- bot_hom.distrib_lattice
- set.Ici.distrib_lattice
- cardinal.distrib_lattice
- filter.distrib_lattice
- nonneg.distrib_lattice
- real.distrib_lattice
- nnreal.distrib_lattice
- ennreal.distrib_lattice
- topological_space.open_nhds_of.distrib_lattice
- topological_space.compacts.distrib_lattice
Instances of other typeclasses for distrib_lattice
- distrib_lattice.has_sizeof_inst
Equations
- order_dual.distrib_lattice α = {sup := lattice.sup (order_dual.lattice α), le := lattice.le (order_dual.lattice α), lt := lattice.lt (order_dual.lattice α), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := lattice.inf (order_dual.lattice α), inf_le_left := _, inf_le_right := _, le_inf := _, le_sup_inf := _}
Prove distributivity of an existing lattice from the dual distributive law.
Equations
- distrib_lattice.of_inf_sup_le inf_sup_le = {sup := lattice.sup _inst_1, le := lattice.le _inst_1, lt := lattice.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := lattice.inf _inst_1, inf_le_left := _, inf_le_right := _, le_inf := _, le_sup_inf := _}
Lattices derived from linear orders #
Equations
- linear_order.to_lattice = {sup := linear_order.max o, le := linear_order.le o, lt := linear_order.lt o, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := linear_order.min o, inf_le_left := _, inf_le_right := _, le_inf := _}
A lattice with total order is a linear order.
See note [reducible non-instances].
Equations
- lattice.to_linear_order α = {le := lattice.le _inst_1, lt := lattice.lt _inst_1, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_total := _, decidable_le := _inst_3, decidable_eq := _inst_2, decidable_lt := _inst_4, max := has_sup.sup (semilattice_sup.to_has_sup α), max_def := _, min := has_inf.inf (semilattice_inf.to_has_inf α), min_def := _}
Equations
- linear_order.to_distrib_lattice = {sup := lattice.sup linear_order.to_lattice, le := lattice.le linear_order.to_lattice, lt := lattice.lt linear_order.to_lattice, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := lattice.inf linear_order.to_lattice, inf_le_left := _, inf_le_right := _, le_inf := _, le_sup_inf := _}
Dual order #
Function lattices #
Equations
- pi.semilattice_sup = {sup := has_sup.sup pi.has_sup, le := partial_order.le pi.partial_order, lt := partial_order.lt pi.partial_order, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _}
Equations
- pi.semilattice_inf = {inf := has_inf.inf pi.has_inf, le := partial_order.le pi.partial_order, lt := partial_order.lt pi.partial_order, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, inf_le_left := _, inf_le_right := _, le_inf := _}
Equations
- pi.lattice = {sup := semilattice_sup.sup pi.semilattice_sup, le := semilattice_sup.le pi.semilattice_sup, lt := semilattice_sup.lt pi.semilattice_sup, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := semilattice_inf.inf pi.semilattice_inf, inf_le_left := _, inf_le_right := _, le_inf := _}
Equations
- pi.distrib_lattice = {sup := lattice.sup pi.lattice, le := lattice.le pi.lattice, lt := lattice.lt pi.lattice, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := lattice.inf pi.lattice, inf_le_left := _, inf_le_right := _, le_inf := _, le_sup_inf := _}
Monotone functions and lattices #
Pointwise maximum of two monotone functions is a monotone function.
Pointwise minimum of two monotone functions is a monotone function.
Pointwise supremum of two monotone functions is a monotone function.
Pointwise infimum of two monotone functions is a monotone function.
Pointwise maximum of two monotone functions is a monotone function.
Pointwise minimum of two monotone functions is a monotone function.
Pointwise maximum of two monotone functions is a monotone function.
Pointwise minimum of two monotone functions is a monotone function.
Pointwise supremum of two antitone functions is a antitone function.
Pointwise infimum of two antitone functions is a antitone function.
Pointwise maximum of two antitone functions is a antitone function.
Pointwise minimum of two antitone functions is a antitone function.
Products of (semi-)lattices #
Equations
- prod.semilattice_sup α β = {sup := has_sup.sup (prod.has_sup α β), le := partial_order.le (prod.partial_order α β), lt := partial_order.lt (prod.partial_order α β), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _}
Equations
- prod.semilattice_inf α β = {inf := has_inf.inf (prod.has_inf α β), le := partial_order.le (prod.partial_order α β), lt := partial_order.lt (prod.partial_order α β), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, inf_le_left := _, inf_le_right := _, le_inf := _}
Equations
- prod.lattice α β = {sup := semilattice_sup.sup (prod.semilattice_sup α β), le := semilattice_inf.le (prod.semilattice_inf α β), lt := semilattice_inf.lt (prod.semilattice_inf α β), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := semilattice_inf.inf (prod.semilattice_inf α β), inf_le_left := _, inf_le_right := _, le_inf := _}
Equations
- prod.distrib_lattice α β = {sup := lattice.sup (prod.lattice α β), le := lattice.le (prod.lattice α β), lt := lattice.lt (prod.lattice α β), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := lattice.inf (prod.lattice α β), inf_le_left := _, inf_le_right := _, le_inf := _, le_sup_inf := _}
Subtypes of (semi-)lattices #
A subtype forms a ⊔-semilattice if ⊔ preserves the property.
See note [reducible non-instances].
Equations
- subtype.semilattice_sup Psup = {sup := λ (x y : {x // P x}), ⟨x.val ⊔ y.val, _⟩, le := partial_order.le (subtype.partial_order P), lt := partial_order.lt (subtype.partial_order P), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _}
A subtype forms a ⊓-semilattice if ⊓ preserves the property.
See note [reducible non-instances].
Equations
- subtype.semilattice_inf Pinf = {inf := λ (x y : {x // P x}), ⟨x.val ⊓ y.val, _⟩, le := partial_order.le (subtype.partial_order P), lt := partial_order.lt (subtype.partial_order P), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, inf_le_left := _, inf_le_right := _, le_inf := _}
A subtype forms a lattice if ⊔ and ⊓ preserve the property.
See note [reducible non-instances].
Equations
- subtype.lattice Psup Pinf = {sup := semilattice_sup.sup (subtype.semilattice_sup Psup), le := semilattice_inf.le (subtype.semilattice_inf Pinf), lt := semilattice_inf.lt (subtype.semilattice_inf Pinf), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := semilattice_inf.inf (subtype.semilattice_inf Pinf), inf_le_left := _, inf_le_right := _, le_inf := _}
A type endowed with ⊔ is a semilattice_sup, if it admits an injective map that
preserves ⊔ to a semilattice_sup.
See note [reducible non-instances].
Equations
- function.injective.semilattice_sup f hf_inj map_sup = {sup := has_sup.sup _inst_1, le := partial_order.le (partial_order.lift f hf_inj), lt := partial_order.lt (partial_order.lift f hf_inj), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _}
A type endowed with ⊓ is a semilattice_inf, if it admits an injective map that
preserves ⊓ to a semilattice_inf.
See note [reducible non-instances].
Equations
- function.injective.semilattice_inf f hf_inj map_inf = {inf := has_inf.inf _inst_1, le := partial_order.le (partial_order.lift f hf_inj), lt := partial_order.lt (partial_order.lift f hf_inj), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, inf_le_left := _, inf_le_right := _, le_inf := _}
A type endowed with ⊔ and ⊓ is a lattice, if it admits an injective map that
preserves ⊔ and ⊓ to a lattice.
See note [reducible non-instances].
Equations
- function.injective.lattice f hf_inj map_sup map_inf = {sup := semilattice_sup.sup (function.injective.semilattice_sup f hf_inj map_sup), le := semilattice_sup.le (function.injective.semilattice_sup f hf_inj map_sup), lt := semilattice_sup.lt (function.injective.semilattice_sup f hf_inj map_sup), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := semilattice_inf.inf (function.injective.semilattice_inf f hf_inj map_inf), inf_le_left := _, inf_le_right := _, le_inf := _}
A type endowed with ⊔ and ⊓ is a distrib_lattice, if it admits an injective map that
preserves ⊔ and ⊓ to a distrib_lattice.
See note [reducible non-instances].
Equations
- function.injective.distrib_lattice f hf_inj map_sup map_inf = {sup := lattice.sup (function.injective.lattice f hf_inj map_sup map_inf), le := lattice.le (function.injective.lattice f hf_inj map_sup map_inf), lt := lattice.lt (function.injective.lattice f hf_inj map_sup map_inf), le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, le_sup_left := _, le_sup_right := _, sup_le := _, inf := lattice.inf (function.injective.lattice f hf_inj map_sup map_inf), inf_le_left := _, inf_le_right := _, le_inf := _, le_sup_inf := _}