Library Coq.Structures.OrderedTypeAlt

Require Import OrderedType.

An alternative (but equivalent) presentation for an Ordered Type

inferface.
NB: comparison, defined in Datatypes.v is Eq|Lt|Gt whereas compare, defined in OrderedType.v is EQ _ | LT _ | GT _

Module Type OrderedTypeAlt.

 Parameter t : Type.

 Parameter compare : t -> t -> comparison.

 Infix "?=" := compare (at level 70, no associativity).

 Parameter compare_sym :
   forall x y, (y?=x) = CompOpp (x?=y).
 Parameter compare_trans :
   forall c x y z, (x?=y) = c -> (y?=z) = c -> (x?=z) = c.

End OrderedTypeAlt.

From this new presentation to the original one.

Module OrderedType_from_Alt (O:OrderedTypeAlt) <: OrderedType.
 Import O.

 Definition t := t.

 Definition eq x y := (x?=y) = Eq.
 Definition lt x y := (x?=y) = Lt.

 Lemma eq_refl : forall x, eq x x.

 Lemma eq_sym : forall x y, eq x y -> eq y x.

 Definition eq_trans := (compare_trans Eq).

 Definition lt_trans := (compare_trans Lt).

 Lemma lt_not_eq : forall x y, lt x y -> ~eq x y.

 Definition compare : forall x y, Compare lt eq x y.

 Definition eq_dec : forall x y, { eq x y } + { ~ eq x y }.

End OrderedType_from_Alt.

From the original presentation to this alternative one.

Module OrderedType_to_Alt (O:OrderedType) <: OrderedTypeAlt.
 Import O.
 Module MO:=OrderedTypeFacts(O).
 Import MO.

 Definition t := t.

 Definition compare x y := match compare x y with
   | LT _ => Lt
   | EQ _ => Eq
   | GT _ => Gt
  end.

 Infix "?=" := compare (at level 70, no associativity).

 Lemma compare_sym :
   forall x y, (y?=x) = CompOpp (x?=y).

 Lemma compare_trans :
   forall c x y z, (x?=y) = c -> (y?=z) = c -> (x?=z) = c.

End OrderedType_to_Alt.