From 0d22cb1002a93dc1508dd39a711a9fa08fa37273 Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Tue, 5 May 2020 13:35:36 +0200
Subject: [PATCH] Notation for binder_insert.

---
 theories/logrel/environments.v      |  5 ++++-
 theories/logrel/term_typing_rules.v | 13 ++++++-------
 2 files changed, 10 insertions(+), 8 deletions(-)

diff --git a/theories/logrel/environments.v b/theories/logrel/environments.v
index faade4b..e4ab014 100644
--- a/theories/logrel/environments.v
+++ b/theories/logrel/environments.v
@@ -1,6 +1,9 @@
 From actris.logrel Require Export term_types.
 From iris.proofmode Require Import tactics.
 
+Notation "<![ b := x ]!>" :=
+  (binder_insert b x%lty) (at level 5, right associativity).
+
 Definition env_ltyped {Σ} (Γ : gmap string (ltty Σ))
     (vs : gmap string val) : iProp Σ :=
   ([∗ map] i ↦ A ∈ Γ, ∃ v, ⌜vs !! i = Some v⌝ ∗ ltty_car A v)%I.
@@ -32,7 +35,7 @@ Section env.
   Lemma env_ltyped_insert Γ vs x A v :
     ltty_car A v -∗
     env_ltyped Γ vs -∗
-    env_ltyped (binder_insert x A Γ) (binder_insert x v vs).
+    env_ltyped (<![x:=A]!> Γ) (<![x:=v]!> vs).
   Proof.
     destruct x as [|x]=> /=; first by auto.
     iIntros "HA HΓ".
diff --git a/theories/logrel/term_typing_rules.v b/theories/logrel/term_typing_rules.v
index b5d3585..178d73f 100644
--- a/theories/logrel/term_typing_rules.v
+++ b/theories/logrel/term_typing_rules.v
@@ -72,7 +72,7 @@ Section properties.
   Qed.
 
   Lemma ltyped_lam Γ Γ' x e A1 A2 :
-    (binder_insert x A1 Γ ⊨ e : A2 ⫤ Γ') -∗
+    (<![x:=A1]!> Γ ⊨ e : A2 ⫤ Γ') -∗
     Γ ⊨ (λ: x, e) : A1 ⊸ A2 ⫤ ∅.
   Proof.
     iIntros "#He" (vs) "!> HΓ /=".
@@ -89,7 +89,7 @@ Section properties.
   (* Typing rule for introducing copyable functions *)
   Lemma ltyped_rec Γ Γ' Γ'' f x e A1 A2 :
     env_copy Γ Γ' -∗
-    (binder_insert f (A1 → A2)%lty (binder_insert x A1 Γ') ⊨ e : A2 ⫤ Γ'') -∗
+    (<![f:=A1 → A2]!> $ <![x:=A1]!> Γ' ⊨ e : A2 ⫤ Γ'') -∗
     Γ ⊨ (rec: f x := e) : A1 → A2 ⫤ ∅.
   Proof.
     iIntros "#Hcopy #He". iIntros (vs) "!> HΓ /=". iApply wp_fupd. wp_pures.
@@ -98,8 +98,7 @@ Section properties.
     iModIntro. iSplitL; last by iApply env_ltyped_empty.
     iLöb as "IH".
     iIntros (v) "!> HA1". wp_pures. set (r := RecV f x _).
-    iDestruct ("He" $!(binder_insert f r (binder_insert x v vs))
-                  with "[HΓ HA1]") as "He'".
+    iDestruct ("He" $! (<![f:=r]!> $ <![x:=v]!> vs) with "[HΓ HA1]") as "He'".
     { iApply (env_ltyped_insert with "IH").
       iApply (env_ltyped_insert with "HA1 HΓ"). }
     iDestruct (wp_wand _ _ _ _ (ltty_car A2) with "He' []") as "He'".
@@ -112,7 +111,7 @@ Section properties.
   Qed.
 
   Lemma ltyped_let Γ1 Γ2 Γ3 x e1 e2 A1 A2 :
-    (Γ1 ⊨ e1 : A1 ⫤ Γ2) -∗ (binder_insert x A1 Γ2 ⊨ e2 : A2 ⫤ Γ3) -∗
+    (Γ1 ⊨ e1 : A1 ⫤ Γ2) -∗ (<![x:=A1]!> Γ2 ⊨ e2 : A2 ⫤ Γ3) -∗
     Γ1 ⊨ (let: x:=e1 in e2) : A2 ⫤ binder_delete x Γ3.
   Proof.
     iIntros "#He1 #He2 !>". iIntros (vs) "HΓ1"=> /=.
@@ -243,7 +242,7 @@ Section properties.
 
   Lemma ltyped_unpack {k} Γ1 Γ2 Γ3 (x : string) e (C : lty Σ k → ltty Σ) A :
     Γ1 !! x = Some (lty_exist C) →
-    (∀ X, binder_insert x (C X) Γ1 ⊨ e : A ⫤ Γ2) -∗
+    (∀ X, <![x:=C X]!> Γ1 ⊨ e : A ⫤ Γ2) -∗
     (Γ1 ⊨ e : A ⫤ Γ2).
   Proof.
     iIntros (Hx) "#He". iIntros (vs) "!> HΓ".
@@ -443,7 +442,7 @@ Section properties.
     Lemma ltyped_recv_texist {kt} Γ1 Γ2 (xc : string) (x : binder) (e : expr)
         (A : ltys Σ kt → ltty Σ) (S : ltys Σ kt → lsty Σ) (B : ltty Σ) :
       (∀ Ys,
-        binder_insert x (A Ys) (<[xc:=(chan (S Ys))%lty]> Γ1) ⊨ e : B ⫤ Γ2) -∗
+        <![x:=A Ys]!> $ <[xc:=(chan (S Ys))%lty]> Γ1 ⊨ e : B ⫤ Γ2) -∗
       <[xc:=(chan (<??.. Xs> TY A Xs; S Xs))%lty]> Γ1 ⊨
         (let: x := recv xc in e) : B ⫤ binder_delete x Γ2.
     Proof.
-- 
GitLab