diff --git a/theories/algebra/agree.v b/theories/algebra/agree.v
index 846d12bbf5a9871b8f45c5b1859d82d1b5c842ce..f1157730930fc20da533689ef3dffbe58d0ff432 100644
--- a/theories/algebra/agree.v
+++ b/theories/algebra/agree.v
@@ -149,7 +149,7 @@ Proof. rewrite /CMRATotal; eauto. Qed.
 Global Instance agree_persistent (x : agree A) : Persistent x.
 Proof. by constructor. Qed.
 
-Global Instance agree_discrete : Discrete A → CMRADiscrete agreeR.
+Global Instance agree_ofe_discrete : OFEDiscrete A → CMRADiscrete agreeR.
 Proof.
   intros HD. split.
   - intros x y [H H'] n; split=> a; setoid_rewrite <-(timeless_iff_0 _ _); auto.
diff --git a/theories/algebra/auth.v b/theories/algebra/auth.v
index 0ac73f70e4c73610c06d49fdb4a36f928667f0e5..cb65e14e8cc0a9caa5426a59b003d4cd9f7ff9e2 100644
--- a/theories/algebra/auth.v
+++ b/theories/algebra/auth.v
@@ -52,7 +52,7 @@ Qed.
 Global Instance Auth_timeless a b :
   Timeless a → Timeless b → Timeless (Auth a b).
 Proof. by intros ?? [??] [??]; split; apply: timeless. Qed.
-Global Instance auth_discrete : Discrete A → Discrete authC.
+Global Instance auth_ofe_discrete : OFEDiscrete A → OFEDiscrete authC.
 Proof. intros ? [??]; apply _. Qed.
 Global Instance auth_leibniz : LeibnizEquiv A → LeibnizEquiv (auth A).
 Proof. by intros ? [??] [??] [??]; f_equal/=; apply leibniz_equiv. Qed.
diff --git a/theories/algebra/cmra.v b/theories/algebra/cmra.v
index 85c785d24d6ff2e084411bfd2781874d4d2bebc2..fb78065ee47d6e1745423e5ed6a403f3693ba9a3 100644
--- a/theories/algebra/cmra.v
+++ b/theories/algebra/cmra.v
@@ -240,7 +240,7 @@ End ucmra_mixin.
 
 (** * Discrete CMRAs *)
 Class CMRADiscrete (A : cmraT) := {
-  cmra_discrete :> Discrete A;
+  cmra_discrete_ofe_discrete :> OFEDiscrete A;
   cmra_discrete_valid (x : A) : ✓{0} x → ✓ x
 }.
 Hint Mode CMRADiscrete ! : typeclass_instances.
@@ -554,7 +554,7 @@ Proof.
   split; first by rewrite cmra_valid_validN.
   eauto using cmra_discrete_valid, cmra_validN_le with lia.
 Qed.
-Lemma cmra_discrete_included_iff `{Discrete A} n x y : x ≼ y ↔ x ≼{n} y.
+Lemma cmra_discrete_included_iff `{OFEDiscrete A} n x y : x ≼ y ↔ x ≼{n} y.
 Proof.
   split; first by apply cmra_included_includedN.
   intros [z ->%(timeless_iff _ _)]; eauto using cmra_included_l.
@@ -597,7 +597,9 @@ Lemma id_free_l x `{!IdFree x} y : ✓x → y ⋅ x ≡ x → False.
 Proof. rewrite comm. eauto using id_free_r. Qed.
 Lemma discrete_id_free x `{CMRADiscrete A}:
   (∀ y, ✓ x → x ⋅ y ≡ x → False) → IdFree x.
-Proof. repeat intro. eauto using cmra_discrete_valid, cmra_discrete, timeless. Qed.
+Proof.
+  intros Hx y ??. apply (Hx y), (timeless _); eauto using cmra_discrete_valid.
+Qed.
 Global Instance id_free_op_r x y : IdFree y → Cancelable x → IdFree (x ⋅ y).
 Proof.
   intros ?? z ? Hid%symmetry. revert Hid. rewrite -assoc=>/(cancelableN x) ?.
diff --git a/theories/algebra/csum.v b/theories/algebra/csum.v
index 41643dfc98abf09ecccaa6dcc5784fd7ba1d183d..a4185fee3d26c076745c7d94e59b5cf3836c1e55 100644
--- a/theories/algebra/csum.v
+++ b/theories/algebra/csum.v
@@ -96,7 +96,7 @@ Next Obligation.
   + rewrite (conv_compl n (csum_chain_r c b')) /=. destruct (c n); naive_solver.
 Qed.
 
-Global Instance csum_discrete : Discrete A → Discrete B → Discrete csumC.
+Global Instance csum_ofe_discrete : OFEDiscrete A → OFEDiscrete B → OFEDiscrete csumC.
 Proof. by inversion_clear 3; constructor; apply (timeless _). Qed.
 Global Instance csum_leibniz :
   LeibnizEquiv A → LeibnizEquiv B → LeibnizEquiv (csumC A B).
diff --git a/theories/algebra/excl.v b/theories/algebra/excl.v
index 1c71576e7004518a3f745da96d5981d45d24aa1c..e7f2385ed86de99d6dbf2d4604d6adf348f88801 100644
--- a/theories/algebra/excl.v
+++ b/theories/algebra/excl.v
@@ -59,7 +59,7 @@ Proof.
   - by intros []; constructor.
 Qed.
 
-Global Instance excl_discrete : Discrete A → Discrete exclC.
+Global Instance excl_ofe_discrete : OFEDiscrete A → OFEDiscrete exclC.
 Proof. by inversion_clear 2; constructor; apply (timeless _). Qed.
 Global Instance excl_leibniz : LeibnizEquiv A → LeibnizEquiv (excl A).
 Proof. by destruct 2; f_equal; apply leibniz_equiv. Qed.
@@ -91,7 +91,7 @@ Proof.
 Qed.
 Canonical Structure exclR := CMRAT (excl A) excl_cmra_mixin.
 
-Global Instance excl_cmra_discrete : Discrete A → CMRADiscrete exclR.
+Global Instance excl_cmra_discrete : OFEDiscrete A → CMRADiscrete exclR.
 Proof. split. apply _. by intros []. Qed.
 
 (** Internalized properties *)
diff --git a/theories/algebra/gmap.v b/theories/algebra/gmap.v
index 2c31f3ee618e91f794a3c0bca0d6693f9597784f..9ae5fba83ae34fe7392ffa3a909a7eb2caa2fdf7 100644
--- a/theories/algebra/gmap.v
+++ b/theories/algebra/gmap.v
@@ -37,7 +37,7 @@ Next Obligation.
   by rewrite conv_compl /=; apply reflexive_eq.
 Qed.
 
-Global Instance gmap_discrete : Discrete A → Discrete gmapC.
+Global Instance gmap_ofe_discrete : OFEDiscrete A → OFEDiscrete gmapC.
 Proof. intros ? m m' ? i. by apply (timeless _). Qed.
 (* why doesn't this go automatic? *)
 Global Instance gmapC_leibniz: LeibnizEquiv A → LeibnizEquiv gmapC.
diff --git a/theories/algebra/list.v b/theories/algebra/list.v
index 5f42e37e0e6ee17f7ee4180133df99044878b5b1..c008b004c4539c5c970e0fb5bccb27a1036fb64a 100644
--- a/theories/algebra/list.v
+++ b/theories/algebra/list.v
@@ -77,7 +77,7 @@ Next Obligation.
   by rewrite Hcn.
 Qed.
 
-Global Instance list_discrete : Discrete A → Discrete listC.
+Global Instance list_ofe_discrete : OFEDiscrete A → OFEDiscrete listC.
 Proof. induction 2; constructor; try apply (timeless _); auto. Qed.
 
 Global Instance nil_timeless : Timeless (@nil A).
diff --git a/theories/algebra/ofe.v b/theories/algebra/ofe.v
index 09992145e4c8b7a37df90e332cdc880d8143e516..afe5aa35a12ff4de6c73e9eeb900caa3ff134d0f 100644
--- a/theories/algebra/ofe.v
+++ b/theories/algebra/ofe.v
@@ -107,8 +107,8 @@ Arguments timeless {_} _ {_} _ _.
 Hint Mode Timeless + ! : typeclass_instances.
 Instance: Params (@Timeless) 1.
 
-Class Discrete (A : ofeT) := discrete_timeless (x : A) :> Timeless x.
-Hint Mode Discrete ! : typeclass_instances.
+Class OFEDiscrete (A : ofeT) := ofe_discrete_timeless (x : A) :> Timeless x.
+Hint Mode OFEDiscrete ! : typeclass_instances.
 
 (** OFEs with a completion *)
 Record chain (A : ofeT) := {
@@ -653,7 +653,7 @@ Section unit.
   Global Program Instance unit_cofe : Cofe unitC := { compl x := () }.
   Next Obligation. by repeat split; try exists 0. Qed.
 
-  Global Instance unit_discrete_cofe : Discrete unitC.
+  Global Instance unit_ofe_discrete : OFEDiscrete unitC.
   Proof. done. Qed.
 End unit.
 
@@ -686,7 +686,7 @@ Section product.
   Global Instance prod_timeless (x : A * B) :
     Timeless (x.1) → Timeless (x.2) → Timeless x.
   Proof. by intros ???[??]; split; apply (timeless _). Qed.
-  Global Instance prod_discrete_cofe : Discrete A → Discrete B → Discrete prodC.
+  Global Instance prod_ofe_discrete : OFEDiscrete A → OFEDiscrete B → OFEDiscrete prodC.
   Proof. intros ?? [??]; apply _. Qed.
 End product.
 
@@ -868,7 +868,7 @@ Section sum.
   Proof. inversion_clear 2; constructor; by apply (timeless _). Qed.
   Global Instance inr_timeless (y : B) : Timeless y → Timeless (inr y).
   Proof. inversion_clear 2; constructor; by apply (timeless _). Qed.
-  Global Instance sum_discrete_ofe : Discrete A → Discrete B → Discrete sumC.
+  Global Instance sum_ofe_discrete : OFEDiscrete A → OFEDiscrete B → OFEDiscrete sumC.
   Proof. intros ?? [?|?]; apply _. Qed.
 End sum.
 
@@ -923,7 +923,7 @@ Section discrete_ofe.
     - done.
   Qed.
 
-  Global Instance discrete_discrete_ofe : Discrete (OfeT A discrete_ofe_mixin).
+  Global Instance discrete_ofe_discrete : OFEDiscrete (OfeT A discrete_ofe_mixin).
   Proof. by intros x y. Qed.
 
   Global Program Instance discrete_cofe : Cofe (OfeT A discrete_ofe_mixin) :=
@@ -992,7 +992,7 @@ Section option.
     destruct (c n); naive_solver.
   Qed.
 
-  Global Instance option_discrete : Discrete A → Discrete optionC.
+  Global Instance option_ofe_discrete : OFEDiscrete A → OFEDiscrete optionC.
   Proof. destruct 2; constructor; by apply (timeless _). Qed.
 
   Global Instance Some_ne : NonExpansive (@Some A).
@@ -1226,7 +1226,7 @@ Section sigma.
 
   Global Instance sig_timeless (x : sig P) :  Timeless (`x) → Timeless x.
   Proof. intros ? y. rewrite sig_dist_alt sig_equiv_alt. apply (timeless _). Qed.
-  Global Instance sig_discrete_ofe : Discrete A → Discrete sigC.
+  Global Instance sig_ofe_discrete : OFEDiscrete A → OFEDiscrete sigC.
   Proof. intros ??. apply _. Qed.
 End sigma.
 
diff --git a/theories/algebra/vector.v b/theories/algebra/vector.v
index 5f1b29130dea34af2d5077c10cbbdeda92776f4e..d6a91347dace78f25b8c2d4a2355e0466dd58d0c 100644
--- a/theories/algebra/vector.v
+++ b/theories/algebra/vector.v
@@ -29,7 +29,7 @@ Section ofe.
     intros ?? v' ?. inv_vec v'=>x' v'. inversion_clear 1.
     constructor. by apply timeless. change (v ≡ v'). by apply timeless.
   Qed.
-  Global Instance vec_discrete_cofe m : Discrete A → Discrete (vecC m).
+  Global Instance vec_ofe_discrete m : OFEDiscrete A → OFEDiscrete (vecC m).
   Proof. intros ? v. induction v; apply _. Qed.
 End ofe.