aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/QArith/QOrderedType.v
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2009-11-03 08:24:06 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2009-11-03 08:24:06 +0000
commit4f0ad99adb04e7f2888e75f2a10e8c916dde179b (patch)
tree4b52d7436fe06f4b2babfd5bfed84762440e7de7 /theories/QArith/QOrderedType.v
parent4e68924f48d3f6d5ffdf1cd394b590b5a6e15ea1 (diff)
OrderedType implementation for various numerical datatypes + min/max structures
- A richer OrderedTypeFull interface : OrderedType + predicate "le" - Implementations {Nat,N,P,Z,Q}OrderedType.v, also providing "order" tactics - By the way: as suggested by S. Lescuyer, specification of compare is now inductive - GenericMinMax: axiomatisation + properties of min and max out of OrderedTypeFull structures. - MinMax.v, {Z,P,N,Q}minmax.v are specialization of GenericMinMax, with also some domain-specific results, and compatibility layer with already existing results. - Some ML code of plugins had to be adapted, otherwise wrong "eq", "lt" or simimlar constants were found by functions like coq_constant. - Beware of the aliasing problems: for instance eq:=@eq t instead of eq:=@eq M.t in Make_UDT made (r)omega stopped working (Z_as_OT.t instead of Z in statement of Zmax_spec). - Some Morphism declaration are now ambiguous: switch to new syntax anyway. - Misc adaptations of FSets/MSets - Classes/RelationPairs.v: from two relations over A and B, we inspect relations over A*B and their properties in terms of classes. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12461 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories/QArith/QOrderedType.v')
-rw-r--r--theories/QArith/QOrderedType.v68
1 files changed, 68 insertions, 0 deletions
diff --git a/theories/QArith/QOrderedType.v b/theories/QArith/QOrderedType.v
new file mode 100644
index 000000000..3d83eac38
--- /dev/null
+++ b/theories/QArith/QOrderedType.v
@@ -0,0 +1,68 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import QArith_base
+ DecidableType2 OrderedType2 OrderedType2Facts.
+
+Local Open Scope Q_scope.
+
+(** * DecidableType structure for rational numbers *)
+
+Module Q_as_DT <: DecidableType.
+ Definition t := Q.
+ Definition eq := Qeq.
+ Instance eq_equiv : Equivalence Qeq.
+ Definition eq_dec := Qeq_dec.
+
+ (** The next fields are not mandatory, but allow [Q_as_DT] to be
+ also a [DecidableTypeOrig]. *)
+ Definition eq_refl := Qeq_refl.
+ Definition eq_sym := Qeq_sym.
+ Definition eq_trans := eq_trans.
+End Q_as_DT.
+
+
+(** * OrderedType structure for rational numbers *)
+
+Module Q_as_OT <: OrderedTypeFull.
+ Include Q_as_DT.
+ Definition lt := Qlt.
+ Definition le := Qle.
+ Definition compare := Qcompare.
+
+ Instance lt_strorder : StrictOrder Qlt.
+ Proof. split; [ exact Qlt_irrefl | exact Qlt_trans ]. Qed.
+
+ Instance lt_compat : Proper (Qeq==>Qeq==>iff) Qlt.
+ Proof. auto with *. Qed.
+
+ Definition le_lteq := Qle_lteq.
+
+ Lemma Qcompare_antisym : forall x y, CompOpp (x ?= y) = (y ?= x).
+ Proof.
+ unfold "?=". intros. apply Zcompare_antisym.
+ Qed.
+
+ Lemma compare_spec : forall x y, Cmp Qeq Qlt x y (Qcompare x y).
+ Proof.
+ intros.
+ destruct (x ?= y) as [ ]_eqn:H; constructor; auto.
+ rewrite Qeq_alt; auto.
+ rewrite Qlt_alt, <- Qcompare_antisym, H; auto.
+ Qed.
+
+End Q_as_OT.
+
+
+(** * An [order] tactic for [Q] numbers *)
+
+Module QOrder := OTF_to_OrderTac Q_as_OT.
+Ltac q_order := QOrder.order.
+
+(** Note that [q_order] is domain-agnostic: it will not prove
+ [1<=2] or [x<=x+x], but rather things like [x<=y -> y<=x -> x==y]. *)