diff --git a/tests/proofmode.v b/tests/proofmode.v
index adfc177806e26664ff802370167e7eb40a3d4001..fe6b8969ac8d5844a71cc8a415e37140eea87183 100644
--- a/tests/proofmode.v
+++ b/tests/proofmode.v
@@ -6,6 +6,12 @@ Section tests.
 Context {PROP : sbi}.
 Implicit Types P Q R : PROP.
 
+Lemma test_eauto_emp_isplit_biwand P : emp ⊢ P ∗-∗ P.
+Proof. eauto 6. Qed.
+
+Lemma test_eauto_isplit_biwand P : (P ∗-∗ P)%I.
+Proof. eauto. Qed.
+
 Check "demo_0".
 Lemma demo_0 P Q : □ (P ∨ Q) -∗ (∀ x, ⌜x = 0⌝ ∨ ⌜x = 1⌝) → (Q ∨ P).
 Proof.
@@ -158,7 +164,7 @@ Proof.
   iIntros "H".
   let H1 := iFresh in
   let H2 := iFresh in
-  let pat :=constr:(IList [cons (IIdent H1) (cons (IIdent H2) nil)]) in 
+  let pat :=constr:(IList [cons (IIdent H1) (cons (IIdent H2) nil)]) in
   iDestruct "H" as pat.
   iFrame.
 Qed.
diff --git a/theories/proofmode/ltac_tactics.v b/theories/proofmode/ltac_tactics.v
index e716667414e4648b965d65db33d753c57cf488b8..6edfd45160554634552608a4cbf16473baa94828 100644
--- a/theories/proofmode/ltac_tactics.v
+++ b/theories/proofmode/ltac_tactics.v
@@ -3153,6 +3153,7 @@ Tactic Notation "iAccu" :=
 
 (** Automation *)
 Hint Extern 0 (_ ⊢ _) => iStartProof : core.
+Hint Extern 0 (bi_emp_valid _) => iStartProof : core.
 
 (* Make sure that by and done solve trivial things in proof mode *)
 Hint Extern 0 (envs_entails _ _) => iPureIntro; try done : core.