summaryrefslogtreecommitdiff
path: root/theories/FSets/OrderedTypeAlt.v
blob: 9bcfbfc7ee25574dfe7fe045730e0a3b0f7f2dad (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
(***********************************************************************)
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team    *)
(* <O___,, *        INRIA-Rocquencourt  &  LRI-CNRS-Orsay              *)
(*   \VV/  *************************************************************)
(*    //   *      This file is distributed under the terms of the      *)
(*         *       GNU Lesser General Public License Version 2.1       *)
(***********************************************************************)

(* Finite sets library.  
 * Authors: Pierre Letouzey and Jean-Christophe Filliâtre 
 * Institution: LRI, CNRS UMR 8623 - Université Paris Sud
 *              91405 Orsay, France *)

(* $Id: OrderedTypeAlt.v 8773 2006-04-29 14:31:32Z letouzey $ *)

Require Import OrderedType.

(** * An alternative (but equivalent) presentation for an Ordered Type inferface. *)

(** NB: [comparison], defined in [theories/Init/datatypes.v] is [Eq|Lt|Gt]
whereas [compare], defined in [theories/FSets/OrderedType.v] is [EQ _ | LT _ | GT _ ] 
*)

Module Type OrderedTypeAlt.

 Parameter t : Set.
 
 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.
 Proof.
 intro x.
 unfold eq.
 assert (H:=compare_sym x x).
 destruct (x ?= x); simpl in *; try discriminate; auto.
 Qed.

 Lemma eq_sym : forall x y, eq x y -> eq y x.
 Proof. 
 unfold eq; intros.
 rewrite compare_sym.
 rewrite H; simpl; auto.
 Qed.

 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.
 Proof.
 unfold eq, lt; intros.
 rewrite H; discriminate.
 Qed.

 Definition compare : forall x y, Compare lt eq x y.
 Proof.
 intros.
 case_eq (x ?= y); intros.
 apply EQ; auto.
 apply LT; auto.
 apply GT; red.
 rewrite compare_sym; rewrite H; auto.
 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).

 Lemma compare_sym : 
   forall x y, (y?=x) = CompOpp (x?=y).
 Proof.
 intros x y.
 unfold compare.
 destruct (O.compare y x); elim_comp; simpl; auto.
 Qed.
 
 Lemma compare_trans : 
   forall c x y z, (x?=y) = c -> (y?=z) = c -> (x?=z) = c.
 Proof.
 intros c x y z.
 destruct c; unfold compare.
 destruct (O.compare x y); intros; try discriminate.
 destruct (O.compare y z); intros; try discriminate.
 elim_comp; auto.
 destruct (O.compare x y); intros; try discriminate.
 destruct (O.compare y z); intros; try discriminate.
 elim_comp; auto.
 destruct (O.compare x y); intros; try discriminate.
 destruct (O.compare y z); intros; try discriminate.
 elim_comp; auto.
 Qed.

End OrderedType_to_Alt.