From 9c41c8ad53b0017d3631b6b421d06bb32adaaca3 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Mon, 21 Sep 2020 00:30:50 +0200 Subject: [PATCH] Lookup function on env. --- theories/logrel/environments.v | 22 ++++++++++++++++++---- theories/logrel/lib/mutex.v | 8 ++++---- theories/logrel/session_typing_rules.v | 16 ++++++++-------- theories/logrel/term_typing_rules.v | 20 ++++++++++---------- 4 files changed, 40 insertions(+), 26 deletions(-) diff --git a/theories/logrel/environments.v b/theories/logrel/environments.v index 4ea545f..bd6d74d 100644 --- a/theories/logrel/environments.v +++ b/theories/logrel/environments.v @@ -57,6 +57,12 @@ Arguments env_filter_ne _ !_ !_ / : simpl nomatch. factored out. *) Arguments filter _ _ _ _ _ !_ / : simpl nomatch. +Instance env_lookup {Σ} : Lookup string (ltty Σ) (env Σ) := λ x Γ, + match env_filter_eq x Γ with + | [EnvItem _ A] => Some A + | _ => None + end. + Definition env_cons {Σ} (b : binder) (A : ltty Σ) (Γ : env Σ) : env Σ := if b is BNamed x then EnvItem x A :: env_filter_ne x Γ else Γ. @@ -84,10 +90,6 @@ Section env. - rewrite filter_cons_True /=; last naive_solver. by rewrite -Permutation_middle -IH. Qed. - Lemma env_filter_eq_perm' Γ Γ' x : - env_filter_eq x Γ = Γ' → - Γ ≡ₚ Γ' ++ env_filter_ne x Γ. - Proof. intros <-. apply env_filter_eq_perm. Qed. Lemma env_filter_ne_anon Γ : env_filter_ne BAnon Γ = Γ. Proof. induction Γ as [|[y A] Γ IH]; by f_equal/=. Qed. @@ -104,6 +106,18 @@ Section env. env_filter_ne x (EnvItem y A :: Γ) = EnvItem y A :: env_filter_ne x Γ. Proof. intros. rewrite /env_filter_ne filter_cons_True; naive_solver. Qed. + Lemma env_lookup_perm Γ x A: + Γ !! x = Some A → + Γ ≡ₚ EnvItem x A :: env_filter_ne x Γ. + Proof. + rewrite /lookup /env_lookup=> ?. + destruct (env_filter_eq x Γ) as [|[x' ?] [|??]] eqn:Hx; simplify_eq/=. + assert (EnvItem x' A ∈ env_filter_eq x Γ) + as [? _]%elem_of_list_filter; simplify_eq/=. + { rewrite Hx. set_solver. } + by rewrite {1}(env_filter_eq_perm Γ x') Hx. + Qed. + (** env typing *) Global Instance env_ltyped_Permutation vs : Proper ((≡ₚ) ==> (⊣⊢)) (@env_ltyped Σ vs). diff --git a/theories/logrel/lib/mutex.v b/theories/logrel/lib/mutex.v index 548959d..ed39324 100644 --- a/theories/logrel/lib/mutex.v +++ b/theories/logrel/lib/mutex.v @@ -118,10 +118,10 @@ Section rules. Qed. Lemma ltyped_mutex_acquire Γ (x : string) A : - env_filter_eq x Γ = [EnvItem x (mutex A)] → + Γ !! x = Some (mutex A)%lty → ⊢ Γ ⊨ mutex_acquire x : A ⫤ env_cons x (mutex_guard A) Γ. Proof. - iIntros (HΓx%env_filter_eq_perm' vs) "!> HΓ /=". rewrite {1}HΓx /=. + iIntros (HΓx%env_lookup_perm vs) "!> HΓ /=". rewrite {1}HΓx /=. iDestruct (env_ltyped_cons with "HΓ") as (vl Hvs) "[Hlock HΓ]"; rewrite Hvs. iDestruct "Hlock" as (γ l lk ->) "#Hlock". rewrite /mutex_acquire. wp_apply (acquire_spec with "Hlock"); iIntros "[Hlocked Hinner]". @@ -131,11 +131,11 @@ Section rules. Qed. Lemma ltyped_mutex_release Γ Γ' (x : string) e A : - env_filter_eq x Γ' = [EnvItem x (mutex_guard A)] → + Γ' !! x = Some (mutex_guard A)%lty → (Γ ⊨ e : A ⫤ Γ') -∗ Γ ⊨ mutex_release x e : () ⫤ env_cons x (mutex A) Γ'. Proof. - iIntros (HΓx%env_filter_eq_perm') "#He". iIntros (vs) "!> HΓ /=". + iIntros (HΓx%env_lookup_perm) "#He". iIntros (vs) "!> HΓ /=". wp_apply (wp_wand with "(He HΓ)"). iIntros (v) "[HA HΓ']". rewrite {2}HΓx /=. iDestruct (env_ltyped_cons with "HΓ'") as (vl Hvs) "[Hguard HΓ']"; rewrite Hvs. diff --git a/theories/logrel/session_typing_rules.v b/theories/logrel/session_typing_rules.v index 791ea33..bb1342b 100644 --- a/theories/logrel/session_typing_rules.v +++ b/theories/logrel/session_typing_rules.v @@ -25,11 +25,11 @@ Section session_typing_rules. Qed. Lemma ltyped_send Γ Γ' (x : string) e A S : - env_filter_eq x Γ' = [EnvItem x (chan (<!!> TY A; S))] → + Γ' !! x = Some (chan (<!!> TY A; S))%lty → (Γ ⊨ e : A ⫤ Γ') -∗ Γ ⊨ send x e : () ⫤ env_cons x (chan S) Γ'. Proof. - iIntros (HΓx%env_filter_eq_perm') "#He !>". iIntros (vs) "HΓ /=". + iIntros (HΓx%env_lookup_perm) "#He !>". iIntros (vs) "HΓ /=". wp_apply (wp_wand with "(He HΓ)"); iIntros (v) "[HA HΓ']". rewrite {2}HΓx /=. iDestruct (env_ltyped_cons with "HΓ'") as (c Hvs) "[Hc HΓ']". rewrite Hvs. @@ -49,7 +49,7 @@ Section session_typing_rules. Lemma ltyped_recv_texist {kt} Γ1 Γ2 M x (xc : string) (e : expr) (A : kt -k> ltty Σ) (S : kt -k> lsty Σ) (B : ltty Σ) : - env_filter_eq xc Γ1 = [EnvItem xc (chan (<??> M))] → + Γ1 !! xc = Some (chan (<??> M))%lty → LtyMsgTele M A S → (∀ Ys, env_cons x (ktele_app A Ys) (env_cons xc (chan (ktele_app S Ys)) Γ1) ⊨ e : B ⫤ Γ2) -∗ @@ -57,7 +57,7 @@ Section session_typing_rules. env_filter_eq x (env_filter_ne xc Γ1) ++ env_filter_ne x Γ2. Proof. rewrite /LtyMsgTele. - iIntros (HΓxc%env_filter_eq_perm' HM) "#He !>". iIntros (vs) "HΓ1 /=". + iIntros (HΓxc%env_lookup_perm HM) "#He !>". iIntros (vs) "HΓ1 /=". rewrite {2}HΓxc /=. iDestruct (env_ltyped_cons with "HΓ1") as (c Hvs) "[Hc HΓ1]". rewrite Hvs. rewrite {2}(env_filter_eq_perm (env_filter_ne xc Γ1) x). @@ -79,10 +79,10 @@ Section session_typing_rules. Qed. Lemma ltyped_recv Γ (x : string) A S : - env_filter_eq x Γ = [EnvItem x (chan (<??> TY A; S))] → + Γ !! x = Some (chan (<??> TY A; S))%lty → ⊢ Γ ⊨ recv x : A ⫤ env_cons x (chan S) Γ. Proof. - iIntros (HΓx%env_filter_eq_perm') "!>". iIntros (vs) "HΓ /=". + iIntros (HΓx%env_lookup_perm) "!>". iIntros (vs) "HΓ /=". rewrite {1}HΓx /=. iDestruct (env_ltyped_cons with "HΓ") as (c Hvs) "[Hc HΓ]". rewrite Hvs. wp_recv (v) as "HA". iFrame "HA". iApply env_ltyped_cons; eauto with iFrame. @@ -91,11 +91,11 @@ Section session_typing_rules. Definition select : val := λ: "c" "i", send "c" "i". Lemma ltyped_select Γ (x : string) (i : Z) (S : lsty Σ) Ss : + Γ !! x = Some (chan (lty_select Ss))%lty → Ss !! i = Some S → - env_filter_eq x Γ = [EnvItem x (chan (lty_select Ss))] → ⊢ Γ ⊨ select x #i : () ⫤ env_cons x (chan S) Γ. Proof. - iIntros (Hin HΓx%env_filter_eq_perm'); iIntros "!>" (vs) "HΓ /=". + iIntros (HΓx%env_lookup_perm Hin); iIntros "!>" (vs) "HΓ /=". rewrite {1}HΓx /=. iDestruct (env_ltyped_cons with "HΓ") as (c Hvs) "[Hc HΓ]". rewrite Hvs. rewrite /select. wp_send with "[]"; [by eauto|]. iSplit; [done|]. diff --git a/theories/logrel/term_typing_rules.v b/theories/logrel/term_typing_rules.v index a8fcbbe..fe4d68a 100644 --- a/theories/logrel/term_typing_rules.v +++ b/theories/logrel/term_typing_rules.v @@ -26,10 +26,10 @@ Section term_typing_rules. (** Variable properties *) Lemma ltyped_var Γ (x : string) A : - env_filter_eq x Γ = [EnvItem x A] → + Γ !! x = Some A → ⊢ Γ ⊨ x : A ⫤ env_cons x (copy- A) Γ. Proof. - iIntros (HΓx%env_filter_eq_perm') "!>"; iIntros (vs) "HΓ /=". + iIntros (HΓx%env_lookup_perm) "!>"; iIntros (vs) "HΓ /=". rewrite {1}HΓx /=. iDestruct (env_ltyped_cons with "HΓ") as (v Hvs) "[HA HΓ]". rewrite Hvs. iAssert (ltty_car (copy- A) v)%lty as "#HAm"; [by iApply coreP_intro|]. @@ -196,10 +196,10 @@ Section term_typing_rules. Qed. Lemma ltyped_fst Γ A1 A2 (x : string) : - env_filter_eq x Γ = [EnvItem x (A1 * A2)] → + Γ !! x = Some (A1 * A2)%lty → ⊢ Γ ⊨ Fst x : A1 ⫤ env_cons x (copy- A1 * A2) Γ. Proof. - iIntros (HΓx%env_filter_eq_perm' vs) "!> HΓ /=". rewrite {1}HΓx /=. + iIntros (HΓx%env_lookup_perm vs) "!> HΓ /=". rewrite {1}HΓx /=. iDestruct (env_ltyped_cons with "HΓ") as (v Hvs) "[HA HΓ]"; rewrite Hvs. iDestruct "HA" as (v1 v2 ->) "[HA1 HA2]". wp_pures. iAssert (ltty_car (copy- A1) v1)%lty as "#HA1m"; [by iApply coreP_intro|]. @@ -208,10 +208,10 @@ Section term_typing_rules. Qed. Lemma ltyped_snd Γ A1 A2 (x : string) : - env_filter_eq x Γ = [EnvItem x (A1 * A2)] → + Γ !! x = Some (A1 * A2)%lty → ⊢ Γ ⊨ Snd x : A2 ⫤ env_cons x (A1 * copy- A2) Γ. Proof. - iIntros (HΓx%env_filter_eq_perm' vs) "!> HΓ /=". rewrite {1}HΓx /=. + iIntros (HΓx%env_lookup_perm vs) "!> HΓ /=". rewrite {1}HΓx /=. iDestruct (env_ltyped_cons with "HΓ") as (v Hvs) "[HA HΓ]"; rewrite Hvs. iDestruct "HA" as (v1 v2 ->) "[HA1 HA2]". wp_pures. iAssert (ltty_car (copy- A2) v2)%lty as "#HA2m"; [by iApply coreP_intro|]. @@ -317,10 +317,10 @@ Section term_typing_rules. Qed. Lemma ltyped_load Γ (x : string) A : - env_filter_eq x Γ = [EnvItem x (ref_uniq A)] → + Γ !! x = Some (ref_uniq A)%lty → ⊢ Γ ⊨ ! x : A ⫤ env_cons x (ref_uniq (copy- A)) Γ. Proof. - iIntros (HΓx%env_filter_eq_perm' vs) "!> HΓ /=". rewrite {1}HΓx /=. + iIntros (HΓx%env_lookup_perm vs) "!> HΓ /=". rewrite {1}HΓx /=. iDestruct (env_ltyped_cons with "HΓ") as (v Hvs) "[HA HΓ]"; rewrite Hvs. iDestruct "HA" as (l w ->) "[? HA]". wp_load. iAssert (ltty_car (copy- A) w)%lty as "#HAm"; [by iApply coreP_intro|]. @@ -329,11 +329,11 @@ Section term_typing_rules. Qed. Lemma ltyped_store Γ Γ' (x : string) e A B : - env_filter_eq x Γ' = [EnvItem x (ref_uniq A)] → + Γ' !! x = Some (ref_uniq A)%lty → (Γ ⊨ e : B ⫤ Γ') -∗ Γ ⊨ x <- e : () ⫤ env_cons x (ref_uniq B) Γ'. Proof. - iIntros (HΓx%env_filter_eq_perm') "#He"; iIntros (vs) "!> HΓ /=". + iIntros (HΓx%env_lookup_perm) "#He"; iIntros (vs) "!> HΓ /=". wp_apply (wp_wand with "(He HΓ)"). iIntros (v) "[HB HΓ']". rewrite {2}HΓx /=. iDestruct (env_ltyped_cons with "HΓ'") as (vl Hvs) "[HA HΓ']"; rewrite Hvs. -- GitLab