diff --git a/theories/base.v b/theories/base.v
index 1a1473881e596470b363d140b0b68df4ff4a1f14..c3009cd38339f6209649e1f7edb131b7e6cbc4f9 100644
--- a/theories/base.v
+++ b/theories/base.v
@@ -150,13 +150,30 @@ Hint Extern 0 (_ ≡ _) => symmetry; assumption.
 (** * Type classes *)
 (** ** Decidable propositions *)
 (** This type class by (Spitters/van der Weegen, 2011) collects decidable
-propositions. For example to declare a parameter expressing decidable equality
-on a type [A] we write [`{∀ x y : A, Decision (x = y)}] and use it by writing
-[decide (x = y)]. *)
+propositions. *)
 Class Decision (P : Prop) := decide : {P} + {¬P}.
 Hint Mode Decision ! : typeclass_instances.
-Arguments decide _ {_} : assert.
-Notation EqDecision A := (∀ x y : A, Decision (x = y)).
+Arguments decide _ {_} : simpl never, assert.
+
+(** Although [RelDecision R] is just [∀ x y, Decision (R x y)], we make this
+an explicit class instead of a notation for two reasons:
+
+- It allows us to control [Hint Mode] more precisely. In particular, if it were
+  defined as a notation, the above [Hint Mode] for [Decision] would not prevent
+  diverging instance search when looking for [RelDecision (@eq ?A)], which would
+  result in it looking for [Decision (@eq ?A x y)], i.e. an instance where the
+  head position of [Decision] is not en evar.
+- We use it to avoid inefficient computation due to eager evaluation of
+  propositions by [vm_compute]. This inefficiency arises for example if
+  [(x = y) := (f x = f y)]. Since [decide (x = y)] evaluates to
+  [decide (f x = f y)], this would then lead to evaluation of [f x] and [f y].
+  Using the [RelDecision], the [f] is hidden under a lambda, which prevents
+  unnecessary evaluation. *)
+Class RelDecision {A B} (R : A → B → Prop) :=
+  decide_rel x y :> Decision (R x y).
+Hint Mode RelDecision ! ! ! : typeclass_instances.
+Arguments decide_rel {_ _} _ {_} _ _ : simpl never, assert.
+Notation EqDecision A := (RelDecision (@eq A)).
 
 (** ** Inhabited types *)
 (** This type class collects types that are inhabited. *)
diff --git a/theories/bset.v b/theories/bset.v
index f3f0d828c0d9e5b1c4d9c0055689d519c81e7166..1c17ee660e15cf56fe6600088f18726395718c7c 100644
--- a/theories/bset.v
+++ b/theories/bset.v
@@ -29,7 +29,8 @@ Proof.
   - intros X Y x; unfold elem_of, bset_elem_of; simpl. 
     destruct (bset_car X x), (bset_car Y x); simpl; tauto.
 Qed.
-Instance bset_elem_of_dec {A} x (X : bset A) : Decision (x ∈ X) := _.
+Instance bset_elem_of_dec {A} : RelDecision (@elem_of _ (bset A) _).
+Proof. refine (λ x X, cast_if (decide (bset_car X x))); done. Defined.
 
 Typeclasses Opaque bset_elem_of.
 Global Opaque bset_empty bset_singleton bset_union
diff --git a/theories/coPset.v b/theories/coPset.v
index 4af9f0233e92a20b4d420be0b68ec9d769b87a0d..6110295d5b4a4c081643facae3c62d5592a6415c 100644
--- a/theories/coPset.v
+++ b/theories/coPset.v
@@ -190,17 +190,18 @@ Proof.
   intros p; apply eq_bool_prop_intro, (HXY p).
 Qed.
 
-Instance coPset_elem_of_dec (p : positive) (X : coPset) : Decision (p ∈ X) := _.
-Instance coPset_equiv_dec (X Y : coPset) : Decision (X ≡ Y).
-Proof. refine (cast_if (decide (X = Y))); abstract (by fold_leibniz). Defined.
-Instance mapset_disjoint_dec (X Y : coPset) : Decision (X ⊥ Y).
+Instance coPset_elem_of_dec : RelDecision (@elem_of _ coPset _).
+Proof. solve_decision. Defined.
+Instance coPset_equiv_dec : RelDecision (@equiv coPset _).
+Proof. refine (λ X Y, cast_if (decide (X = Y))); abstract (by fold_leibniz). Defined.
+Instance mapset_disjoint_dec : RelDecision (@disjoint coPset _).
 Proof.
- refine (cast_if (decide (X ∩ Y = ∅)));
+ refine (λ X Y, cast_if (decide (X ∩ Y = ∅)));
   abstract (by rewrite disjoint_intersection_L).
 Defined.
-Instance mapset_subseteq_dec (X Y : coPset) : Decision (X ⊆ Y).
+Instance mapset_subseteq_dec : RelDecision (@subseteq coPset _).
 Proof.
- refine (cast_if (decide (X ∪ Y = Y))); abstract (by rewrite subseteq_union_L).
+ refine (λ X Y, cast_if (decide (X ∪ Y = Y))); abstract (by rewrite subseteq_union_L).
 Defined.
 
 (** * Top *)
diff --git a/theories/collections.v b/theories/collections.v
index 134af648961b502059787fe706048e0cb79dcfd7..00522a5cd3be4e1bdb9f97e28edd344b04908e3b 100644
--- a/theories/collections.v
+++ b/theories/collections.v
@@ -527,7 +527,7 @@ Section simple_collection.
   End leibniz.
 
   Section dec.
-    Context `{∀ (X Y : C), Decision (X ≡ Y)}.
+    Context `{!RelDecision (@equiv C _)}.
     Lemma collection_subseteq_inv X Y : X ⊆ Y → X ⊂ Y ∨ X ≡ Y.
     Proof. destruct (decide (X ≡ Y)); [by right|left;set_solver]. Qed.
     Lemma collection_not_subset_inv X Y : X ⊄ Y → X ⊈ Y ∨ X ≡ Y.
@@ -692,7 +692,7 @@ Section collection.
   End leibniz.
 
   Section dec.
-    Context `{∀ (x : A) (X : C), Decision (x ∈ X)}.
+    Context `{!RelDecision (@elem_of A C _)}.
     Lemma not_elem_of_intersection x X Y : x ∉ X ∩ Y ↔ x ∉ X ∨ x ∉ Y.
     Proof. rewrite elem_of_intersection. destruct (decide (x ∈ X)); tauto. Qed.
     Lemma not_elem_of_difference x X Y : x ∉ X ∖ Y ↔ x ∉ X ∨ x ∈ Y.
diff --git a/theories/countable.v b/theories/countable.v
index 4c2f6f03bd900780900c55b91efeaf2700b69c20..de849d78142f8db75187ec29e084a175304fdb10 100644
--- a/theories/countable.v
+++ b/theories/countable.v
@@ -10,6 +10,7 @@ Class Countable A `{EqDecision A} := {
   decode : positive → option A;
   decode_encode x : decode (encode x) = Some x
 }.
+Hint Mode Countable ! - : typeclass_instances.
 Arguments encode : simpl never.
 Arguments decode : simpl never.
 
@@ -36,8 +37,8 @@ Section choice.
   Context `{Countable A} (P : A → Prop).
 
   Inductive choose_step: relation positive :=
-    | choose_step_None {p} : decode p = None → choose_step (Psucc p) p
-    | choose_step_Some {p x} :
+    | choose_step_None {p} : decode (A:=A) p = None → choose_step (Psucc p) p
+    | choose_step_Some {p} {x : A} :
        decode p = Some x → ¬P x → choose_step (Psucc p) p.
   Lemma choose_step_acc : (∃ x, P x) → Acc choose_step 1%positive.
   Proof.
@@ -320,7 +321,7 @@ Arguments GenNode {_} _ _ : assert.
 Instance gen_tree_dec `{EqDecision T} : EqDecision (gen_tree T).
 Proof.
  refine (
-  fix go t1 t2 :=
+  fix go t1 t2 := let _ : EqDecision _ := @go in
   match t1, t2 with
   | GenLeaf x1, GenLeaf x2 => cast_if (decide (x1 = x2))
   | GenNode n1 ts1, GenNode n2 ts2 =>
diff --git a/theories/decidable.v b/theories/decidable.v
index ddbac7e67db2245d50589259d8d76cffdb92c3b8..cd2ad8d4a457d88635f2ffc9e3b79f546e2838de 100644
--- a/theories/decidable.v
+++ b/theories/decidable.v
@@ -6,8 +6,6 @@ type class. *)
 From stdpp Require Export proof_irrel.
 Set Default Proof Using "Type*".
 
-Hint Extern 200 (Decision _) => progress (lazy beta) : typeclass_instances.
-
 Lemma dec_stable `{Decision P} : ¬¬P → P.
 Proof. firstorder. Qed.
 
@@ -16,17 +14,6 @@ Proof. destruct b. left; constructor. right. intros []. Qed.
 Instance: Inj (=) (↔) Is_true.
 Proof. intros [] []; simpl; intuition. Qed.
 
-(** We introduce [decide_rel] to avoid inefficienct computation due to eager
-evaluation of propositions by [vm_compute]. This inefficiency occurs if
-[(x = y) := (f x = f y)] as [decide (x = y)] evaluates to [decide (f x = f y)]
-which then might lead to evaluation of [f x] and [f y]. Using [decide_rel]
-we hide [f] under a lambda abstraction to avoid this unnecessary evaluation. *)
-Definition decide_rel {A B} (R : A → B → Prop) {dec : ∀ x y, Decision (R x y)}
-  (x : A) (y : B) : Decision (R x y) := dec x y.
-Lemma decide_rel_correct {A B} (R : A → B → Prop) `{∀ x y, Decision (R x y)}
-  (x : A) (y : B) : decide_rel R x y = decide (R x y).
-Proof. reflexivity. Qed.
-
 Lemma decide_True {A} `{Decision P} (x y : A) :
   P → (if decide P then x else y) = x.
 Proof. destruct (decide P); tauto. Qed.
@@ -75,9 +62,10 @@ Ltac solve_trivial_decision :=
   | |- Decision (?P) => apply _
   | |- sumbool ?P (¬?P) => change (Decision P); apply _
   end.
-Ltac solve_decision := intros; first
-  [ solve_trivial_decision
-  | unfold Decision; decide equality; solve_trivial_decision ].
+Ltac solve_decision :=
+  unfold EqDecision; intros; first
+    [ solve_trivial_decision
+    | unfold Decision; decide equality; solve_trivial_decision ].
 
 (** The following combinators are useful to create Decision proofs in
 combination with the [refine] tactic. *)
diff --git a/theories/fin_collections.v b/theories/fin_collections.v
index ddb1220c05c4c63ee3f8eab8b0905dfcc5c7a040..15f04654fa64f60753b6f50aa269e8524d5fdac2 100644
--- a/theories/fin_collections.v
+++ b/theories/fin_collections.v
@@ -23,9 +23,9 @@ Implicit Types X Y : C.
 Lemma fin_collection_finite X : set_finite X.
 Proof. by exists (elements X); intros; rewrite elem_of_elements. Qed.
 
-Instance elem_of_dec_slow (x : A) (X : C) : Decision (x ∈ X) | 100.
+Instance elem_of_dec_slow : RelDecision (@elem_of A C _) | 100.
 Proof.
-  refine (cast_if (decide_rel (∈) x (elements X)));
+  refine (λ x X, cast_if (decide_rel (∈) x (elements X)));
     by rewrite <-(elem_of_elements _).
 Defined.
 
diff --git a/theories/fin_maps.v b/theories/fin_maps.v
index fcaba6ce290f1c215813c8dbad6536d1a3eb3b6c..6e6d1ddce9a785296f500d7d41f8cb1659e44bdb 100644
--- a/theories/fin_maps.v
+++ b/theories/fin_maps.v
@@ -1238,10 +1238,9 @@ Proof.
     destruct (m1 !! i), (m2 !! i); simplify_eq/=; auto;
       by eapply bool_decide_unpack, Hm.
 Qed.
-Global Instance map_relation_dec `{∀ x y, Decision (R x y), ∀ x, Decision (P x),
-  ∀ y, Decision (Q y)} (m1 : M A) (m2 : M B) : Decision (map_relation R P Q m1 m2).
+Global Instance map_relation_dec : RelDecision (map_relation (M:=M) R P Q).
 Proof.
-  refine (cast_if (decide (map_Forall (λ _, Is_true) (merge f m1 m2))));
+  refine (λ m1 m2, cast_if (decide (map_Forall (λ _, Is_true) (merge f m1 m2))));
     abstract by rewrite map_relation_alt.
 Defined.
 (** Due to the finiteness of finite maps, we can extract a witness if the
diff --git a/theories/finite.v b/theories/finite.v
index fef53c3b759164e0c9ab6dcf934b573a09cef9b6..aba0ae9267dcf7b63379cf6d7fcd9c1d1aa8e3e1 100644
--- a/theories/finite.v
+++ b/theories/finite.v
@@ -8,6 +8,7 @@ Class Finite A `{EqDecision A} := {
   NoDup_enum : NoDup enum;
   elem_of_enum x : x ∈ enum
 }.
+Hint Mode Finite ! - : typeclass_instances.
 Arguments enum : clear implicits.
 Arguments enum _ {_ _} : assert.
 Arguments NoDup_enum : clear implicits.
diff --git a/theories/gmultiset.v b/theories/gmultiset.v
index f61416e9627f0953bcdbdb3f6b78942a2991b04b..8d84b49ef7a741ef98705482e66207d9819f7b98 100644
--- a/theories/gmultiset.v
+++ b/theories/gmultiset.v
@@ -7,18 +7,14 @@ Record gmultiset A `{Countable A} := GMultiSet { gmultiset_car : gmap A nat }.
 Arguments GMultiSet {_ _ _} _ : assert.
 Arguments gmultiset_car {_ _ _} _ : assert.
 
-Lemma gmultiset_eq_dec `{Countable A} : EqDecision (gmultiset A).
+Instance gmultiset_eq_dec `{Countable A} : EqDecision (gmultiset A).
 Proof. solve_decision. Defined.
-Hint Extern 1 (Decision (@eq (gmultiset _) _ _)) =>
-  eapply @gmultiset_eq_dec : typeclass_instances.
 
-Program Definition gmultiset_countable `{Countable A} :
+Program Instance gmultiset_countable `{Countable A} :
     Countable (gmultiset A) := {|
   encode X := encode (gmultiset_car X);  decode p := GMultiSet <$> decode p
 |}.
 Next Obligation. intros A ?? [X]; simpl. by rewrite decode_encode. Qed.
-Hint Extern 1 (Countable (gmultiset _)) =>
-  eapply @gmultiset_countable : typeclass_instances.
 
 Section definitions.
   Context `{Countable A}.
@@ -102,8 +98,8 @@ Proof.
       by split; auto with lia.
   - intros X Y x. rewrite !elem_of_multiplicity, multiplicity_union. omega.
 Qed.
-Global Instance gmultiset_elem_of_dec x X : Decision (x ∈ X).
-Proof. unfold elem_of, gmultiset_elem_of. apply _. Defined.
+Global Instance gmultiset_elem_of_dec : RelDecision (@elem_of _ (gmultiset A) _).
+Proof. refine (λ x X, cast_if (decide (0 < multiplicity x X))); done. Defined.
 
 (* Algebraic laws *)
 Global Instance gmultiset_comm : Comm (@eq (gmultiset A)) (∪).
@@ -144,8 +140,9 @@ Qed.
 Lemma gmultiset_elements_empty_inv X : elements X = [] → X = ∅.
 Proof.
   destruct X as [X]; unfold elements, gmultiset_elements; simpl.
-  intros; apply (f_equal GMultiSet). destruct (map_to_list X)
-    as [|[]] eqn:?; naive_solver eauto using map_to_list_empty_inv.
+  intros; apply (f_equal GMultiSet). destruct (map_to_list X) as [|[]] eqn:?.
+  - by apply map_to_list_empty_inv.
+  - naive_solver.
 Qed.
 Lemma gmultiset_elements_empty' X : elements X = [] ↔ X = ∅.
 Proof.
@@ -249,9 +246,9 @@ Proof.
   apply forall_proper; intros x. unfold multiplicity.
   destruct (gmultiset_car X !! x), (gmultiset_car Y !! x); naive_solver omega.
 Qed.
-Global Instance gmultiset_subseteq_dec X Y : Decision (X ⊆ Y).
+Global Instance gmultiset_subseteq_dec : RelDecision (@subseteq (gmultiset A) _).
 Proof.
- refine (cast_if (decide (map_relation (≤)
+ refine (λ X Y, cast_if (decide (map_relation (≤)
    (λ _, False) (λ _, True) (gmultiset_car X) (gmultiset_car Y))));
    by rewrite gmultiset_subseteq_alt.
 Defined.
diff --git a/theories/list.v b/theories/list.v
index a258bea8e2959d1c4d7010aeb01aa51a708dcef4..d29ab1c69ba78ece3fb79bdfb569dbb49037774b 100644
--- a/theories/list.v
+++ b/theories/list.v
@@ -311,13 +311,13 @@ Instance list_subseteq {A} : SubsetEq (list A) := λ l1 l2, ∀ x, x ∈ l1 →
 
 Section list_set.
   Context `{dec : EqDecision A}.
-  Global Instance elem_of_list_dec (x : A) : ∀ l : list A, Decision (x ∈ l).
+  Global Instance elem_of_list_dec : RelDecision (@elem_of A (list A) _).
   Proof.
    refine (
-    fix go l :=
+    fix go x l :=
     match l return Decision (x ∈ l) with
     | [] => right _
-    | y :: l => cast_if_or (decide (x = y)) (go l)
+    | y :: l => cast_if_or (decide (x = y)) (go x l)
     end); clear go dec; subst; try (by constructor); abstract by inversion 1.
   Defined.
   Fixpoint remove_dups (l : list A) : list A :=
@@ -1505,9 +1505,9 @@ Proof.
   - intros ?. by eexists [].
   - intros ???[k1->] [k2->]. exists (k2 ++ k1). by rewrite (assoc_L (++)).
 Qed.
-Global Instance prefix_dec `{!EqDecision A} : ∀ l1 l2,
-    Decision (l1 `prefix_of` l2) := fix go l1 l2 :=
-  match l1, l2 return { l1 `prefix_of` l2 } + { ¬l1 `prefix_of` l2 } with
+Global Instance prefix_dec `{!EqDecision A} : RelDecision prefix :=
+  fix go l1 l2 :=
+  match l1, l2  with
   | [], _ => left (prefix_nil _)
   | _, [] => right (prefix_nil_not _ _)
   | x :: l1, y :: l2 =>
@@ -1639,10 +1639,9 @@ Lemma suffix_length l1 l2 : l1 `suffix_of` l2 → length l1 ≤ length l2.
 Proof. intros [? ->]. rewrite app_length. lia. Qed.
 Lemma suffix_cons_not x l : ¬x :: l `suffix_of` l.
 Proof. intros [??]. discriminate_list. Qed.
-Global Instance suffix_dec `{!EqDecision A} l1 l2 :
-  Decision (l1 `suffix_of` l2).
+Global Instance suffix_dec `{!EqDecision A} : RelDecision (@suffix A).
 Proof.
-  refine (cast_if (decide_rel prefix (reverse l1) (reverse l2)));
+  refine (λ l1 l2, cast_if (decide_rel prefix (reverse l1) (reverse l2)));
    abstract (by rewrite suffix_prefix_reverse).
 Defined.
 
@@ -2087,14 +2086,14 @@ Section submseteq_dec.
       destruct (list_remove x l2) as [k'|] eqn:?; intros; simplify_eq.
       rewrite submseteq_cons_l. eauto using list_remove_Some.
   Qed.
-  Global Instance submseteq_dec l1 l2 : Decision (l1 ⊆+ l2).
+  Global Instance submseteq_dec : RelDecision (submseteq : relation (list A)).
   Proof.
-   refine (cast_if (decide (is_Some (list_remove_list l1 l2))));
+   refine (λ l1 l2, cast_if (decide (is_Some (list_remove_list l1 l2))));
     abstract (rewrite list_remove_list_submseteq; tauto).
   Defined.
-  Global Instance Permutation_dec l1 l2 : Decision (l1 ≡ₚ l2).
+  Global Instance Permutation_dec : RelDecision (Permutation : relation (list A)).
   Proof.
-   refine (cast_if_and
+   refine (λ l1 l2, cast_if_and
     (decide (length l1 = length l2)) (decide (l1 ⊆+ l2)));
     abstract (rewrite Permutation_alt; tauto).
   Defined.
@@ -2621,7 +2620,7 @@ Section Forall2.
   Qed.
 
   Global Instance Forall2_dec `{dec : ∀ x y, Decision (P x y)} :
-    ∀ l k, Decision (Forall2 P l k).
+    RelDecision (Forall2 P).
   Proof.
    refine (
     fix go l k : Decision (Forall2 P l k) :=
diff --git a/theories/mapset.v b/theories/mapset.v
index bb7c479d8678dcdadcadaa54e83f1eafbf133148..6e9f72cbc2c50ecc4e7bea1a79576ae700ce11a9 100644
--- a/theories/mapset.v
+++ b/theories/mapset.v
@@ -76,18 +76,18 @@ Section deciders.
     match X1, X2 with Mapset m1, Mapset m2 => cast_if (decide (m1 = m2)) end);
     abstract congruence.
   Defined.
-  Global Instance mapset_equiv_dec (X1 X2 : mapset M) : Decision (X1 ≡ X2) | 1.
-  Proof. refine (cast_if (decide (X1 = X2))); abstract (by fold_leibniz). Defined.
-  Global Instance mapset_elem_of_dec x (X : mapset M) : Decision (x ∈ X) | 1.
-  Proof. solve_decision. Defined.
-  Global Instance mapset_disjoint_dec (X1 X2 : mapset M) : Decision (X1 ⊥ X2).
+  Global Instance mapset_equiv_dec : RelDecision (@equiv (mapset M)_) | 1.
+  Proof. refine (λ X1 X2, cast_if (decide (X1 = X2))); abstract (by fold_leibniz). Defined.
+  Global Instance mapset_elem_of_dec : RelDecision (@elem_of _ (mapset M) _) | 1.
+  Proof. refine (λ x X, cast_if (decide (mapset_car X !! x = Some ()))); done. Defined.
+  Global Instance mapset_disjoint_dec : RelDecision (@disjoint (mapset M) _).
   Proof.
-   refine (cast_if (decide (X1 ∩ X2 = ∅)));
+   refine (λ X1 X2, cast_if (decide (X1 ∩ X2 = ∅)));
     abstract (by rewrite disjoint_intersection_L).
   Defined.
-  Global Instance mapset_subseteq_dec (X1 X2 : mapset M) : Decision (X1 ⊆ X2).
+  Global Instance mapset_subseteq_dec : RelDecision (@subseteq (mapset M) _).
   Proof.
-   refine (cast_if (decide (X1 ∪ X2 = X2)));
+   refine (λ X1 X2, cast_if (decide (X1 ∪ X2 = X2)));
     abstract (by rewrite subseteq_union_L).
   Defined.
 End deciders.
diff --git a/theories/numbers.v b/theories/numbers.v
index c451c91d80884a902c5b6063cf852edb67a8e029..029154f2f7ab4c48e41e47e5201831eaa0d3920f 100644
--- a/theories/numbers.v
+++ b/theories/numbers.v
@@ -36,8 +36,8 @@ Infix "`max`" := Nat.max (at level 35) : nat_scope.
 Infix "`min`" := Nat.min (at level 35) : nat_scope.
 
 Instance nat_eq_dec: EqDecision nat := eq_nat_dec.
-Instance nat_le_dec: ∀ x y : nat, Decision (x ≤ y) := le_dec.
-Instance nat_lt_dec: ∀ x y : nat, Decision (x < y) := lt_dec.
+Instance nat_le_dec: RelDecision le := le_dec.
+Instance nat_lt_dec: RelDecision lt := lt_dec.
 Instance nat_inhabited: Inhabited nat := populate 0%nat.
 Instance S_inj: Inj (=) (=) S.
 Proof. by injection 1. Qed.
@@ -77,9 +77,9 @@ Proof. intros. destruct (Nat_mul_split_l n x2 x1 y2 y1); auto with lia. Qed.
 Notation lcm := Nat.lcm.
 Notation divide := Nat.divide.
 Notation "( x | y )" := (divide x y) : nat_scope.
-Instance Nat_divide_dec x y : Decision (x | y).
+Instance Nat_divide_dec : RelDecision Nat.divide.
 Proof.
-  refine (cast_if (decide (lcm x y = y))); by rewrite Nat.divide_lcm_iff.
+  refine (λ x y, cast_if (decide (lcm x y = y))); by rewrite Nat.divide_lcm_iff.
 Defined.
 Instance: PartialOrder divide.
 Proof.
@@ -130,6 +130,10 @@ Arguments Pos.of_nat : simpl never.
 Arguments Pmult : simpl never.
 
 Instance positive_eq_dec: EqDecision positive := Pos.eq_dec.
+Instance positive_le_dec: RelDecision Pos.le.
+Proof. refine (λ x y, decide ((x ?= y) ≠ Gt)). Defined.
+Instance positive_lt_dec: RelDecision Pos.lt.
+Proof. refine (λ x y, decide ((x ?= y) = Lt)). Defined.
 Instance positive_inhabited: Inhabited positive := populate 1.
 
 Instance maybe_xO : Maybe xO := λ p, match p with p~0 => Some p | _ => None end.
@@ -218,10 +222,10 @@ Instance Npos_inj : Inj (=) (=) Npos.
 Proof. by injection 1. Qed.
 
 Instance N_eq_dec: EqDecision N := N.eq_dec.
-Program Instance N_le_dec (x y : N) : Decision (x ≤ y)%N :=
+Program Instance N_le_dec : RelDecision N.le := λ x y,
   match Ncompare x y with Gt => right _ | _ => left _ end.
 Solve Obligations with naive_solver.
-Program Instance N_lt_dec (x y : N) : Decision (x < y)%N :=
+Program Instance N_lt_dec : RelDecision N.lt := λ x y,
   match Ncompare x y with Lt => left _ | _ => right _ end.
 Solve Obligations with naive_solver.
 Instance N_inhabited: Inhabited N := populate 1%N.
@@ -259,8 +263,8 @@ Instance Z_of_nat_inj : Inj (=) (=) Z.of_nat.
 Proof. intros n1 n2. apply Nat2Z.inj. Qed.
 
 Instance Z_eq_dec: EqDecision Z := Z.eq_dec.
-Instance Z_le_dec: ∀ x y : Z, Decision (x ≤ y) := Z_le_dec.
-Instance Z_lt_dec: ∀ x y : Z, Decision (x < y) := Z_lt_dec.
+Instance Z_le_dec: RelDecision Z.le := Z_le_dec.
+Instance Z_lt_dec: RelDecision Z.lt := Z_lt_dec.
 Instance Z_inhabited: Inhabited Z := populate 1.
 Instance Z_le_po : PartialOrder (≤).
 Proof.
@@ -365,11 +369,11 @@ Hint Extern 1 (_ ≤ _) => reflexivity || discriminate.
 Arguments Qred : simpl never.
 
 Instance Qc_eq_dec: EqDecision Qc := Qc_eq_dec.
-Program Instance Qc_le_dec (x y : Qc) : Decision (x ≤ y) :=
+Program Instance Qc_le_dec: RelDecision Qcle := λ x y,
   if Qclt_le_dec y x then right _ else left _.
 Next Obligation. intros x y; apply Qclt_not_le. Qed.
 Next Obligation. done. Qed.
-Program Instance Qc_lt_dec (x y : Qc) : Decision (x < y) :=
+Program Instance Qc_lt_dec: RelDecision Qclt := λ x y,
   if Qclt_le_dec x y then left _ else right _.
 Solve Obligations with done.
 Next Obligation. intros x y; apply Qcle_not_lt. Qed.
diff --git a/theories/orders.v b/theories/orders.v
index 9fe161a524f046ccfbe5445d1b90ecfa89dd3ec7..b7fc6948705a38ea62ba103046eb8170db1cc500 100644
--- a/theories/orders.v
+++ b/theories/orders.v
@@ -41,15 +41,16 @@ Section orders.
     split; try apply _.
     eauto using strict_transitive_r, strict_include.
   Qed.
-  Global Instance preorder_subset_dec_slow `{∀ X Y, Decision (X ⊆ Y)}
-    (X Y : A) : Decision (X ⊂ Y) | 100 := _.
+  Global Instance preorder_subset_dec_slow `{!RelDecision R} :
+    RelDecision (strict R) | 100.
+  Proof. solve_decision. Defined.
   Lemma strict_spec_alt `{!AntiSymm (=) R} X Y : X ⊂ Y ↔ X ⊆ Y ∧ X ≠ Y.
   Proof.
     split.
     - intros [? HYX]. split. done. by intros <-.
     - intros [? HXY]. split. done. by contradict HXY; apply (anti_symm R).
   Qed.
-  Lemma po_eq_dec `{!PartialOrder R, ∀ X Y, Decision (X ⊆ Y)} : EqDecision A.
+  Lemma po_eq_dec `{!PartialOrder R, !RelDecision R} : EqDecision A.
   Proof.
     refine (λ X Y, cast_if_and (decide (X ⊆ Y)) (decide (Y ⊆ X)));
      abstract (rewrite anti_symm_iff; tauto).
@@ -76,8 +77,8 @@ Section strict_orders.
   Lemma strict_anti_symm `{!StrictOrder R} X Y :
     X ⊂ Y → Y ⊂ X → False.
   Proof. intros. apply (irreflexivity R X). by trans Y. Qed.
-  Global Instance trichotomyT_dec `{!TrichotomyT R, !StrictOrder R} X Y :
-      Decision (X ⊂ Y) :=
+  Global Instance trichotomyT_dec `{!TrichotomyT R, !StrictOrder R} :
+      RelDecision R := λ X Y,
     match trichotomyT R X Y with
     | inleft (left H) => left H
     | inleft (right H) => right (irreflexive_eq _ _ H)
diff --git a/theories/sorting.v b/theories/sorting.v
index ff33637e1e3695f1ca6c3eedad70e2320347c53f..9576fdbc145b8c9ea35ebf23d01b49c8072e34d0 100644
--- a/theories/sorting.v
+++ b/theories/sorting.v
@@ -15,7 +15,7 @@ Section merge_sort.
     | [], _ => l2
     | _, [] => l1
     | x1 :: l1, x2 :: l2 =>
-       if decide_rel R x1 x2 then x1 :: list_merge l1 (x2 :: l2)
+       if decide (R x1 x2) then x1 :: list_merge l1 (x2 :: l2)
        else x2 :: list_merge_aux l2
     end.
   Global Arguments list_merge !_ !_ / : assert.