Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)
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.

 

forall x : O.t, eq x x

forall x : O.t, eq x x
x:O.t

eq x x
x:O.t

(x ?= x) = Eq
x:O.t
H:(x ?= x) = CompOpp (x ?= x)

(x ?= x) = Eq
destruct (x ?= x); simpl in *; try discriminate; auto. Qed.

forall x y : O.t, eq x y -> eq y x

forall x y : O.t, eq x y -> eq y x
x, y:O.t
H:(x ?= y) = Eq

(y ?= x) = Eq
x, y:O.t
H:(x ?= y) = Eq

CompOpp (x ?= y) = Eq
rewrite H; simpl; auto. Qed. Definition eq_trans := (compare_trans Eq). Definition lt_trans := (compare_trans Lt).

forall x y : O.t, lt x y -> ~ eq x y

forall x y : O.t, lt x y -> ~ eq x y
x, y:O.t
H:(x ?= y) = Lt

(x ?= y) <> Eq
rewrite H; discriminate. Qed.

forall x y : O.t, Compare lt eq x y

forall x y : O.t, Compare lt eq x y
x, y:O.t

Compare lt eq x y
x, y:O.t
H:(x ?= y) = Eq

Compare lt eq x y
x, y:O.t
H:(x ?= y) = Lt
Compare lt eq x y
x, y:O.t
H:(x ?= y) = Gt
Compare lt eq x y
x, y:O.t
H:(x ?= y) = Lt

Compare lt eq x y
x, y:O.t
H:(x ?= y) = Gt
Compare lt eq x y
x, y:O.t
H:(x ?= y) = Gt

Compare lt eq x y
x, y:O.t
H:(x ?= y) = Gt

(y ?= x) = Lt
rewrite compare_sym; rewrite H; auto. Defined.

forall x y : O.t, {eq x y} + {~ eq x y}

forall x y : O.t, {eq x y} + {~ eq x y}
x, y:O.t

{(x ?= y) = Eq} + {(x ?= y) <> Eq}
case (x ?= y); [ left | right | right ]; auto; discriminate. Defined. 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).

 

forall x y : O.t, (y ?= x) = CompOpp (x ?= y)

forall x y : O.t, (y ?= x) = CompOpp (x ?= y)
x, y:O.t

match O.compare y x with | LT _ => Lt | EQ _ => Eq | GT _ => Gt end = CompOpp match O.compare x y with | LT _ => Lt | EQ _ => Eq | GT _ => Gt end
destruct O.compare; elim_comp; simpl; auto. Qed.

forall (c : comparison) (x y z : O.t), (x ?= y) = c -> (y ?= z) = c -> (x ?= z) = c

forall (c : comparison) (x y z : O.t), (x ?= y) = c -> (y ?= z) = c -> (x ?= z) = c
c:comparison
x, y, z:O.t

(x ?= y) = c -> (y ?= z) = c -> (x ?= z) = c
destruct c; unfold compare; do 2 (destruct O.compare; intros; try discriminate); elim_comp; auto. Qed. End OrderedType_to_Alt.