diff --git a/theories/algebra/auth.v b/theories/algebra/auth.v
index 49467a7703537cae0736836e341a18c01566f53c..0824f520216ecda77b1f93b5c7aa5dff3a5fcb04 100644
--- a/theories/algebra/auth.v
+++ b/theories/algebra/auth.v
@@ -286,6 +286,8 @@ Lemma auth_frag_op a b : â—¯ (a â‹… b) = â—¯ a â‹… â—¯ b.
 Proof. done. Qed.
 Lemma auth_frag_mono a b : a ≼ b → ◯ a ≼ ◯ b.
 Proof. intros [c ->]. rewrite auth_frag_op. apply cmra_included_l. Qed.
+Lemma auth_frag_core a : core (â—¯ a) = â—¯ (core a).
+Proof. done. Qed.
 
 Global Instance auth_frag_sep_homomorphism :
   MonoidHomomorphism op op (≡) (@auth_frag A).
diff --git a/theories/algebra/cmra.v b/theories/algebra/cmra.v
index 54a984004e7c94a5331b0c97ae5dfe37e9ea4d81..32c7d214f95d30318fe464fc0ecb8c712b592ff2 100644
--- a/theories/algebra/cmra.v
+++ b/theories/algebra/cmra.v
@@ -1221,6 +1221,9 @@ Section prod.
   Proof. done. Qed.
   Lemma pair_valid (a : A) (b : B) : ✓ (a, b) ↔ ✓ a ∧ ✓ b.
   Proof. done. Qed.
+  Lemma pair_included (a a' : A) (b b' : B) :
+    (a, b) ≼ (a', b') ↔ a ≼ a' ∧ b ≼ b'.
+  Proof. apply prod_included. Qed.
 
   Lemma pair_core `{!CmraTotal A, !CmraTotal B} (a : A) (b : B) :
     core (a, b) = (core a, core b).
diff --git a/theories/algebra/gmap.v b/theories/algebra/gmap.v
index 51ae2084dc7c7148caac26bf9b7e991719c6253c..a94bce9fba35adc6f8861f5f4229071d0811718b 100644
--- a/theories/algebra/gmap.v
+++ b/theories/algebra/gmap.v
@@ -292,6 +292,9 @@ Lemma core_singleton' (i : K) (x : A) cx :
 Proof.
   intros (cx'&?&->)%equiv_Some_inv_r'. by rewrite (core_singleton _ _ cx').
 Qed.
+Lemma singleton_core_total `{!CmraTotal A} (i : K) (x : A) :
+  core ({[ i := x ]} : gmap K A) = {[ i := core x ]}.
+Proof. apply core_singleton. rewrite cmra_pcore_core //. Qed.
 Lemma op_singleton (i : K) (x y : A) :
   {[ i := x ]} â‹… {[ i := y ]} = ({[ i := x â‹… y ]} : gmap K A).
 Proof. by apply (merge_singleton _ _ _ x y). Qed.