diff --git a/barrier/client.v b/barrier/client.v
index fa09898b66a5882de377e05980eee6a716e6ee0d..7dc6dd2f98ff44c57e3b5ce488d71531feec1047 100644
--- a/barrier/client.v
+++ b/barrier/client.v
@@ -29,11 +29,11 @@ Section ClosedProofs.
 
   Lemma client_safe_closed σ : {{ ownP σ : iProp }} client {{ λ v, True }}.
   Proof.
-    apply ht_alt. rewrite (heap_alloc ⊤ (nroot ..: "Barrier")); last first.
+    apply ht_alt. rewrite (heap_alloc ⊤ (nroot .@ "Barrier")); last first.
     { (* FIXME Really?? set_solver takes forever on "⊆ ⊤"?!? *)
       by move=>? _. }
     apply wp_strip_pvs, exist_elim=> ?. rewrite and_elim_l.
-    rewrite -(client_safe (nroot ..: "Heap" ) (nroot ..: "Barrier" )) //.
+    rewrite -(client_safe (nroot .@ "Heap" ) (nroot .@ "Barrier" )) //.
     (* This, too, should be automated. *)
     by apply ndot_ne_disjoint.
   Qed.
diff --git a/program_logic/ghost_ownership.v b/program_logic/ghost_ownership.v
index f6fc64e5813aa567be66ef059f90e9971c66684a..623d794895cd5de4139763e593daec14ce81afa5 100644
--- a/program_logic/ghost_ownership.v
+++ b/program_logic/ghost_ownership.v
@@ -23,7 +23,7 @@ Class inG (Λ : language) (Σ : iFunctorG) (A : cmraT) := InG {
 
 Definition to_globalF `{inG Λ Σ A} (γ : gname) (a : A) : iGst Λ (globalF Σ) :=
   iprod_singleton inG_id {[ γ := cmra_transport inG_prf a ]}.
-Definition own `{inG Λ Σ A} (γ : gname) (a : A) : iProp Λ (globalF Σ) :=
+Definition own `{inG Λ Σ A} (γ : gname) (a : A) : iPropG Λ Σ :=
   ownG (to_globalF γ a).
 Instance: Params (@to_globalF) 5.
 Instance: Params (@own) 5.
diff --git a/program_logic/namespaces.v b/program_logic/namespaces.v
index 39f83876e405433a067adb32748ecc06ecab4d25..3fb896e995536819181e66e5558f39a4bc296210 100644
--- a/program_logic/namespaces.v
+++ b/program_logic/namespaces.v
@@ -7,8 +7,8 @@ Definition ndot `{Countable A} (N : namespace) (x : A) : namespace :=
   encode x :: N.
 Coercion nclose (N : namespace) : coPset := coPset_suffixes (encode N).
 
-Infix "..:" := ndot (at level 19, left associativity) : C_scope.
-Notation "(..:)" := ndot (only parsing) : C_scope.
+Infix ".@" := ndot (at level 19, left associativity) : C_scope.
+Notation "(.@)" := ndot (only parsing) : C_scope.
 
 Instance ndot_inj `{Countable A} : Inj2 (=) (=) (=) (@ndot A _ _).
 Proof. by intros N1 x1 N2 x2 ?; simplify_eq. Qed.
@@ -16,13 +16,13 @@ Lemma nclose_nroot : nclose nroot = ⊤.
 Proof. by apply (sig_eq_pi _). Qed.
 Lemma encode_nclose N : encode N ∈ nclose N.
 Proof. by apply elem_coPset_suffixes; exists xH; rewrite (left_id_L _ _). Qed.
-Lemma nclose_subseteq `{Countable A} N x : nclose (N ..: x) ⊆ nclose N.
+Lemma nclose_subseteq `{Countable A} N x : nclose (N .@ x) ⊆ nclose N.
 Proof.
   intros p; rewrite /nclose !elem_coPset_suffixes; intros [q ->].
-  destruct (list_encode_suffix N (N ..: x)) as [q' ?]; [by exists [encode x]|].
+  destruct (list_encode_suffix N (N .@ x)) as [q' ?]; [by exists [encode x]|].
   by exists (q ++ q')%positive; rewrite <-(assoc_L _); f_equal.
 Qed.
-Lemma ndot_nclose `{Countable A} N x : encode (N ..: x) ∈ nclose N.
+Lemma ndot_nclose `{Countable A} N x : encode (N .@ x) ∈ nclose N.
 Proof. apply nclose_subseteq with x, encode_nclose. Qed.
 
 Instance ndisjoint : Disjoint namespace := λ N1 N2,
@@ -36,16 +36,16 @@ Section ndisjoint.
   Global Instance ndisjoint_comm : Comm iff ndisjoint.
   Proof. intros N1 N2. rewrite /disjoint /ndisjoint; naive_solver. Qed.
 
-  Lemma ndot_ne_disjoint N (x y : A) : x ≠ y → N ..: x ⊥ N ..: y.
-  Proof. intros Hxy. exists (N ..: x), (N ..: y); naive_solver. Qed.
+  Lemma ndot_ne_disjoint N (x y : A) : x ≠ y → N .@ x ⊥ N .@ y.
+  Proof. intros Hxy. exists (N .@ x), (N .@ y); naive_solver. Qed.
 
-  Lemma ndot_preserve_disjoint_l N1 N2 x : N1 ⊥ N2 → N1 ..: x ⊥ N2.
+  Lemma ndot_preserve_disjoint_l N1 N2 x : N1 ⊥ N2 → N1 .@ x ⊥ N2.
   Proof.
     intros (N1' & N2' & Hpr1 & Hpr2 & Hl & Hne). exists N1', N2'.
     split_and?; try done; []. by apply suffix_of_cons_r.
   Qed.
 
-  Lemma ndot_preserve_disjoint_r N1 N2 x : N1 ⊥ N2 → N1 ⊥ N2 ..: x .
+  Lemma ndot_preserve_disjoint_r N1 N2 x : N1 ⊥ N2 → N1 ⊥ N2 .@ x .
   Proof. rewrite ![N1 ⊥ _]comm. apply ndot_preserve_disjoint_l. Qed.
 
   Lemma ndisj_disjoint N1 N2 : N1 ⊥ N2 → nclose N1 ∩ nclose N2 = ∅.