aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--src/Algebra.v8
-rw-r--r--src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v2
-rw-r--r--src/Util/Decidable.v29
3 files changed, 24 insertions, 15 deletions
diff --git a/src/Algebra.v b/src/Algebra.v
index 15e0bcd38..e45ab8c33 100644
--- a/src/Algebra.v
+++ b/src/Algebra.v
@@ -1,12 +1,18 @@
Require Import Coq.Classes.Morphisms. Require Coq.Setoids.Setoid.
Require Import Crypto.Util.Tactics Crypto.Tactics.Nsatz.
+Require Import Crypto.Util.Decidable.
Local Close Scope nat_scope. Local Close Scope type_scope. Local Close Scope core_scope.
+Notation is_eq_dec := (DecidableRel _) (only parsing).
+Notation "@ 'is_eq_dec' T R" := (DecidableRel (R:T->T->Prop))
+ (at level 10, T at level 8, R at level 8, only parsing).
+Notation eq_dec x y := (@dec (_ x y) _) (only parsing).
+
Section Algebra.
Context {T:Type} {eq:T->T->Prop}.
Local Infix "=" := eq : type_scope. Local Notation "a <> b" := (not (a = b)) : type_scope.
- Class is_eq_dec := { eq_dec : forall x y : T, {x=y} + {x<>y} }.
+ Local Notation is_eq_dec := (@is_eq_dec T eq).
Section SingleOperation.
Context {op:T->T->T}.
diff --git a/src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v b/src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v
index 476592b36..e3809bb8a 100644
--- a/src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v
+++ b/src/CompleteEdwardsCurve/CompleteEdwardsCurveTheorems.v
@@ -119,7 +119,7 @@ Module E.
| |- _ => split
| |- Feq _ _ => field_algebra
| |- _ <> 0 => expand_opp; solve [nsatz_nonzero|eauto 6]
- | |- {_}+{_} => eauto 15 using decide_and, @eq_dec with typeclass_instances
+ | |- Decidable.Decidable _ => solve [ typeclasses eauto ]
end.
Ltac bash := repeat bash_step.
diff --git a/src/Util/Decidable.v b/src/Util/Decidable.v
index 726b52b6b..9ab05699a 100644
--- a/src/Util/Decidable.v
+++ b/src/Util/Decidable.v
@@ -32,24 +32,24 @@ Local Ltac solve_decidable_transparent := solve_decidable_transparent_with first
Local Hint Extern 0 => solve [ solve_decidable_transparent ] : typeclass_instances.
-Global Instance dec_True : Decidable True := left I.
-Global Instance dec_False : Decidable False := right (fun x => x).
-Global Instance dec_or {A B} `{Decidable A, Decidable B} : Decidable (A \/ B). exact _. Defined.
-Global Instance dec_and {A B} `{Decidable A, Decidable B} : Decidable (A /\ B). exact _. Defined.
-Global Instance dec_impl {A B} `{Decidable (B \/ ~A)} : Decidable (A -> B) | 3. exact _. Defined.
-Global Instance dec_impl_simple {A B} `{Decidable A, Decidable B} : Decidable (A -> B). exact _. Defined.
-Global Instance dec_iff {A B} `{Decidable A, Decidable B} : Decidable (A <-> B). exact _. Defined.
+Global Instance dec_True : Decidable True | 10 := left I.
+Global Instance dec_False : Decidable False | 10 := right (fun x => x).
+Global Instance dec_or {A B} `{Decidable A, Decidable B} : Decidable (A \/ B) | 10. exact _. Defined.
+Global Instance dec_and {A B} `{Decidable A, Decidable B} : Decidable (A /\ B) | 10. exact _. Defined.
+Global Instance dec_impl {A B} `{Decidable (B \/ ~A)} : Decidable (A -> B) | 10. exact _. Defined.
+Global Instance dec_impl_simple {A B} `{Decidable A, Decidable B} : Decidable (A -> B) | 10. exact _. Defined.
+Global Instance dec_iff {A B} `{Decidable A, Decidable B} : Decidable (A <-> B) | 10. exact _. Defined.
Lemma dec_not {A} `{Decidable A} : Decidable (~A).
Proof. solve_decidable_transparent. Defined.
(** Disallow infinite loops of dec_not *)
Hint Extern 0 (Decidable (~?A)) => apply (@dec_not A) : typeclass_instances.
-Global Instance dec_eq_unit : DecidableRel (@eq unit). exact _. Defined.
-Global Instance dec_eq_bool : DecidableRel (@eq bool). exact _. Defined.
-Global Instance dec_eq_Empty_set : DecidableRel (@eq Empty_set). exact _. Defined.
-Global Instance dec_eq_nat : DecidableRel (@eq nat). exact _. Defined.
-Global Instance dec_eq_prod {A B} `{DecidableRel (@eq A), DecidableRel (@eq B)} : DecidableRel (@eq (A * B)). exact _. Defined.
-Global Instance dec_eq_sum {A B} `{DecidableRel (@eq A), DecidableRel (@eq B)} : DecidableRel (@eq (A + B)). exact _. Defined.
+Global Instance dec_eq_unit : DecidableRel (@eq unit) | 10. exact _. Defined.
+Global Instance dec_eq_bool : DecidableRel (@eq bool) | 10. exact _. Defined.
+Global Instance dec_eq_Empty_set : DecidableRel (@eq Empty_set) | 10. exact _. Defined.
+Global Instance dec_eq_nat : DecidableRel (@eq nat) | 10. exact _. Defined.
+Global Instance dec_eq_prod {A B} `{DecidableRel (@eq A), DecidableRel (@eq B)} : DecidableRel (@eq (A * B)) | 10. exact _. Defined.
+Global Instance dec_eq_sum {A B} `{DecidableRel (@eq A), DecidableRel (@eq B)} : DecidableRel (@eq (A + B)) | 10. exact _. Defined.
Lemma Decidable_respects_iff A B (H : A <-> B) : (Decidable A -> Decidable B) * (Decidable B -> Decidable A).
Proof. solve_decidable_transparent. Defined.
@@ -59,3 +59,6 @@ Proof. solve_decidable_transparent. Defined.
Lemma Decidable_iff_to_flip_impl A B (H : A <-> B) : Decidable B -> Decidable A.
Proof. solve_decidable_transparent. Defined.
+
+(** For dubious compatibility with [eauto using]. *)
+Hint Extern 2 (Decidable _) => progress unfold Decidable : typeclass_instances core.