diff --git a/modures/cmra.v b/modures/cmra.v
index 4c43104751729d221b344ccb2d57989b43357c7c..8b146841172fb2f3c2291dc313a0224fa546e165 100644
--- a/modures/cmra.v
+++ b/modures/cmra.v
@@ -318,17 +318,17 @@ Proof.
   * by intros Hx z ?; exists y; split; [done|apply (Hx z)].
   * by intros Hx z n ?; destruct (Hx z n) as (?&<-&?).
 Qed.
-Lemma ra_updateP_id (P : A → Prop) x : P x → x ⇝: P.
+Lemma cmra_updateP_id (P : A → Prop) x : P x → x ⇝: P.
 Proof. by intros ? z n ?; exists x. Qed.
-Lemma ra_updateP_compose (P Q : A → Prop) x :
+Lemma cmra_updateP_compose (P Q : A → Prop) x :
   x ⇝: P → (∀ y, P y → y ⇝: Q) → x ⇝: Q.
 Proof.
   intros Hx Hy z n ?. destruct (Hx z n) as (y&?&?); auto. by apply (Hy y).
 Qed.
-Lemma ra_updateP_weaken (P Q : A → Prop) x : x ⇝: P → (∀ y, P y → Q y) → x ⇝: Q.
-Proof. eauto using ra_updateP_compose, ra_updateP_id. Qed.
+Lemma cmra_updateP_weaken (P Q : A → Prop) x : x ⇝: P → (∀ y, P y → Q y) → x ⇝: Q.
+Proof. eauto using cmra_updateP_compose, cmra_updateP_id. Qed.
 
-Lemma ra_updateP_op (P1 P2 Q : A → Prop) x1 x2 :
+Lemma cmra_updateP_op (P1 P2 Q : A → Prop) x1 x2 :
   x1 ⇝: P1 → x2 ⇝: P2 → (∀ y1 y2, P1 y1 → P2 y2 → Q (y1 ⋅ y2)) → x1 ⋅ x2 ⇝: Q.
 Proof.
   intros Hx1 Hx2 Hy z n ?.
@@ -337,9 +337,12 @@ Proof.
     first by rewrite associative (commutative _ x2) -associative.
   exists (y1 â‹… y2); split; last rewrite (commutative _ y1) -associative; auto.
 Qed.
-Lemma ra_update_op x1 x2 y1 y2 : x1 ⇝ y1 → x2 ⇝ y2 → x1 ⋅ x2 ⇝ y1 ⋅ y2.
+Lemma cmra_updateP_op' (P1 P2 : A → Prop) x1 x2 :
+  x1 ⇝: P1 → x2 ⇝: P2 → x1 ⋅ x2 ⇝: λ y, ∃ y1 y2, y = y1 ⋅ y2 ∧ P1 y1 ∧ P2 y2.
+Proof. eauto 10 using cmra_updateP_op. Qed.
+Lemma cmra_update_op x1 x2 y1 y2 : x1 ⇝ y1 → x2 ⇝ y2 → x1 ⋅ x2 ⇝ y1 ⋅ y2.
 Proof.
-  rewrite !cmra_update_updateP; eauto using ra_updateP_op with congruence.
+  rewrite !cmra_update_updateP; eauto using cmra_updateP_op with congruence.
 Qed.
 End cmra.
 
@@ -487,13 +490,16 @@ Section prod.
   Qed.
   Lemma prod_update x y : x.1 ⇝ y.1 → x.2 ⇝ y.2 → x ⇝ y.
   Proof. intros ?? z n [??]; split; simpl in *; auto. Qed.
-  Lemma prod_updateP (P : A * B → Prop) P1 P2 x :
-    x.1 ⇝: P1 → x.2 ⇝: P2 → (∀ a b, P1 a → P2 b → P (a,b)) → x ⇝: P.
+  Lemma prod_updateP P1 P2 (Q : A * B → Prop)  x :
+    x.1 ⇝: P1 → x.2 ⇝: P2 → (∀ a b, P1 a → P2 b → Q (a,b)) → x ⇝: Q.
   Proof.
     intros Hx1 Hx2 HP z n [??]; simpl in *.
     destruct (Hx1 (z.1) n) as (a&?&?), (Hx2 (z.2) n) as (b&?&?); auto.
     exists (a,b); repeat split; auto.
   Qed.
+  Lemma prod_updateP' P1 P2 x :
+    x.1 ⇝: P1 → x.2 ⇝: P2 → x ⇝: λ y, P1 (y.1) ∧ P2 (y.2).
+  Proof. eauto using prod_updateP. Qed.
 End prod.
 Arguments prodRA : clear implicits.
 
diff --git a/modures/option.v b/modures/option.v
index 5536885af6828659bcd12532b9b6714d9fe4ef6a..5275e5b31e50a330a54316ad839a53187609568e 100644
--- a/modures/option.v
+++ b/modures/option.v
@@ -147,10 +147,12 @@ Proof.
   destruct (Hx (unit x) n) as (y'&?&?); rewrite ?cmra_unit_r; auto.
   by exists (Some y'); split; [auto|apply cmra_validN_op_l with (unit x)].
 Qed.
+Lemma option_updateP' (P : A → Prop) x :
+  x ⇝: P → Some x ⇝: λ y, default False y P.
+Proof. eauto using option_updateP. Qed.
 Lemma option_update x y : x ⇝ y → Some x ⇝ Some y.
 Proof.
-  intros; apply cmra_update_updateP, (option_updateP (y=)); [|naive_solver].
-  by apply cmra_update_updateP.
+  rewrite !cmra_update_updateP; eauto using option_updateP with congruence.
 Qed.
 End cmra.