Orders #
Defines classes for preorders, partial orders, and linear orders and proves some basic lemmas about them.
Unbundled classes #
IsAntisymm X r
means the binary relation r
on X
is antisymmetric.
- antisymm : ∀ (a b : α), r a b → r b a → a = b
Instances
IsPreorder X r
means that the binary relation r
on X
is a pre-order, that is, reflexive
and transitive.
Instances
IsPartialOrder X r
means that the binary relation r
on X
is a partial order, that is,
IsPreorder X r
and IsAntisymm X r
.
Instances
IsLinearOrder X r
means that the binary relation r
on X
is a linear order, that is,
IsPartialOrder X r
and IsTotal X r
.
Instances
IsStrictWeakOrder X lt
means that the binary relation lt
on X
is a strict weak order,
that is, IsStrictOrder X lt
and ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a
.
- trans : ∀ (a b c : α), lt a b → lt b c → lt a c
Instances
IsTrichotomous X lt
means that the binary relation lt
on X
is trichotomous, that is,
either lt a b
or a = b
or lt b a
for any a
and b
.
Instances
IsStrictTotalOrder X lt
means that the binary relation lt
on X
is a strict total order,
that is, IsTrichotomous X lt
and IsStrictOrder X lt
.
- trichotomous : ∀ (a b : α), lt a b ∨ a = b ∨ lt b a
- trans : ∀ (a b c : α), lt a b → lt b c → lt a c
Instances
Equality is an equivalence relation.
IsTrans
as a definition, suitable for use in proofs.
Equations
- Transitive r = ∀ ⦃x y z : α⦄, r x y → r y z → r x z
Instances For
IsIrrefl
as a definition, suitable for use in proofs.
Equations
- Irreflexive r = ∀ (x : α), ¬r x x
Instances For
IsAntisymm
as a definition, suitable for use in proofs.
Equations
- AntiSymmetric r = ∀ ⦃x y : α⦄, r x y → r y x → x = y
Instances For
Minimal and maximal #
Upper and lower sets #
The type of upper sets of an order.
An upper set in an order α
is a set such that any element greater than one of its members is
also a member. Also called up-set, upward-closed set.
- carrier : Set α
The carrier of an
UpperSet
. - upper' : IsUpperSet self.carrier
The carrier of an
UpperSet
is an upper set.
Instances For
The type of lower sets of an order.
A lower set in an order α
is a set such that any element less than one of its members is also
a member. Also called down-set, downward-closed set.
- carrier : Set α
The carrier of a
LowerSet
. - lower' : IsLowerSet self.carrier
The carrier of a
LowerSet
is a lower set.