From ad13438ba48cf8b76cb11dfcfb57b80bc3bbf19b Mon Sep 17 00:00:00 2001
From: Ralf Jung <jung@mpi-sws.org>
Date: Tue, 27 Oct 2020 16:40:34 +0100
Subject: [PATCH] change gen_inv_heap notation to get out of the way of
 persistent points-to

---
 theories/base_logic/lib/gen_inv_heap.v | 16 ++++++++--------
 1 file changed, 8 insertions(+), 8 deletions(-)

diff --git a/theories/base_logic/lib/gen_inv_heap.v b/theories/base_logic/lib/gen_inv_heap.v
index a43b372b5..4ff3256eb 100644
--- a/theories/base_logic/lib/gen_inv_heap.v
+++ b/theories/base_logic/lib/gen_inv_heap.v
@@ -67,12 +67,12 @@ Section definitions.
 
 End definitions.
 
-Local Notation "l '↦□' I" := (inv_mapsto l I%stdpp%type)
-  (at level 20, format "l  '↦□'  I") : bi_scope.
-
 Local Notation "l '↦_' I v" := (inv_mapsto_own l v I%stdpp%type)
   (at level 20, I at level 9, format "l  '↦_' I  v") : bi_scope.
 
+Local Notation "l '↦_' I □" := (inv_mapsto l I%stdpp%type)
+  (at level 20, I at level 9, format "l  '↦_' I  '□'") : bi_scope.
+
 (* [inv_heap_inv] has no parameters to infer the types from, so we need to
    make them explicit. *)
 Arguments inv_heap_inv _ _ {_ _ _ _ _ _}.
@@ -133,7 +133,7 @@ Section inv_heap.
   (** * Helpers *)
 
   Lemma inv_mapsto_lookup_Some l h I :
-    l ↦□ I -∗ own (inv_heap_name gG) (● to_inv_heap h) -∗
+    l ↦_I □ -∗ own (inv_heap_name gG) (● to_inv_heap h) -∗
     ⌜∃ v I', h !! l = Some (v, I') ∧ ∀ w, I w ↔ I' w ⌝.
   Proof.
     iIntros "Hl_inv Hâ—¯".
@@ -179,10 +179,10 @@ Section inv_heap.
   Global Instance inv_heap_inv_persistent : Persistent (inv_heap_inv L V).
   Proof. apply _. Qed.
 
-  Global Instance inv_mapsto_persistent l I : Persistent (l ↦□ I).
+  Global Instance inv_mapsto_persistent l I : Persistent (l ↦_I □).
   Proof. apply _. Qed.
 
-  Global Instance inv_mapsto_timeless l I : Timeless (l ↦□ I).
+  Global Instance inv_mapsto_timeless l I : Timeless (l ↦_I □).
   Proof. apply _. Qed.
 
   Global Instance inv_mapsto_own_timeless l v I : Timeless (l ↦_I v).
@@ -215,7 +215,7 @@ Section inv_heap.
       + iModIntro. by rewrite /inv_mapsto_own to_inv_heap_singleton.
   Qed.
 
-  Lemma inv_mapsto_own_inv l v I : l ↦_I v -∗ l ↦□ I.
+  Lemma inv_mapsto_own_inv l v I : l ↦_I v -∗ l ↦_I □.
   Proof.
     apply own_mono, auth_frag_mono. rewrite singleton_included pair_included.
     right. split; [apply: ucmra_unit_least|done].
@@ -266,7 +266,7 @@ Section inv_heap.
 
   Lemma inv_mapsto_acc l I E :
     ↑inv_heapN ⊆ E →
-    inv_heap_inv L V -∗ l ↦□ I ={E, E ∖ ↑inv_heapN}=∗
+    inv_heap_inv L V -∗ l ↦_I □ ={E, E ∖ ↑inv_heapN}=∗
     ∃ v, ⌜I v⌝ ∗ l ↦ v ∗ (l ↦ v ={E ∖ ↑inv_heapN, E}=∗ ⌜True⌝).
   Proof.
     iIntros (HN) "#Hinv Hl_inv".
-- 
GitLab