Library Coq.Sets.Partial_Order


Require Export Ensembles.
Require Export Relations_1.

Section Partial_orders.
  Variable U : Type.

  Definition Carrier := Ensemble U.

  Definition Rel := Relation U.

  Variable p : PO.

  Definition Strict_Rel_of : Rel := fun x y:U => Rel_of p x y /\ x <> y.

  Inductive covers (y x:U) : Prop :=
    Definition_of_covers :
    Strict_Rel_of x y ->
    ~ (exists z : _, Strict_Rel_of x z /\ Strict_Rel_of z y) ->
    covers y x.

End Partial_orders.

Hint Unfold Carrier_of Rel_of Strict_Rel_of: sets.
Hint Resolve Definition_of_covers: sets.

Section Partial_order_facts.
  Variable U : Type.
  Variable D : PO U.

  Lemma Strict_Rel_Transitive_with_Rel :
    forall x y z:U,
      Strict_Rel_of U D x y -> @Rel_of U D y z -> Strict_Rel_of U D x z.

  Lemma Strict_Rel_Transitive_with_Rel_left :
    forall x y z:U,
      @Rel_of U D x y -> Strict_Rel_of U D y z -> Strict_Rel_of U D x z.

  Lemma Strict_Rel_Transitive : Transitive U (Strict_Rel_of U D).
End Partial_order_facts.