aboutsummaryrefslogtreecommitdiffstats
path: root/lib/Ordered.v
diff options
context:
space:
mode:
authorVincent Laporte <vbgl@users.noreply.github.com>2019-10-02 19:23:51 +0200
committerXavier Leroy <xavierleroy@users.noreply.github.com>2019-10-02 19:23:51 +0200
commitb7374d225af55ecc6f5d6aa8f3684bfae99ff465 (patch)
tree75a9858d151984c2033c66fa171bed3305fa20ce /lib/Ordered.v
parentca5f8a7629a6e31cc287139ad0a69b8154514260 (diff)
downloadcompcert-kvx-b7374d225af55ecc6f5d6aa8f3684bfae99ff465.tar.gz
compcert-kvx-b7374d225af55ecc6f5d6aa8f3684bfae99ff465.zip
Make explicit the use of hints from OrderedType (#316)
Some hints will move from the core database to the `ordered_type` database (see https://github.com/coq/coq/pull/9772). This commit prepares for this move by adding `with ordered_type` to the invocations of `auto` and `eauto` that use the hints in question.
Diffstat (limited to 'lib/Ordered.v')
-rw-r--r--lib/Ordered.v10
1 files changed, 6 insertions, 4 deletions
diff --git a/lib/Ordered.v b/lib/Ordered.v
index bcf24cbd..1adbd330 100644
--- a/lib/Ordered.v
+++ b/lib/Ordered.v
@@ -21,6 +21,8 @@ Require Import Coqlib.
Require Import Maps.
Require Import Integers.
+Create HintDb ordered_type.
+
(** The ordered type of positive numbers *)
Module OrderedPositive <: OrderedType.
@@ -173,17 +175,17 @@ Definition eq (x y: t) :=
Lemma eq_refl : forall x : t, eq x x.
Proof.
- intros; split; auto.
+ intros; split; auto with ordered_type.
Qed.
Lemma eq_sym : forall x y : t, eq x y -> eq y x.
Proof.
- unfold eq; intros. intuition auto.
+ unfold eq; intros. intuition auto with ordered_type.
Qed.
Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
Proof.
- unfold eq; intros. intuition eauto.
+ unfold eq; intros. intuition eauto with ordered_type.
Qed.
Definition lt (x y: t) :=
@@ -201,7 +203,7 @@ Proof.
case (A.compare (fst x) (fst z)); intro.
assumption.
generalize (A.lt_not_eq H2); intro. elim H5.
- apply A.eq_trans with (fst z). auto. auto.
+ apply A.eq_trans with (fst z). auto. auto with ordered_type.
generalize (@A.lt_not_eq (fst z) (fst y)); intro.
elim H5. apply A.lt_trans with (fst x); auto.
apply A.eq_sym; auto.