Skip to content
Snippets Groups Projects
Commit 76e934c0 authored by Daniël Louwrink's avatar Daniël Louwrink
Browse files

Merge branch 'daniel/update_typing_rules' into 'master'

Update typing rules

See merge request iris/actris!15
parents 01ec9471 35a4e878
No related branches found
No related tags found
No related merge requests found
......@@ -27,5 +27,6 @@ theories/logrel/operators.v
theories/logrel/term_typing_judgment.v
theories/logrel/subtyping_rules.v
theories/logrel/term_typing_rules.v
theories/logrel/lib/mutex.v
theories/logrel/examples/double.v
theories/logrel/examples/pair.v
From iris.base_logic.lib Require Import invariants.
From iris.heap_lang Require Export spin_lock.
From actris.logrel Require Export term_types term_typing_judgment subtyping.
From actris.logrel Require Import environments.
From actris.channel Require Import proofmode.
From iris.heap_lang Require Import metatheory.
(** Mutex functions *)
Definition mutex_alloc : val := λ: "x", (newlock #(), ref "x").
Definition mutex_acquire : val := λ: "x", acquire (Fst "x");; ! (Snd "x").
Definition mutex_release : val :=
λ: "guard" "inner", Snd "guard" <- "inner";; release (Fst "guard").
(** Semantic types *)
Definition lty_mutex `{heapG Σ, lockG Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w,
(γ : gname) (l : loc) (lk : val),
w = PairV lk #l
is_lock γ lk ( v_inner, l v_inner ltty_car A v_inner))%I.
Definition lty_mutex_guard `{heapG Σ, lockG Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w,
(γ : gname) (l : loc) (lk : val) (v : val),
w = PairV lk #l
is_lock γ lk ( v_inner, l v_inner ltty_car A v_inner)
spin_lock.locked γ l v)%I.
Instance: Params (@lty_mutex) 3 := {}.
Instance: Params (@lty_mutex_guard) 3 := {}.
Notation "'mutex' A" := (lty_mutex A) (at level 10) : lty_scope.
Notation "'mutex_guard' A" := (lty_mutex_guard A) (at level 10) : lty_scope.
Section properties.
Context `{heapG Σ, lockG Σ}.
Implicit Types A : ltty Σ.
Global Instance lty_mutex_contractive : Contractive lty_mutex.
Proof. solve_contractive. Qed.
Global Instance lty_mutex_ne : NonExpansive lty_mutex.
Proof. solve_proper. Qed.
Global Instance lty_mutex_guard_contractive :
Contractive lty_mutex_guard.
Proof. solve_contractive. Qed.
Global Instance lty_mutex_guard_ne : NonExpansive lty_mutex_guard.
Proof. solve_proper. Qed.
Lemma lty_le_mutex A1 A2 :
(A1 <:> A2) -∗
mutex A1 <: mutex A2.
Proof.
iIntros "#[Hle1 Hle2]" (v) "!>". iDestruct 1 as (γ l lk ->) "Hinv".
iExists γ, l, lk. iSplit; first done.
iApply (spin_lock.is_lock_iff with "Hinv").
iIntros "!> !>". iSplit.
- iDestruct 1 as (v) "[Hl HA]". iExists v. iFrame "Hl". by iApply "Hle1".
- iDestruct 1 as (v) "[Hl HA]". iExists v. iFrame "Hl". by iApply "Hle2".
Qed.
Lemma lty_copyable_mutex A :
lty_copyable (mutex A).
Proof. iIntros (v) "!> #Hv !>". iFrame "Hv". Qed.
Lemma lty_le_mutex_guard A1 A2 :
(A1 <:> A2) -∗
mutex_guard A1 <: mutex_guard A2.
Proof.
iIntros "#[Hle1 Hle2]" (v) "!>".
iDestruct 1 as (γ l lk w ->) "[Hinv [Hlock Hinner]]".
iExists γ, l, lk, w. iSplit; first done.
iFrame "Hlock Hinner". iApply (spin_lock.is_lock_iff with "Hinv").
iIntros "!> !>". iSplit.
- iDestruct 1 as (v) "[Hl HA]". iExists v. iFrame "Hl". by iApply "Hle1".
- iDestruct 1 as (v) "[Hl HA]". iExists v. iFrame "Hl". by iApply "Hle2".
Qed.
End properties.
Section rules.
Context `{heapG Σ, lockG Σ}.
(** Mutex properties *)
Lemma ltyped_mutex_alloc A :
mutex_alloc : A mutex A.
Proof.
iIntros (vs) "!> HΓ /=". iApply wp_value.
iSplitL; last by iApply env_ltyped_empty.
iIntros "!>" (v) "Hv". rewrite /mutex_alloc. wp_pures.
wp_alloc l as "Hl".
wp_bind (newlock _).
set (N := nroot .@ "makelock").
iAssert ( inner, l inner ltty_car A inner)%I with "[Hl Hv]" as "Hlock".
{ iExists v. iFrame "Hl Hv". }
wp_apply (newlock_spec with "Hlock").
iIntros (lk γ) "Hlock".
wp_pures. iExists γ, l, lk. iSplit=> //.
Qed.
Lemma ltyped_mutex_acquire Γ (x : string) A :
Γ !! x = Some (mutex A)%lty
Γ mutex_acquire x : A <[x := (mutex_guard A)%lty]> Γ.
Proof.
iIntros (Hx vs) "!> HΓ /=".
iDestruct (env_ltyped_lookup with "HΓ") as (v Hv) "[HA HΓ]"; first done; rewrite Hv.
rewrite /mutex_acquire. wp_pures.
iDestruct "HA" as (γ l lk ->) "#Hlock".
wp_bind (acquire _).
wp_apply (acquire_spec with "Hlock"). iIntros "[Hlocked Hinner]".
iDestruct "Hinner" as (v) "[Hl HA]".
wp_pures. wp_load.
iFrame "HA".
iAssert (ltty_car (mutex_guard A)%lty (lk, #l)) with "[Hlocked Hl]" as "Hguard".
{ iExists γ, l, lk, v. iSplit=>//. iFrame "Hlocked Hl Hlock". }
iDestruct (env_ltyped_insert _ _ x with "Hguard HΓ") as "HΓ".
rewrite /binder_insert insert_delete (insert_id _ _ _ Hv).
iFrame "HΓ".
Qed.
Lemma ltyped_mutex_release Γ Γ' (x : string) e A :
Γ' !! x = Some (mutex_guard A)%lty
(Γ e : A Γ') -∗
Γ mutex_release x e : () <[x := (mutex A)%lty]> Γ'.
Proof.
iIntros (Hx) "#He". iIntros (vs) "!> HΓ /=".
wp_bind (subst_map vs e).
iApply (wp_wand with "(He HΓ)"). iIntros (v) "[HA HΓ']".
iDestruct (env_ltyped_lookup with "HΓ'") as (g Hg) "[Hguard HΓ']"; first done; rewrite Hg.
iDestruct "Hguard" as (γ l lk inner ->) "(#Hlock & Hlocked & Hinner)".
rewrite /mutex_release.
wp_pures. wp_store. wp_pures.
iAssert ( inner, l inner ltty_car A inner)%I
with "[Hinner HA]" as "Hinner".
{ iExists v. iFrame "Hinner HA". }
wp_apply (release_spec γ _ ( inner, l inner ltty_car A inner)%I
with "[Hlocked Hinner]").
{ iFrame "Hlock Hlocked".
iDestruct "Hinner" as (w) "[Hl HA]". eauto with iFrame. }
iIntros "_". wp_pures.
iSplit=> //.
iAssert (ltty_car (mutex A)%lty (lk, #l)) with "[Hlock]" as "Hmutex".
{ iExists γ, l, lk. iSplit=>//. }
iDestruct (env_ltyped_insert _ _ x with "Hmutex HΓ'") as "HΓ'".
rewrite /binder_insert insert_delete (insert_id _ _ _ Hg).
iFrame "HΓ'".
Qed.
End rules.
......@@ -171,7 +171,7 @@ Section subtyping_rules.
Qed.
Lemma lty_le_exist C1 C2 :
( A, C1 A <: C2 A) -∗
( A, C1 A <: C2 A) -∗
( A, C1 A) <: ( A, C2 A).
Proof.
iIntros "#Hle" (v) "!>". iDestruct 1 as (A) "H". iExists A. by iApply "Hle".
......@@ -187,7 +187,7 @@ Section subtyping_rules.
Qed.
Lemma lty_copyable_exist (C : ltty Σ ltty Σ) :
( M, lty_copyable (C M)) -∗ lty_copyable (lty_exist C).
( M, lty_copyable (C M)) -∗ lty_copyable (lty_exist C).
Proof.
iIntros "#Hle". rewrite /lty_copyable /tc_opaque.
iApply lty_le_r; last by iApply lty_le_exist_copy.
......@@ -236,33 +236,6 @@ Section subtyping_rules.
lty_copyable (ref_shr A).
Proof. iIntros (v) "!> #Hv !>". iFrame "Hv". Qed.
Lemma lty_le_mutex A1 A2 :
(A1 <:> A2) -∗
mutex A1 <: mutex A2.
Proof.
iIntros "#[Hle1 Hle2]" (v) "!>". iDestruct 1 as (γ l lk ->) "Hinv".
iExists γ, l, lk. iSplit; first done.
iApply (spin_lock.is_lock_iff with "Hinv").
iIntros "!> !>". iSplit.
- iDestruct 1 as (v) "[Hl HA]". iExists v. iFrame "Hl". by iApply "Hle1".
- iDestruct 1 as (v) "[Hl HA]". iExists v. iFrame "Hl". by iApply "Hle2".
Qed.
Lemma lty_copyable_mutex A :
lty_copyable (mutex A).
Proof. iIntros (v) "!> #Hv !>". iFrame "Hv". Qed.
Lemma lty_le_mutexguard A1 A2 :
(A1 <:> A2) -∗
mutexguard A1 <: mutexguard A2.
Proof.
iIntros "#[Hle1 Hle2]" (v) "!>".
iDestruct 1 as (γ l lk w ->) "[Hinv [Hlock Hinner]]".
iExists γ, l, lk, w. iSplit; first done.
iFrame "Hlock Hinner". iApply (spin_lock.is_lock_iff with "Hinv").
iIntros "!> !>". iSplit.
- iDestruct 1 as (v) "[Hl HA]". iExists v. iFrame "Hl". by iApply "Hle1".
- iDestruct 1 as (v) "[Hl HA]". iExists v. iFrame "Hl". by iApply "Hle2".
Qed.
Lemma lty_le_chan S1 S2 :
(S1 <: S2) -∗ chan S1 <: chan S2.
......
......@@ -7,7 +7,7 @@ From actris.channel Require Export channel.
Definition lty_any {Σ} : ltty Σ := Ltty (λ w, True%I).
Definition lty_copy {Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w, ltty_car A w)%I.
Definition lty_copy_inv {Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w, coreP (ltty_car A w)).
Definition lty_copy_minus {Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w, coreP (ltty_car A w)).
Definition lty_copyable {Σ} (A : ltty Σ) : iProp Σ :=
tc_opaque (A <: lty_copy A)%I.
......@@ -27,7 +27,7 @@ Definition lty_sum {Σ} (A1 A2 : ltty Σ) : ltty Σ := Ltty (λ w,
Definition lty_forall `{heapG Σ} {k} (C : lty Σ k ltty Σ) : ltty Σ :=
Ltty (λ w, A, WP w #() {{ ltty_car (C A) }})%I.
Definition lty_exist {Σ k} (C : lty Σ k ltty Σ) : ltty Σ :=
Ltty (λ w, A, ltty_car (C A) w)%I.
Ltty (λ w, A, ltty_car (C A) w)%I.
Definition lty_ref_mut `{heapG Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w,
(l : loc) (v : val), w = #l l v ltty_car A v)%I.
......@@ -35,21 +35,11 @@ Definition ref_shrN := nroot .@ "shr_ref".
Definition lty_ref_shr `{heapG Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w,
l : loc, w = #l inv (ref_shrN .@ l) ( v, l v ltty_car A v))%I.
Definition lty_mutex `{heapG Σ, lockG Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w,
(γ : gname) (l : loc) (lk : val),
w = PairV lk #l
is_lock γ lk ( v_inner, l v_inner ltty_car A v_inner))%I.
Definition lty_mutexguard `{heapG Σ, lockG Σ} (A : ltty Σ) : ltty Σ := Ltty (λ w,
(γ : gname) (l : loc) (lk : val) (v : val),
w = PairV lk #l
is_lock γ lk ( v_inner, l v_inner ltty_car A v_inner)
spin_lock.locked γ l v)%I.
Definition lty_chan `{heapG Σ, chanG Σ} (P : lsty Σ) : ltty Σ :=
Ltty (λ w, w lsty_car P)%I.
Instance: Params (@lty_copy) 1 := {}.
Instance: Params (@lty_copy_inv) 1 := {}.
Instance: Params (@lty_copy_minus) 1 := {}.
Instance: Params (@lty_copyable) 1 := {}.
Instance: Params (@lty_arr) 2 := {}.
Instance: Params (@lty_prod) 1 := {}.
......@@ -58,14 +48,12 @@ Instance: Params (@lty_forall) 2 := {}.
Instance: Params (@lty_sum) 1 := {}.
Instance: Params (@lty_ref_mut) 2 := {}.
Instance: Params (@lty_ref_shr) 2 := {}.
Instance: Params (@lty_mutex) 3 := {}.
Instance: Params (@lty_mutexguard) 3 := {}.
Instance: Params (@lty_chan) 3 := {}.
Notation any := lty_any.
Notation "()" := lty_unit : lty_scope.
Notation "'copy' A" := (lty_copy A) (at level 10) : lty_scope.
Notation "'copy-' A" := (lty_copy_inv A) (at level 10) : lty_scope.
Notation "'copy-' A" := (lty_copy_minus A) (at level 10) : lty_scope.
Notation "A ⊸ B" := (lty_arr A B)
(at level 99, B at level 200, right associativity) : lty_scope.
......@@ -81,9 +69,6 @@ Notation "∃ A1 .. An , C" :=
Notation "'ref_mut' A" := (lty_ref_mut A) (at level 10) : lty_scope.
Notation "'ref_shr' A" := (lty_ref_shr A) (at level 10) : lty_scope.
Notation "'mutex' A" := (lty_mutex A) (at level 10) : lty_scope.
Notation "'mutexguard' A" := (lty_mutexguard A) (at level 10) : lty_scope.
Notation "'chan' A" := (lty_chan A) (at level 10) : lty_scope.
Section term_types.
......@@ -92,7 +77,7 @@ Section term_types.
Global Instance lty_copy_ne : NonExpansive (@lty_copy Σ).
Proof. solve_proper. Qed.
Global Instance lty_copy_inv_ne : NonExpansive (@lty_copy_inv Σ).
Global Instance lty_copy_minus_ne : NonExpansive (@lty_copy_minus Σ).
Proof. solve_proper. Qed.
Global Instance lty_copyable_plain A : Plain (lty_copyable A).
......@@ -130,9 +115,6 @@ Section term_types.
Global Instance lty_forall_ne `{heapG Σ} k n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@lty_forall Σ _ k).
Proof. solve_proper. Qed.
Global Instance lty_exist_contractive k n :
Proper (pointwise_relation _ (dist_later n) ==> dist n) (@lty_exist Σ k).
Proof. solve_contractive. Qed.
Global Instance lty_exist_ne k n :
Proper (pointwise_relation _ (dist n) ==> dist n) (@lty_exist Σ k).
Proof. solve_proper. Qed.
......@@ -147,17 +129,6 @@ Section term_types.
Global Instance lty_ref_shr_ne `{heapG Σ} : NonExpansive lty_ref_shr.
Proof. solve_proper. Qed.
Global Instance lty_mutex_contractive `{heapG Σ, lockG Σ} : Contractive lty_mutex.
Proof. solve_contractive. Qed.
Global Instance lty_mutex_ne `{heapG Σ, lockG Σ} : NonExpansive lty_mutex.
Proof. solve_proper. Qed.
Global Instance lty_mutexguard_contractive `{heapG Σ, lockG Σ} :
Contractive lty_mutexguard.
Proof. solve_contractive. Qed.
Global Instance lty_mutexguard_ne `{heapG Σ, lockG Σ} : NonExpansive lty_mutexguard.
Proof. solve_proper. Qed.
Global Instance lty_chan_ne `{heapG Σ, chanG Σ} : NonExpansive lty_chan.
Proof. solve_proper. Qed.
End term_types.
......@@ -14,19 +14,6 @@ Notation "Γ ⊨ e : A ⫤ Γ'" := (ltyped Γ Γ' e A)
Notation "Γ ⊨ e : A" := (Γ e : A Γ)
(at level 100, e at next level, A at level 200).
Lemma ltyped_frame `{!heapG Σ} (Γ Γ' Γ1 Γ1' Γ2 : gmap string (ltty Σ)) e A :
env_split Γ Γ1 Γ2 -∗
env_split Γ' Γ1' Γ2 -∗
(Γ1 e : A Γ1') -∗
Γ e : A Γ'.
Proof.
iIntros "#Hsplit #Hsplit' #Htyped !>" (vs) "Henv".
iDestruct ("Hsplit" with "Henv") as "[Henv1 Henv2]".
iApply (wp_wand with "(Htyped Henv1)").
iIntros (v) "[$ Henv1']".
iApply "Hsplit'". iFrame "Henv1' Henv2".
Qed.
Lemma ltyped_safety `{heapPreG Σ} e σ es σ' e' :
( `{heapG Σ}, A Γ', e : A Γ')
rtc erased_step ([e], σ) (es, σ') e' es
......
This diff is collapsed.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment