diff --git a/theories/proofmode/classes.v b/theories/proofmode/classes.v
index b9dca5801196141828c6411cd2fe00d2ff7a124d..1ce3603bb590e9c9fd4a0e37660c3e66fc6bf49f 100644
--- a/theories/proofmode/classes.v
+++ b/theories/proofmode/classes.v
@@ -571,11 +571,8 @@ Hint Mode IntoAcc + - ! - - - - - - - - : typeclass_instances.
    still benefit from [iInv].
 
    TODO: Add support for a binder (like accessors have it).
-
-   This is defined on sbi instead of bi as typeclass search otherwise
-   fails (e.g. in the iInv as used in cancelable_invariants.v).
 *)
-Class ElimInv {PROP : sbi} {X : Type} (φ : Prop)
+Class ElimInv {PROP : bi} {X : Type} (φ : Prop)
       (Pinv Pin : PROP) (Pout : X → PROP) (Pclose : option (X → PROP))
       (Q : PROP) (Q' : X → PROP) :=
   elim_inv : φ → Pinv ∗ Pin ∗ (∀ x, Pout x ∗ (default (λ _, emp) Pclose id) x -∗ Q' x) ⊢ Q.
diff --git a/theories/proofmode/coq_tactics.v b/theories/proofmode/coq_tactics.v
index 9617346e101ad2895d03123de70f8adf7906682d..0801c091b666c63a83aa01c4d679ab614202d16d 100644
--- a/theories/proofmode/coq_tactics.v
+++ b/theories/proofmode/coq_tactics.v
@@ -1093,6 +1093,26 @@ Proof.
   rewrite envs_entails_eq => <-. by setoid_rewrite <-envs_incr_counter_equiv.
 Qed.
 
+(** * Invariants *)
+Lemma tac_inv_elim {X : Type} Δ Δ' i j φ p Pinv Pin Pout (Pclose : option (X → PROP))
+      Q (Q' : X → PROP) :
+  envs_lookup_delete false i Δ = Some (p, Pinv, Δ') →
+  ElimInv φ Pinv Pin Pout Pclose Q Q' →
+  φ →
+  (∀ R, ∃ Δ'',
+    envs_app false (Esnoc Enil j
+                    (Pin -∗ (∀ x, Pout x -∗ default (Q' x) Pclose (λ Pclose, Pclose x -∗ Q' x)) -∗ R)%I) Δ'
+      = Some Δ'' ∧
+    envs_entails Δ'' R) →
+  envs_entails Δ Q.
+Proof.
+  rewrite envs_entails_eq=> /envs_lookup_delete_Some [? ->] Hinv ? /(_ Q) [Δ'' [? <-]].
+  rewrite (envs_lookup_sound' _ false) // envs_app_singleton_sound //; simpl.
+  apply wand_elim_r', wand_mono; last done. apply wand_intro_r, wand_intro_r.
+  rewrite intuitionistically_if_elim -assoc. destruct Pclose; simpl in *.
+  - setoid_rewrite wand_curry. auto.
+  - setoid_rewrite <-(right_id emp%I _ (Pout _)). auto.
+Qed.
 
 End bi_tactics.
 
@@ -1324,27 +1344,6 @@ Implicit Types Γ : env PROP.
 Implicit Types Δ : envs PROP.
 Implicit Types P Q : PROP.
 
-(** * Invariants *)
-Lemma tac_inv_elim {X : Type} Δ Δ' i j φ p Pinv Pin Pout (Pclose : option (X → PROP))
-      Q (Q' : X → PROP) :
-  envs_lookup_delete false i Δ = Some (p, Pinv, Δ') →
-  ElimInv φ Pinv Pin Pout Pclose Q Q' →
-  φ →
-  (∀ R, ∃ Δ'',
-    envs_app false (Esnoc Enil j
-                    (Pin -∗ (∀ x, Pout x -∗ default (Q' x) Pclose (λ Pclose, Pclose x -∗ Q' x)) -∗ R)%I) Δ'
-      = Some Δ'' ∧
-    envs_entails Δ'' R) →
-  envs_entails Δ Q.
-Proof.
-  rewrite envs_entails_eq=> /envs_lookup_delete_Some [? ->] Hinv ? /(_ Q) [Δ'' [? <-]].
-  rewrite (envs_lookup_sound' _ false) // envs_app_singleton_sound //; simpl.
-  apply wand_elim_r', wand_mono; last done. apply wand_intro_r, wand_intro_r.
-  rewrite intuitionistically_if_elim -assoc. destruct Pclose; simpl in *.
-  - setoid_rewrite wand_curry. auto.
-  - setoid_rewrite <-(right_id emp%I _ (Pout _)). auto.
-Qed.
-
 (** * Rewriting *)
 Lemma tac_rewrite Δ i p Pxy d Q :
   envs_lookup i Δ = Some (p, Pxy) →