From a8591b70819f987a064faa12d266aeb577082cd7 Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Wed, 17 Feb 2016 00:46:40 +0100
Subject: [PATCH] =?UTF-8?q?Use=20{[=5F=20:=3D=20=5F]}=20for=20singleton=20?=
 =?UTF-8?q?map=20so=20we=20can=20use=20=E2=86=A6=20for=20maps=20to.?=
MIME-Version: 1.0
Content-Type: text/plain; charset=UTF-8
Content-Transfer-Encoding: 8bit

The singleton maps notation is now also more consistent with the
insert <[_ := _]> _ notation for maps.
---
 algebra/cmra_big_op.v           |  2 +-
 algebra/fin_maps.v              | 28 +++++++--------
 algebra/upred_big_op.v          |  2 +-
 barrier/barrier.v               |  2 +-
 heap_lang/heap.v                | 29 ++++++++-------
 prelude/base.v                  |  3 +-
 prelude/fin_map_dom.v           |  4 +--
 prelude/fin_maps.v              | 64 ++++++++++++++++-----------------
 prelude/hashset.v               |  2 +-
 prelude/mapset.v                |  2 +-
 program_logic/ghost_ownership.v |  2 +-
 program_logic/ownership.v       |  3 +-
 program_logic/pviewshifts.v     |  2 +-
 program_logic/wsat.v            |  2 +-
 14 files changed, 72 insertions(+), 75 deletions(-)

diff --git a/algebra/cmra_big_op.v b/algebra/cmra_big_op.v
index 6c6da1654..4740440b3 100644
--- a/algebra/cmra_big_op.v
+++ b/algebra/cmra_big_op.v
@@ -55,7 +55,7 @@ Lemma big_opM_delete (m : M A) i x :
 Proof.
   intros. by rewrite -{2}(insert_delete m i x) // big_opM_insert ?lookup_delete.
 Qed.
-Lemma big_opM_singleton i x : big_opM ({[i ↦ x]} : M A) ≡ x.
+Lemma big_opM_singleton i x : big_opM ({[i := x]} : M A) ≡ x.
 Proof.
   rewrite -insert_empty big_opM_insert /=; last auto using lookup_empty.
   by rewrite big_opM_empty right_id.
diff --git a/algebra/fin_maps.v b/algebra/fin_maps.v
index c9971cc93..b3cfb69eb 100644
--- a/algebra/fin_maps.v
+++ b/algebra/fin_maps.v
@@ -84,7 +84,7 @@ Proof.
   by apply (timeless _); rewrite -Hm lookup_insert_ne.
 Qed.
 Global Instance map_singleton_timeless i x :
-  Timeless x → Timeless ({[ i ↦ x ]} : gmap K A) := _.
+  Timeless x → Timeless ({[ i := x ]} : gmap K A) := _.
 End cofe.
 
 Arguments mapC _ {_ _} _.
@@ -196,16 +196,16 @@ Lemma map_insert_validN n m i x : ✓{n} x → ✓{n} m → ✓{n} <[i:=x]>m.
 Proof. by intros ?? j; destruct (decide (i = j)); simplify_map_equality. Qed.
 Lemma map_insert_valid m i x : ✓ x → ✓ m → ✓ <[i:=x]>m.
 Proof. intros ?? n j; apply map_insert_validN; auto. Qed.
-Lemma map_singleton_validN n i x : ✓{n} ({[ i ↦ x ]} : gmap K A) ↔ ✓{n} x.
+Lemma map_singleton_validN n i x : ✓{n} ({[ i := x ]} : gmap K A) ↔ ✓{n} x.
 Proof.
   split; [|by intros; apply map_insert_validN, cmra_empty_valid].
   by move=>/(_ i); simplify_map_equality.
 Qed.
-Lemma map_singleton_valid i x : ✓ ({[ i ↦ x ]} : gmap K A) ↔ ✓ x.
+Lemma map_singleton_valid i x : ✓ ({[ i := x ]} : gmap K A) ↔ ✓ x.
 Proof. split; intros ? n; eapply map_singleton_validN; eauto. Qed.
 
 Lemma map_insert_singleton_opN n m i x :
-  m !! i = None ∨ m !! i ≡{n}≡ Some (unit x) → <[i:=x]> m ≡{n}≡ {[ i ↦ x ]} ⋅ m.
+  m !! i = None ∨ m !! i ≡{n}≡ Some (unit x) → <[i:=x]> m ≡{n}≡ {[ i := x ]} ⋅ m.
 Proof.
   intros Hi j; destruct (decide (i = j)) as [->|];
     [|by rewrite lookup_op lookup_insert_ne // lookup_singleton_ne // left_id].
@@ -213,20 +213,20 @@ Proof.
   by destruct Hi as [->| ->]; constructor; rewrite ?cmra_unit_r.
 Qed.
 Lemma map_insert_singleton_op m i x :
-  m !! i = None ∨ m !! i ≡ Some (unit x) → <[i:=x]> m ≡ {[ i ↦ x ]} ⋅ m.
+  m !! i = None ∨ m !! i ≡ Some (unit x) → <[i:=x]> m ≡ {[ i := x ]} ⋅ m.
 Proof.
   rewrite !equiv_dist; naive_solver eauto using map_insert_singleton_opN.
 Qed.
 
 Lemma map_unit_singleton (i : K) (x : A) :
-  unit ({[ i ↦ x ]} : gmap K A) = {[ i ↦ unit x ]}.
+  unit ({[ i := x ]} : gmap K A) = {[ i := unit x ]}.
 Proof. apply map_fmap_singleton. Qed.
 Lemma map_op_singleton (i : K) (x y : A) :
-  {[ i ↦ x ]} ⋅ {[ i ↦ y ]} = ({[ i ↦ x ⋅ y ]} : gmap K A).
+  {[ i := x ]} â‹… {[ i := y ]} = ({[ i := x â‹… y ]} : gmap K A).
 Proof. by apply (merge_singleton _ _ _ x y). Qed.
 
 Lemma singleton_includedN n m i x :
-  {[ i ↦ x ]} ≼{n} m ↔ ∃ y, m !! i ≡{n}≡ Some y ∧ x ≼ y.
+  {[ i := x ]} ≼{n} m ↔ ∃ y, m !! i ≡{n}≡ Some y ∧ x ≼ y.
   (* not m !! i = Some y ∧ x ≼{n} y to deal with n = 0 *)
 Proof.
   split.
@@ -264,23 +264,23 @@ Proof.
 Qed.
 
 Lemma map_singleton_updateP (P : A → Prop) (Q : gmap K A → Prop) i x :
-  x ~~>: P → (∀ y, P y → Q {[ i ↦ y ]}) → {[ i ↦ x ]} ~~>: Q.
+  x ~~>: P → (∀ y, P y → Q {[ i := y ]}) → {[ i := x ]} ~~>: Q.
 Proof. apply map_insert_updateP. Qed.
 Lemma map_singleton_updateP' (P : A → Prop) i x :
-  x ~~>: P → {[ i ↦ x ]} ~~>: λ m, ∃ y, m = {[ i ↦ y ]} ∧ P y.
+  x ~~>: P → {[ i := x ]} ~~>: λ m, ∃ y, m = {[ i := y ]} ∧ P y.
 Proof. apply map_insert_updateP'. Qed.
-Lemma map_singleton_update i (x y : A) : x ~~> y → {[ i ↦ x ]} ~~> {[ i ↦ y ]}.
+Lemma map_singleton_update i (x y : A) : x ~~> y → {[ i := x ]} ~~> {[ i := y ]}.
 Proof. apply map_insert_update. Qed.
 
 Lemma map_singleton_updateP_empty `{Empty A, !CMRAIdentity A}
     (P : A → Prop) (Q : gmap K A → Prop) i :
-  ∅ ~~>: P → (∀ y, P y → Q {[ i ↦ y ]}) → ∅ ~~>: Q.
+  ∅ ~~>: P → (∀ y, P y → Q {[ i := y ]}) → ∅ ~~>: Q.
 Proof.
   intros Hx HQ gf n Hg.
   destruct (Hx (from_option ∅ (gf !! i)) n) as (y&?&Hy).
   { move:(Hg i). rewrite !left_id.
     case _: (gf !! i); simpl; auto using cmra_empty_valid. }
-  exists {[ i ↦ y ]}; split; first by auto.
+  exists {[ i := y ]}; split; first by auto.
   intros i'; destruct (decide (i' = i)) as [->|].
   - rewrite lookup_op lookup_singleton.
     move:Hy; case _: (gf !! i); first done.
@@ -288,7 +288,7 @@ Proof.
   - move:(Hg i'). by rewrite !lookup_op lookup_singleton_ne // !left_id.
 Qed.
 Lemma map_singleton_updateP_empty' `{Empty A, !CMRAIdentity A} (P: A → Prop) i :
-  ∅ ~~>: P → ∅ ~~>: λ m, ∃ y, m = {[ i ↦ y ]} ∧ P y.
+  ∅ ~~>: P → ∅ ~~>: λ m, ∃ y, m = {[ i := y ]} ∧ P y.
 Proof. eauto using map_singleton_updateP_empty. Qed.
 
 Section freshness.
diff --git a/algebra/upred_big_op.v b/algebra/upred_big_op.v
index 25d744197..05310f67e 100644
--- a/algebra/upred_big_op.v
+++ b/algebra/upred_big_op.v
@@ -77,7 +77,7 @@ Section fin_map.
   Proof.
     intros ?; by rewrite /uPred_big_sep /uPred_big_sepM map_to_list_insert.
   Qed.
-  Lemma big_sepM_singleton i x : (Π★{map {[i ↦ x]}} P)%I ≡ (P i x)%I.
+  Lemma big_sepM_singleton i x : (Π★{map {[i := x]}} P)%I ≡ (P i x)%I.
   Proof.
     rewrite -insert_empty big_sepM_insert/=; last auto using lookup_empty.
     by rewrite big_sepM_empty right_id.
diff --git a/barrier/barrier.v b/barrier/barrier.v
index f6f4ff826..aae49ccac 100644
--- a/barrier/barrier.v
+++ b/barrier/barrier.v
@@ -116,7 +116,7 @@ Section proof.
   Local Notation state_to_val s :=
     (match s with State Low _ => 0 | State High _ => 1 end).
   Definition barrier_inv (l : loc) (P : iProp) (s : stateT) : iProp :=
-    (l !=> '(state_to_val s) ★
+    (l ↦ '(state_to_val s) ★
      match s with State Low I' => waiting P I' | State High I' => ress I' end
     )%I.
 
diff --git a/heap_lang/heap.v b/heap_lang/heap.v
index e43497074..bc1a843d4 100644
--- a/heap_lang/heap.v
+++ b/heap_lang/heap.v
@@ -20,14 +20,13 @@ Definition to_heap : state → heapRA := fmap Excl.
 Definition of_heap : heapRA → state := omap (maybe Excl).
 
 Definition heap_mapsto `{heapG Σ} (l : loc) (v: val) : iPropG heap_lang Σ :=
-  auth_own heap_name {[ l ↦ Excl v ]}.
+  auth_own heap_name {[ l := Excl v ]}.
 Definition heap_inv `{i : heapG Σ} (h : heapRA) : iPropG heap_lang Σ :=
   ownP (of_heap h).
 Definition heap_ctx `{i : heapG Σ} (N : namespace) : iPropG heap_lang Σ :=
   auth_ctx heap_name N heap_inv.
 
-(* FIXME: ↦ is already used for the singleton empty map. Resolve that... *)
-Notation "l !=> v" := (heap_mapsto l v) (at level 20) : uPred_scope.
+Notation "l ↦ v" := (heap_mapsto l v) (at level 20) : uPred_scope.
 
 Section heap.
   Context {Σ : iFunctorG}.
@@ -56,7 +55,7 @@ Section heap.
     by case: (h !! l)=> [[]|]; auto.
   Qed.
   Lemma heap_singleton_inv_l h l v :
-    ✓ ({[l ↦ Excl v]} ⋅ h) → h !! l = None ∨ h !! l ≡ Some ExclUnit.
+    ✓ ({[l := Excl v]} ⋅ h) → h !! l = None ∨ h !! l ≡ Some ExclUnit.
   Proof.
     move=> /(_ O l). rewrite lookup_op lookup_singleton.
     by case: (h !! l)=> [[]|]; auto.
@@ -86,7 +85,7 @@ Section heap.
   Proof. intros h1 h2. by fold_leibniz=> ->. Qed.
 
   (** General properties of mapsto *)
-  Lemma heap_mapsto_disjoint l v1 v2 : (l !=> v1 ★ l !=> v2)%I ⊑ False.
+  Lemma heap_mapsto_disjoint l v1 v2 : (l ↦ v1 ★ l ↦ v2)%I ⊑ False.
   Proof.
     rewrite /heap_mapsto -auto_own_op auto_own_valid map_op_singleton.
     rewrite map_validI (forall_elim l) lookup_singleton.
@@ -97,7 +96,7 @@ Section heap.
   Lemma wp_alloc N E e v P Q :
     to_val e = Some v → nclose N ⊆ E →
     P ⊑ heap_ctx N →
-    P ⊑ (▷ ∀ l, l !=> v -★ Q (LocV l)) →
+    P ⊑ (▷ ∀ l, l ↦ v -★ Q (LocV l)) →
     P ⊑ wp E (Alloc e) Q.
   Proof.
     rewrite /heap_ctx /heap_inv /heap_mapsto=> ?? Hctx HP.
@@ -112,7 +111,7 @@ Section heap.
     apply sep_mono_r; rewrite HP; apply later_mono.
     apply forall_mono=> l; apply wand_intro_l.
     rewrite always_and_sep_l -assoc; apply const_elim_sep_l=> ?.
-    rewrite -(exist_intro (op {[ l ↦ Excl v ]})).
+    rewrite -(exist_intro (op {[ l := Excl v ]})).
     repeat erewrite <-exist_intro by apply _; simpl.
     rewrite -of_heap_insert left_id right_id !assoc.
     apply sep_mono_l.
@@ -124,12 +123,12 @@ Section heap.
   Lemma wp_load N E l v P Q :
     nclose N ⊆ E →
     P ⊑ heap_ctx N →
-    P ⊑ (▷ l !=> v ★ ▷ (l !=> v -★ Q v)) →
+    P ⊑ (▷ l ↦ v ★ ▷ (l ↦ v -★ Q v)) →
     P ⊑ wp E (Load (Loc l)) Q.
   Proof.
     rewrite /heap_ctx /heap_inv /heap_mapsto=>HN ? HPQ.
     apply (auth_fsa' heap_inv (wp_fsa _) id)
-      with N heap_name {[ l ↦ Excl v ]}; simpl; eauto with I.
+      with N heap_name {[ l := Excl v ]}; simpl; eauto with I.
     rewrite HPQ{HPQ}; apply sep_mono_r, forall_intro=> h; apply wand_intro_l.
     rewrite -assoc; apply const_elim_sep_l=> ?.
     rewrite {1}[(â–·ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs.
@@ -143,12 +142,12 @@ Section heap.
   Lemma wp_store N E l v' e v P Q :
     to_val e = Some v → nclose N ⊆ E → 
     P ⊑ heap_ctx N →
-    P ⊑ (▷ l !=> v' ★ ▷ (l !=> v -★ Q (LitV LitUnit))) →
+    P ⊑ (▷ l ↦ v' ★ ▷ (l ↦ v -★ Q (LitV LitUnit))) →
     P ⊑ wp E (Store (Loc l) e) Q.
   Proof.
     rewrite /heap_ctx /heap_inv /heap_mapsto=>? HN ? HPQ.
     apply (auth_fsa' heap_inv (wp_fsa _) (alter (λ _, Excl v) l))
-      with N heap_name {[ l ↦ Excl v' ]}; simpl; eauto with I.
+      with N heap_name {[ l := Excl v' ]}; simpl; eauto with I.
     rewrite HPQ{HPQ}; apply sep_mono_r, forall_intro=> h; apply wand_intro_l.
     rewrite -assoc; apply const_elim_sep_l=> ?.
     rewrite {1}[(â–·ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs.
@@ -164,12 +163,12 @@ Section heap.
     to_val e1 = Some v1 → to_val e2 = Some v2 → v' ≠ v1 →
     nclose N ⊆ E →
     P ⊑ heap_ctx N →
-    P ⊑ (▷ l !=> v' ★ ▷ (l !=> v' -★ Q (LitV (LitBool false)))) →
+    P ⊑ (▷ l ↦ v' ★ ▷ (l ↦ v' -★ Q (LitV (LitBool false)))) →
     P ⊑ wp E (Cas (Loc l) e1 e2) Q.
   Proof.
     rewrite /heap_ctx /heap_inv /heap_mapsto=>??? HN ? HPQ.
     apply (auth_fsa' heap_inv (wp_fsa _) id)
-      with N heap_name {[ l ↦ Excl v' ]}; simpl; eauto 10 with I.
+      with N heap_name {[ l := Excl v' ]}; simpl; eauto 10 with I.
     rewrite HPQ{HPQ}; apply sep_mono_r, forall_intro=> h; apply wand_intro_l.
     rewrite -assoc; apply const_elim_sep_l=> ?.
     rewrite {1}[(â–·ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs.
@@ -184,12 +183,12 @@ Section heap.
     to_val e1 = Some v1 → to_val e2 = Some v2 →
     nclose N ⊆ E →
     P ⊑ heap_ctx N →
-    P ⊑ (▷ l !=> v1 ★ ▷ (l !=> v2 -★ Q (LitV (LitBool true)))) →
+    P ⊑ (▷ l ↦ v1 ★ ▷ (l ↦ v2 -★ Q (LitV (LitBool true)))) →
     P ⊑ wp E (Cas (Loc l) e1 e2) Q.
   Proof.
     rewrite /heap_ctx /heap_inv /heap_mapsto=> ?? HN ? HPQ.
     apply (auth_fsa' heap_inv (wp_fsa _) (alter (λ _, Excl v2) l))
-      with N heap_name {[ l ↦ Excl v1 ]}; simpl; eauto 10 with I.
+      with N heap_name {[ l := Excl v1 ]}; simpl; eauto 10 with I.
     rewrite HPQ{HPQ}; apply sep_mono_r, forall_intro=> h; apply wand_intro_l.
     rewrite -assoc; apply const_elim_sep_l=> ?.
     rewrite {1}[(â–·ownP _)%I]pvs_timeless pvs_frame_r; apply wp_strip_pvs.
diff --git a/prelude/base.v b/prelude/base.v
index 08e1f021e..41fb9c5b9 100644
--- a/prelude/base.v
+++ b/prelude/base.v
@@ -433,7 +433,7 @@ Arguments lookup _ _ _ _ !_ !_ / : simpl nomatch.
 (** The singleton map *)
 Class SingletonM K A M := singletonM: K → A → M.
 Instance: Params (@singletonM) 5.
-Notation "{[ x ↦ y ]}" := (singletonM x y) (at level 1) : C_scope.
+Notation "{[ k := a ]}" := (singletonM k a) (at level 1) : C_scope.
 
 (** The function insert [<[k:=a]>m] should update the element at key [k] with
 value [a] in [m]. *)
@@ -628,7 +628,6 @@ Class Lattice A `{SubsetEq A, Union A, Intersection A} : Prop := {
 (** ** Axiomatization of collections *)
 (** The class [SimpleCollection A C] axiomatizes a collection of type [C] with
 elements of type [A]. *)
-Instance: Params (@map) 3.
 Class SimpleCollection A C `{ElemOf A C,
     Empty C, Singleton A C, Union C} : Prop := {
   not_elem_of_empty (x : A) : x ∉ ∅;
diff --git a/prelude/fin_map_dom.v b/prelude/fin_map_dom.v
index 66e2da568..76b90521f 100644
--- a/prelude/fin_map_dom.v
+++ b/prelude/fin_map_dom.v
@@ -61,7 +61,7 @@ Proof. rewrite (dom_insert _). solve_elem_of. Qed.
 Lemma dom_insert_subseteq_compat_l {A} (m : M A) i x X :
   X ⊆ dom D m → X ⊆ dom D (<[i:=x]>m).
 Proof. intros. transitivity (dom D m); eauto using dom_insert_subseteq. Qed.
-Lemma dom_singleton {A} (i : K) (x : A) : dom D {[i ↦ x]} ≡ {[ i ]}.
+Lemma dom_singleton {A} (i : K) (x : A) : dom D {[i := x]} ≡ {[ i ]}.
 Proof. rewrite <-insert_empty, dom_insert, dom_empty; solve_elem_of. Qed.
 Lemma dom_delete {A} (m : M A) i : dom D (delete i m) ≡ dom D m ∖ {[ i ]}.
 Proof.
@@ -123,7 +123,7 @@ Lemma dom_alter_L {A} f (m : M A) i : dom D (alter f i m) = dom D m.
 Proof. unfold_leibniz; apply dom_alter. Qed.
 Lemma dom_insert_L {A} (m : M A) i x : dom D (<[i:=x]>m) = {[ i ]} ∪ dom D m.
 Proof. unfold_leibniz; apply dom_insert. Qed.
-Lemma dom_singleton_L {A} (i : K) (x : A) : dom D {[i ↦ x]} = {[ i ]}.
+Lemma dom_singleton_L {A} (i : K) (x : A) : dom D {[i := x]} = {[ i ]}.
 Proof. unfold_leibniz; apply dom_singleton. Qed.
 Lemma dom_delete_L {A} (m : M A) i : dom D (delete i m) = dom D m ∖ {[ i ]}.
 Proof. unfold_leibniz; apply dom_delete. Qed.
diff --git a/prelude/fin_maps.v b/prelude/fin_maps.v
index 5aee3f6c1..bf5d5d49e 100644
--- a/prelude/fin_maps.v
+++ b/prelude/fin_maps.v
@@ -348,7 +348,7 @@ Proof.
 Qed.
 Lemma delete_empty {A} i : delete i (∅ : M A) = ∅.
 Proof. rewrite <-(partial_alter_self ∅) at 2. by rewrite lookup_empty. Qed.
-Lemma delete_singleton {A} i (x : A) : delete i {[i ↦ x]} = ∅.
+Lemma delete_singleton {A} i (x : A) : delete i {[i := x]} = ∅.
 Proof. setoid_rewrite <-partial_alter_compose. apply delete_empty. Qed.
 Lemma delete_commute {A} (m : M A) i j :
   delete i (delete j m) = delete j (delete i m).
@@ -482,39 +482,39 @@ Proof.
   * eauto using insert_delete_subset.
   * by rewrite lookup_delete.
 Qed.
-Lemma insert_empty {A} i (x : A) : <[i:=x]>∅ = {[i ↦ x]}.
+Lemma insert_empty {A} i (x : A) : <[i:=x]>∅ = {[i := x]}.
 Proof. done. Qed.
 
 (** ** Properties of the singleton maps *)
 Lemma lookup_singleton_Some {A} i j (x y : A) :
-  {[i ↦ x]} !! j = Some y ↔ i = j ∧ x = y.
+  {[i := x]} !! j = Some y ↔ i = j ∧ x = y.
 Proof.
   rewrite <-insert_empty,lookup_insert_Some, lookup_empty; intuition congruence.
 Qed.
-Lemma lookup_singleton_None {A} i j (x : A) : {[i ↦ x]} !! j = None ↔ i ≠ j.
+Lemma lookup_singleton_None {A} i j (x : A) : {[i := x]} !! j = None ↔ i ≠ j.
 Proof. rewrite <-insert_empty,lookup_insert_None, lookup_empty; tauto. Qed.
-Lemma lookup_singleton {A} i (x : A) : {[i ↦ x]} !! i = Some x.
+Lemma lookup_singleton {A} i (x : A) : {[i := x]} !! i = Some x.
 Proof. by rewrite lookup_singleton_Some. Qed.
-Lemma lookup_singleton_ne {A} i j (x : A) : i ≠ j → {[i ↦ x]} !! j = None.
+Lemma lookup_singleton_ne {A} i j (x : A) : i ≠ j → {[i := x]} !! j = None.
 Proof. by rewrite lookup_singleton_None. Qed.
-Lemma map_non_empty_singleton {A} i (x : A) : {[i ↦ x]} ≠ ∅.
+Lemma map_non_empty_singleton {A} i (x : A) : {[i := x]} ≠ ∅.
 Proof.
   intros Hix. apply (f_equal (!! i)) in Hix.
   by rewrite lookup_empty, lookup_singleton in Hix.
 Qed.
-Lemma insert_singleton {A} i (x y : A) : <[i:=y]>{[i ↦ x]} = {[i ↦ y]}.
+Lemma insert_singleton {A} i (x y : A) : <[i:=y]>{[i := x]} = {[i := y]}.
 Proof.
   unfold singletonM, map_singleton, insert, map_insert.
   by rewrite <-partial_alter_compose.
 Qed.
-Lemma alter_singleton {A} (f : A → A) i x : alter f i {[i ↦ x]} = {[i ↦ f x]}.
+Lemma alter_singleton {A} (f : A → A) i x : alter f i {[i := x]} = {[i := f x]}.
 Proof.
   intros. apply map_eq. intros i'. destruct (decide (i = i')) as [->|?].
   * by rewrite lookup_alter, !lookup_singleton.
   * by rewrite lookup_alter_ne, !lookup_singleton_ne.
 Qed.
 Lemma alter_singleton_ne {A} (f : A → A) i j x :
-  i ≠ j → alter f i {[j ↦ x]} = {[j ↦ x]}.
+  i ≠ j → alter f i {[j := x]} = {[j := x]}.
 Proof.
   intros. apply map_eq; intros i'. by destruct (decide (i = i')) as [->|?];
     rewrite ?lookup_alter, ?lookup_singleton_ne, ?lookup_alter_ne by done.
@@ -538,12 +538,12 @@ Proof.
   * by rewrite lookup_omap, !lookup_insert.
   * by rewrite lookup_omap, !lookup_insert_ne, lookup_omap by done.
 Qed.
-Lemma map_fmap_singleton {A B} (f : A → B) i x : f <$> {[i ↦ x]} = {[i ↦ f x]}.
+Lemma map_fmap_singleton {A B} (f : A → B) i x : f <$> {[i := x]} = {[i := f x]}.
 Proof.
   by unfold singletonM, map_singleton; rewrite fmap_insert, map_fmap_empty.
 Qed.
 Lemma omap_singleton {A B} (f : A → option B) i x y :
-  f x = Some y → omap f {[ i ↦ x ]} = {[ i ↦ y ]}.
+  f x = Some y → omap f {[ i := x ]} = {[ i := y ]}.
 Proof.
   intros. unfold singletonM, map_singleton.
   by erewrite omap_insert, omap_empty by eauto.
@@ -898,7 +898,7 @@ Lemma insert_merge m1 m2 i x y z :
   <[i:=x]>(merge f m1 m2) = merge f (<[i:=y]>m1) (<[i:=z]>m2).
 Proof. by intros; apply partial_alter_merge. Qed.
 Lemma merge_singleton i x y z :
-  f (Some y) (Some z) = Some x → merge f {[i ↦ y]} {[i ↦ z]} = {[i ↦ x]}.
+  f (Some y) (Some z) = Some x → merge f {[i := y]} {[i := z]} = {[i := x]}.
 Proof.
   intros. by erewrite <-!insert_empty, <-insert_merge, merge_empty by eauto.
 Qed.
@@ -1000,23 +1000,23 @@ Proof. rewrite map_disjoint_spec, eq_None_not_Some. intros ?? [??]; eauto. Qed.
 Lemma map_disjoint_Some_r {A} (m1 m2 : M A) i x:
   m1 ⊥ₘ m2 → m2 !! i = Some x → m1 !! i = None.
 Proof. rewrite (symmetry_iff map_disjoint). apply map_disjoint_Some_l. Qed.
-Lemma map_disjoint_singleton_l {A} (m: M A) i x : {[i↦x]} ⊥ₘ m ↔ m !! i = None.
+Lemma map_disjoint_singleton_l {A} (m: M A) i x : {[i:=x]} ⊥ₘ m ↔ m !! i = None.
 Proof.
   split; [|rewrite !map_disjoint_spec].
-  * intro. apply (map_disjoint_Some_l {[i ↦ x]} _ _ x);
+  * intro. apply (map_disjoint_Some_l {[i := x]} _ _ x);
       auto using lookup_singleton.
   * intros ? j y1 y2. destruct (decide (i = j)) as [->|].
     + rewrite lookup_singleton. intuition congruence.
     + by rewrite lookup_singleton_ne.
 Qed.
 Lemma map_disjoint_singleton_r {A} (m : M A) i x :
-  m ⊥ₘ {[i ↦ x]} ↔ m !! i = None.
+  m ⊥ₘ {[i := x]} ↔ m !! i = None.
 Proof. by rewrite (symmetry_iff map_disjoint), map_disjoint_singleton_l. Qed.
 Lemma map_disjoint_singleton_l_2 {A} (m : M A) i x :
-  m !! i = None → {[i ↦ x]} ⊥ₘ m.
+  m !! i = None → {[i := x]} ⊥ₘ m.
 Proof. by rewrite map_disjoint_singleton_l. Qed.
 Lemma map_disjoint_singleton_r_2 {A} (m : M A) i x :
-  m !! i = None → m ⊥ₘ {[i ↦ x]}.
+  m !! i = None → m ⊥ₘ {[i := x]}.
 Proof. by rewrite map_disjoint_singleton_r. Qed.
 Lemma map_disjoint_delete_l {A} (m1 m2 : M A) i : m1 ⊥ₘ m2 → delete i m1 ⊥ₘ m2.
 Proof.
@@ -1233,7 +1233,7 @@ Proof. by rewrite map_disjoint_union_l. Qed.
 Lemma map_disjoint_union_r_2 {A} (m1 m2 m3 : M A) :
   m1 ⊥ₘ m2 → m1 ⊥ₘ m3 → m1 ⊥ₘ m2 ∪ m3.
 Proof. by rewrite map_disjoint_union_r. Qed.
-Lemma insert_union_singleton_l {A} (m : M A) i x : <[i:=x]>m = {[i ↦ x]} ∪ m.
+Lemma insert_union_singleton_l {A} (m : M A) i x : <[i:=x]>m = {[i := x]} ∪ m.
 Proof.
   apply map_eq. intros j. apply option_eq. intros y.
   rewrite lookup_union_Some_raw.
@@ -1242,7 +1242,7 @@ Proof.
   * rewrite !lookup_singleton_ne, lookup_insert_ne; intuition congruence.
 Qed.
 Lemma insert_union_singleton_r {A} (m : M A) i x :
-  m !! i = None → <[i:=x]>m = m ∪ {[i ↦ x]}.
+  m !! i = None → <[i:=x]>m = m ∪ {[i := x]}.
 Proof.
   intro. rewrite insert_union_singleton_l, map_union_comm; [done |].
   by apply map_disjoint_singleton_l.
@@ -1446,15 +1446,15 @@ End theorems.
 (** * Tactics *)
 (** The tactic [decompose_map_disjoint] simplifies occurrences of [disjoint]
 in the hypotheses that involve the empty map [∅], the union [(∪)] or insert
-[<[_:=_]>] operation, the singleton [{[_↦ _]}] map, and disjointness of lists of
+[<[_:=_]>] operation, the singleton [{[_:= _]}] map, and disjointness of lists of
 maps. This tactic does not yield any information loss as all simplifications
 performed are reversible. *)
 Ltac decompose_map_disjoint := repeat
   match goal with
   | H : _ ∪ _ ⊥ₘ _ |- _ => apply map_disjoint_union_l in H; destruct H
   | H : _ ⊥ₘ _ ∪ _ |- _ => apply map_disjoint_union_r in H; destruct H
-  | H : {[ _ ↦ _ ]} ⊥ₘ _ |- _ => apply map_disjoint_singleton_l in H
-  | H : _ ⊥ₘ {[ _ ↦ _ ]} |- _ =>  apply map_disjoint_singleton_r in H
+  | H : {[ _ := _ ]} ⊥ₘ _ |- _ => apply map_disjoint_singleton_l in H
+  | H : _ ⊥ₘ {[ _ := _ ]} |- _ =>  apply map_disjoint_singleton_r in H
   | H : <[_:=_]>_ ⊥ₘ _ |- _ => apply map_disjoint_insert_l in H; destruct H
   | H : _ ⊥ₘ <[_:=_]>_ |- _ => apply map_disjoint_insert_r in H; destruct H
   | H : ⋃ _ ⊥ₘ _ |- _ => apply map_disjoint_union_list_l in H
@@ -1478,9 +1478,9 @@ Ltac solve_map_disjoint :=
 Hint Extern 1 (_ ⊥ₘ _) => done : map_disjoint.
 Hint Extern 2 (∅ ⊥ₘ _) => apply map_disjoint_empty_l : map_disjoint.
 Hint Extern 2 (_ ⊥ₘ ∅) => apply map_disjoint_empty_r : map_disjoint.
-Hint Extern 2 ({[ _ ↦ _ ]} ⊥ₘ _) =>
+Hint Extern 2 ({[ _ := _ ]} ⊥ₘ _) =>
   apply map_disjoint_singleton_l_2 : map_disjoint.
-Hint Extern 2 (_ ⊥ₘ {[ _ ↦ _ ]}) =>
+Hint Extern 2 (_ ⊥ₘ {[ _ := _ ]}) =>
   apply map_disjoint_singleton_r_2 : map_disjoint.
 Hint Extern 2 (_ ∪ _ ⊥ₘ _) => apply map_disjoint_union_l_2 : map_disjoint.
 Hint Extern 2 (_ ⊥ₘ _ ∪ _) => apply map_disjoint_union_r_2 : map_disjoint.
@@ -1512,7 +1512,7 @@ Tactic Notation "simpl_map" "by" tactic3(tac) := repeat
     rewrite lookup_alter in H || rewrite lookup_alter_ne in H by tac
   | H : context[ (delete _ _) !! _] |- _ =>
     rewrite lookup_delete in H || rewrite lookup_delete_ne in H by tac
-  | H : context[ {[ _ ↦ _ ]} !! _ ] |- _ =>
+  | H : context[ {[ _ := _ ]} !! _ ] |- _ =>
     rewrite lookup_singleton in H || rewrite lookup_singleton_ne in H by tac
   | H : context[ (_ <$> _) !! _ ] |- _ => rewrite lookup_fmap in H
   | H : context[ (omap _ _) !! _ ] |- _ => rewrite lookup_omap in H
@@ -1529,7 +1529,7 @@ Tactic Notation "simpl_map" "by" tactic3(tac) := repeat
     rewrite lookup_alter || rewrite lookup_alter_ne by tac
   | |- context[ (delete _ _) !! _ ] =>
     rewrite lookup_delete || rewrite lookup_delete_ne by tac
-  | |- context[ {[ _ ↦ _ ]} !! _ ] =>
+  | |- context[ {[ _ := _ ]} !! _ ] =>
     rewrite lookup_singleton || rewrite lookup_singleton_ne by tac
   | |- context[ (_ <$> _) !! _ ] => rewrite lookup_fmap
   | |- context[ (omap _ _) !! _ ] => rewrite lookup_omap
@@ -1546,7 +1546,7 @@ Tactic Notation "simpl_map" := simpl_map by eauto with simpl_map map_disjoint.
 
 Hint Extern 80 ((_ ∪ _) !! _ = Some _) => apply lookup_union_Some_l : simpl_map.
 Hint Extern 81 ((_ ∪ _) !! _ = Some _) => apply lookup_union_Some_r : simpl_map.
-Hint Extern 80 ({[ _↦_ ]} !! _ = Some _) => apply lookup_singleton : simpl_map.
+Hint Extern 80 ({[ _:=_ ]} !! _ = Some _) => apply lookup_singleton : simpl_map.
 Hint Extern 80 (<[_:=_]> _ !! _ = Some _) => apply lookup_insert : simpl_map.
 
 (** Now we take everything together and also discharge conflicting look ups,
@@ -1558,8 +1558,8 @@ Tactic Notation "simplify_map_equality" "by" tactic3(tac) :=
   | _ => progress simpl_map by tac
   | _ => progress simplify_equality
   | _ => progress simpl_option by tac
-  | H : {[ _ ↦ _ ]} !! _ = None |- _ => rewrite lookup_singleton_None in H
-  | H : {[ _ ↦ _ ]} !! _ = Some _ |- _ =>
+  | H : {[ _ := _ ]} !! _ = None |- _ => rewrite lookup_singleton_None in H
+  | H : {[ _ := _ ]} !! _ = Some _ |- _ =>
     rewrite lookup_singleton_Some in H; destruct H
   | H1 : ?m1 !! ?i = Some ?x, H2 : ?m2 !! ?i = Some ?y |- _ =>
     let H3 := fresh in
@@ -1572,8 +1572,8 @@ Tactic Notation "simplify_map_equality" "by" tactic3(tac) :=
     apply map_union_cancel_l in H; [|by tac|by tac]
   | H : _ ∪ ?m = _ ∪ ?m |- _ =>
     apply map_union_cancel_r in H; [|by tac|by tac]
-  | H : {[?i ↦ ?x]} = ∅ |- _ => by destruct (map_non_empty_singleton i x)
-  | H : ∅ = {[?i ↦ ?x]} |- _ => by destruct (map_non_empty_singleton i x)
+  | H : {[?i := ?x]} = ∅ |- _ => by destruct (map_non_empty_singleton i x)
+  | H : ∅ = {[?i := ?x]} |- _ => by destruct (map_non_empty_singleton i x)
   | H : ?m !! ?i = Some _, H2 : ?m !! ?j = None |- _ =>
      unless (i ≠ j) by done;
      assert (i ≠ j) by (by intros ?; simplify_equality)
diff --git a/prelude/hashset.v b/prelude/hashset.v
index 148d431b2..b6f3f85f9 100644
--- a/prelude/hashset.v
+++ b/prelude/hashset.v
@@ -23,7 +23,7 @@ Instance hashset_elem_of: ElemOf A (hashset hash) := λ x m, ∃ l,
 Program Instance hashset_empty: Empty (hashset hash) := Hashset ∅ _.
 Next Obligation. by intros n X; simpl_map. Qed.
 Program Instance hashset_singleton: Singleton A (hashset hash) := λ x,
-  Hashset {[ hash x ↦ [x] ]} _.
+  Hashset {[ hash x := [x] ]} _.
 Next Obligation.
   intros x n l [<- <-]%lookup_singleton_Some.
   rewrite Forall_singleton; auto using NoDup_singleton.
diff --git a/prelude/mapset.v b/prelude/mapset.v
index 630e40547..e0f3dff01 100644
--- a/prelude/mapset.v
+++ b/prelude/mapset.v
@@ -17,7 +17,7 @@ Instance mapset_elem_of: ElemOf K (mapset M) := λ x X,
   mapset_car X !! x = Some ().
 Instance mapset_empty: Empty (mapset M) := Mapset ∅.
 Instance mapset_singleton: Singleton K (mapset M) := λ x,
-  Mapset {[ x ↦ () ]}.
+  Mapset {[ x := () ]}.
 Instance mapset_union: Union (mapset M) := λ X1 X2,
   let (m1) := X1 in let (m2) := X2 in Mapset (m1 ∪ m2).
 Instance mapset_intersection: Intersection (mapset M) := λ X1 X2,
diff --git a/program_logic/ghost_ownership.v b/program_logic/ghost_ownership.v
index cfcae20e5..de694e399 100644
--- a/program_logic/ghost_ownership.v
+++ b/program_logic/ghost_ownership.v
@@ -17,7 +17,7 @@ Class inG (Λ : language) (Σ : gid → iFunctor) (A : cmraT) := InG {
 }.
 
 Definition to_globalF `{inG Λ Σ A} (γ : gname) (a : A) : iGst Λ (globalF Σ) :=
-  iprod_singleton inG_id {[ γ ↦ cmra_transport inG_prf a ]}.
+  iprod_singleton inG_id {[ γ := cmra_transport inG_prf a ]}.
 Definition own `{inG Λ Σ A} (γ : gname) (a : A) : iProp Λ (globalF Σ) :=
   ownG (to_globalF γ a).
 Instance: Params (@to_globalF) 5.
diff --git a/program_logic/ownership.v b/program_logic/ownership.v
index 7327ff100..c13c8fea3 100644
--- a/program_logic/ownership.v
+++ b/program_logic/ownership.v
@@ -1,7 +1,7 @@
 From program_logic Require Export model.
 
 Definition ownI {Λ Σ} (i : positive) (P : iProp Λ Σ) : iProp Λ Σ :=
-  uPred_ownM (Res {[ i ↦ to_agree (Next (iProp_unfold P)) ]} ∅ ∅).
+  uPred_ownM (Res {[ i := to_agree (Next (iProp_unfold P)) ]} ∅ ∅).
 Arguments ownI {_ _} _ _%I.
 Definition ownP {Λ Σ} (σ: state Λ) : iProp Λ Σ := uPred_ownM (Res ∅ (Excl σ) ∅).
 Definition ownG {Λ Σ} (m: iGst Λ Σ) : iProp Λ Σ := uPred_ownM (Res ∅ ∅ (Some m)).
@@ -66,7 +66,6 @@ Proof. apply (uPred.always_entails_r _ _), ownG_valid. Qed.
 Global Instance ownG_timeless m : Timeless m → TimelessP (ownG m).
 Proof. rewrite /ownG; apply _. Qed.
 
-
 (* inversion lemmas *)
 Lemma ownI_spec r n i P :
   ✓{n} r →
diff --git a/program_logic/pviewshifts.v b/program_logic/pviewshifts.v
index 6127a0a17..1bc285f97 100644
--- a/program_logic/pviewshifts.v
+++ b/program_logic/pviewshifts.v
@@ -121,7 +121,7 @@ Proof.
   intros ? r [|n] ? HP rf [|k] Ef σ ???; try lia.
   destruct (wsat_alloc k E Ef σ rf P r) as (i&?&?&?); auto.
   { apply uPred_weaken with r n; eauto. }
-  exists (Res {[ i ↦ to_agree (Next (iProp_unfold P)) ]} ∅ ∅).
+  exists (Res {[ i := to_agree (Next (iProp_unfold P)) ]} ∅ ∅).
   by split; [by exists i; split; rewrite /uPred_holds /=|].
 Qed.
 
diff --git a/program_logic/wsat.v b/program_logic/wsat.v
index c602d2912..d25dbd25c 100644
--- a/program_logic/wsat.v
+++ b/program_logic/wsat.v
@@ -137,7 +137,7 @@ Qed.
 Lemma wsat_alloc n E1 E2 σ r P rP :
   ¬set_finite E1 → P n rP → wsat (S n) (E1 ∪ E2) σ (rP ⋅ r) →
   ∃ i, wsat (S n) (E1 ∪ E2) σ
-         (Res {[i ↦ to_agree (Next (iProp_unfold P))]} ∅ ∅ ⋅ r) ∧
+         (Res {[i := to_agree (Next (iProp_unfold P))]} ∅ ∅ ⋅ r) ∧
        wld r !! i = None ∧ i ∈ E1.
 Proof.
   intros HE1 ? [rs [Hval Hσ HE Hwld]].
-- 
GitLab