diff --git a/theories/algebra/local_updates.v b/theories/algebra/local_updates.v index 78cf251331231b7d58a43e823bef4c9f58e56822..93f185ff38674f14528be5db66c23f3301b8981a 100644 --- a/theories/algebra/local_updates.v +++ b/theories/algebra/local_updates.v @@ -59,6 +59,14 @@ Section updates. destruct mz as [z|]; [|done]. by destruct (id_freeN_r n n x z). Qed. + Lemma core_id_local_update x y z `{!CoreId y} : + y ≼ x → (x, z) ~l~> (x, z ⋅ y). + Proof. + intros Hincl n mf ? Heq; simpl in *; split; first done. + rewrite [x]core_id_extract // Heq. destruct mf as [f|]; last done. + simpl. rewrite -assoc [f ⋅ y]comm assoc //. + Qed. + Lemma local_update_discrete `{!CmraDiscrete A} (x y x' y' : A) : (x,y) ~l~> (x',y') ↔ ∀ mz, ✓ x → x ≡ y ⋅? mz → ✓ x' ∧ x' ≡ y' ⋅? mz. Proof.